From d6311461f836919a215761bdb1ba3f2f04b5e312 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 3 Aug 2024 17:12:42 +0000 Subject: [PATCH 001/359] feat(Finset/Image): generalize `disjoint_range_addLeftEmbedding` (#15423) and `disjoint_range_addRightEmbedding`. --- Mathlib/Data/Finset/Image.lean | 15 ++++++++++----- 1 file changed, 10 insertions(+), 5 deletions(-) diff --git a/Mathlib/Data/Finset/Image.lean b/Mathlib/Data/Finset/Image.lean index a5dcb5282ec0d..e0e8779501a94 100644 --- a/Mathlib/Data/Finset/Image.lean +++ b/Mathlib/Data/Finset/Image.lean @@ -255,11 +255,16 @@ theorem map_nontrivial : (s.map f).Nontrivial ↔ s.Nontrivial := theorem attach_map_val {s : Finset α} : s.attach.map (Embedding.subtype _) = s := eq_of_veq <| by rw [map_val, attach_val]; exact Multiset.attach_map_val _ -theorem disjoint_range_addLeftEmbedding (a b : ℕ) : - Disjoint (range a) (map (addLeftEmbedding a) (range b)) := by simp [disjoint_left]; omega - -theorem disjoint_range_addRightEmbedding (a b : ℕ) : - Disjoint (range a) (map (addRightEmbedding a) (range b)) := by simp [disjoint_left]; omega +theorem disjoint_range_addLeftEmbedding (a : ℕ) (s : Finset ℕ): + Disjoint (range a) (map (addLeftEmbedding a) s) := by + simp_rw [disjoint_left, mem_map, mem_range, addLeftEmbedding_apply] + rintro _ h ⟨l, -, rfl⟩ + omega + +theorem disjoint_range_addRightEmbedding (a : ℕ) (s : Finset ℕ) : + Disjoint (range a) (map (addRightEmbedding a) s) := by + rw [← addLeftEmbedding_eq_addRightEmbedding] + apply disjoint_range_addLeftEmbedding theorem map_disjiUnion {f : α ↪ β} {s : Finset α} {t : β → Finset γ} {h} : (s.map f).disjiUnion t h = From 10a631f1d3d9d99b990791ca0a251a5e78a15d84 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 3 Aug 2024 20:26:03 +0000 Subject: [PATCH 002/359] chore(Fintype/Perm): use `getElem` in a proof (#15475) --- Mathlib/Data/Fintype/Perm.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/Data/Fintype/Perm.lean b/Mathlib/Data/Fintype/Perm.lean index 16d2d765a1ed2..4e2bc19dc273d 100644 --- a/Mathlib/Data/Fintype/Perm.lean +++ b/Mathlib/Data/Fintype/Perm.lean @@ -99,18 +99,18 @@ theorem nodup_permsOfList : ∀ {l : List α}, l.Nodup → (permsOfList l).Nodup have hln' : (permsOfList l).Nodup := nodup_permsOfList hl' have hmeml : ∀ {f : Perm α}, f ∈ permsOfList l → f a = a := fun {f} hf => not_not.1 (mt (mem_of_mem_permsOfList hf _) (nodup_cons.1 hl).1) - rw [permsOfList, List.nodup_append, List.nodup_bind, pairwise_iff_get] + rw [permsOfList, List.nodup_append, List.nodup_bind, pairwise_iff_getElem] refine ⟨?_, ⟨⟨?_,?_ ⟩, ?_⟩⟩ · exact hln' · exact fun _ _ => hln'.map fun _ _ => mul_left_cancel - · intros i j hij x hx₁ hx₂ + · intros i j hi hj hij x hx₁ hx₂ let ⟨f, hf⟩ := List.mem_map.1 hx₁ let ⟨g, hg⟩ := List.mem_map.1 hx₂ - have hix : x a = List.get l i := by + have hix : x a = l[i] := by rw [← hf.2, mul_apply, hmeml hf.1, swap_apply_left] - have hiy : x a = List.get l j := by + have hiy : x a = l[j] := by rw [← hg.2, mul_apply, hmeml hg.1, swap_apply_left] - have hieqj : i = j := nodup_iff_injective_get.1 hl' (hix.symm.trans hiy) + have hieqj : i = j := hl'.getElem_inj_iff.1 (hix.symm.trans hiy) exact absurd hieqj (_root_.ne_of_lt hij) · intros f hf₁ hf₂ let ⟨x, hx, hx'⟩ := List.mem_bind.1 hf₂ From f9701f4274237958e1e6f07a50320685e1c27e02 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Sun, 4 Aug 2024 10:11:14 +0000 Subject: [PATCH 003/359] feat(NumberTheory/Padics): Q_p and Z_p have an ultrametric (#15078) --- Mathlib/NumberTheory/Padics/PadicIntegers.lean | 2 ++ Mathlib/NumberTheory/Padics/PadicNumbers.lean | 4 ++++ Mathlib/Topology/MetricSpace/Ultra/Basic.lean | 3 +++ 3 files changed, 9 insertions(+) diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 01caec9bff7eb..9bcf9525b8ada 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -197,6 +197,8 @@ variable (p : ℕ) [Fact p.Prime] instance : MetricSpace ℤ_[p] := Subtype.metricSpace +instance : IsUltrametricDist ℤ_[p] := IsUltrametricDist.subtype _ + instance completeSpace : CompleteSpace ℤ_[p] := have : IsClosed { x : ℚ_[p] | ‖x‖ ≤ 1 } := isClosed_le continuous_norm continuous_const this.completeSpace_coe diff --git a/Mathlib/NumberTheory/Padics/PadicNumbers.lean b/Mathlib/NumberTheory/Padics/PadicNumbers.lean index 476ed6962d109..3297fc932dd26 100644 --- a/Mathlib/NumberTheory/Padics/PadicNumbers.lean +++ b/Mathlib/NumberTheory/Padics/PadicNumbers.lean @@ -7,6 +7,7 @@ import Mathlib.RingTheory.Valuation.Basic import Mathlib.NumberTheory.Padics.PadicNorm import Mathlib.Analysis.Normed.Field.Basic import Mathlib.Tactic.Peel +import Mathlib.Topology.MetricSpace.Ultra.Basic /-! # p-adic numbers @@ -706,6 +707,9 @@ variable (p : ℕ) [Fact p.Prime] instance : Dist ℚ_[p] := ⟨fun x y ↦ padicNormE (x - y : ℚ_[p])⟩ +instance : IsUltrametricDist ℚ_[p] := + ⟨fun x y z ↦ by simpa [dist] using padicNormE.nonarchimedean' (x - y) (y - z)⟩ + instance metricSpace : MetricSpace ℚ_[p] where dist_self := by simp [dist] dist := dist diff --git a/Mathlib/Topology/MetricSpace/Ultra/Basic.lean b/Mathlib/Topology/MetricSpace/Ultra/Basic.lean index 25c85fe172bba..5b00adb8cfda9 100644 --- a/Mathlib/Topology/MetricSpace/Ultra/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Ultra/Basic.lean @@ -50,6 +50,9 @@ lemma dist_triangle_max : dist x z ≤ max (dist x y) (dist y z) := namespace IsUltrametricDist +instance subtype (p : X → Prop) : IsUltrametricDist (Subtype p) := + ⟨fun _ _ _ ↦ by simpa [Subtype.dist_eq] using dist_triangle_max _ _ _⟩ + lemma ball_eq_of_mem {x y : X} {r : ℝ} (h : y ∈ ball x r) : ball x r = ball y r := by ext a simp_rw [mem_ball] at h ⊢ From 240d6edbfe88a5f704bf04ee6b0d2f4c814994f6 Mon Sep 17 00:00:00 2001 From: Alessandro Iraci Date: Sun, 4 Aug 2024 10:22:11 +0000 Subject: [PATCH 004/359] feat: Complete homogeneous and monomial symmetric polynomials (#12572) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Implement complete homogeneous and monomial symmetric polynomials. Co-authored-by: Nirvana Coppola Co-authored-by: Yaël Dillies --- .../MvPolynomial/NewtonIdentities.lean | 2 +- .../RingTheory/MvPolynomial/Symmetric.lean | 170 ++++++++++++++---- 2 files changed, 140 insertions(+), 32 deletions(-) diff --git a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean index 3f216f57d4abe..d9f1a7bdaee22 100644 --- a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean @@ -198,7 +198,7 @@ private theorem esymm_to_weight (k : ℕ) : k * esymm σ R k = private theorem esymm_mul_psum_summand_to_weight (k : ℕ) (a : ℕ × ℕ) (ha : a ∈ antidiagonal k) : ∑ A ∈ powersetCard a.fst univ, ∑ j, weight σ R k (A, j) = (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by - simp only [esymm, psum_def, weight, ← mul_assoc, mul_sum] + simp only [esymm, psum, weight, ← mul_assoc, mul_sum] rw [sum_comm] refine sum_congr rfl fun x _ ↦ ?_ rw [sum_mul] diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 8b564402a7a67..49097e138f0bb 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -4,14 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Hanting Zhang, Johan Commelin -/ import Mathlib.Algebra.Algebra.Subalgebra.Basic -import Mathlib.Algebra.MvPolynomial.Rename import Mathlib.Algebra.MvPolynomial.CommRing +import Mathlib.Combinatorics.Enumerative.Partition /-! # Symmetric Polynomials and Elementary Symmetric Polynomials -This file defines symmetric `MvPolynomial`s and elementary symmetric `MvPolynomial`s. -We also prove some basic facts about them. +This file defines symmetric `MvPolynomial`s and the bases of elementary, complete homogeneous, +power sum, and monomial symmetric `MvPolynomial`s. We also prove some basic facts about them. ## Main declarations @@ -21,15 +21,24 @@ We also prove some basic facts about them. * `MvPolynomial.esymm` +* `MvPolynomial.hsymm` + * `MvPolynomial.psum` +* `MvPolynomial.msymm` + ## Notation + `esymm σ R n` is the `n`th elementary symmetric polynomial in `MvPolynomial σ R`. ++ `hsymm σ R n` is the `n`th complete homogeneous symmetric polynomial in `MvPolynomial σ R`. + + `psum σ R n` is the degree-`n` power sum in `MvPolynomial σ R`, i.e. the sum of monomials `(X i)^n` over `i ∈ σ`. ++ `msymm σ R μ` is the monomial symmetric polynomial whose exponents set are the parts + of `μ ⊢ n` in `MvPolynomial σ R`. + As in other polynomial files, we typically use the notation: + `σ τ : Type*` (indexing the variables) @@ -66,25 +75,20 @@ end Multiset namespace MvPolynomial -variable {σ : Type*} {R : Type*} -variable {τ : Type*} {S : Type*} +variable {σ τ : Type*} {R S : Type*} /-- A `MvPolynomial φ` is symmetric if it is invariant under permutations of its variables by the `rename` operation -/ def IsSymmetric [CommSemiring R] (φ : MvPolynomial σ R) : Prop := ∀ e : Perm σ, rename e φ = φ -variable (σ R) - /-- The subalgebra of symmetric `MvPolynomial`s. -/ -def symmetricSubalgebra [CommSemiring R] : Subalgebra R (MvPolynomial σ R) where +def symmetricSubalgebra (σ R : Type*) [CommSemiring R] : Subalgebra R (MvPolynomial σ R) where carrier := setOf IsSymmetric algebraMap_mem' r e := rename_C e r mul_mem' ha hb e := by rw [map_mul, ha, hb] add_mem' ha hb e := by rw [map_add, ha, hb] -variable {σ R} - @[simp] theorem mem_symmetricSubalgebra [CommSemiring R] (p : MvPolynomial σ R) : p ∈ symmetricSubalgebra σ R ↔ p.IsSymmetric := @@ -157,22 +161,29 @@ def renameSymmetricSubalgebra [CommSemiring R] (e : σ ≃ τ) : (AlgHom.ext <| fun p => Subtype.ext <| by simp) (AlgHom.ext <| fun p => Subtype.ext <| by simp) +variable (σ R : Type*) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] + section ElementarySymmetric open Finset -variable (σ R) [CommSemiring R] [CommSemiring S] [Fintype σ] [Fintype τ] - -/-- The `n`th elementary symmetric `MvPolynomial σ R`. -/ +/-- The `n`th elementary symmetric `MvPolynomial σ R`. +It is the sum over all the degree n squarefree monomials in `MvPolynomial σ R`. -/ def esymm (n : ℕ) : MvPolynomial σ R := ∑ t ∈ powersetCard n univ, ∏ i ∈ t, X i +/-- +`esymmPart` is the product of the symmetric polynomials `esymm μᵢ`, +where `μ = (μ₁, μ₂, ...)` is a partition. +-/ +def esymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (esymm σ R)).prod + /-- The `n`th elementary symmetric `MvPolynomial σ R` is obtained by evaluating the `n`th elementary symmetric at the `Multiset` of the monomials -/ theorem esymm_eq_multiset_esymm : esymm σ R = (univ.val.map X).esymm := by exact funext fun n => (esymm_map_val X _ n).symm -theorem aeval_esymm_eq_multiset_esymm [Algebra R S] (f : σ → S) (n : ℕ) : +theorem aeval_esymm_eq_multiset_esymm [Algebra R S] (n : ℕ) (f : σ → S) : aeval f (esymm σ R n) = (univ.val.map f).esymm n := by simp_rw [esymm, aeval_sum, aeval_prod, aeval_X, esymm_map_val] @@ -191,6 +202,15 @@ theorem esymm_eq_sum_monomial (n : ℕ) : theorem esymm_zero : esymm σ R 0 = 1 := by simp only [esymm, powersetCard_zero, sum_singleton, prod_empty] +@[simp] +theorem esymm_one : esymm σ R 1 = ∑ i, X i := by simp [esymm, powersetCard_one] + +theorem esymmPart_zero : esymmPart σ R (.indiscrete 0) = 1 := by simp [esymmPart] + +@[simp] +theorem esymmPart_indiscrete (n : ℕ) : esymmPart σ R (.indiscrete n) = esymm σ R n := by + cases n <;> simp [esymmPart] + theorem map_esymm (n : ℕ) (f : R →+* S) : map f (esymm σ R n) = esymm σ S n := by simp_rw [esymm, map_sum, map_prod, map_X] @@ -209,7 +229,7 @@ theorem esymm_isSymmetric (n : ℕ) : IsSymmetric (esymm σ R n) := by intro rw [rename_esymm] -theorem support_esymm'' (n : ℕ) [DecidableEq σ] [Nontrivial R] : +theorem support_esymm'' [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm σ R n).support = (powersetCard n (univ : Finset σ)).biUnion fun t => (Finsupp.single (∑ i ∈ t, Finsupp.single i 1) (1 : R)).support := by @@ -233,21 +253,19 @@ theorem support_esymm'' (n : ℕ) [DecidableEq σ] [Nontrivial R] : exact absurd this hst.symm all_goals intro x y; simp [Finsupp.support_single_disjoint] -theorem support_esymm' (n : ℕ) [DecidableEq σ] [Nontrivial R] : - (esymm σ R n).support = - (powersetCard n (univ : Finset σ)).biUnion fun t => {∑ i ∈ t, Finsupp.single i 1} := by +theorem support_esymm' [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm σ R n).support = + (powersetCard n (univ : Finset σ)).biUnion fun t => {∑ i ∈ t, Finsupp.single i 1} := by rw [support_esymm''] congr funext exact Finsupp.support_single_ne_zero _ one_ne_zero -theorem support_esymm (n : ℕ) [DecidableEq σ] [Nontrivial R] : - (esymm σ R n).support = - (powersetCard n (univ : Finset σ)).image fun t => ∑ i ∈ t, Finsupp.single i 1 := by +theorem support_esymm [DecidableEq σ] [Nontrivial R] (n : ℕ) : (esymm σ R n).support = + (powersetCard n (univ : Finset σ)).image fun t => ∑ i ∈ t, Finsupp.single i 1 := by rw [support_esymm'] exact biUnion_singleton -theorem degrees_esymm [Nontrivial R] (n : ℕ) (hpos : 0 < n) (hn : n ≤ Fintype.card σ) : +theorem degrees_esymm [Nontrivial R] {n : ℕ} (hpos : 0 < n) (hn : n ≤ Fintype.card σ) : (esymm σ R n).degrees = (univ : Finset σ).val := by classical have : @@ -268,25 +286,71 @@ theorem degrees_esymm [Nontrivial R] (n : ℕ) (hpos : 0 < n) (hn : n ≤ Fintyp end ElementarySymmetric +section CompleteHomogeneousSymmetric + +open Finset Multiset Sym + +variable [DecidableEq σ] [DecidableEq τ] + +/-- The `n`th complete homogeneous symmetric `MvPolynomial σ R`. +It is the sum over all the degree n monomials in `MvPolynomial σ R`. -/ +def hsymm (n : ℕ) : MvPolynomial σ R := ∑ s : Sym σ n, (s.1.map X).prod + +/-- `hsymmPart` is the product of the symmetric polynomials `hsymm μᵢ`, +where `μ = (μ₁, μ₂, ...)` is a partition. -/ +def hsymmPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (hsymm σ R)).prod + +@[simp] +theorem hsymm_zero : hsymm σ R 0 = 1 := by simp [hsymm, eq_nil_of_card_zero] + +@[simp] +theorem hsymm_one : hsymm σ R 1 = ∑ i, X i := by + symm + apply Fintype.sum_equiv oneEquiv + simp only [oneEquiv_apply, Multiset.map_singleton, Multiset.prod_singleton, implies_true] + +theorem hsymmPart_zero : hsymmPart σ R (.indiscrete 0) = 1 := by simp [hsymmPart] + +@[simp] +theorem hsymmPart_indiscrete (n : ℕ) : hsymmPart σ R (.indiscrete n) = hsymm σ R n := by + cases n <;> simp [hsymmPart] + +theorem map_hsymm (n : ℕ) (f : R →+* S) : map f (hsymm σ R n) = hsymm σ S n := by + simp [hsymm, ← Multiset.prod_hom'] + +theorem rename_hsymm (n : ℕ) (e : σ ≃ τ) : rename e (hsymm σ R n) = hsymm τ R n := by + simp_rw [hsymm, map_sum, ← prod_hom', rename_X] + apply Fintype.sum_equiv (equivCongr e) + simp + +theorem hsymm_isSymmetric (n : ℕ) : IsSymmetric (hsymm σ R n) := rename_hsymm _ _ n + +end CompleteHomogeneousSymmetric + section PowerSum open Finset -variable (σ R) [CommSemiring R] [Fintype σ] [Fintype τ] - -/-- The degree-`n` power sum -/ +/-- The degree-`n` power sum symmetric `MvPolynomial σ R`. +It is the sum over all the `n`-th powers of the variables. -/ def psum (n : ℕ) : MvPolynomial σ R := ∑ i, X i ^ n -lemma psum_def (n : ℕ) : psum σ R n = ∑ i, X i ^ n := rfl +/-- `psumPart` is the product of the symmetric polynomials `psum μᵢ`, +where `μ = (μ₁, μ₂, ...)` is a partition. -/ +def psumPart {n : ℕ} (μ : n.Partition) : MvPolynomial σ R := (μ.parts.map (psum σ R)).prod + +@[simp] +theorem psum_zero : psum σ R 0 = Fintype.card σ := by simp [psum] + +@[simp] +theorem psum_one : psum σ R 1 = ∑ i, X i := by simp [psum] @[simp] -theorem psum_zero : psum σ R 0 = Fintype.card σ := by - simp only [psum, _root_.pow_zero, ← cast_card] - exact rfl +theorem psumPart_zero : psumPart σ R (.indiscrete 0) = 1 := by simp [psumPart] @[simp] -theorem psum_one : psum σ R 1 = ∑ i, X i := by - simp only [psum, _root_.pow_one] +theorem psumPart_indiscrete {n : ℕ} (npos : n ≠ 0) : + psumPart σ R (.indiscrete n) = psum σ R n := by simp [psumPart, npos] @[simp] theorem rename_psum (n : ℕ) (e : σ ≃ τ) : rename e (psum σ R n) = psum τ R n := by @@ -296,4 +360,48 @@ theorem psum_isSymmetric (n : ℕ) : IsSymmetric (psum σ R n) := rename_psum _ end PowerSum +section MonomialSymmetric + +variable [DecidableEq σ] [DecidableEq τ] {n : ℕ} + +/-- The monomial symmetric `MvPolynomial σ R` with exponent set μ. +It is the sum over all the monomials in `MvPolynomial σ R` such that +the multiset of exponents is equal to the multiset of parts of μ. -/ +def msymm (μ : n.Partition) : MvPolynomial σ R := + ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod + +@[simp] +theorem msymm_zero : msymm σ R (.indiscrete 0) = 1 := by + rw [msymm, Fintype.sum_subsingleton _ ⟨(Sym.nil : Sym σ 0), by rfl⟩] + simp + +@[simp] +theorem msymm_one : msymm σ R (.indiscrete 1) = ∑ i, X i := by + have : (fun (x : Sym σ 1) ↦ x ∈ Set.univ) = + (fun x ↦ Nat.Partition.ofSym x = Nat.Partition.indiscrete 1) := by + simp_rw [Set.mem_univ, Nat.Partition.ofSym_one] + symm + rw [Fintype.sum_equiv (Equiv.trans Sym.oneEquiv (Equiv.Set.univ (Sym σ 1)).symm) + _ (fun s ↦ (s.1.1.map X).prod)] + · apply Fintype.sum_equiv (Equiv.subtypeEquivProp this) + intro x + congr + · intro x + rw [← Multiset.prod_singleton (X x), ← Multiset.map_singleton] + congr + +@[simp] +theorem rename_msymm (μ : n.Partition) (e : σ ≃ τ) : + rename e (msymm σ R μ) = msymm τ R μ := by + rw [msymm, map_sum] + apply Fintype.sum_equiv (Nat.Partition.ofSymShapeEquiv μ e) + intro + rw [← Multiset.prod_hom, Multiset.map_map, Nat.Partition.ofSymShapeEquiv] + simp + +theorem msymm_isSymmetric (μ : n.Partition) : IsSymmetric (msymm σ R μ) := + rename_msymm _ _ μ + +end MonomialSymmetric + end MvPolynomial From d027c3a3153b809e85efbf24d8ceb33eca4a6457 Mon Sep 17 00:00:00 2001 From: Patrick Massot Date: Sun, 4 Aug 2024 10:22:13 +0000 Subject: [PATCH 005/359] feat: add simp lemmas for explicit matrix traces (#15330) --- Mathlib/LinearAlgebra/Matrix/Trace.lean | 21 ++++++++++++++++----- 1 file changed, 16 insertions(+), 5 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/Trace.lean b/Mathlib/LinearAlgebra/Matrix/Trace.lean index 44f96e45688ad..1e018cf22f592 100644 --- a/Mathlib/LinearAlgebra/Matrix/Trace.lean +++ b/Mathlib/LinearAlgebra/Matrix/Trace.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ import Mathlib.Data.Matrix.Block import Mathlib.Data.Matrix.RowCol +import Mathlib.Data.Matrix.Notation /-! # Trace of a matrix @@ -178,13 +179,10 @@ section Fin variable [AddCommMonoid R] -/-! ### Special cases for `Fin n` - -While `simp [Fin.sum_univ_succ]` can prove these, we include them for convenience and consistency -with `Matrix.det_fin_two` etc. +/-! ### Special cases for `Fin n` for low values of `n` -/ - +@[simp] theorem trace_fin_zero (A : Matrix (Fin 0) (Fin 0) R) : trace A = 0 := rfl @@ -198,6 +196,19 @@ theorem trace_fin_three (A : Matrix (Fin 3) (Fin 3) R) : trace A = A 0 0 + A 1 1 rw [← add_zero (A 2 2), add_assoc] rfl +@[simp] +theorem trace_fin_one_of (a : R) : trace !![a] = a := + trace_fin_one _ + +@[simp] +theorem trace_fin_two_of (a b c d : R) : trace !![a, b; c, d] = a + d := + trace_fin_two _ + +@[simp] +theorem trace_fin_three_of (a b c d e f g h i : R) : + trace !![a, b, c; d, e, f; g, h, i] = a + e + i := + trace_fin_three _ + end Fin end Matrix From 2a92c208ab1496acb010448c144a0d15dc6d2989 Mon Sep 17 00:00:00 2001 From: grunweg Date: Sun, 4 Aug 2024 10:22:14 +0000 Subject: [PATCH 006/359] doc(Cyclotomic/Basic): fix theorem docstring (#15337) --- Mathlib/NumberTheory/Cyclotomic/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/Cyclotomic/Basic.lean b/Mathlib/NumberTheory/Cyclotomic/Basic.lean index 3dd2d52f0f8cc..08a0a83d2f9f7 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Basic.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Basic.lean @@ -415,7 +415,7 @@ theorem splits_X_pow_sub_one [H : IsCyclotomicExtension S K L] (hS : n ∈ S) : obtain ⟨z, hz⟩ := ((isCyclotomicExtension_iff _ _ _).1 H).1 hS exact X_pow_sub_one_splits hz -/-- A cyclotomic extension splits `cyclotomic n K` if `n ∈ S` and `ne_zero (n : K)`. -/ +/-- A cyclotomic extension splits `cyclotomic n K` if `n ∈ S`. -/ theorem splits_cyclotomic [IsCyclotomicExtension S K L] (hS : n ∈ S) : Splits (algebraMap K L) (cyclotomic n K) := by refine splits_of_splits_of_dvd _ (X_pow_sub_C_ne_zero n.pos _) (splits_X_pow_sub_one K L hS) ?_ From cbbb5d7d849bf4f1bacde6d121ba73238602d1be Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Calle=20S=C3=B6nne?= Date: Sun, 4 Aug 2024 10:58:26 +0000 Subject: [PATCH 007/359] feat(Limits/Shapes/Pullback): add more pasting lemmas (#14414) This PR does the following: 1. Rewrites the current pasting lemmas so that they take any `PullbackCone` as input, and not just one obtained through `PullbackCone.mk`. 2. Adds vertical pasting lemmas 3. Adds some missing pasting lemmas in terms of the `HasPullbacks`/`HasPushouts` API (e.g. `pullbackRightPullbackSndIso`). This contribution was inspired by the AIM workshop "Formalizing algebraic geometry" in June 2024. --- Mathlib/AlgebraicGeometry/Pullbacks.lean | 34 +- .../Limits/Shapes/Diagonal.lean | 64 +- .../Limits/Shapes/Pullback/Assoc.lean | 61 +- .../Limits/Shapes/Pullback/CommSq.lean | 9 +- .../Limits/Shapes/Pullback/Mono.lean | 10 +- .../Limits/Shapes/Pullback/Pasting.lean | 662 ++++++++++++++---- 6 files changed, 592 insertions(+), 248 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/Pullbacks.lean b/Mathlib/AlgebraicGeometry/Pullbacks.lean index c6a2251ef7e71..27e3ad47bf323 100644 --- a/Mathlib/AlgebraicGeometry/Pullbacks.lean +++ b/Mathlib/AlgebraicGeometry/Pullbacks.lean @@ -438,18 +438,8 @@ instance base_affine_hasPullback {C : CommRingCat} {X Y : Scheme} (f : X ⟶ Spe instance left_affine_comp_pullback_hasPullback {X Y Z : Scheme} (f : X ⟶ Z) (g : Y ⟶ Z) (i : Z.affineCover.J) : HasPullback ((Z.affineCover.pullbackCover f).map i ≫ f) g := by - let Xᵢ := pullback f (Z.affineCover.map i) - let Yᵢ := pullback g (Z.affineCover.map i) - let W := pullback (pullback.snd _ _ : Yᵢ ⟶ _) (pullback.snd _ _ : Xᵢ ⟶ _) - have := - bigSquareIsPullback (pullback.fst _ _ : W ⟶ _) (pullback.fst _ _ : Yᵢ ⟶ _) - (pullback.snd _ _ : Xᵢ ⟶ _) (Z.affineCover.map i) (pullback.snd _ _) - (pullback.snd _ _) g pullback.condition.symm - pullback.condition.symm (PullbackCone.isLimitOfFlip <| pullbackIsPullback _ _) - (PullbackCone.isLimitOfFlip <| pullbackIsPullback _ _) - have : HasPullback (pullback.snd _ _ ≫ Z.affineCover.map i : Xᵢ ⟶ _) g := ⟨⟨⟨_, this⟩⟩⟩ - rw [← pullback.condition] at this - exact this + simp only [OpenCover.pullbackCover_obj, OpenCover.pullbackCover_map, pullback.condition] + exact hasPullback_assoc_symm f (Z.affineCover.map i) (Z.affineCover.map i) g instance {X Y Z : Scheme} (f : X ⟶ Z) (g : Y ⟶ Z) : HasPullback f g := hasPullback_of_cover (Z.affineCover.pullbackCover f) f g @@ -512,22 +502,15 @@ def openCoverOfLeftRight (𝒰X : X.OpenCover) (𝒰Y : Y.OpenCover) (f : X ⟶ def openCoverOfBase' (𝒰 : OpenCover Z) (f : X ⟶ Z) (g : Y ⟶ Z) : OpenCover (pullback f g) := by apply (openCoverOfLeft (𝒰.pullbackCover f) f g).bind intro i - let Xᵢ := pullback f (𝒰.map i) - let Yᵢ := pullback g (𝒰.map i) - let W := pullback (pullback.snd _ _ : Yᵢ ⟶ _) (pullback.snd _ _ : Xᵢ ⟶ _) - have := - bigSquareIsPullback (pullback.fst _ _ : W ⟶ _) (pullback.fst _ _ : Yᵢ ⟶ _) - (pullback.snd _ _ : Xᵢ ⟶ _) (𝒰.map i) (pullback.snd _ _) (pullback.snd _ _) g - pullback.condition.symm pullback.condition.symm - (PullbackCone.isLimitOfFlip <| pullbackIsPullback _ _) - (PullbackCone.isLimitOfFlip <| pullbackIsPullback _ _) + have := PullbackCone.flipIsLimit <| + pasteVertIsPullback rfl (pullbackIsPullback g (𝒰.map i)) + (pullbackIsPullback (pullback.snd g (𝒰.map i)) (pullback.snd f (𝒰.map i))) refine @openCoverOfIsIso (f := (pullbackSymmetry _ _).hom ≫ (limit.isoLimitCone ⟨_, this⟩).inv ≫ - pullback.map _ _ _ _ (𝟙 _) (𝟙 _) (𝟙 _) ?_ ?_) ?_ + pullback.map _ _ _ _ (𝟙 _) (𝟙 _) (𝟙 _) ?_ ?_) inferInstance · simp [← pullback.condition] · simp only [Category.comp_id, Category.id_comp] - · infer_instance /-- Given an open cover `{ Zᵢ }` of `Z`, then `X ×[Z] Y` is covered by `Xᵢ ×[Zᵢ] Yᵢ`, where `Xᵢ = X ×[Z] Zᵢ` and `Yᵢ = Y ×[Z] Zᵢ` is the preimage of `Zᵢ` in `X` and `Y`. -/ @@ -547,8 +530,9 @@ def openCoverOfBase (𝒰 : OpenCover Z) (f : X ⟶ Z) (g : Y ⟶ Z) : OpenCover apply pullback.hom_ext <;> dsimp <;> · simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, Category.assoc, limit.lift_π_assoc, cospan_left, Category.comp_id, limit.isoLimitCone_inv_π, - limit.isoLimitCone_inv_π_assoc, pullbackSymmetry_hom_comp_fst_assoc, - pullbackSymmetry_hom_comp_snd_assoc] + limit.isoLimitCone_inv_π_assoc, PullbackCone.flip_pt, PullbackCone.π_app_left, + PullbackCone.π_app_right, PullbackCone.flip_fst, PullbackCone.flip_snd, + pullbackSymmetry_hom_comp_snd_assoc, pullbackSymmetry_hom_comp_fst_assoc] rfl end Pullback diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean b/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean index 4c7ff0cbdedd3..9ec24870220e9 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean @@ -99,7 +99,34 @@ theorem pullback_diagonal_map_snd_snd_fst : variable [HasPullback i₁ i₂] -set_option maxHeartbeats 400000 in +/-- The underlying map of `pullbackDiagonalIso` -/ +abbrev pullbackDiagonalMapIso.hom : + pullback (diagonal f) + (map (i₁ ≫ snd _ _) (i₂ ≫ snd _ _) f f (i₁ ≫ fst _ _) (i₂ ≫ fst _ _) i + (by simp only [Category.assoc, condition]) + (by simp only [Category.assoc, condition])) ⟶ + pullback i₁ i₂ := + pullback.lift (pullback.snd _ _ ≫ pullback.fst _ _) (pullback.snd _ _ ≫ pullback.snd _ _) (by + ext + · simp only [Category.assoc, pullback_diagonal_map_snd_fst_fst, + pullback_diagonal_map_snd_snd_fst] + · simp only [Category.assoc, condition]) + +/-- The underlying inverse of `pullbackDiagonalIso` -/ +abbrev pullbackDiagonalMapIso.inv : pullback i₁ i₂ ⟶ + pullback (diagonal f) + (map (i₁ ≫ snd _ _) (i₂ ≫ snd _ _) f f (i₁ ≫ fst _ _) (i₂ ≫ fst _ _) i + (by simp only [Category.assoc, condition]) + (by simp only [Category.assoc, condition])) := + pullback.lift (pullback.fst _ _ ≫ i₁ ≫ pullback.fst _ _) + (pullback.map _ _ _ _ (𝟙 _) (𝟙 _) (pullback.snd _ _) (Category.id_comp _).symm + (Category.id_comp _).symm) (by + ext + · simp only [Category.assoc, diagonal_fst, Category.comp_id, limit.lift_π, + PullbackCone.mk_pt, PullbackCone.mk_π_app, limit.lift_π_assoc, cospan_left] + · simp only [condition_assoc, Category.assoc, diagonal_snd, Category.comp_id, limit.lift_π, + PullbackCone.mk_pt, PullbackCone.mk_π_app, limit.lift_π_assoc, cospan_right]) + /-- This iso witnesses the fact that given `f : X ⟶ Y`, `i : U ⟶ Y`, and `i₁ : V₁ ⟶ X ×[Y] U`, `i₂ : V₂ ⟶ X ×[Y] U`, the diagram @@ -120,52 +147,39 @@ def pullbackDiagonalMapIso : (by simp only [Category.assoc, condition]) (by simp only [Category.assoc, condition])) ≅ pullback i₁ i₂ where - hom := - pullback.lift (pullback.snd _ _ ≫ pullback.fst _ _) (pullback.snd _ _ ≫ pullback.snd _ _) (by - ext - · simp [Category.assoc, pullback_diagonal_map_snd_fst_fst, pullback_diagonal_map_snd_snd_fst] - · simp [Category.assoc, pullback.condition, pullback.condition_assoc]) - inv := - pullback.lift (pullback.fst _ _ ≫ i₁ ≫ pullback.fst _ _) - (pullback.map _ _ _ _ (𝟙 _) (𝟙 _) (pullback.snd _ _) (Category.id_comp _).symm - (Category.id_comp _).symm) (by - ext - · simp only [Category.assoc, diagonal_fst, Category.comp_id, limit.lift_π, - PullbackCone.mk_pt, PullbackCone.mk_π_app, limit.lift_π_assoc, cospan_left] - · simp only [condition_assoc, Category.assoc, diagonal_snd, Category.comp_id, - limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, - limit.lift_π_assoc, cospan_right]) + hom := pullbackDiagonalMapIso.hom f i i₁ i₂ + inv := pullbackDiagonalMapIso.inv f i i₁ i₂ @[reassoc (attr := simp)] -theorem pullbackDiagonalMapIso_hom_fst : +theorem pullbackDiagonalMapIso.hom_fst : (pullbackDiagonalMapIso f i i₁ i₂).hom ≫ pullback.fst _ _ = pullback.snd _ _ ≫ pullback.fst _ _ := by delta pullbackDiagonalMapIso - simp + simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app] @[reassoc (attr := simp)] -theorem pullbackDiagonalMapIso_hom_snd : +theorem pullbackDiagonalMapIso.hom_snd : (pullbackDiagonalMapIso f i i₁ i₂).hom ≫ pullback.snd _ _ = pullback.snd _ _ ≫ pullback.snd _ _ := by delta pullbackDiagonalMapIso - simp + simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app] @[reassoc (attr := simp)] -theorem pullbackDiagonalMapIso_inv_fst : +theorem pullbackDiagonalMapIso.inv_fst : (pullbackDiagonalMapIso f i i₁ i₂).inv ≫ pullback.fst _ _ = pullback.fst _ _ ≫ i₁ ≫ pullback.fst _ _ := by delta pullbackDiagonalMapIso - simp + simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app] @[reassoc (attr := simp)] -theorem pullbackDiagonalMapIso_inv_snd_fst : +theorem pullbackDiagonalMapIso.inv_snd_fst : (pullbackDiagonalMapIso f i i₁ i₂).inv ≫ pullback.snd _ _ ≫ pullback.fst _ _ = pullback.fst _ _ := by delta pullbackDiagonalMapIso simp @[reassoc (attr := simp)] -theorem pullbackDiagonalMapIso_inv_snd_snd : +theorem pullbackDiagonalMapIso.inv_snd_snd : (pullbackDiagonalMapIso f i i₁ i₂).inv ≫ pullback.snd _ _ ≫ pullback.snd _ _ = pullback.snd _ _ := by delta pullbackDiagonalMapIso @@ -179,7 +193,7 @@ theorem pullback_fst_map_snd_isPullback : (map (i₁ ≫ snd _ _) (i₂ ≫ snd _ _) f f (i₁ ≫ fst _ _) (i₂ ≫ fst _ _) i (by simp [condition]) (by simp [condition])) := IsPullback.of_iso_pullback ⟨by ext <;> simp [condition_assoc]⟩ - (pullbackDiagonalMapIso f i i₁ i₂).symm (pullbackDiagonalMapIso_inv_fst f i i₁ i₂) + (pullbackDiagonalMapIso f i i₁ i₂).symm (pullbackDiagonalMapIso.inv_fst f i i₁ i₂) (by aesop_cat) end diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Assoc.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Assoc.lean index 49acc342db45e..b6008f36b9768 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Assoc.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Assoc.lean @@ -103,10 +103,8 @@ local notation "l₂'" => (pullback.snd f₁ (g₃ ≫ f₂)) /-- `(X₁ ×[Y₁] X₂) ×[Y₂] X₃` is the pullback `(X₁ ×[Y₁] X₂) ×[X₂] (X₂ ×[Y₂] X₃)`. -/ def pullbackPullbackLeftIsPullback [HasPullback (g₂ ≫ f₃) f₄] : IsLimit (PullbackCone.mk l₁ l₂ (show l₁ ≫ g₂ = l₂ ≫ g₃ from (pullback.lift_fst _ _ _).symm)) := by - apply leftSquareIsPullback - · exact pullbackIsPullback f₃ f₄ - convert pullbackIsPullback (g₂ ≫ f₃) f₄ - rw [pullback.lift_snd] + apply leftSquareIsPullback _ rfl (pullbackIsPullback f₃ f₄) + simpa [PullbackCone.pasteHoriz] using PullbackCone.mkSelfIsLimit (pullbackIsPullback _ f₄) /-- `(X₁ ×[Y₁] X₂) ×[Y₂] X₃` is the pullback `X₁ ×[Y₁] (X₂ ×[Y₂] X₃)`. -/ def pullbackAssocIsPullback [HasPullback (g₂ ≫ f₃) f₄] : @@ -114,14 +112,8 @@ def pullbackAssocIsPullback [HasPullback (g₂ ≫ f₃) f₄] : (PullbackCone.mk (l₁ ≫ g₁) l₂ (show (l₁ ≫ g₁) ≫ f₁ = l₂ ≫ g₃ ≫ f₂ by rw [pullback.lift_fst_assoc, Category.assoc, Category.assoc, pullback.condition])) := by - apply PullbackCone.isLimitOfFlip - apply bigSquareIsPullback - · apply PullbackCone.isLimitOfFlip - exact pullbackIsPullback f₁ f₂ - · apply PullbackCone.isLimitOfFlip - apply pullbackPullbackLeftIsPullback - · exact pullback.lift_fst _ _ _ - · exact pullback.condition.symm + simpa using pasteVertIsPullback rfl (pullbackIsPullback _ _) + (pullbackPullbackLeftIsPullback f₁ f₂ f₃ f₄) theorem hasPullback_assoc [HasPullback (g₂ ≫ f₃) f₄] : HasPullback f₁ (g₃ ≫ f₂) := ⟨⟨⟨_, pullbackAssocIsPullback f₁ f₂ f₃ f₄⟩⟩⟩ @@ -129,14 +121,8 @@ theorem hasPullback_assoc [HasPullback (g₂ ≫ f₃) f₄] : HasPullback f₁ /-- `X₁ ×[Y₁] (X₂ ×[Y₂] X₃)` is the pullback `(X₁ ×[Y₁] X₂) ×[X₂] (X₂ ×[Y₂] X₃)`. -/ def pullbackPullbackRightIsPullback [HasPullback f₁ (g₃ ≫ f₂)] : IsLimit (PullbackCone.mk l₁' l₂' (show l₁' ≫ g₂ = l₂' ≫ g₃ from pullback.lift_snd _ _ _)) := by - apply PullbackCone.isLimitOfFlip - apply leftSquareIsPullback - · apply PullbackCone.isLimitOfFlip - exact pullbackIsPullback f₁ f₂ - · apply PullbackCone.isLimitOfFlip - exact IsLimit.ofIsoLimit (pullbackIsPullback f₁ (g₃ ≫ f₂)) - (PullbackCone.ext (Iso.refl _) (by simp) (by simp)) - · exact pullback.condition.symm + apply topSquareIsPullback _ rfl (pullbackIsPullback f₁ f₂) + simpa [PullbackCone.pasteVert] using PullbackCone.mkSelfIsLimit (pullbackIsPullback _ _) /-- `X₁ ×[Y₁] (X₂ ×[Y₂] X₃)` is the pullback `(X₁ ×[Y₁] X₂) ×[Y₂] X₃`. -/ def pullbackAssocSymmIsPullback [HasPullback f₁ (g₃ ≫ f₂)] : @@ -144,9 +130,8 @@ def pullbackAssocSymmIsPullback [HasPullback f₁ (g₃ ≫ f₂)] : (PullbackCone.mk l₁' (l₂' ≫ g₄) (show l₁' ≫ g₂ ≫ f₃ = (l₂' ≫ g₄) ≫ f₄ by rw [pullback.lift_snd_assoc, Category.assoc, Category.assoc, pullback.condition])) := by - apply bigSquareIsPullback - · exact pullbackIsPullback f₃ f₄ - · apply pullbackPullbackRightIsPullback + simpa [PullbackCone.pasteHoriz] using pasteHorizIsPullback rfl + (pullbackIsPullback f₃ f₄) (pullbackPullbackRightIsPullback _ _ _ _) theorem hasPullback_assoc_symm [HasPullback f₁ (g₃ ≫ f₂)] : HasPullback (g₂ ≫ f₃) f₄ := ⟨⟨⟨_, pullbackAssocSymmIsPullback f₁ f₂ f₃ f₄⟩⟩⟩ @@ -288,13 +273,9 @@ local notation "l₂'" => def pushoutPushoutLeftIsPushout [HasPushout (g₃ ≫ f₂) g₄] : IsColimit (PushoutCocone.mk l₁' l₂' (show f₂ ≫ l₁' = f₃ ≫ l₂' from (pushout.inl_desc _ _ _).symm)) := by - apply PushoutCocone.isColimitOfFlip - apply rightSquareIsPushout - · apply PushoutCocone.isColimitOfFlip - exact pushoutIsPushout g₃ g₄ - · exact IsColimit.ofIsoColimit (PushoutCocone.flipIsColimit (pushoutIsPushout (g₃ ≫ f₂) g₄)) - (PushoutCocone.ext (Iso.refl _) (by simp) (by simp)) - · exact pushout.condition.symm + apply botSquareIsPushout _ rfl (pushoutIsPushout _ g₄) + simpa [PushoutCocone.pasteVert] using + PushoutCocone.mkSelfIsColimit (pushoutIsPushout (g₃ ≫ f₂) g₄) /-- `(X₁ ⨿[Z₁] X₂) ⨿[Z₂] X₃` is the pushout `X₁ ⨿[Z₁] (X₂ ⨿[Z₂] X₃)`. -/ def pushoutAssocIsPushout [HasPushout (g₃ ≫ f₂) g₄] : @@ -302,9 +283,8 @@ def pushoutAssocIsPushout [HasPushout (g₃ ≫ f₂) g₄] : (PushoutCocone.mk (f₁ ≫ l₁') l₂' (show g₁ ≫ f₁ ≫ l₁' = (g₂ ≫ f₃) ≫ l₂' by rw [Category.assoc, pushout.inl_desc, pushout.condition_assoc])) := by - apply bigSquareIsPushout - · apply pushoutPushoutLeftIsPushout - · exact pushoutIsPushout _ _ + simpa using pasteHorizIsPushout rfl (pushoutIsPushout g₁ g₂) + (pushoutPushoutLeftIsPushout g₁ g₂ g₃ g₄) theorem hasPushout_assoc [HasPushout (g₃ ≫ f₂) g₄] : HasPushout g₁ (g₂ ≫ f₃) := ⟨⟨⟨_, pushoutAssocIsPushout g₁ g₂ g₃ g₄⟩⟩⟩ @@ -312,10 +292,8 @@ theorem hasPushout_assoc [HasPushout (g₃ ≫ f₂) g₄] : HasPushout g₁ (g /-- `X₁ ⨿[Z₁] (X₂ ⨿[Z₂] X₃)` is the pushout `(X₁ ⨿[Z₁] X₂) ×[X₂] (X₂ ⨿[Z₂] X₃)`. -/ def pushoutPushoutRightIsPushout [HasPushout g₁ (g₂ ≫ f₃)] : IsColimit (PushoutCocone.mk l₁ l₂ (show f₂ ≫ l₁ = f₃ ≫ l₂ from pushout.inr_desc _ _ _)) := by - apply rightSquareIsPushout - · exact pushoutIsPushout _ _ - · convert pushoutIsPushout g₁ (g₂ ≫ f₃) - rw [pushout.inl_desc] + apply rightSquareIsPushout _ rfl (pushoutIsPushout _ _) + simpa [PushoutCocone.pasteHoriz] using PushoutCocone.mkSelfIsColimit (pushoutIsPushout _ _) /-- `X₁ ⨿[Z₁] (X₂ ⨿[Z₂] X₃)` is the pushout `(X₁ ⨿[Z₁] X₂) ⨿[Z₂] X₃`. -/ def pushoutAssocSymmIsPushout [HasPushout g₁ (g₂ ≫ f₃)] : @@ -323,14 +301,7 @@ def pushoutAssocSymmIsPushout [HasPushout g₁ (g₂ ≫ f₃)] : (PushoutCocone.mk l₁ (f₄ ≫ l₂) (show (g₃ ≫ f₂) ≫ l₁ = g₄ ≫ f₄ ≫ l₂ by rw [Category.assoc, pushout.inr_desc, pushout.condition_assoc])) := by - apply PushoutCocone.isColimitOfFlip - apply bigSquareIsPushout - · apply PushoutCocone.isColimitOfFlip - apply pushoutPushoutRightIsPushout - · apply PushoutCocone.isColimitOfFlip - exact pushoutIsPushout g₃ g₄ - · exact pushout.condition.symm - · exact (pushout.inr_desc _ _ _).symm + simpa using pasteVertIsPushout rfl (pushoutIsPushout _ _) (pushoutPushoutRightIsPushout _ _ _ _) theorem hasPushout_assoc_symm [HasPushout g₁ (g₂ ≫ f₃)] : HasPushout (g₃ ≫ f₂) g₄ := ⟨⟨⟨_, pushoutAssocSymmIsPushout g₁ g₂ g₃ g₄⟩⟩⟩ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean index 3bc2f8b2e9d02..665866bf2883a 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean @@ -609,7 +609,7 @@ theorem paste_vert {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁ {h₃₁ : X₃₁ ⟶ X₃₂} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₂₁ : X₂₁ ⟶ X₃₁} {v₂₂ : X₂₂ ⟶ X₃₂} (s : IsPullback h₁₁ v₁₁ v₁₂ h₂₁) (t : IsPullback h₂₁ v₂₁ v₂₂ h₃₁) : IsPullback h₁₁ (v₁₁ ≫ v₂₁) (v₁₂ ≫ v₂₂) h₃₁ := - of_isLimit (bigSquareIsPullback _ _ _ _ _ _ _ s.w t.w t.isLimit s.isLimit) + of_isLimit (pasteHorizIsPullback rfl t.isLimit s.isLimit) /-- Paste two pullback squares "horizontally" to obtain another pullback square. @@ -648,7 +648,7 @@ theorem of_bot {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁₁ : {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₂₁ : X₂₁ ⟶ X₃₁} {v₂₂ : X₂₂ ⟶ X₃₂} (s : IsPullback h₁₁ (v₁₁ ≫ v₂₁) (v₁₂ ≫ v₂₂) h₃₁) (p : h₁₁ ≫ v₁₂ = v₁₁ ≫ h₂₁) (t : IsPullback h₂₁ v₂₁ v₂₂ h₃₁) : IsPullback h₁₁ v₁₁ v₁₂ h₂₁ := - of_isLimit (leftSquareIsPullback _ _ _ _ _ _ _ p t.w t.isLimit s.isLimit) + of_isLimit (leftSquareIsPullback (PullbackCone.mk h₁₁ _ p) rfl t.isLimit s.isLimit) /-- Given a pullback square assembled from a commuting square on the left and a pullback square on the right, the left square is a pullback square. @@ -913,7 +913,7 @@ theorem paste_vert {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁ {h₃₁ : X₃₁ ⟶ X₃₂} {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₂₁ : X₂₁ ⟶ X₃₁} {v₂₂ : X₂₂ ⟶ X₃₂} (s : IsPushout h₁₁ v₁₁ v₁₂ h₂₁) (t : IsPushout h₂₁ v₂₁ v₂₂ h₃₁) : IsPushout h₁₁ (v₁₁ ≫ v₂₁) (v₁₂ ≫ v₂₂) h₃₁ := - of_isColimit (bigSquareIsPushout _ _ _ _ _ _ _ s.w t.w t.isColimit s.isColimit) + of_isColimit (pasteHorizIsPushout rfl s.isColimit t.isColimit) /-- Paste two pushout squares "horizontally" to obtain another pushout square. @@ -952,7 +952,8 @@ theorem of_top {X₁₁ X₁₂ X₂₁ X₂₂ X₃₁ X₃₂ : C} {h₁₁ : {v₁₁ : X₁₁ ⟶ X₂₁} {v₁₂ : X₁₂ ⟶ X₂₂} {v₂₁ : X₂₁ ⟶ X₃₁} {v₂₂ : X₂₂ ⟶ X₃₂} (s : IsPushout h₁₁ (v₁₁ ≫ v₂₁) (v₁₂ ≫ v₂₂) h₃₁) (p : h₂₁ ≫ v₂₂ = v₂₁ ≫ h₃₁) (t : IsPushout h₁₁ v₁₁ v₁₂ h₂₁) : IsPushout h₂₁ v₂₁ v₂₂ h₃₁ := - of_isColimit (rightSquareIsPushout _ _ _ _ _ _ _ t.w p t.isColimit s.isColimit) + of_isColimit <| rightSquareIsPushout + (PushoutCocone.mk _ _ p) (cocone_inr _) t.isColimit s.isColimit /-- Given a pushout square assembled from a pushout square on the left and a commuting square on the right, the right square is a pushout square. diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean index 6e410d8ba12e2..17d8d712cac9a 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Mono.lean @@ -212,9 +212,8 @@ instance fst_iso_of_mono_eq [Mono f] : IsIso (pullback.fst f f) := by · simp · simp [fst_eq_snd_of_mono_eq] -instance snd_iso_of_mono_eq [Mono f] : IsIso (pullback.snd f f) := by - rw [← fst_eq_snd_of_mono_eq] - infer_instance +instance snd_iso_of_mono_eq [Mono f] : IsIso (pullback.snd f f) := + fst_eq_snd_of_mono_eq f ▸ fst_iso_of_mono_eq f end @@ -377,9 +376,8 @@ instance inl_iso_of_epi_eq [Epi f] : IsIso (pushout.inl _ _ : _ ⟶ pushout f f) · simp · simp [inl_eq_inr_of_epi_eq] -instance inr_iso_of_epi_eq [Epi f] : IsIso (pushout.inr f f) := by - rw [← inl_eq_inr_of_epi_eq] - infer_instance +instance inr_iso_of_epi_eq [Epi f] : IsIso (pushout.inr f f) := + inl_eq_inr_of_epi_eq f ▸ inl_iso_of_epi_eq f end diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Pasting.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Pasting.lean index f3e8489e37826..adb760c265945 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Pasting.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/Pasting.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2018 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Andrew Yang +Authors: Andrew Yang, Calle Sönne -/ import Mathlib.CategoryTheory.Limits.Shapes.Pullback.HasPullback @@ -21,10 +21,12 @@ if the right square is a pullback, then the left square is a pullback iff the bi pullback. ## Main results -* `bigSquareIsPullback` shows that the big square is a pullback if both the small squares are. +* `pasteHorizIsPullback` shows that the big square is a pullback if both the small squares are. * `leftSquareIsPullback` shows that the left square is a pullback if the other two are. * `pullbackRightPullbackFstIso` shows, using the `pullback` API, that -`W ×[X] (X ×[Z] Y) ≅ W ×[Z] Y`. + `W ×[X] (X ×[Z] Y) ≅ W ×[Z] Y`. +* `pullbackLeftPullbackSndIso` shows, using the `pullback` API, that + `(X ×[Z] Y) ×[Y] W ≅ X ×[Z] W`. -/ @@ -39,13 +41,9 @@ namespace CategoryTheory.Limits variable {C : Type u} [Category.{v} C] section PasteLemma +section PastePullbackHoriz -variable {X₁ X₂ X₃ Y₁ Y₂ Y₃ : C} (f₁ : X₁ ⟶ X₂) (f₂ : X₂ ⟶ X₃) (g₁ : Y₁ ⟶ Y₂) (g₂ : Y₂ ⟶ Y₃) -variable (i₁ : X₁ ⟶ Y₁) (i₂ : X₂ ⟶ Y₂) (i₃ : X₃ ⟶ Y₃) -variable (h₁ : i₁ ≫ g₁ = f₁ ≫ i₂) (h₂ : i₂ ≫ g₂ = f₂ ≫ i₃) - -/-- Given - +/- Let's consider the following diagram ``` X₁ - f₁ -> X₂ - f₂ -> X₃ | | | @@ -53,36 +51,26 @@ i₁ i₂ i₃ ↓ ↓ ↓ Y₁ - g₁ -> Y₂ - g₂ -> Y₃ ``` - -Then the big square is a pullback if both the small squares are. +where `t₁` denotes the cone corresponding to the left square, and `t₂` denotes the cone +corresponding to the right square. -/ -def bigSquareIsPullback (H : IsLimit (PullbackCone.mk _ _ h₂)) - (H' : IsLimit (PullbackCone.mk _ _ h₁)) : - IsLimit - (PullbackCone.mk _ _ - (show i₁ ≫ g₁ ≫ g₂ = (f₁ ≫ f₂) ≫ i₃ by - rw [← Category.assoc, h₁, Category.assoc, h₂, Category.assoc])) := by - fapply PullbackCone.isLimitAux' - intro s - have : (s.fst ≫ g₁) ≫ g₂ = s.snd ≫ i₃ := by rw [← s.condition, Category.assoc] - rcases PullbackCone.IsLimit.lift' H (s.fst ≫ g₁) s.snd this with ⟨l₁, hl₁, hl₁'⟩ - rcases PullbackCone.IsLimit.lift' H' s.fst l₁ hl₁.symm with ⟨l₂, hl₂, hl₂'⟩ - use l₂ - use hl₂ - use - show l₂ ≫ f₁ ≫ f₂ = s.snd by - rw [← hl₁', ← hl₂', Category.assoc] - rfl - intro m hm₁ hm₂ - apply PullbackCone.IsLimit.hom_ext H' - · erw [hm₁, hl₂] - · apply PullbackCone.IsLimit.hom_ext H - · erw [Category.assoc, ← h₁, ← Category.assoc, hm₁, ← hl₂, Category.assoc, Category.assoc, h₁] - rfl - · erw [Category.assoc, hm₂, ← hl₁', ← hl₂'] -/-- Given +variable {X₃ Y₁ Y₂ Y₃ : C} {g₁ : Y₁ ⟶ Y₂} {g₂ : Y₂ ⟶ Y₃} {i₃ : X₃ ⟶ Y₃} (t₂ : PullbackCone g₂ i₃) +variable {i₂ : t₂.pt ⟶ Y₂} (t₁ : PullbackCone g₁ i₂) (hi₂ : i₂ = t₂.fst) + +local notation "f₂" => t₂.snd +local notation "X₁" => t₁.pt +local notation "i₁" => t₁.fst +local notation "f₁" => t₁.snd +/-- The `PullbackCone` obtained by pasting two `PullbackCone`'s horizontally -/ +abbrev PullbackCone.pasteHoriz : PullbackCone (g₁ ≫ g₂) i₃ := + PullbackCone.mk i₁ (f₁ ≫ f₂) + (by rw [reassoc_of% t₁.condition, Category.assoc, ← t₂.condition, ← hi₂]) + +variable {t₁} {t₂} + +/-- Given ``` X₁ - f₁ -> X₂ - f₂ -> X₃ | | | @@ -90,37 +78,27 @@ i₁ i₂ i₃ ↓ ↓ ↓ Y₁ - g₁ -> Y₂ - g₂ -> Y₃ ``` - -Then the big square is a pushout if both the small squares are. +Then the big square is a pullback if both the small squares are. -/ -def bigSquareIsPushout (H : IsColimit (PushoutCocone.mk _ _ h₂)) - (H' : IsColimit (PushoutCocone.mk _ _ h₁)) : - IsColimit - (PushoutCocone.mk _ _ - (show i₁ ≫ g₁ ≫ g₂ = (f₁ ≫ f₂) ≫ i₃ by - rw [← Category.assoc, h₁, Category.assoc, h₂, Category.assoc])) := by - fapply PushoutCocone.isColimitAux' +def pasteHorizIsPullback (H : IsLimit t₂) (H' : IsLimit t₁) : IsLimit (t₂.pasteHoriz t₁ hi₂) := by + apply PullbackCone.isLimitAux' intro s - have : i₁ ≫ s.inl = f₁ ≫ f₂ ≫ s.inr := by rw [s.condition, Category.assoc] - rcases PushoutCocone.IsColimit.desc' H' s.inl (f₂ ≫ s.inr) this with ⟨l₁, hl₁, hl₁'⟩ - rcases PushoutCocone.IsColimit.desc' H l₁ s.inr hl₁' with ⟨l₂, hl₂, hl₂'⟩ - use l₂ - use - show (g₁ ≫ g₂) ≫ l₂ = s.inl by - rw [← hl₁, ← hl₂, Category.assoc] - rfl - use hl₂' + -- Obtain the lift from lifting from both the small squares consecutively. + obtain ⟨l₂, hl₂, hl₂'⟩ := PullbackCone.IsLimit.lift' H (s.fst ≫ g₁) s.snd + (by rw [← s.condition, Category.assoc]) + obtain ⟨l₁, hl₁, hl₁'⟩ := PullbackCone.IsLimit.lift' H' s.fst l₂ (by rw [← hl₂, hi₂]) + refine ⟨l₁, hl₁, by simp [reassoc_of% hl₁', hl₂'], ?_⟩ + -- Uniqueness also follows from the universal property of both the small squares. intro m hm₁ hm₂ - apply PushoutCocone.IsColimit.hom_ext H - · apply PushoutCocone.IsColimit.hom_ext H' - · erw [← Category.assoc, hm₁, hl₂, hl₁] - · erw [← Category.assoc, h₂, Category.assoc, hm₂, ← hl₂', ← Category.assoc, ← Category.assoc, ← - h₂] - rfl - · erw [hm₂, hl₂'] + apply PullbackCone.IsLimit.hom_ext H' (by simpa [hl₁] using hm₁) + apply PullbackCone.IsLimit.hom_ext H + · dsimp at hm₁ + rw [Category.assoc, ← hi₂, ← t₁.condition, reassoc_of% hm₁, hl₁', hi₂, hl₂] + · simpa [hl₁', hl₂'] using hm₂ -/-- Given +variable (t₁) +/-- Given ``` X₁ - f₁ -> X₂ - f₂ -> X₃ | | | @@ -128,104 +106,359 @@ i₁ i₂ i₃ ↓ ↓ ↓ Y₁ - g₁ -> Y₂ - g₂ -> Y₃ ``` - Then the left square is a pullback if the right square and the big square are. -/ -def leftSquareIsPullback (H : IsLimit (PullbackCone.mk _ _ h₂)) - (H' : - IsLimit - (PullbackCone.mk _ _ - (show i₁ ≫ g₁ ≫ g₂ = (f₁ ≫ f₂) ≫ i₃ by - rw [← Category.assoc, h₁, Category.assoc, h₂, Category.assoc]))) : - IsLimit (PullbackCone.mk _ _ h₁) := by - fapply PullbackCone.isLimitAux' +def leftSquareIsPullback (H : IsLimit t₂) (H' : IsLimit (t₂.pasteHoriz t₁ hi₂)) : IsLimit t₁ := by + apply PullbackCone.isLimitAux' intro s - have : s.fst ≫ g₁ ≫ g₂ = (s.snd ≫ f₂) ≫ i₃ := by - rw [← Category.assoc, s.condition, Category.assoc, Category.assoc, h₂] - rcases PullbackCone.IsLimit.lift' H' s.fst (s.snd ≫ f₂) this with ⟨l₁, hl₁, hl₁'⟩ - use l₁ - use hl₁ - constructor + -- Obtain the induced morphism from the universal property of the big square + obtain ⟨l, hl, hl'⟩ := PullbackCone.IsLimit.lift' H' s.fst (s.snd ≫ f₂) + (by rw [Category.assoc, ← t₂.condition, reassoc_of% s.condition, ← hi₂]) + refine ⟨l, hl, ?_, ?_⟩ + -- To check that `l` is compatible with the projections, we use the universal property of `t₂` · apply PullbackCone.IsLimit.hom_ext H - · erw [Category.assoc, ← h₁, ← Category.assoc, hl₁, s.condition] - rfl - · erw [Category.assoc, hl₁'] - rfl + · simp [← s.condition, ← hl, ← t₁.condition, ← hi₂] + · simpa using hl' + -- Uniqueness of the lift follows from the universal property of the big square · intro m hm₁ hm₂ apply PullbackCone.IsLimit.hom_ext H' - · erw [hm₁, hl₁] - · erw [hl₁', ← hm₂] - exact (Category.assoc _ _ _).symm + · simpa [hm₁] using hl.symm + · simpa [← hm₂] using hl'.symm + +/-- Given that the right square is a pullback, the pasted square is a pullback iff the left +square is. -/ +def pasteHorizIsPullbackEquiv (H : IsLimit t₂) : IsLimit (t₂.pasteHoriz t₁ hi₂) ≃ IsLimit t₁ where + toFun H' := leftSquareIsPullback t₁ _ H H' + invFun H' := pasteHorizIsPullback _ H H' + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ + +end PastePullbackHoriz + +section PastePullbackVert + +/- Let's consider the following diagram +``` +Y₃ - i₃ -> X₃ +| | +g₂ f₂ +∨ ∨ +Y₂ - i₂ -> X₂ +| | +g₁ f₁ +∨ ∨ +Y₁ - i₁ -> X₁ +``` +Let `t₁` denote the cone corresponding to the bottom square, and `t₂` denote the cone corresponding +to the top square. + +-/ +variable {X₁ X₂ X₃ Y₁ : C} {f₁ : X₂ ⟶ X₁} {f₂ : X₃ ⟶ X₂} {i₁ : Y₁ ⟶ X₁} +variable (t₁ : PullbackCone i₁ f₁) {i₂ : t₁.pt ⟶ X₂} (t₂ : PullbackCone i₂ f₂) + (hi₂ : i₂ = t₁.snd) + +local notation "Y₂" => t₁.pt +local notation "g₁" => t₁.fst +local notation "i₂" => t₁.snd +local notation "Y₃" => t₂.pt +local notation "g₂" => t₂.fst +local notation "i₃" => t₂.snd + +/-- The `PullbackCone` obtained by pasting two `PullbackCone`'s vertically -/ +abbrev PullbackCone.pasteVert : PullbackCone i₁ (f₂ ≫ f₁) := + PullbackCone.mk (g₂ ≫ g₁) i₃ + (by rw [← reassoc_of% t₂.condition, Category.assoc, t₁.condition, ← hi₂]) + +/-- Pasting two pullback cones vertically is isomorphic to the pullback cone obtained by flipping +them, pasting horizontally, and then flipping the result again. -/ +def PullbackCone.pasteVertFlip : (t₁.pasteVert t₂ hi₂).flip ≅ (t₁.flip.pasteHoriz t₂.flip hi₂) := + PullbackCone.ext (Iso.refl _) (by simp) (by simp) + +variable {t₁} {t₂} /-- Given +``` +Y₃ - i₃ -> X₃ +| | +g₂ f₂ +∨ ∨ +Y₂ - i₂ -> X₂ +| | +g₁ f₁ +∨ ∨ +Y₁ - i₁ -> X₁ +``` +The big square is a pullback if both the small squares are. +-/ +def pasteVertIsPullback (H₁ : IsLimit t₁) (H₂ : IsLimit t₂) : IsLimit (t₁.pasteVert t₂ hi₂) := by + apply PullbackCone.isLimitOfFlip <| IsLimit.ofIsoLimit _ (t₁.pasteVertFlip t₂ hi₂).symm + exact pasteHorizIsPullback hi₂ (PullbackCone.flipIsLimit H₁) (PullbackCone.flipIsLimit H₂) + +variable (t₂) +/-- Given +``` +Y₃ - i₃ -> X₃ +| | +g₂ f₂ +∨ ∨ +Y₂ - i₂ -> X₂ +| | +g₁ f₁ +∨ ∨ +Y₁ - i₁ -> X₁ +``` +The top square is a pullback if the bottom square and the big square are. +-/ +def topSquareIsPullback (H₁ : IsLimit t₁) (H₂ : IsLimit (t₁.pasteVert t₂ hi₂)) : IsLimit t₂ := + PullbackCone.isLimitOfFlip + (leftSquareIsPullback _ hi₂ (PullbackCone.flipIsLimit H₁) (PullbackCone.flipIsLimit H₂)) + +/-- Given that the bottom square is a pullback, the pasted square is a pullback iff the top +square is. -/ +def pasteVertIsPullbackEquiv (H : IsLimit t₁) : IsLimit (t₁.pasteVert t₂ hi₂) ≃ IsLimit t₂ where + toFun H' := topSquareIsPullback t₂ _ H H' + invFun H' := pasteVertIsPullback _ H H' + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ + +end PastePullbackVert + +section PastePushoutHoriz + +/- Let's consider the following diagram ``` X₁ - f₁ -> X₂ - f₂ -> X₃ | | | i₁ i₂ i₃ -↓ ↓ ↓ +∨ ∨ ∨ Y₁ - g₁ -> Y₂ - g₂ -> Y₃ ``` +where `t₁` denotes the left pushout cocone, and `t₂` denotes the right pushout cocone. +-/ +variable {X₁ X₂ X₃ Y₁ : C} {f₁ : X₁ ⟶ X₂} {f₂ : X₂ ⟶ X₃} {i₁ : X₁ ⟶ Y₁} +variable (t₁ : PushoutCocone i₁ f₁) {i₂ : X₂ ⟶ t₁.pt} (t₂ : PushoutCocone i₂ f₂) +variable (hi₂ : i₂ = t₁.inr) + +local notation "Y₂" => t₁.pt +local notation "g₁" => t₁.inl +local notation "i₂" => t₁.inr +local notation "Y₃" => t₂.pt +local notation "g₂" => t₂.inl +local notation "i₃" => t₂.inr + +/-- The pushout cocone obtained by pasting two pushout cocones horizontally. -/ +abbrev PushoutCocone.pasteHoriz : PushoutCocone i₁ (f₁ ≫ f₂) := + PushoutCocone.mk (g₁ ≫ g₂) i₃ + (by rw [reassoc_of% t₁.condition, Category.assoc, ← t₂.condition, ← hi₂]) + +variable {t₁} {t₂} + +/-- Given +``` +X₁ - f₁ -> X₂ - f₂ -> X₃ +| | | +i₁ i₂ i₃ +∨ ∨ ∨ +Y₁ - g₁ -> Y₂ - g₂ -> Y₃ +``` +Then the big square is a pushout if both the small squares are. +-/ +def pasteHorizIsPushout (H : IsColimit t₁) (H' : IsColimit t₂) : + IsColimit (t₁.pasteHoriz t₂ hi₂) := by + apply PushoutCocone.isColimitAux' + intro s + -- Obtain the induced map from descending from both the small squares consecutively. + obtain ⟨l₁, hl₁, hl₁'⟩ := PushoutCocone.IsColimit.desc' H s.inl (f₂ ≫ s.inr) + (by rw [s.condition, Category.assoc]) + obtain ⟨l₂, hl₂, hl₂'⟩ := PushoutCocone.IsColimit.desc' H' l₁ s.inr (by rw [← hl₁', hi₂]) + refine ⟨l₂, by simp [hl₂, hl₁], hl₂', ?_⟩ + -- Uniqueness also follows from the universal property of both the small squares. + intro m hm₁ hm₂ + apply PushoutCocone.IsColimit.hom_ext H' _ (by simpa [hl₂'] using hm₂) + simp only [PushoutCocone.mk_pt, PushoutCocone.mk_ι_app, Category.assoc] at hm₁ hm₂ + apply PushoutCocone.IsColimit.hom_ext H + · rw [hm₁, ← hl₁, hl₂] + · rw [← hi₂, reassoc_of% t₂.condition, reassoc_of% t₂.condition, hm₂, hl₂'] + +variable (t₂) + +/-- Given + +X₁ - f₁ -> X₂ - f₂ -> X₃ +| | | +i₁ i₂ i₃ +∨ ∨ ∨ +Y₁ - g₁ -> Y₂ - g₂ -> Y₃ Then the right square is a pushout if the left square and the big square are. -/ -def rightSquareIsPushout (H : IsColimit (PushoutCocone.mk _ _ h₁)) - (H' : - IsColimit - (PushoutCocone.mk _ _ - (show i₁ ≫ g₁ ≫ g₂ = (f₁ ≫ f₂) ≫ i₃ by - rw [← Category.assoc, h₁, Category.assoc, h₂, Category.assoc]))) : - IsColimit (PushoutCocone.mk _ _ h₂) := by - fapply PushoutCocone.isColimitAux' +def rightSquareIsPushout (H : IsColimit t₁) (H' : IsColimit (t₁.pasteHoriz t₂ hi₂)) : + IsColimit t₂ := by + apply PushoutCocone.isColimitAux' intro s - have : i₁ ≫ g₁ ≫ s.inl = (f₁ ≫ f₂) ≫ s.inr := by - rw [Category.assoc, ← s.condition, ← Category.assoc, ← Category.assoc, h₁] - rcases PushoutCocone.IsColimit.desc' H' (g₁ ≫ s.inl) s.inr this with ⟨l₁, hl₁, hl₁'⟩ - dsimp at * - use l₁ - refine ⟨?_, ?_, ?_⟩ - · apply PushoutCocone.IsColimit.hom_ext H - · erw [← Category.assoc, hl₁] - rfl - · erw [← Category.assoc, h₂, Category.assoc, hl₁', s.condition] - · exact hl₁' + -- Obtain the induced morphism from the universal property of the big square + obtain ⟨l, hl, hl'⟩ := PushoutCocone.IsColimit.desc' H' (g₁ ≫ s.inl) s.inr + (by rw [reassoc_of% t₁.condition, ← hi₂, s.condition, Category.assoc]) + refine ⟨l, ?_, hl', ?_⟩ + -- To check that `l` is compatible with the projections, we use the universal property of `t₁` + · simp at hl hl' + apply PushoutCocone.IsColimit.hom_ext H hl + rw [← Category.assoc, ← hi₂, t₂.condition, s.condition, Category.assoc, hl'] + -- Uniqueness of the lift follows from the universal property of the big square · intro m hm₁ hm₂ apply PushoutCocone.IsColimit.hom_ext H' - · erw [hl₁, Category.assoc, hm₁] - · erw [hm₂, hl₁'] + · simpa [← hm₁] using hl.symm + · simpa [← hm₂] using hl'.symm + +/-- Given that the left square is a pushout, the pasted square is a pushout iff the right square is. +-/ +def pasteHorizIsPushoutEquiv (H : IsColimit t₁) : + IsColimit (t₁.pasteHoriz t₂ hi₂) ≃ IsColimit t₂ where + toFun H' := rightSquareIsPushout t₂ _ H H' + invFun H' := pasteHorizIsPushout _ H H' + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ + +end PastePushoutHoriz + +section PastePushoutVert + +/- Let's consider the following diagram +``` +Y₃ - i₃ -> X₃ +| | +g₂ f₂ +∨ ∨ +Y₂ - i₂ -> X₂ +| | +g₁ f₁ +∨ ∨ +Y₁ - i₁ -> X₁ +``` +Let `t₁` denote the cone corresponding to the bottom square, and `t₂` denote the cone corresponding +to the top square. +-/ +variable {Y₃ Y₂ Y₁ X₃ : C} {g₂ : Y₃ ⟶ Y₂} {g₁ : Y₂ ⟶ Y₁} {i₃ : Y₃ ⟶ X₃} +variable (t₁ : PushoutCocone g₂ i₃) {i₂ : Y₂ ⟶ t₁.pt} (t₂ : PushoutCocone g₁ i₂) + (hi₂ : i₂ = t₁.inl) + +local notation "X₂" => t₁.pt +local notation "f₂" => t₁.inr +local notation "i₂" => t₁.inl +local notation "X₁" => t₂.pt +local notation "f₁" => t₂.inr +local notation "i₁" => t₂.inl + +/-- The `PullbackCone` obtained by pasting two `PullbackCone`'s vertically -/ +abbrev PushoutCocone.pasteVert : PushoutCocone (g₂ ≫ g₁) i₃ := + PushoutCocone.mk i₁ (f₂ ≫ f₁) + (by rw [← reassoc_of% t₁.condition, Category.assoc, t₂.condition, ← hi₂]) + +/-- Pasting two pushout cocones vertically is isomorphic to the pushout cocone obtained by flipping +them, pasting horizontally, and then flipping the result again. -/ +def PushoutCocone.pasteVertFlip : (t₁.pasteVert t₂ hi₂).flip ≅ (t₁.flip.pasteHoriz t₂.flip hi₂) := + PushoutCocone.ext (Iso.refl _) (by simp) (by simp) + +variable {t₁} {t₂} + +/-- Given +``` +Y₃ - i₃ -> X₃ +| | +g₂ f₂ +∨ ∨ +Y₂ - i₂ -> X₂ +| | +g₁ f₁ +∨ ∨ +Y₁ - i₁ -> X₁ +``` +The big square is a pushout if both the small squares are. +-/ +def pasteVertIsPushout (H₁ : IsColimit t₁) (H₂ : IsColimit t₂) : + IsColimit (t₁.pasteVert t₂ hi₂) := by + apply PushoutCocone.isColimitOfFlip <| IsColimit.ofIsoColimit _ (t₁.pasteVertFlip t₂ hi₂).symm + exact pasteHorizIsPushout hi₂ (PushoutCocone.flipIsColimit H₁) (PushoutCocone.flipIsColimit H₂) + +variable (t₂) + +/-- Given +``` +Y₃ - i₃ -> X₃ +| | +g₂ f₂ +∨ ∨ +Y₂ - i₂ -> X₂ +| | +g₁ f₁ +∨ ∨ +Y₁ - i₁ -> X₁ +``` +The bottom square is a pushout if the top square and the big square are. +-/ +def botSquareIsPushout (H₁ : IsColimit t₁) (H₂ : IsColimit (t₁.pasteVert t₂ hi₂)) : IsColimit t₂ := + PushoutCocone.isColimitOfFlip + (rightSquareIsPushout _ hi₂ (PushoutCocone.flipIsColimit H₁) (PushoutCocone.flipIsColimit H₂)) + +/-- Given that the top square is a pushout, the pasted square is a pushout iff the bottom square is. +-/ +def pasteVertIsPushoutEquiv (H : IsColimit t₁) : + IsColimit (t₁.pasteVert t₂ hi₂) ≃ IsColimit t₂ where + toFun H' := botSquareIsPushout t₂ _ H H' + invFun H' := pasteVertIsPushout _ H H' + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ + +end PastePushoutVert end PasteLemma section +/- Let's consider the following diagram of pullbacks +``` +W ×[X] (X ×[Z] Y) --snd--> X ×[Z] Y --snd--> Y + | | | + fst fst g + v v v + W --------- f' ---------> X ---- f ---> Y +``` +In this section we show that `W ×[X] (X ×[Z] Y) ≅ W ×[Z] Y`. +-/ variable {W X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) (f' : W ⟶ X) variable [HasPullback f g] [HasPullback f' (pullback.fst f g)] -variable [HasPullback (f' ≫ f) g] + +instance hasPullbackHorizPaste : HasPullback (f' ≫ f) g := + HasLimit.mk { + cone := (pullback.cone f g).pasteHoriz (pullback.cone f' (pullback.fst f g)) rfl + isLimit := pasteHorizIsPullback rfl (pullback.isLimit f g) + (pullback.isLimit f' (pullback.fst f g)) + } /-- The canonical isomorphism `W ×[X] (X ×[Z] Y) ≅ W ×[Z] Y` -/ noncomputable def pullbackRightPullbackFstIso : - pullback f' (pullback.fst f g) ≅ pullback (f' ≫ f) g := by - let this := - bigSquareIsPullback (pullback.snd _ _ : pullback f' (pullback.fst f g) ⟶ _) - (pullback.snd _ _) f' f (pullback.fst _ _) (pullback.fst _ _) g - pullback.condition pullback.condition - (pullbackIsPullback _ _) (pullbackIsPullback _ _) - exact (this.conePointUniqueUpToIso (pullbackIsPullback _ _) : _) + pullback f' (pullback.fst f g) ≅ pullback (f' ≫ f) g := + IsLimit.conePointUniqueUpToIso + (pasteHorizIsPullback rfl (pullback.isLimit f g) (pullback.isLimit f' (pullback.fst f g))) + (pullback.isLimit (f' ≫ f) g) @[reassoc (attr := simp)] theorem pullbackRightPullbackFstIso_hom_fst : - (pullbackRightPullbackFstIso f g f').hom ≫ pullback.fst _ _ = pullback.fst _ _ := + (pullbackRightPullbackFstIso f g f').hom ≫ pullback.fst (f' ≫ f) g = + pullback.fst f' (pullback.fst f g) := IsLimit.conePointUniqueUpToIso_hom_comp _ _ WalkingCospan.left @[reassoc (attr := simp)] theorem pullbackRightPullbackFstIso_hom_snd : (pullbackRightPullbackFstIso f g f').hom ≫ pullback.snd _ _ = - pullback.snd _ _ ≫ pullback.snd _ _ := + pullback.snd f' (pullback.fst f g) ≫ pullback.snd f g := IsLimit.conePointUniqueUpToIso_hom_comp _ _ WalkingCospan.right @[reassoc (attr := simp)] theorem pullbackRightPullbackFstIso_inv_fst : - (pullbackRightPullbackFstIso f g f').inv ≫ pullback.fst _ _ = pullback.fst _ _ := + (pullbackRightPullbackFstIso f g f').inv ≫ pullback.fst f' (pullback.fst f g) = + pullback.fst (f' ≫ f) g := IsLimit.conePointUniqueUpToIso_inv_comp _ _ WalkingCospan.left @[reassoc (attr := simp)] @@ -239,58 +472,201 @@ theorem pullbackRightPullbackFstIso_inv_snd_fst : (pullbackRightPullbackFstIso f g f').inv ≫ pullback.snd _ _ ≫ pullback.fst _ _ = pullback.fst _ _ ≫ f' := by rw [← pullback.condition] - exact pullbackRightPullbackFstIso_inv_fst_assoc _ _ _ _ + exact pullbackRightPullbackFstIso_inv_fst_assoc f g f' _ + +end +section +/- Let's consider the following diagram of pullbacks +``` +(X ×[Z] Y) ×[Y] W --snd--> W + | | + fst g' + v v + (X ×[Z] Y) --- snd ---> Y + | | + fst g + v v + X -------- f --------> Z + +``` +In this section we show that `(X ×[Z] Y) ×[Y] W ≅ X ×[Z] W`. +-/ + + +variable {W X Y Z : C} (f : X ⟶ Z) (g : Y ⟶ Z) (g' : W ⟶ Y) +variable [HasPullback f g] [HasPullback (pullback.snd f g) g'] + +instance hasPullbackVertPaste : HasPullback f (g' ≫ g) := + HasLimit.mk { + cone := (pullback.cone f g).pasteVert (pullback.cone (pullback.snd f g) g') rfl + isLimit := pasteVertIsPullback rfl (pullback.isLimit f g) + (pullback.isLimit (pullback.snd f g) g') + } + +/-- The canonical isomorphism `(X ×[Z] Y) ×[Y] W ≅ X ×[Z] W` -/ +def pullbackLeftPullbackSndIso : + pullback (pullback.snd f g) g' ≅ pullback f (g' ≫ g) := + IsLimit.conePointUniqueUpToIso + (pasteVertIsPullback rfl (pullback.isLimit f g) (pullback.isLimit (pullback.snd f g) g')) + (pullback.isLimit f (g' ≫ g)) + +@[reassoc (attr := simp)] +theorem pullbackLeftPullbackSndIso_hom_fst : + (pullbackLeftPullbackSndIso f g g').hom ≫ pullback.fst _ _ = + pullback.fst _ _ ≫ pullback.fst _ _ := + IsLimit.conePointUniqueUpToIso_hom_comp _ _ WalkingCospan.left + +@[reassoc (attr := simp)] +theorem pullbackLeftPullbackSndIso_hom_snd : + (pullbackLeftPullbackSndIso f g g').hom ≫ pullback.snd _ _ = pullback.snd _ _ := + IsLimit.conePointUniqueUpToIso_hom_comp _ _ WalkingCospan.right + +@[reassoc (attr := simp)] +theorem pullbackLeftPullbackSndIso_inv_fst : + (pullbackLeftPullbackSndIso f g g').inv ≫ pullback.fst _ _ ≫ pullback.fst _ _ = + pullback.fst _ _ := + IsLimit.conePointUniqueUpToIso_inv_comp _ _ WalkingCospan.left + +@[reassoc (attr := simp)] +theorem pullbackLeftPullbackSndIso_inv_snd_snd : + (pullbackLeftPullbackSndIso f g g').inv ≫ pullback.snd _ _ = pullback.snd _ _ := + IsLimit.conePointUniqueUpToIso_inv_comp _ _ WalkingCospan.right + +@[reassoc (attr := simp)] +theorem pullbackLeftPullbackSndIso_inv_fst_snd : + (pullbackLeftPullbackSndIso f g g').inv ≫ pullback.fst _ _ ≫ pullback.snd _ _ = + pullback.snd _ _ ≫ g' := by + rw [pullback.condition] + exact pullbackLeftPullbackSndIso_inv_snd_snd_assoc f g g' g' end section +/- Let's consider the following diagram of pushouts +``` +X ---- g ----> Z ----- g' -----> W +| | | +f inr inr +v v v +Y - inl -> Y ⨿[X] Z --inl--> (Y ⨿[X] Z) ⨿[Z] W +``` +In this section we show that `(Y ⨿[X] Z) ⨿[Z] W ≅ Y ⨿[X] W`. +-/ variable {W X Y Z : C} (f : X ⟶ Y) (g : X ⟶ Z) (g' : Z ⟶ W) -variable [HasPushout f g] [HasPushout (pushout.inr _ _ : _ ⟶ pushout f g) g'] -variable [HasPushout f (g ≫ g')] +variable [HasPushout f g] [HasPushout (pushout.inr f g) g'] -/-- The canonical isomorphism `(Y ⨿[X] Z) ⨿[Z] W ≅ Y ×[X] W` -/ +instance : HasPushout f (g ≫ g') := + HasColimit.mk { + cocone := (pushout.cocone f g).pasteHoriz (pushout.cocone (pushout.inr f g) g') rfl + isColimit := pasteHorizIsPushout rfl (pushout.isColimit f g) + (pushout.isColimit (pushout.inr f g) g') + } + +/-- The canonical isomorphism `(Y ⨿[X] Z) ⨿[Z] W ≅ Y ⨿[X] W` -/ noncomputable def pushoutLeftPushoutInrIso : - pushout (pushout.inr _ _ : _ ⟶ pushout f g) g' ≅ pushout f (g ≫ g') := - ((bigSquareIsPushout g g' _ _ f _ _ pushout.condition pushout.condition (pushoutIsPushout _ _) - (pushoutIsPushout _ _)).coconePointUniqueUpToIso - (pushoutIsPushout _ _) : - _) + pushout (pushout.inr f g) g' ≅ pushout f (g ≫ g') := + IsColimit.coconePointUniqueUpToIso + (pasteHorizIsPushout rfl (pushout.isColimit f g) (pushout.isColimit (pushout.inr f g) g')) + (pushout.isColimit f (g ≫ g')) @[reassoc (attr := simp)] theorem inl_pushoutLeftPushoutInrIso_inv : - pushout.inl _ _ ≫ (pushoutLeftPushoutInrIso f g g').inv = pushout.inl _ _ ≫ pushout.inl _ _ := - ((bigSquareIsPushout g g' _ _ f _ _ pushout.condition pushout.condition (pushoutIsPushout _ _) - (pushoutIsPushout _ _)).comp_coconePointUniqueUpToIso_inv - (pushoutIsPushout _ _) WalkingSpan.left : - _) + (pushout.inl f (g ≫ g')) ≫ (pushoutLeftPushoutInrIso f g g').inv = + (pushout.inl f g) ≫ (pushout.inl (pushout.inr f g) g') := + IsColimit.comp_coconePointUniqueUpToIso_inv _ _ WalkingSpan.left @[reassoc (attr := simp)] theorem inr_pushoutLeftPushoutInrIso_hom : - pushout.inr _ _ ≫ (pushoutLeftPushoutInrIso f g g').hom = pushout.inr _ _ := - ((bigSquareIsPushout g g' _ _ f _ _ pushout.condition pushout.condition (pushoutIsPushout _ _) - (pushoutIsPushout _ _)).comp_coconePointUniqueUpToIso_hom - (pushoutIsPushout _ _) WalkingSpan.right : - _) + (pushout.inr (pushout.inr f g) g') ≫ (pushoutLeftPushoutInrIso f g g').hom = + (pushout.inr f (g ≫ g')) := + IsColimit.comp_coconePointUniqueUpToIso_hom (pasteHorizIsPushout _ _ _) _ WalkingSpan.right @[reassoc (attr := simp)] theorem inr_pushoutLeftPushoutInrIso_inv : - pushout.inr _ _ ≫ (pushoutLeftPushoutInrIso f g g').inv = pushout.inr _ _ := by - rw [Iso.comp_inv_eq, inr_pushoutLeftPushoutInrIso_hom] + (pushout.inr f (g ≫ g')) ≫ (pushoutLeftPushoutInrIso f g g').inv = + (pushout.inr (pushout.inr f g) g') := + IsColimit.comp_coconePointUniqueUpToIso_inv _ _ WalkingSpan.right @[reassoc (attr := simp)] theorem inl_inl_pushoutLeftPushoutInrIso_hom : - pushout.inl _ _ ≫ pushout.inl _ _ ≫ (pushoutLeftPushoutInrIso f g g').hom = - pushout.inl _ _ := by - rw [← Category.assoc, ← Iso.eq_comp_inv, inl_pushoutLeftPushoutInrIso_inv] + (pushout.inl f g) ≫ (pushout.inl (pushout.inr f g) g') ≫ + (pushoutLeftPushoutInrIso f g g').hom = (pushout.inl f (g ≫ g')) := by + rw [← Category.assoc] + apply IsColimit.comp_coconePointUniqueUpToIso_hom (pasteHorizIsPushout _ _ _) _ WalkingSpan.left @[reassoc (attr := simp)] theorem inr_inl_pushoutLeftPushoutInrIso_hom : - pushout.inr _ _ ≫ pushout.inl _ _ ≫ (pushoutLeftPushoutInrIso f g g').hom = - g' ≫ pushout.inr _ _ := by + pushout.inr f g ≫ pushout.inl (pushout.inr f g) g' ≫ (pushoutLeftPushoutInrIso f g g').hom = + g' ≫ pushout.inr f (g ≫ g') := by rw [← Category.assoc, ← Iso.eq_comp_inv, Category.assoc, inr_pushoutLeftPushoutInrIso_inv, pushout.condition] +end +section + +/- Let's consider the diagram of pushouts +``` +X ---- g ----> Z +| | +f inr +v v +Y - inl -> Y ⨿[X] Z +| | +f' inr +v v +W - inl -> W ⨿[Y] (Y ⨿[X] Z) +``` + +In this section we will construct the isomorphism `W ⨿[Y] (Y ⨿[X] Z) ≅ W ⨿[X] Z`. +-/ + +variable {W X Y Z : C} (f : X ⟶ Y) (g : X ⟶ Z) (f' : Y ⟶ W) +variable [HasPushout f g] [HasPushout f' (pushout.inl f g)] + +instance hasPushoutVertPaste : HasPushout (f ≫ f') g := + HasColimit.mk { + cocone := (pushout.cocone f g).pasteVert (pushout.cocone f' (pushout.inl f g)) rfl + isColimit := pasteVertIsPushout rfl (pushout.isColimit f g) + (pushout.isColimit f' (pushout.inl f g)) + } + +/-- The canonical isomorphism `W ⨿[Y] (Y ⨿[X] Z) ≅ W ⨿[X] Z` -/ +noncomputable def pushoutRightPushoutInlIso : + pushout f' (pushout.inl f g) ≅ pushout (f ≫ f') g := + IsColimit.coconePointUniqueUpToIso + (pasteVertIsPushout rfl (pushout.isColimit f g) (pushout.isColimit f' (pushout.inl f g))) + (pushout.isColimit (f ≫ f') g) + +@[reassoc (attr := simp)] +theorem inl_pushoutRightPushoutInlIso_inv : + pushout.inl _ _ ≫ (pushoutRightPushoutInlIso f g f').inv = pushout.inl _ _ := + IsColimit.comp_coconePointUniqueUpToIso_inv _ _ WalkingSpan.left + +@[reassoc (attr := simp)] +theorem inr_inr_pushoutRightPushoutInlIso_hom : + pushout.inr _ _ ≫ pushout.inr _ _ ≫ (pushoutRightPushoutInlIso f g f').hom = + pushout.inr _ _ := by + rw [← Category.assoc] + apply IsColimit.comp_coconePointUniqueUpToIso_hom (pasteVertIsPushout rfl _ _) _ WalkingSpan.right + +@[reassoc (attr := simp)] +theorem inr_pushoutRightPushoutInlIso_inv : + pushout.inr _ _ ≫ (pushoutRightPushoutInlIso f g f').inv = + pushout.inr _ _ ≫ pushout.inr _ _ := + IsColimit.comp_coconePointUniqueUpToIso_inv _ _ WalkingSpan.right + +@[reassoc (attr := simp)] +theorem inl_pushoutRightPushoutInlIso_hom : + pushout.inl _ _ ≫ (pushoutRightPushoutInlIso f g f').hom = pushout.inl _ _ := + IsColimit.comp_coconePointUniqueUpToIso_hom (pasteVertIsPushout rfl _ _) _ WalkingSpan.left + +@[reassoc (attr := simp)] +theorem inr_inl_pushoutRightPushoutInlIso_hom : + pushout.inl _ _ ≫ pushout.inr _ _ ≫ (pushoutRightPushoutInlIso f g f').hom = + f' ≫ pushout.inl _ _ := by + rw [← Category.assoc, ← pushout.condition, Category.assoc, inl_pushoutRightPushoutInlIso_hom] + end end CategoryTheory.Limits From 3bd273a62d48234b900297712e314aacf179a525 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Sun, 4 Aug 2024 10:58:27 +0000 Subject: [PATCH 008/359] feat(GroupTheory/Index): finite quotient lemmas (#14921) Add some lemmas about finiteness of `MulAction.orbitRel.Quotient` for actions by a subgroup of finite index. From AperiodicMonotilesLean. --- Mathlib/GroupTheory/Index.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 4dfb748379a3c..e7f2ad5c5c1d1 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -379,6 +379,20 @@ noncomputable def fintypeOfIndexNeZero (hH : H.index ≠ 0) : Fintype (G ⧸ H) theorem one_lt_index_of_ne_top [Finite (G ⧸ H)] (hH : H ≠ ⊤) : 1 < H.index := Nat.one_lt_iff_ne_zero_and_ne_one.mpr ⟨index_ne_zero_of_finite, mt index_eq_one.mp hH⟩ +@[to_additive] +lemma finite_quotient_of_finite_quotient_of_index_ne_zero {X : Type*} [MulAction G X] + [Finite <| MulAction.orbitRel.Quotient G X] (hi : H.index ≠ 0) : + Finite <| MulAction.orbitRel.Quotient H X := by + have := fintypeOfIndexNeZero hi + exact MulAction.finite_quotient_of_finite_quotient_of_finite_quotient + +@[to_additive] +lemma finite_quotient_of_pretransitive_of_index_ne_zero {X : Type*} [MulAction G X] + [MulAction.IsPretransitive G X] (hi : H.index ≠ 0) : + Finite <| MulAction.orbitRel.Quotient H X := by + have := (MulAction.pretransitive_iff_subsingleton_quotient G X).1 inferInstance + exact finite_quotient_of_finite_quotient_of_index_ne_zero hi + section FiniteIndex variable (H K) From 7445812e2c064d18ac074eef0789a09f09d9490d Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sun, 4 Aug 2024 10:58:29 +0000 Subject: [PATCH 009/359] feat: if one set is clopen in another, then it is the union of connected components in the larger set (#14927) --- Mathlib/Topology/Connected/Clopen.lean | 26 ++++++++++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/Mathlib/Topology/Connected/Clopen.lean b/Mathlib/Topology/Connected/Clopen.lean index a9fa5d0abcd16..ba60382112ad4 100644 --- a/Mathlib/Topology/Connected/Clopen.lean +++ b/Mathlib/Topology/Connected/Clopen.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov -/ +import Mathlib.Data.Set.Subset import Mathlib.Topology.Clopen import Mathlib.Topology.Connected.Basic @@ -121,6 +122,16 @@ theorem IsClopen.eq_univ [PreconnectedSpace α] {s : Set α} (h' : IsClopen s) ( s = univ := (isClopen_iff.mp h').resolve_left h.ne_empty +open Set.Notation in +lemma isClopen_preimage_val {X : Type*} [TopologicalSpace X] {u v : Set X} + (hu : IsOpen u) (huv : Disjoint (frontier u) v) : IsClopen (v ↓∩ u) := by + refine ⟨?_, isOpen_induced hu (f := Subtype.val)⟩ + refine isClosed_induced_iff.mpr ⟨closure u, isClosed_closure, ?_⟩ + apply image_val_injective + simp only [Subtype.image_preimage_coe] + rw [closure_eq_self_union_frontier, inter_union_distrib_left, inter_comm _ (frontier u), + huv.inter_eq, union_empty] + section disjoint_subsets variable [PreconnectedSpace α] @@ -348,6 +359,21 @@ theorem IsClopen.biUnion_connectedComponent_eq {Z : Set α} (h : IsClopen Z) : Subset.antisymm (iUnion₂_subset fun _ => h.connectedComponent_subset) fun _ h => mem_iUnion₂_of_mem h mem_connectedComponent +open Set.Notation in +/-- If `u v : Set X` and `u ⊆ v` is clopen in `v`, then `u` is the union of the connected +components of `v` in `X` which intersect `u`. -/ +lemma IsClopen.biUnion_connectedComponentIn {X : Type*} [TopologicalSpace X] {u v : Set X} + (hu : IsClopen (v ↓∩ u)) (huv₁ : u ⊆ v) : + u = ⋃ x ∈ u, connectedComponentIn v x := by + have := congr(((↑) : Set v → Set X) $(hu.biUnion_connectedComponent_eq.symm)) + simp only [Subtype.image_preimage_coe, mem_preimage, iUnion_coe_set, image_val_iUnion, + inter_eq_right.mpr huv₁] at this + nth_rw 1 [this] + congr! 2 with x hx + simp only [← connectedComponentIn_eq_image] + exact le_antisymm (iUnion_subset fun _ ↦ le_rfl) <| + iUnion_subset fun hx ↦ subset_iUnion₂_of_subset (huv₁ hx) hx le_rfl + /-- The preimage of a connected component is preconnected if the function has connected fibers and a subset is closed iff the preimage is. -/ theorem preimage_connectedComponent_connected [TopologicalSpace β] {f : α → β} From f758d8b81f6b4d6070b2b78840b764a5d9b5957a Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sun, 4 Aug 2024 10:58:30 +0000 Subject: [PATCH 010/359] chore: move more files into `Analysis.Normed.Algebra` (#15168) --- Mathlib.lean | 6 +++--- Mathlib/Analysis/CstarAlgebra/Exponential.lean | 2 +- .../{NormedSpace => Normed/Algebra}/Exponential.lean | 2 +- .../{NormedSpace => Normed/Algebra}/MatrixExponential.lean | 2 +- .../Algebra}/QuaternionExponential.lean | 2 +- Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 2 +- Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean | 2 +- Mathlib/Analysis/SpecialFunctions/Exponential.lean | 4 ++-- 8 files changed, 11 insertions(+), 11 deletions(-) rename Mathlib/Analysis/{NormedSpace => Normed/Algebra}/Exponential.lean (99%) rename Mathlib/Analysis/{NormedSpace => Normed/Algebra}/MatrixExponential.lean (99%) rename Mathlib/Analysis/{NormedSpace => Normed/Algebra}/QuaternionExponential.lean (99%) diff --git a/Mathlib.lean b/Mathlib.lean index a3c2bcc5b7ab3..82042139dae44 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1136,6 +1136,9 @@ import Mathlib.Analysis.MeanInequalitiesPow import Mathlib.Analysis.MellinInversion import Mathlib.Analysis.MellinTransform import Mathlib.Analysis.Normed.Algebra.Basic +import Mathlib.Analysis.Normed.Algebra.Exponential +import Mathlib.Analysis.Normed.Algebra.MatrixExponential +import Mathlib.Analysis.Normed.Algebra.QuaternionExponential import Mathlib.Analysis.Normed.Algebra.Spectrum import Mathlib.Analysis.Normed.Algebra.TrivSqZeroExt import Mathlib.Analysis.Normed.Algebra.Unitization @@ -1206,7 +1209,6 @@ import Mathlib.Analysis.NormedSpace.Connected import Mathlib.Analysis.NormedSpace.ContinuousAffineMap import Mathlib.Analysis.NormedSpace.DualNumber import Mathlib.Analysis.NormedSpace.ENorm -import Mathlib.Analysis.NormedSpace.Exponential import Mathlib.Analysis.NormedSpace.Extend import Mathlib.Analysis.NormedSpace.Extr import Mathlib.Analysis.NormedSpace.FunctionSeries @@ -1217,7 +1219,6 @@ import Mathlib.Analysis.NormedSpace.HomeomorphBall import Mathlib.Analysis.NormedSpace.IndicatorFunction import Mathlib.Analysis.NormedSpace.Int import Mathlib.Analysis.NormedSpace.MStructure -import Mathlib.Analysis.NormedSpace.MatrixExponential import Mathlib.Analysis.NormedSpace.MazurUlam import Mathlib.Analysis.NormedSpace.Multilinear.Basic import Mathlib.Analysis.NormedSpace.Multilinear.Curry @@ -1232,7 +1233,6 @@ import Mathlib.Analysis.NormedSpace.OperatorNorm.Prod import Mathlib.Analysis.NormedSpace.PiTensorProduct.InjectiveSeminorm import Mathlib.Analysis.NormedSpace.PiTensorProduct.ProjectiveSeminorm import Mathlib.Analysis.NormedSpace.Pointwise -import Mathlib.Analysis.NormedSpace.QuaternionExponential import Mathlib.Analysis.NormedSpace.RCLike import Mathlib.Analysis.NormedSpace.Real import Mathlib.Analysis.NormedSpace.RieszLemma diff --git a/Mathlib/Analysis/CstarAlgebra/Exponential.lean b/Mathlib/Analysis/CstarAlgebra/Exponential.lean index e953c4fb5f0ff..d7bfd49513760 100644 --- a/Mathlib/Analysis/CstarAlgebra/Exponential.lean +++ b/Mathlib/Analysis/CstarAlgebra/Exponential.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.NormedSpace.Exponential +import Mathlib.Analysis.Normed.Algebra.Exponential /-! # The exponential map from selfadjoint to unitary In this file, we establish various properties related to the map `fun a ↦ exp ℂ A (I • a)` diff --git a/Mathlib/Analysis/NormedSpace/Exponential.lean b/Mathlib/Analysis/Normed/Algebra/Exponential.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/Exponential.lean rename to Mathlib/Analysis/Normed/Algebra/Exponential.lean index afb211322d7cf..62c52c3b53ebb 100644 --- a/Mathlib/Analysis/NormedSpace/Exponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/Exponential.lean @@ -105,7 +105,7 @@ variable {𝔸} It is defined as the sum of the `FormalMultilinearSeries` `expSeries 𝕂 𝔸`. Note that when `𝔸 = Matrix n n 𝕂`, this is the **Matrix Exponential**; see -[`Analysis.NormedSpace.MatrixExponential`](./MatrixExponential) for lemmas specific to that +[`Analysis.Normed.Algebra.MatrixExponential`](./MatrixExponential) for lemmas specific to that case. -/ noncomputable def exp (x : 𝔸) : 𝔸 := (expSeries 𝕂 𝔸).sum x diff --git a/Mathlib/Analysis/NormedSpace/MatrixExponential.lean b/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/MatrixExponential.lean rename to Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean index dc8026fb80f4b..fb4f9d8aca29b 100644 --- a/Mathlib/Analysis/NormedSpace/MatrixExponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean @@ -3,7 +3,7 @@ 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.Analysis.NormedSpace.Exponential +import Mathlib.Analysis.Normed.Algebra.Exponential import Mathlib.Analysis.Matrix import Mathlib.LinearAlgebra.Matrix.ZPow import Mathlib.LinearAlgebra.Matrix.Hermitian diff --git a/Mathlib/Analysis/NormedSpace/QuaternionExponential.lean b/Mathlib/Analysis/Normed/Algebra/QuaternionExponential.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/QuaternionExponential.lean rename to Mathlib/Analysis/Normed/Algebra/QuaternionExponential.lean index 9bc1e3fecc9b9..c202556f03cf7 100644 --- a/Mathlib/Analysis/NormedSpace/QuaternionExponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/QuaternionExponential.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.Analysis.Quaternion -import Mathlib.Analysis.NormedSpace.Exponential +import Mathlib.Analysis.Normed.Algebra.Exponential import Mathlib.Analysis.SpecialFunctions.Trigonometric.Series /-! diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index 1d83a00d0260b..cbe8ee3895f27 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -9,7 +9,7 @@ import Mathlib.Analysis.Complex.Liouville import Mathlib.Analysis.Complex.Polynomial.Basic import Mathlib.Analysis.Analytic.RadiusLiminf import Mathlib.Topology.Algebra.Module.CharacterSpace -import Mathlib.Analysis.NormedSpace.Exponential +import Mathlib.Analysis.Normed.Algebra.Exponential import Mathlib.Analysis.Normed.Algebra.UnitizationL1 import Mathlib.Tactic.ContinuousFunctionalCalculus diff --git a/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean index eaa57d0c34175..730859031052a 100644 --- a/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Analysis.NormedSpace.Exponential +import Mathlib.Analysis.Normed.Algebra.Exponential import Mathlib.Analysis.Normed.Lp.ProdLp import Mathlib.Topology.Instances.TrivSqZeroExt diff --git a/Mathlib/Analysis/SpecialFunctions/Exponential.lean b/Mathlib/Analysis/SpecialFunctions/Exponential.lean index 11335991607be..b41bef6f28d1f 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exponential.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exponential.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker, Eric Wieser -/ -import Mathlib.Analysis.NormedSpace.Exponential +import Mathlib.Analysis.Normed.Algebra.Exponential import Mathlib.Analysis.Calculus.FDeriv.Analytic import Mathlib.Topology.MetricSpace.CauSeqFilter @@ -12,7 +12,7 @@ import Mathlib.Topology.MetricSpace.CauSeqFilter In this file, we prove basic properties about the derivative of the exponential map `exp 𝕂` in a Banach algebra `𝔸` over a field `𝕂`. We keep them separate from the main file -`Analysis/NormedSpace/Exponential` in order to minimize dependencies. +`Analysis.Normed.Algebra.Exponential` in order to minimize dependencies. ## Main results From 64580bf1ede64f4a5ca522fd31f3b8bb0a040de7 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Sun, 4 Aug 2024 10:58:31 +0000 Subject: [PATCH 011/359] chore: Address todo in `Data/Matrix/Rank.lean` (#15191) This PR addresses a todo in `Data/Matrix/Rank.lean`, generalizing from `LinearOrderedField` to `Field` as suggested by this old Zulip thread: https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/row.20rank.20equals.20column.20rank/near/350462992 --- Mathlib/Data/Matrix/Rank.lean | 29 ++++++++++++----------------- 1 file changed, 12 insertions(+), 17 deletions(-) diff --git a/Mathlib/Data/Matrix/Rank.lean b/Mathlib/Data/Matrix/Rank.lean index 5b1d17408efaa..e62b41b4656d9 100644 --- a/Mathlib/Data/Matrix/Rank.lean +++ b/Mathlib/Data/Matrix/Rank.lean @@ -7,6 +7,7 @@ import Mathlib.LinearAlgebra.Determinant import Mathlib.LinearAlgebra.FiniteDimensional import Mathlib.LinearAlgebra.Matrix.Diagonal import Mathlib.LinearAlgebra.Matrix.DotProduct +import Mathlib.LinearAlgebra.Matrix.Dual /-! # Rank of matrices @@ -18,12 +19,6 @@ This definition does not depend on the choice of basis, see `Matrix.rank_eq_finr * `Matrix.rank`: the rank of a matrix -## TODO - -* Do a better job of generalizing over `ℚ`, `ℝ`, and `ℂ` in `Matrix.rank_transpose` and - `Matrix.rank_conjTranspose`. See - [this Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/row.20rank.20equals.20column.20rank/near/350462992). - -/ @@ -250,22 +245,22 @@ theorem rank_transpose_mul_self (A : Matrix m n R) : (Aᵀ * A).rank = A.rank := · rw [ker_mulVecLin_transpose_mul_self] · simp only [LinearMap.finrank_range_add_finrank_ker] -/-- TODO: prove this in greater generality. -/ +end LinearOrderedField + @[simp] -theorem rank_transpose (A : Matrix m n R) : Aᵀ.rank = A.rank := - le_antisymm ((rank_transpose_mul_self _).symm.trans_le <| rank_mul_le_left _ _) - ((rank_transpose_mul_self _).symm.trans_le <| rank_mul_le_left _ _) +theorem rank_transpose [Field R] [Fintype m] (A : Matrix m n R) : Aᵀ.rank = A.rank := by + classical + rw [Aᵀ.rank_eq_finrank_range_toLin (Pi.basisFun R n).dualBasis (Pi.basisFun R m).dualBasis, + toLin_transpose, ← LinearMap.dualMap_def, LinearMap.finrank_range_dualMap_eq_finrank_range, + toLin_eq_toLin', toLin'_apply', rank] @[simp] -theorem rank_self_mul_transpose (A : Matrix m n R) : (A * Aᵀ).rank = A.rank := by +theorem rank_self_mul_transpose [LinearOrderedField R] [Fintype m] (A : Matrix m n R) : + (A * Aᵀ).rank = A.rank := by simpa only [rank_transpose, transpose_transpose] using rank_transpose_mul_self Aᵀ -end LinearOrderedField - -/-- The rank of a matrix is the rank of the space spanned by its rows. - -TODO: prove this in a generality that works for `ℂ` too, not just `ℚ` and `ℝ`. -/ -theorem rank_eq_finrank_span_row [LinearOrderedField R] [Finite m] (A : Matrix m n R) : +/-- The rank of a matrix is the rank of the space spanned by its rows. -/ +theorem rank_eq_finrank_span_row [Field R] [Finite m] (A : Matrix m n R) : A.rank = finrank R (Submodule.span R (Set.range A)) := by cases nonempty_fintype m rw [← rank_transpose, rank_eq_finrank_span_cols, transpose_transpose] From 79df738051ef1376db92aa53f64d151b5a62e884 Mon Sep 17 00:00:00 2001 From: Lorenzo Luccioli Date: Sun, 4 Aug 2024 10:58:32 +0000 Subject: [PATCH 012/359] feat (MeasureTheory): `Measure.add_sub_cancel` (#15468) Add `Measure.add_sub_cancel`. --- Mathlib/MeasureTheory/Function/Jacobian.lean | 2 +- Mathlib/MeasureTheory/Measure/Sub.lean | 6 ++++++ 2 files changed, 7 insertions(+), 1 deletion(-) diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index e5f3f8dc8e980..459316708b3a6 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -505,7 +505,7 @@ theorem _root_.ApproximatesLinearOn.norm_fderiv_sub_le {A : E →L[ℝ] E} {δ : exact ⟨a, az, by simp only [ha, add_neg_cancel_left]⟩ have norm_a : ‖a‖ ≤ ‖z‖ + ε := calc - ‖a‖ = ‖z + (a - z)‖ := by simp only [add_sub_cancel] + ‖a‖ = ‖z + (a - z)‖ := by simp only [_root_.add_sub_cancel] _ ≤ ‖z‖ + ‖a - z‖ := norm_add_le _ _ _ ≤ ‖z‖ + ε := add_le_add_left (mem_closedBall_iff_norm.1 az) _ -- use the approximation properties to control `(f' x - A) a`, and then `(f' x - A) z` as `z` is diff --git a/Mathlib/MeasureTheory/Measure/Sub.lean b/Mathlib/MeasureTheory/Measure/Sub.lean index 1090310c2ed02..956465498e954 100644 --- a/Mathlib/MeasureTheory/Measure/Sub.lean +++ b/Mathlib/MeasureTheory/Measure/Sub.lean @@ -90,6 +90,12 @@ theorem sub_add_cancel_of_le [IsFiniteMeasure ν] (h₁ : ν ≤ μ) : μ - ν + ext1 s h_s_meas rw [add_apply, sub_apply h_s_meas h₁, tsub_add_cancel_of_le (h₁ s)] +@[simp] +lemma add_sub_cancel [IsFiniteMeasure ν] : μ + ν - ν = μ := by + ext1 s hs + rw [sub_apply hs (Measure.le_add_left (le_refl _)), add_apply, + ENNReal.add_sub_cancel_right (measure_ne_top ν s)] + theorem restrict_sub_eq_restrict_sub_restrict (h_meas_s : MeasurableSet s) : (μ - ν).restrict s = μ.restrict s - ν.restrict s := by repeat rw [sub_def] From 16631a0912659f022f2f0506fe2d0a3de7109371 Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Sun, 4 Aug 2024 11:42:40 +0000 Subject: [PATCH 013/359] refactor(Combinatorics/SimpleGraph): remove sndOfNotNil in favour of getVert 1 (#15237) This allows to leave out the not-nil hypothesis, in addition to not needing it for the definition of tail. --- .../Combinatorics/SimpleGraph/Acyclic.lean | 9 +- .../SimpleGraph/Connectivity/Subgraph.lean | 8 +- .../SimpleGraph/Hamiltonian.lean | 19 +-- Mathlib/Combinatorics/SimpleGraph/Path.lean | 2 +- Mathlib/Combinatorics/SimpleGraph/Walk.lean | 125 +++++++++++------- 5 files changed, 100 insertions(+), 63 deletions(-) diff --git a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean index 6b8117d6925da..2cb9bcfec2287 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean @@ -170,11 +170,12 @@ lemma IsTree.card_edgeFinset [Fintype V] [Fintype G.edgeSet] (hG : G.IsTree) : · exact (congrArg (·.fst) h) · have h1 : ((f a).firstDart <| not_nil_of_ne (by simpa using ha)).snd = b := congrArg (·.snd) h - have h3 := congrArg length (hf' _ (((f _).tail _).copy h1 rfl) ?_) - · rw [length_copy, ← add_left_inj 1, length_tail_add_one] at h3 + have h3 := congrArg length (hf' _ ((f _).tail.copy h1 rfl) ?_) + · rw [length_copy, ← add_left_inj 1, + length_tail_add_one (not_nil_of_ne (by simpa using ha))] at h3 omega · simp only [ne_eq, eq_mp_eq_cast, id_eq, isPath_copy] - exact (hf _).tail _ + exact (hf _).tail (not_nil_of_ne (by simpa using ha)) case surj => simp only [mem_edgeFinset, Finset.mem_compl, Finset.mem_singleton, Sym2.forall, mem_edgeSet] intros x y h @@ -188,7 +189,7 @@ lemma IsTree.card_edgeFinset [Fintype V] [Fintype G.edgeSet] (hG : G.IsTree) : length_cons, length_nil] at h' simp [Nat.le_zero, Nat.one_ne_zero] at h' rw [← hf' _ (.cons h.symm (f x)) ((cons_isPath_iff _ _).2 ⟨hf _, fun hy => ?contra⟩)] - · rfl + · simp only [firstDart_toProd, getVert_cons_succ, getVert_zero, Prod.swap_prod_mk] case contra => suffices (f x).takeUntil y hy = .cons h .nil by rw [← take_spec _ hy] at h' diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity/Subgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity/Subgraph.lean index e73b22eb31d5a..b19fdc540c1bb 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity/Subgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity/Subgraph.lean @@ -189,9 +189,9 @@ theorem toSubgraph_adj_getVert {u v} (w : G.Walk u v) {i : ℕ} (hi : i < w.leng | nil => cases hi | cons hxy i' ih => cases i - · simp only [Walk.toSubgraph, Walk.getVert_zero, zero_add, cons_getVert_succ, Subgraph.sup_adj, + · simp only [Walk.toSubgraph, Walk.getVert_zero, zero_add, getVert_cons_succ, Subgraph.sup_adj, subgraphOfAdj_adj, true_or] - · simp only [Walk.toSubgraph, cons_getVert_succ, Subgraph.sup_adj, subgraphOfAdj_adj, Sym2.eq, + · simp only [Walk.toSubgraph, getVert_cons_succ, Subgraph.sup_adj, subgraphOfAdj_adj, Sym2.eq, Sym2.rel_iff', Prod.mk.injEq, Prod.swap_prod_mk] right exact ih (Nat.succ_lt_succ_iff.mp hi) @@ -211,7 +211,7 @@ theorem toSubgraph_adj_iff {u v u' v'} (w : G.Walk u v) : cases hadj with | inl hl => use 0 - simp only [Walk.getVert_zero, zero_add, cons_getVert_succ] + simp only [Walk.getVert_zero, zero_add, getVert_cons_succ] refine ⟨?_, by simp only [length_cons, Nat.zero_lt_succ]⟩ simp only [Sym2.eq, Sym2.rel_iff', Prod.mk.injEq, Prod.swap_prod_mk] cases hl with @@ -220,7 +220,7 @@ theorem toSubgraph_adj_iff {u v u' v'} (w : G.Walk u v) : | inr hr => obtain ⟨i, hi⟩ := (toSubgraph_adj_iff _).mp hr use i + 1 - simp only [cons_getVert_succ] + simp only [getVert_cons_succ] constructor · exact hi.1 · simp only [Walk.length_cons, add_lt_add_iff_right, Nat.add_lt_add_right hi.2 1] diff --git a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean index aad8e884da180..ce4ffb9224f02 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean @@ -67,7 +67,7 @@ end /-- A hamiltonian cycle is a cycle that visits every vertex once. -/ structure IsHamiltonianCycle (p : G.Walk a a) extends p.IsCycle : Prop := - isHamiltonian_tail : (p.tail toIsCycle.not_nil).IsHamiltonian + isHamiltonian_tail : p.tail.IsHamiltonian variable {p : G.Walk a a} @@ -84,22 +84,23 @@ lemma IsHamiltonianCycle.map {H : SimpleGraph β} (f : G →g H) (hf : Bijective intro x rcases p with (_ | ⟨y, p⟩) · cases hp.ne_nil rfl - simp only [support_cons, List.count_cons, List.map_cons, List.head_cons, hf.injective.eq_iff, - add_tsub_cancel_right] + simp only [map_cons, getVert_cons_succ, tail_cons_eq, support_copy,support_map] + rw [List.count_map_of_injective _ _ hf.injective, ← support_copy, ← tail_cons_eq] exact hp.isHamiltonian_tail _ lemma isHamiltonianCycle_isCycle_and_isHamiltonian_tail : - p.IsHamiltonianCycle ↔ ∃ h : p.IsCycle, (p.tail h.not_nil).IsHamiltonian := + p.IsHamiltonianCycle ↔ p.IsCycle ∧ p.tail.IsHamiltonian := ⟨fun ⟨h, h'⟩ ↦ ⟨h, h'⟩, fun ⟨h, h'⟩ ↦ ⟨h, h'⟩⟩ lemma isHamiltonianCycle_iff_isCycle_and_support_count_tail_eq_one : p.IsHamiltonianCycle ↔ p.IsCycle ∧ ∀ a, (support p).tail.count a = 1 := by - simp only [isHamiltonianCycle_isCycle_and_isHamiltonian_tail, IsHamiltonian, support_tail, - exists_prop] + simp (config := { contextual := true }) [isHamiltonianCycle_isCycle_and_isHamiltonian_tail, + IsHamiltonian, support_tail, IsCycle.not_nil, exists_prop] /-- A hamiltonian cycle visits every vertex. -/ lemma IsHamiltonianCycle.mem_support (hp : p.IsHamiltonianCycle) (b : α) : - b ∈ p.support := List.mem_of_mem_tail <| support_tail p _ ▸ hp.isHamiltonian_tail.mem_support _ + b ∈ p.support := + List.mem_of_mem_tail <| support_tail p hp.1.not_nil ▸ hp.isHamiltonian_tail.mem_support _ /-- The length of a hamiltonian cycle is the number of vertices. -/ lemma IsHamiltonianCycle.length_eq [Fintype α] (hp : p.IsHamiltonianCycle) : @@ -110,11 +111,11 @@ lemma IsHamiltonianCycle.length_eq [Fintype α] (hp : p.IsHamiltonianCycle) : lemma IsHamiltonianCycle.count_support_self (hp : p.IsHamiltonianCycle) : p.support.count a = 2 := by - rw [support_eq_cons, List.count_cons_self, ← support_tail, hp.isHamiltonian_tail] + rw [support_eq_cons, List.count_cons_self, ← support_tail _ hp.1.not_nil, hp.isHamiltonian_tail] lemma IsHamiltonianCycle.support_count_of_ne (hp : p.IsHamiltonianCycle) (h : a ≠ b) : p.support.count b = 1 := by - rw [← cons_support_tail p, List.count_cons_of_ne h.symm, hp.isHamiltonian_tail] + rw [← cons_support_tail p hp.1.not_nil, List.count_cons_of_ne h.symm, hp.isHamiltonian_tail] end Walk diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index 7dd0f77cbd996..0c7b176088328 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -236,7 +236,7 @@ theorem cons_isCycle_iff {u v : V} (p : G.Walk v u) (h : G.Adj u v) : have : p.support.Nodup → p.edges.Nodup := edges_nodup_of_support_nodup tauto -lemma IsPath.tail {p : G.Walk u v} (hp : p.IsPath) (hp' : ¬ p.Nil) : (p.tail hp').IsPath := by +lemma IsPath.tail {p : G.Walk u v} (hp : p.IsPath) (hp' : ¬ p.Nil) : p.tail.IsPath := by rw [Walk.isPath_def] at hp ⊢ rw [← cons_support_tail _ hp', List.nodup_cons] at hp exact hp.2 diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index 2c7a2bc63fe8f..c6972002e049b 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -177,16 +177,21 @@ theorem adj_getVert_succ {u v} (w : G.Walk u v) {i : ℕ} (hi : i < w.length) : · simp [getVert, hxy] · exact ih (Nat.succ_lt_succ_iff.1 hi) +lemma getVert_cons_one {u v w} (q : G.Walk v w) (hadj : G.Adj u v) : + (q.cons hadj).getVert 1 = v := by + have : (q.cons hadj).getVert 1 = q.getVert 0 := rfl + simpa [getVert_zero] using this + @[simp] -lemma cons_getVert_succ {u v w n} (p : G.Walk v w) (h : G.Adj u v) : +lemma getVert_cons_succ {u v w n} (p : G.Walk v w) (h : G.Adj u v) : (p.cons h).getVert (n + 1) = p.getVert n := rfl -lemma cons_getVert {u v w n} (p : G.Walk v w) (h : G.Adj u v) (hn : n ≠ 0) : +lemma getVert_cons {u v w n} (p : G.Walk v w) (h : G.Adj u v) (hn : n ≠ 0) : (p.cons h).getVert n = p.getVert (n - 1) := by obtain ⟨i, hi⟩ : ∃ (i : ℕ), i.succ = n := by use n - 1; exact Nat.succ_pred_eq_of_ne_zero hn rw [← hi] - simp only [Nat.succ_eq_add_one, cons_getVert_succ, Nat.add_sub_cancel] + simp only [Nat.succ_eq_add_one, getVert_cons_succ, Nat.add_sub_cancel] @[simp] theorem cons_append {u v w x : V} (h : G.Adj u v) (p : G.Walk v w) (q : G.Walk w x) : @@ -741,6 +746,9 @@ lemma nil_iff_support_eq {p : G.Walk v w} : p.Nil ↔ p.support = [v] := by lemma nil_iff_length_eq {p : G.Walk v w} : p.Nil ↔ p.length = 0 := by cases p <;> simp +lemma not_nil_iff_lt_length {p : G.Walk v w} : ¬ p.Nil ↔ 0 < p.length := by + cases p <;> simp + lemma not_nil_iff {p : G.Walk v w} : ¬ p.Nil ↔ ∃ (u : V) (h : G.Adj v u) (q : G.Walk u w), p = cons h q := by cases p <;> simp [*] @@ -765,61 +773,77 @@ lemma notNilRec_cons {motive : {u w : V} → (p : G.Walk u w) → ¬ p.Nil → S motive (q.cons h) Walk.not_nil_cons) (h' : G.Adj u v) (q' : G.Walk v w) : @Walk.notNilRec _ _ _ _ _ cons _ _ = cons h' q' := by rfl -/-- The second vertex along a non-nil walk. -/ -def sndOfNotNil (p : G.Walk v w) (hp : ¬ p.Nil) : V := - p.notNilRec (@fun _ u _ _ _ => u) hp +@[simp] lemma adj_getVert_one {p : G.Walk v w} (hp : ¬ p.Nil) : + G.Adj v (p.getVert 1) := by + simpa using adj_getVert_succ p (by simpa [not_nil_iff_lt_length] using hp : 0 < p.length) -@[simp] lemma adj_sndOfNotNil {p : G.Walk v w} (hp : ¬ p.Nil) : - G.Adj v (p.sndOfNotNil hp) := - p.notNilRec (fun h _ => h) hp +/-- The walk obtained by removing the first `n` darts of a walk. -/ +def drop {u v : V} (p : G.Walk u v) (n : ℕ) : G.Walk (p.getVert n) v := + match p, n with + | .nil, _ => .nil + | p, 0 => p.copy (getVert_zero p).symm rfl + | .cons h q, (n + 1) => (q.drop n).copy (getVert_cons_succ _ h).symm rfl /-- The walk obtained by removing the first dart of a non-nil walk. -/ -def tail (p : G.Walk u v) (hp : ¬ p.Nil) : G.Walk (p.sndOfNotNil hp) v := - p.notNilRec (fun _ q => q) hp +def tail (p : G.Walk u v) : G.Walk (p.getVert 1) v := p.drop 1 + +@[simp] +lemma tail_cons_nil (h : G.Adj u v) : (Walk.cons h .nil).tail = .nil := by rfl + +lemma tail_cons_eq (h : G.Adj u v) (p : G.Walk v w) : + (p.cons h).tail = p.copy (getVert_zero p).symm rfl := by + match p with + | .nil => rfl + | .cons h q => rfl /-- The first dart of a walk. -/ @[simps] def firstDart (p : G.Walk v w) (hp : ¬ p.Nil) : G.Dart where fst := v - snd := p.sndOfNotNil hp - adj := p.adj_sndOfNotNil hp + snd := p.getVert 1 + adj := p.adj_getVert_one hp lemma edge_firstDart (p : G.Walk v w) (hp : ¬ p.Nil) : - (p.firstDart hp).edge = s(v, p.sndOfNotNil hp) := rfl + (p.firstDart hp).edge = s(v, p.getVert 1) := rfl variable {x y : V} -- TODO: rename to u, v, w instead? -@[simp] lemma cons_tail_eq (p : G.Walk x y) (hp : ¬ p.Nil) : - cons (p.adj_sndOfNotNil hp) (p.tail hp) = p := - p.notNilRec (fun _ _ => rfl) hp +lemma cons_tail_eq (p : G.Walk x y) (hp : ¬ p.Nil) : + cons (p.adj_getVert_one hp) p.tail = p := by + cases p with + | nil => simp only [nil_nil, not_true_eq_false] at hp + | cons h q => + simp only [getVert_cons_succ, tail_cons_eq, cons_copy, copy_rfl_rfl] @[simp] lemma cons_support_tail (p : G.Walk x y) (hp : ¬p.Nil) : - x :: (p.tail hp).support = p.support := by - rw [← support_cons, cons_tail_eq] + x :: p.tail.support = p.support := by + rw [← support_cons, cons_tail_eq _ hp] @[simp] lemma length_tail_add_one {p : G.Walk x y} (hp : ¬ p.Nil) : - (p.tail hp).length + 1 = p.length := by - rw [← length_cons, cons_tail_eq] + p.tail.length + 1 = p.length := by + rw [← length_cons, cons_tail_eq _ hp] @[simp] lemma nil_copy {x' y' : V} {p : G.Walk x y} (hx : x = x') (hy : y = y') : (p.copy hx hy).Nil = p.Nil := by subst_vars; rfl -@[simp] lemma support_tail (p : G.Walk v v) (hp) : - (p.tail hp).support = p.support.tail := by +@[simp] lemma support_tail (p : G.Walk v v) (hp : ¬ p.Nil) : + p.tail.support = p.support.tail := by rw [← cons_support_tail p hp, List.tail_cons] @[simp] lemma tail_cons {t u v} (p : G.Walk u v) (h : G.Adj t u) : - (p.cons h).tail not_nil_cons = p := by - unfold Walk.tail; simp only [notNilRec_cons] + (p.cons h).tail = p.copy (getVert_zero p).symm rfl := by + match p with + | .nil => rfl + | .cons h q => rfl -lemma tail_support_eq_support_tail (p : G.Walk u v) (hnp : ¬p.Nil) : - (p.tail hnp).support = p.support.tail := - p.notNilRec (by - intro u v w huv q - unfold Walk.tail - simp only [notNilRec_cons, Walk.support_cons, List.tail_cons]) hnp +lemma support_tail_of_not_nil (p : G.Walk u v) (hnp : ¬p.Nil) : + p.tail.support = p.support.tail := by + match p with + | .nil => simp only [nil_nil, not_true_eq_false] at hnp + | .cons h q => + simp only [tail_cons, getVert_cons_succ, support_copy, support_cons, List.tail_cons] /-! ### Walk decompositions -/ @@ -995,17 +1019,23 @@ theorem exists_boundary_dart {u v : V} (p : G.Walk u v) (S : Set V) (uS : u ∈ exact ⟨d, List.Mem.tail _ hd, hcd⟩ · exact ⟨⟨(x, y), a⟩, List.Mem.head _, uS, h⟩ -lemma getVert_tail {u v n} (p : G.Walk u v) (hnp: ¬ p.Nil) : - (p.tail hnp).getVert n = p.getVert (n + 1) := - p.notNilRec (fun _ _ ↦ by simp only [tail_cons, cons_getVert_succ]) hnp - -@[simp] -lemma cons_sndOfNotNil (q : G.Walk v w) (hadj : G.Adj u v) : - (q.cons hadj).sndOfNotNil not_nil_cons = v := by - unfold sndOfNotNil; simp only [notNilRec_cons] - -lemma getVert_one (p : G.Walk u v) (hnp : ¬ p.Nil) : p.getVert 1 = p.sndOfNotNil hnp := - p.notNilRec (fun _ _ ↦ by simp only [cons_getVert_succ, getVert_zero, cons_sndOfNotNil]) hnp +@[simp] lemma getVert_copy {u v w x : V} (p : G.Walk u v) (i : ℕ) (h : u = w) (h' : v = x) : + (p.copy h h').getVert i = p.getVert i := by + subst_vars + match p, i with + | .nil, _ => + rw [getVert_of_length_le _ (by simp only [length_nil, Nat.zero_le] : nil.length ≤ _)] + rw [getVert_of_length_le _ (by simp only [length_copy, length_nil, Nat.zero_le])] + | .cons hadj q, 0 => simp only [copy_rfl_rfl, getVert_zero] + | .cons hadj q, (n + 1) => simp only [copy_cons, getVert_cons_succ]; rfl + +@[simp] lemma getVert_tail {u v n} (p : G.Walk u v) (hnp: ¬ p.Nil) : + p.tail.getVert n = p.getVert (n + 1) := by + match p with + | .nil => rfl + | .cons h q => + simp only [getVert_cons_succ, tail_cons_eq, getVert_cons] + exact getVert_copy q n (getVert_zero q).symm rfl /-- Given a walk `w` and a node in the support, there exists a natural `n`, such that given node is the `n`-th node (zero-indexed) in the walk. In addition, `n` is at most the length of the path. @@ -1032,13 +1062,18 @@ theorem mem_support_iff_exists_getVert {u v w : V} {p : G.Walk v w} : rw [@nil_iff_length_eq] have : 1 ≤ p.length := by omega exact Nat.not_eq_zero_of_lt this - rw [← tail_support_eq_support_tail _ hnp] + rw [← support_tail_of_not_nil _ hnp] rw [mem_support_iff_exists_getVert] use n - 1 - simp only [Nat.sub_le_iff_le_add, length_tail_add_one, getVert_tail] - have : n - 1 + 1 = n := by omega + simp only [Nat.sub_le_iff_le_add] + rw [getVert_tail _ hnp, length_tail_add_one hnp] + have : (n - 1 + 1) = n:= by omega rwa [this] termination_by p.length +decreasing_by +· simp_wf + rw [@Nat.lt_iff_add_one_le] + rw [length_tail_add_one hnp] end Walk From 34cfdc4ea7becb86738b7fe29e743aa69420a14b Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Sun, 4 Aug 2024 11:42:42 +0000 Subject: [PATCH 014/359] chore(ModelTheory): Fix Model Theory docstrings (#15408) Updates the style on docstrings of files in the Model Theory folder to better match the current style guide. --- Mathlib/ModelTheory/Algebra/Field/Basic.lean | 13 +++-- Mathlib/ModelTheory/Algebra/Field/CharP.lean | 5 +- Mathlib/ModelTheory/Algebra/Ring/Basic.lean | 14 +++-- .../Algebra/Ring/FreeCommRing.lean | 2 - Mathlib/ModelTheory/Basic.lean | 18 ++++--- Mathlib/ModelTheory/Bundled.lean | 8 +-- Mathlib/ModelTheory/Complexity.lean | 18 +++---- Mathlib/ModelTheory/Definability.lean | 21 ++++---- Mathlib/ModelTheory/DirectLimit.lean | 8 +-- Mathlib/ModelTheory/ElementaryMaps.lean | 14 ++--- .../ModelTheory/ElementarySubstructures.lean | 6 ++- Mathlib/ModelTheory/Encoding.lean | 26 +++++---- Mathlib/ModelTheory/FinitelyGenerated.lean | 11 ++-- Mathlib/ModelTheory/Fraisse.lean | 51 ++++++++++-------- Mathlib/ModelTheory/Graph.lean | 14 ++--- Mathlib/ModelTheory/LanguageMap.lean | 13 +++-- Mathlib/ModelTheory/Order.lean | 26 ++++----- Mathlib/ModelTheory/Quotients.lean | 9 ++-- Mathlib/ModelTheory/Satisfiability.lean | 43 ++++++++------- Mathlib/ModelTheory/Semantics.lean | 46 ++++++++-------- Mathlib/ModelTheory/Skolem.lean | 8 +-- Mathlib/ModelTheory/Substructures.lean | 34 ++++++------ Mathlib/ModelTheory/Syntax.lean | 53 ++++++++++--------- Mathlib/ModelTheory/Types.lean | 22 ++++---- Mathlib/ModelTheory/Ultraproducts.lean | 15 +++--- 25 files changed, 273 insertions(+), 225 deletions(-) diff --git a/Mathlib/ModelTheory/Algebra/Field/Basic.lean b/Mathlib/ModelTheory/Algebra/Field/Basic.lean index 1ff3ab6cf5133..afa727e1810e7 100644 --- a/Mathlib/ModelTheory/Algebra/Field/Basic.lean +++ b/Mathlib/ModelTheory/Algebra/Field/Basic.lean @@ -11,19 +11,18 @@ import Mathlib.Algebra.Field.MinimalAxioms import Mathlib.Data.Nat.Cast.Order.Ring /-! - # The First Order Theory of Fields This file defines the first order theory of fields as a theory over the language of rings. ## Main definitions -* `FirstOrder.Language.Theory.field` : the theory of fields -* `FirstOrder.Model.fieldOfModelField` : a model of the theory of fields on a type `K` that - already has ring operations. -* `FirstOrder.Model.compatibleRingOfModelField` : shows that the ring operations on `K` given -by `fieldOfModelField` are compatible with the ring operations on `K` given by the -`Language.ring.Structure` instance. +- `FirstOrder.Language.Theory.field` : the theory of fields +- `FirstOrder.Model.fieldOfModelField` : a model of the theory of fields on a type `K` that + already has ring operations. +- `FirstOrder.Model.compatibleRingOfModelField` : shows that the ring operations on `K` given + by `fieldOfModelField` are compatible with the ring operations on `K` given by the + `Language.ring.Structure` instance. -/ variable {K : Type*} diff --git a/Mathlib/ModelTheory/Algebra/Field/CharP.lean b/Mathlib/ModelTheory/Algebra/Field/CharP.lean index 2d3c81932591f..03491be723527 100644 --- a/Mathlib/ModelTheory/Algebra/Field/CharP.lean +++ b/Mathlib/ModelTheory/Algebra/Field/CharP.lean @@ -10,16 +10,15 @@ import Mathlib.ModelTheory.Algebra.Ring.FreeCommRing import Mathlib.ModelTheory.Algebra.Field.Basic /-! - # First order theory of fields This file defines the first order theory of fields of characteristic `p` as a theory over the language of rings ## Main definitions -* `FirstOrder.Language.Theory.fieldOfChar` : the first order theory of fields of characteristic `p` - as a theory over the language of rings +- `FirstOrder.Language.Theory.fieldOfChar` : the first order theory of fields of characteristic `p` + as a theory over the language of rings -/ variable {p : ℕ} {K : Type*} diff --git a/Mathlib/ModelTheory/Algebra/Ring/Basic.lean b/Mathlib/ModelTheory/Algebra/Ring/Basic.lean index 7c3cfdf40dc70..7df3ce6ac5bdb 100644 --- a/Mathlib/ModelTheory/Algebra/Ring/Basic.lean +++ b/Mathlib/ModelTheory/Algebra/Ring/Basic.lean @@ -9,7 +9,6 @@ import Mathlib.ModelTheory.Semantics import Mathlib.Algebra.Ring.Equiv /-! - # First Order Language of Rings This file defines the first order language of rings, as well as defining instance of `Add`, `Mul`, @@ -17,12 +16,12 @@ etc. on terms in the language. ## Main Definitions -* `FirstOrder.Language.ring` : the language of rings, with function symbols `+`, `*`, `-`, `0`, `1` -* `FirstOrder.Ring.CompatibleRing` : A class stating that a type is a `Language.ring.Structure`, and -that this structure is the same as the structure given by the classes `Add`, `Mul`, etc. already on -`R`. -* `FirstOrder.Ring.compatibleRingOfRing` : Given a type `R` with instances for each of the `Ring` -operations, make a `compatibleRing` instance. +- `FirstOrder.Language.ring` : the language of rings, with function symbols `+`, `*`, `-`, `0`, `1` +- `FirstOrder.Ring.CompatibleRing` : A class stating that a type is a `Language.ring.Structure`, and + that this structure is the same as the structure given by the classes `Add`, `Mul`, etc. already + on `R`. +- `FirstOrder.Ring.compatibleRingOfRing` : Given a type `R` with instances for each of the `Ring` + operations, make a `compatibleRing` instance. ## Implementation Notes @@ -38,7 +37,6 @@ a `Language.ring.Structure K` instance and for example an instance of `Theory.fi you must add local instances with definitions like `ModelTheory.Field.fieldOfModelField K` and `FirstOrder.Ring.compatibleRingOfModelField K`. (in `Mathlib/ModelTheory/Algebra/Field/Basic.lean`), depending on the Theory. - -/ variable {α : Type*} diff --git a/Mathlib/ModelTheory/Algebra/Ring/FreeCommRing.lean b/Mathlib/ModelTheory/Algebra/Ring/FreeCommRing.lean index 7a8f5addefcd0..ea7888ce6a2c9 100644 --- a/Mathlib/ModelTheory/Algebra/Ring/FreeCommRing.lean +++ b/Mathlib/ModelTheory/Algebra/Ring/FreeCommRing.lean @@ -8,7 +8,6 @@ import Mathlib.ModelTheory.Algebra.Ring.Basic import Mathlib.RingTheory.FreeCommRing /-! - # Making a term in the language of rings from an element of the FreeCommRing This file defines the function `FirstOrder.Ring.termOfFreeCommRing` which constructs a @@ -16,7 +15,6 @@ This file defines the function `FirstOrder.Ring.termOfFreeCommRing` which constr The theorem `FirstOrder.Ring.realize_termOfFreeCommRing` shows that the term constructed when realized in a ring `R` is equal to the lift of the element of `FreeCommRing α` to `R`. - -/ namespace FirstOrder diff --git a/Mathlib/ModelTheory/Basic.lean b/Mathlib/ModelTheory/Basic.lean index cd8a63175e143..20bd4499e732e 100644 --- a/Mathlib/ModelTheory/Basic.lean +++ b/Mathlib/ModelTheory/Basic.lean @@ -8,33 +8,35 @@ import Mathlib.SetTheory.Cardinal.Basic /-! # Basics on First-Order Structures + This file defines first-order languages and structures in the style of the [Flypitch project](https://flypitch.github.io/), as well as several important maps between structures. ## Main Definitions -* A `FirstOrder.Language` defines a language as a pair of functions from the natural numbers to + +- A `FirstOrder.Language` defines a language as a pair of functions from the natural numbers to `Type l`. One sends `n` to the type of `n`-ary functions, and the other sends `n` to the type of `n`-ary relations. -* A `FirstOrder.Language.Structure` interprets the symbols of a given `FirstOrder.Language` in the +- A `FirstOrder.Language.Structure` interprets the symbols of a given `FirstOrder.Language` in the context of a given type. -* A `FirstOrder.Language.Hom`, denoted `M →[L] N`, is a map from the `L`-structure `M` to the +- A `FirstOrder.Language.Hom`, denoted `M →[L] N`, is a map from the `L`-structure `M` to the `L`-structure `N` that commutes with the interpretations of functions, and which preserves the interpretations of relations (although only in the forward direction). -* A `FirstOrder.Language.Embedding`, denoted `M ↪[L] N`, is an embedding from the `L`-structure `M` +- A `FirstOrder.Language.Embedding`, denoted `M ↪[L] N`, is an embedding from the `L`-structure `M` to the `L`-structure `N` that commutes with the interpretations of functions, and which preserves the interpretations of relations in both directions. -* A `FirstOrder.Language.Equiv`, denoted `M ≃[L] N`, is an equivalence from the `L`-structure `M` +- A `FirstOrder.Language.Equiv`, denoted `M ≃[L] N`, is an equivalence from the `L`-structure `M` to the `L`-structure `N` that commutes with the interpretations of functions, and which preserves the interpretations of relations in both directions. ## References + For the Flypitch project: - [J. Han, F. van Doorn, *A formal proof of the independence of the continuum hypothesis*] -[flypitch_cpp] + [flypitch_cpp] - [J. Han, F. van Doorn, *A formalization of forcing and the unprovability of -the continuum hypothesis*][flypitch_itp] - + the continuum hypothesis*][flypitch_itp] -/ universe u v u' v' w w' diff --git a/Mathlib/ModelTheory/Bundled.lean b/Mathlib/ModelTheory/Bundled.lean index f71b195165196..9b260af7fec41 100644 --- a/Mathlib/ModelTheory/Bundled.lean +++ b/Mathlib/ModelTheory/Bundled.lean @@ -8,15 +8,17 @@ import Mathlib.CategoryTheory.ConcreteCategory.Bundled /-! # Bundled First-Order Structures + This file bundles types together with their first-order structure. ## Main Definitions -* `FirstOrder.Language.Theory.ModelType` is the type of nonempty models of a particular theory. -* `FirstOrder.Language.equivSetoid` is the isomorphism equivalence relation on bundled structures. + +- `FirstOrder.Language.Theory.ModelType` is the type of nonempty models of a particular theory. +- `FirstOrder.Language.equivSetoid` is the isomorphism equivalence relation on bundled structures. ## TODO -* Define category structures on bundled structures and models. +- Define category structures on bundled structures and models. -/ diff --git a/Mathlib/ModelTheory/Complexity.lean b/Mathlib/ModelTheory/Complexity.lean index ce6ee121fcd48..e51cf9755ab1a 100644 --- a/Mathlib/ModelTheory/Complexity.lean +++ b/Mathlib/ModelTheory/Complexity.lean @@ -12,19 +12,19 @@ This file defines quantifier complexity of first-order formulas, and constructs ## Main Definitions -* `FirstOrder.Language.BoundedFormula.IsAtomic` defines atomic formulas - those which are -constructed only from terms and relations. -* `FirstOrder.Language.BoundedFormula.IsQF` defines quantifier-free formulas - those which are -constructed only from atomic formulas and boolean operations. -* `FirstOrder.Language.BoundedFormula.IsPrenex` defines when a formula is in prenex normal form - -when it consists of a series of quantifiers applied to a quantifier-free formula. -* `FirstOrder.Language.BoundedFormula.toPrenex` constructs a prenex normal form of a given formula. +- `FirstOrder.Language.BoundedFormula.IsAtomic` defines atomic formulas - those which are + constructed only from terms and relations. +- `FirstOrder.Language.BoundedFormula.IsQF` defines quantifier-free formulas - those which are + constructed only from atomic formulas and boolean operations. +- `FirstOrder.Language.BoundedFormula.IsPrenex` defines when a formula is in prenex normal form - + when it consists of a series of quantifiers applied to a quantifier-free formula. +- `FirstOrder.Language.BoundedFormula.toPrenex` constructs a prenex normal form of a given formula. ## Main Results -* `FirstOrder.Language.BoundedFormula.realize_toPrenex` shows that the prenex normal form of a -formula has the same realization as the original formula. +- `FirstOrder.Language.BoundedFormula.realize_toPrenex` shows that the prenex normal form of a + formula has the same realization as the original formula. -/ diff --git a/Mathlib/ModelTheory/Definability.lean b/Mathlib/ModelTheory/Definability.lean index 0413ebeacdec9..771c189b32504 100644 --- a/Mathlib/ModelTheory/Definability.lean +++ b/Mathlib/ModelTheory/Definability.lean @@ -9,21 +9,24 @@ import Mathlib.ModelTheory.Semantics /-! # Definable Sets + This file defines what it means for a set over a first-order structure to be definable. ## Main Definitions -* `Set.Definable` is defined so that `A.Definable L s` indicates that the -set `s` of a finite cartesian power of `M` is definable with parameters in `A`. -* `Set.Definable₁` is defined so that `A.Definable₁ L s` indicates that -`(s : Set M)` is definable with parameters in `A`. -* `Set.Definable₂` is defined so that `A.Definable₂ L s` indicates that -`(s : Set (M × M))` is definable with parameters in `A`. -* A `FirstOrder.Language.DefinableSet` is defined so that `L.DefinableSet A α` is the boolean + +- `Set.Definable` is defined so that `A.Definable L s` indicates that the + set `s` of a finite cartesian power of `M` is definable with parameters in `A`. +- `Set.Definable₁` is defined so that `A.Definable₁ L s` indicates that + `(s : Set M)` is definable with parameters in `A`. +- `Set.Definable₂` is defined so that `A.Definable₂ L s` indicates that + `(s : Set (M × M))` is definable with parameters in `A`. +- A `FirstOrder.Language.DefinableSet` is defined so that `L.DefinableSet A α` is the boolean algebra of subsets of `α → M` defined by formulas with parameters in `A`. ## Main Results -* `L.DefinableSet A α` forms a `BooleanAlgebra` -* `Set.Definable.image_comp` shows that definability is closed under projections in finite + +- `L.DefinableSet A α` forms a `BooleanAlgebra` +- `Set.Definable.image_comp` shows that definability is closed under projections in finite dimensions. -/ diff --git a/Mathlib/ModelTheory/DirectLimit.lean b/Mathlib/ModelTheory/DirectLimit.lean index a7123bd8b72db..5f0d653d1499c 100644 --- a/Mathlib/ModelTheory/DirectLimit.lean +++ b/Mathlib/ModelTheory/DirectLimit.lean @@ -10,15 +10,17 @@ import Mathlib.ModelTheory.FinitelyGenerated /-! # Direct Limits of First-Order Structures + This file constructs the direct limit of a directed system of first-order embeddings. ## Main Definitions -* `FirstOrder.Language.DirectLimit G f` is the direct limit of the directed system `f` of + +- `FirstOrder.Language.DirectLimit G f` is the direct limit of the directed system `f` of first-order embeddings between the structures indexed by `G`. -* `FirstOrder.Language.DirectLimit.lift` is the universal property of the direct limit: maps +- `FirstOrder.Language.DirectLimit.lift` is the universal property of the direct limit: maps from the components to another module that respect the directed system structure give rise to a unique map out of the direct limit. -* `FirstOrder.Language.DirectLimit.equiv_lift` is the equivalence between limits of +- `FirstOrder.Language.DirectLimit.equiv_lift` is the equivalence between limits of isomorphic direct systems. -/ diff --git a/Mathlib/ModelTheory/ElementaryMaps.lean b/Mathlib/ModelTheory/ElementaryMaps.lean index 0bb0592b96c09..f9e0c790bdce4 100644 --- a/Mathlib/ModelTheory/ElementaryMaps.lean +++ b/Mathlib/ModelTheory/ElementaryMaps.lean @@ -10,16 +10,18 @@ import Mathlib.ModelTheory.Substructures # Elementary Maps Between First-Order Structures ## Main Definitions -* A `FirstOrder.Language.ElementaryEmbedding` is an embedding that commutes with the + +- A `FirstOrder.Language.ElementaryEmbedding` is an embedding that commutes with the realizations of formulas. -* The `FirstOrder.Language.elementaryDiagram` of a structure is the set of all sentences with +- The `FirstOrder.Language.elementaryDiagram` of a structure is the set of all sentences with parameters that the structure satisfies. -* `FirstOrder.Language.ElementaryEmbedding.ofModelsElementaryDiagram` is the canonical -elementary embedding of any structure into a model of its elementary diagram. +- `FirstOrder.Language.ElementaryEmbedding.ofModelsElementaryDiagram` is the canonical + elementary embedding of any structure into a model of its elementary diagram. ## Main Results -* The Tarski-Vaught Test for embeddings: `FirstOrder.Language.Embedding.isElementary_of_exists` -gives a simple criterion for an embedding to be elementary. + +- The Tarski-Vaught Test for embeddings: `FirstOrder.Language.Embedding.isElementary_of_exists` + gives a simple criterion for an embedding to be elementary. -/ diff --git a/Mathlib/ModelTheory/ElementarySubstructures.lean b/Mathlib/ModelTheory/ElementarySubstructures.lean index 4fcc1c53ad2e9..e0a575f7c8a3d 100644 --- a/Mathlib/ModelTheory/ElementarySubstructures.lean +++ b/Mathlib/ModelTheory/ElementarySubstructures.lean @@ -9,11 +9,13 @@ import Mathlib.ModelTheory.ElementaryMaps # Elementary Substructures ## Main Definitions -* A `FirstOrder.Language.ElementarySubstructure` is a substructure where the realization of each + +- A `FirstOrder.Language.ElementarySubstructure` is a substructure where the realization of each formula agrees with the realization in the larger model. ## Main Results -* The Tarski-Vaught Test for substructures: + +- The Tarski-Vaught Test for substructures: `FirstOrder.Language.Substructure.isElementary_of_exists` gives a simple criterion for a substructure to be elementary. -/ diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index d6e14f6927b01..21a23619eb838 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -8,23 +8,27 @@ import Mathlib.Logic.Small.List import Mathlib.ModelTheory.Syntax import Mathlib.SetTheory.Cardinal.Ordinal -/-! # Encodings and Cardinality of First-Order Syntax +/-! +# Encodings and Cardinality of First-Order Syntax ## Main Definitions -* `FirstOrder.Language.Term.encoding` encodes terms as lists. -* `FirstOrder.Language.BoundedFormula.encoding` encodes bounded formulas as lists. + +- `FirstOrder.Language.Term.encoding` encodes terms as lists. +- `FirstOrder.Language.BoundedFormula.encoding` encodes bounded formulas as lists. ## Main Results -* `FirstOrder.Language.Term.card_le` shows that the number of terms in `L.Term α` is at most -`max ℵ₀ # (α ⊕ Σ i, L.Functions i)`. -* `FirstOrder.Language.BoundedFormula.card_le` shows that the number of bounded formulas in -`Σ n, L.BoundedFormula α n` is at most -`max ℵ₀ (Cardinal.lift.{max u v} #α + Cardinal.lift.{u'} L.card)`. + +- `FirstOrder.Language.Term.card_le` shows that the number of terms in `L.Term α` is at most + `max ℵ₀ # (α ⊕ Σ i, L.Functions i)`. +- `FirstOrder.Language.BoundedFormula.card_le` shows that the number of bounded formulas in + `Σ n, L.BoundedFormula α n` is at most + `max ℵ₀ (Cardinal.lift.{max u v} #α + Cardinal.lift.{u'} L.card)`. ## TODO -* `Primcodable` instances for terms and formulas, based on the `encoding`s -* Computability facts about term and formula operations, to set up a computability approach to -incompleteness + +- `Primcodable` instances for terms and formulas, based on the `encoding`s +- Computability facts about term and formula operations, to set up a computability approach to + incompleteness -/ diff --git a/Mathlib/ModelTheory/FinitelyGenerated.lean b/Mathlib/ModelTheory/FinitelyGenerated.lean index 58f48c2b9aab2..df79d2ed1b59d 100644 --- a/Mathlib/ModelTheory/FinitelyGenerated.lean +++ b/Mathlib/ModelTheory/FinitelyGenerated.lean @@ -7,17 +7,20 @@ import Mathlib.ModelTheory.Substructures /-! # Finitely Generated First-Order Structures + This file defines what it means for a first-order (sub)structure to be finitely or countably generated, similarly to other finitely-generated objects in the algebra library. ## Main Definitions -* `FirstOrder.Language.Substructure.FG` indicates that a substructure is finitely generated. -* `FirstOrder.Language.Structure.FG` indicates that a structure is finitely generated. -* `FirstOrder.Language.Substructure.CG` indicates that a substructure is countably generated. -* `FirstOrder.Language.Structure.CG` indicates that a structure is countably generated. + +- `FirstOrder.Language.Substructure.FG` indicates that a substructure is finitely generated. +- `FirstOrder.Language.Structure.FG` indicates that a structure is finitely generated. +- `FirstOrder.Language.Substructure.CG` indicates that a substructure is countably generated. +- `FirstOrder.Language.Structure.CG` indicates that a structure is countably generated. ## TODO + Develop a more unified definition of finite generation using the theory of closure operators, or use this definition of finite generation to define the others. diff --git a/Mathlib/ModelTheory/Fraisse.lean b/Mathlib/ModelTheory/Fraisse.lean index 58439360a1313..599d3e5bafc19 100644 --- a/Mathlib/ModelTheory/Fraisse.lean +++ b/Mathlib/ModelTheory/Fraisse.lean @@ -18,39 +18,44 @@ ultrahomogeneous structures. To each is associated a unique (up to nonunique iso Fraïssé limit - the countable ultrahomogeneous structure with that age. ## Main Definitions -* `FirstOrder.Language.age` is the class of finitely-generated structures that embed into a -particular structure. -* A class `K` is `FirstOrder.Language.Hereditary` when all finitely-generated -structures that embed into structures in `K` are also in `K`. -* A class `K` has `FirstOrder.Language.JointEmbedding` when for every `M`, `N` in -`K`, there is another structure in `K` into which both `M` and `N` embed. -* A class `K` has `FirstOrder.Language.Amalgamation` when for any pair of embeddings -of a structure `M` in `K` into other structures in `K`, those two structures can be embedded into a -fourth structure in `K` such that the resulting square of embeddings commutes. -* `FirstOrder.Language.IsFraisse` indicates that a class is nonempty, isomorphism-invariant, -essentially countable, and satisfies the hereditary, joint embedding, and amalgamation properties. -* `FirstOrder.Language.IsFraisseLimit` indicates that a structure is a Fraïssé limit for a given -class. + +- `FirstOrder.Language.age` is the class of finitely-generated structures that embed into a + particular structure. +- A class `K` is `FirstOrder.Language.Hereditary` when all finitely-generated + structures that embed into structures in `K` are also in `K`. +- A class `K` has `FirstOrder.Language.JointEmbedding` when for every `M`, `N` in + `K`, there is another structure in `K` into which both `M` and `N` embed. +- A class `K` has `FirstOrder.Language.Amalgamation` when for any pair of embeddings + of a structure `M` in `K` into other structures in `K`, those two structures can be embedded into + a fourth structure in `K` such that the resulting square of embeddings commutes. +- `FirstOrder.Language.IsFraisse` indicates that a class is nonempty, isomorphism-invariant, + essentially countable, and satisfies the hereditary, joint embedding, and amalgamation properties. +- `FirstOrder.Language.IsFraisseLimit` indicates that a structure is a Fraïssé limit for a given + class. ## Main Results -* We show that the age of any structure is isomorphism-invariant and satisfies the hereditary and -joint-embedding properties. -* `FirstOrder.Language.age.countable_quotient` shows that the age of any countable structure is -essentially countable. -* `FirstOrder.Language.exists_countable_is_age_of_iff` gives necessary and sufficient conditions -for a class to be the age of a countable structure in a language with countably many functions. + +- We show that the age of any structure is isomorphism-invariant and satisfies the hereditary and + joint-embedding properties. +- `FirstOrder.Language.age.countable_quotient` shows that the age of any countable structure is + essentially countable. +- `FirstOrder.Language.exists_countable_is_age_of_iff` gives necessary and sufficient conditions + for a class to be the age of a countable structure in a language with countably many functions. ## Implementation Notes -* Classes of structures are formalized with `Set (Bundled L.Structure)`. -* Some results pertain to countable limit structures, others to countably-generated limit -structures. In the case of a language with countably many function symbols, these are equivalent. + +- Classes of structures are formalized with `Set (Bundled L.Structure)`. +- Some results pertain to countable limit structures, others to countably-generated limit + structures. In the case of a language with countably many function symbols, these are equivalent. ## References + - [W. Hodges, *A Shorter Model Theory*][Hodges97] - [K. Tent, M. Ziegler, *A Course in Model Theory*][Tent_Ziegler] ## TODO -* Show existence and uniqueness of Fraïssé limits + +- Show existence and uniqueness of Fraïssé limits -/ diff --git a/Mathlib/ModelTheory/Graph.lean b/Mathlib/ModelTheory/Graph.lean index 030591ae88cec..06519b49d2e04 100644 --- a/Mathlib/ModelTheory/Graph.lean +++ b/Mathlib/ModelTheory/Graph.lean @@ -8,15 +8,17 @@ import Mathlib.Combinatorics.SimpleGraph.Basic /-! # First-Order Structures in Graph Theory + This file defines first-order languages, structures, and theories in graph theory. ## Main Definitions -* `FirstOrder.Language.graph` is the language consisting of a single relation representing -adjacency. -* `SimpleGraph.structure` is the first-order structure corresponding to a given simple graph. -* `FirstOrder.Language.Theory.simpleGraph` is the theory of simple graphs. -* `FirstOrder.Language.simpleGraphOfStructure` gives the simple graph corresponding to a model -of the theory of simple graphs. + +- `FirstOrder.Language.graph` is the language consisting of a single relation representing + adjacency. +- `SimpleGraph.structure` is the first-order structure corresponding to a given simple graph. +- `FirstOrder.Language.Theory.simpleGraph` is the theory of simple graphs. +- `FirstOrder.Language.simpleGraphOfStructure` gives the simple graph corresponding to a model + of the theory of simple graphs. -/ diff --git a/Mathlib/ModelTheory/LanguageMap.lean b/Mathlib/ModelTheory/LanguageMap.lean index 9fe80219394c2..a70065ee3ba7e 100644 --- a/Mathlib/ModelTheory/LanguageMap.lean +++ b/Mathlib/ModelTheory/LanguageMap.lean @@ -7,24 +7,27 @@ import Mathlib.ModelTheory.Basic /-! # Language Maps + Maps between first-order languages in the style of the [Flypitch project](https://flypitch.github.io/), as well as several important maps between structures. ## Main Definitions -* A `FirstOrder.Language.LHom`, denoted `L →ᴸ L'`, is a map between languages, sending the symbols + +- A `FirstOrder.Language.LHom`, denoted `L →ᴸ L'`, is a map between languages, sending the symbols of one to symbols of the same kind and arity in the other. -* A `FirstOrder.Language.LEquiv`, denoted `L ≃ᴸ L'`, is an invertible language homomorphism. -* `FirstOrder.Language.withConstants` is defined so that if `M` is an `L.Structure` and +- A `FirstOrder.Language.LEquiv`, denoted `L ≃ᴸ L'`, is an invertible language homomorphism. +- `FirstOrder.Language.withConstants` is defined so that if `M` is an `L.Structure` and `A : Set M`, `L.withConstants A`, denoted `L[[A]]`, is a language which adds constant symbols for elements of `A` to `L`. ## References + For the Flypitch project: - [J. Han, F. van Doorn, *A formal proof of the independence of the continuum hypothesis*] -[flypitch_cpp] + [flypitch_cpp] - [J. Han, F. van Doorn, *A formalization of forcing and the unprovability of -the continuum hypothesis*][flypitch_itp] + the continuum hypothesis*][flypitch_itp] -/ diff --git a/Mathlib/ModelTheory/Order.lean b/Mathlib/ModelTheory/Order.lean index ac84d3d2c499f..a5c6f4fce9ec4 100644 --- a/Mathlib/ModelTheory/Order.lean +++ b/Mathlib/ModelTheory/Order.lean @@ -7,24 +7,26 @@ import Mathlib.ModelTheory.Semantics /-! # Ordered First-Ordered Structures + This file defines ordered first-order languages and structures, as well as their theories. ## Main Definitions -* `FirstOrder.Language.order` is the language consisting of a single relation representing `≤`. -* `FirstOrder.Language.orderStructure` is the structure on an ordered type, assigning the symbol -representing `≤` to the actual relation `≤`. -* `FirstOrder.Language.IsOrdered` points out a specific symbol in a language as representing `≤`. -* `FirstOrder.Language.OrderedStructure` indicates that the `≤` symbol in an ordered language -is interpreted as the actual relation `≤` in a particular structure. -* `FirstOrder.Language.linearOrderTheory` and similar define the theories of preorders, -partial orders, and linear orders. -* `FirstOrder.Language.dlo` defines the theory of dense linear orders without endpoints, a -particularly useful example in model theory. + +- `FirstOrder.Language.order` is the language consisting of a single relation representing `≤`. +- `FirstOrder.Language.orderStructure` is the structure on an ordered type, assigning the symbol + representing `≤` to the actual relation `≤`. +- `FirstOrder.Language.IsOrdered` points out a specific symbol in a language as representing `≤`. +- `FirstOrder.Language.OrderedStructure` indicates that the `≤` symbol in an ordered language + is interpreted as the actual relation `≤` in a particular structure. +- `FirstOrder.Language.linearOrderTheory` and similar define the theories of preorders, + partial orders, and linear orders. +- `FirstOrder.Language.dlo` defines the theory of dense linear orders without endpoints, a + particularly useful example in model theory. ## Main Results -* `PartialOrder`s model the theory of partial orders, `LinearOrder`s model the theory of -linear orders, and dense linear orders without endpoints model `Language.dlo`. +- `PartialOrder`s model the theory of partial orders, `LinearOrder`s model the theory of + linear orders, and dense linear orders without endpoints model `Language.dlo`. -/ diff --git a/Mathlib/ModelTheory/Quotients.lean b/Mathlib/ModelTheory/Quotients.lean index 9236306e1976f..798514c512fb1 100644 --- a/Mathlib/ModelTheory/Quotients.lean +++ b/Mathlib/ModelTheory/Quotients.lean @@ -8,14 +8,15 @@ import Mathlib.ModelTheory.Semantics /-! # Quotients of First-Order Structures + This file defines prestructures and quotients of first-order structures. ## Main Definitions -* If `s` is a setoid (equivalence relation) on `M`, a `FirstOrder.Language.Prestructure s` is the -data for a first-order structure on `M` that will still be a structure when modded out by `s`. -* The structure `FirstOrder.Language.quotientStructure s` is the resulting structure on -`Quotient s`. +- If `s` is a setoid (equivalence relation) on `M`, a `FirstOrder.Language.Prestructure s` is the + data for a first-order structure on `M` that will still be a structure when modded out by `s`. +- The structure `FirstOrder.Language.quotientStructure s` is the resulting structure on + `Quotient s`. -/ diff --git a/Mathlib/ModelTheory/Satisfiability.lean b/Mathlib/ModelTheory/Satisfiability.lean index 3f2c0e9f416fb..3185e8db3b62c 100644 --- a/Mathlib/ModelTheory/Satisfiability.lean +++ b/Mathlib/ModelTheory/Satisfiability.lean @@ -9,34 +9,37 @@ import Mathlib.ModelTheory.Skolem /-! # First-Order Satisfiability + This file deals with the satisfiability of first-order theories, as well as equivalence over them. ## Main Definitions -* `FirstOrder.Language.Theory.IsSatisfiable`: `T.IsSatisfiable` indicates that `T` has a nonempty -model. -* `FirstOrder.Language.Theory.IsFinitelySatisfiable`: `T.IsFinitelySatisfiable` indicates that -every finite subset of `T` is satisfiable. -* `FirstOrder.Language.Theory.IsComplete`: `T.IsComplete` indicates that `T` is satisfiable and -models each sentence or its negation. -* `FirstOrder.Language.Theory.SemanticallyEquivalent`: `T.SemanticallyEquivalent φ ψ` indicates -that `φ` and `ψ` are equivalent formulas or sentences in models of `T`. -* `Cardinal.Categorical`: A theory is `κ`-categorical if all models of size `κ` are isomorphic. + +- `FirstOrder.Language.Theory.IsSatisfiable`: `T.IsSatisfiable` indicates that `T` has a nonempty + model. +- `FirstOrder.Language.Theory.IsFinitelySatisfiable`: `T.IsFinitelySatisfiable` indicates that + every finite subset of `T` is satisfiable. +- `FirstOrder.Language.Theory.IsComplete`: `T.IsComplete` indicates that `T` is satisfiable and + models each sentence or its negation. +- `FirstOrder.Language.Theory.SemanticallyEquivalent`: `T.SemanticallyEquivalent φ ψ` indicates + that `φ` and `ψ` are equivalent formulas or sentences in models of `T`. +- `Cardinal.Categorical`: A theory is `κ`-categorical if all models of size `κ` are isomorphic. ## Main Results -* The Compactness Theorem, `FirstOrder.Language.Theory.isSatisfiable_iff_isFinitelySatisfiable`, -shows that a theory is satisfiable iff it is finitely satisfiable. -* `FirstOrder.Language.completeTheory.isComplete`: The complete theory of a structure is -complete. -* `FirstOrder.Language.Theory.exists_large_model_of_infinite_model` shows that any theory with an -infinite model has arbitrarily large models. -* `FirstOrder.Language.Theory.exists_elementaryEmbedding_card_eq`: The Upward Löwenheim–Skolem -Theorem: If `κ` is a cardinal greater than the cardinalities of `L` and an infinite `L`-structure -`M`, then `M` has an elementary extension of cardinality `κ`. + +- The Compactness Theorem, `FirstOrder.Language.Theory.isSatisfiable_iff_isFinitelySatisfiable`, + shows that a theory is satisfiable iff it is finitely satisfiable. +- `FirstOrder.Language.completeTheory.isComplete`: The complete theory of a structure is + complete. +- `FirstOrder.Language.Theory.exists_large_model_of_infinite_model` shows that any theory with an + infinite model has arbitrarily large models. +- `FirstOrder.Language.Theory.exists_elementaryEmbedding_card_eq`: The Upward Löwenheim–Skolem + Theorem: If `κ` is a cardinal greater than the cardinalities of `L` and an infinite `L`-structure + `M`, then `M` has an elementary extension of cardinality `κ`. ## Implementation Details -* Satisfiability of an `L.Theory` `T` is defined in the minimal universe containing all the symbols -of `L`. By Löwenheim-Skolem, this is equivalent to satisfiability in any universe. +- Satisfiability of an `L.Theory` `T` is defined in the minimal universe containing all the symbols + of `L`. By Löwenheim-Skolem, this is equivalent to satisfiability in any universe. -/ diff --git a/Mathlib/ModelTheory/Semantics.lean b/Mathlib/ModelTheory/Semantics.lean index 9e968e6b6137f..41d4021996c56 100644 --- a/Mathlib/ModelTheory/Semantics.lean +++ b/Mathlib/ModelTheory/Semantics.lean @@ -9,40 +9,44 @@ import Mathlib.Data.List.ProdSigma /-! # Basics on First-Order Semantics + This file defines the interpretations of first-order terms, formulas, sentences, and theories in a style inspired by the [Flypitch project](https://flypitch.github.io/). ## Main Definitions -* `FirstOrder.Language.Term.realize` is defined so that `t.realize v` is the term `t` evaluated at -variables `v`. -* `FirstOrder.Language.BoundedFormula.Realize` is defined so that `φ.Realize v xs` is the bounded -formula `φ` evaluated at tuples of variables `v` and `xs`. -* `FirstOrder.Language.Formula.Realize` is defined so that `φ.Realize v` is the formula `φ` -evaluated at variables `v`. -* `FirstOrder.Language.Sentence.Realize` is defined so that `φ.Realize M` is the sentence `φ` -evaluated in the structure `M`. Also denoted `M ⊨ φ`. -* `FirstOrder.Language.Theory.Model` is defined so that `T.Model M` is true if and only if every -sentence of `T` is realized in `M`. Also denoted `T ⊨ φ`. + +- `FirstOrder.Language.Term.realize` is defined so that `t.realize v` is the term `t` evaluated at + variables `v`. +- `FirstOrder.Language.BoundedFormula.Realize` is defined so that `φ.Realize v xs` is the bounded + formula `φ` evaluated at tuples of variables `v` and `xs`. +- `FirstOrder.Language.Formula.Realize` is defined so that `φ.Realize v` is the formula `φ` + evaluated at variables `v`. +- `FirstOrder.Language.Sentence.Realize` is defined so that `φ.Realize M` is the sentence `φ` + evaluated in the structure `M`. Also denoted `M ⊨ φ`. +- `FirstOrder.Language.Theory.Model` is defined so that `T.Model M` is true if and only if every + sentence of `T` is realized in `M`. Also denoted `T ⊨ φ`. ## Main Results -* Several results in this file show that syntactic constructions such as `relabel`, `castLE`, -`liftAt`, `subst`, and the actions of language maps commute with realization of terms, formulas, -sentences, and theories. + +- Several results in this file show that syntactic constructions such as `relabel`, `castLE`, + `liftAt`, `subst`, and the actions of language maps commute with realization of terms, formulas, + sentences, and theories. ## Implementation Notes -* Formulas use a modified version of de Bruijn variables. Specifically, a `L.BoundedFormula α n` -is a formula with some variables indexed by a type `α`, which cannot be quantified over, and some -indexed by `Fin n`, which can. For any `φ : L.BoundedFormula α (n + 1)`, we define the formula -`∀' φ : L.BoundedFormula α n` by universally quantifying over the variable indexed by -`n : Fin (n + 1)`. + +- Formulas use a modified version of de Bruijn variables. Specifically, a `L.BoundedFormula α n` + is a formula with some variables indexed by a type `α`, which cannot be quantified over, and some + indexed by `Fin n`, which can. For any `φ : L.BoundedFormula α (n + 1)`, we define the formula + `∀' φ : L.BoundedFormula α n` by universally quantifying over the variable indexed by + `n : Fin (n + 1)`. ## References + For the Flypitch project: - [J. Han, F. van Doorn, *A formal proof of the independence of the continuum hypothesis*] -[flypitch_cpp] + [flypitch_cpp] - [J. Han, F. van Doorn, *A formalization of forcing and the unprovability of -the continuum hypothesis*][flypitch_itp] - + the continuum hypothesis*][flypitch_itp] -/ diff --git a/Mathlib/ModelTheory/Skolem.lean b/Mathlib/ModelTheory/Skolem.lean index db877645459f1..eefcbc6d811c1 100644 --- a/Mathlib/ModelTheory/Skolem.lean +++ b/Mathlib/ModelTheory/Skolem.lean @@ -9,17 +9,19 @@ import Mathlib.ModelTheory.ElementarySubstructures # Skolem Functions and Downward Löwenheim–Skolem ## Main Definitions -* `FirstOrder.Language.skolem₁` is a language consisting of Skolem functions for another language. + +- `FirstOrder.Language.skolem₁` is a language consisting of Skolem functions for another language. ## Main Results -* `FirstOrder.Language.exists_elementarySubstructure_card_eq` is the Downward Löwenheim–Skolem + +- `FirstOrder.Language.exists_elementarySubstructure_card_eq` is the Downward Löwenheim–Skolem theorem: If `s` is a set in an `L`-structure `M` and `κ` an infinite cardinal such that `max (#s, L.card) ≤ κ` and `κ ≤ # M`, then `M` has an elementary substructure containing `s` of cardinality `κ`. ## TODO -* Use `skolem₁` recursively to construct an actual Skolemization of a language. +- Use `skolem₁` recursively to construct an actual Skolemization of a language. -/ diff --git a/Mathlib/ModelTheory/Substructures.lean b/Mathlib/ModelTheory/Substructures.lean index 89a0d7225726a..1d96f21064925 100644 --- a/Mathlib/ModelTheory/Substructures.lean +++ b/Mathlib/ModelTheory/Substructures.lean @@ -9,29 +9,31 @@ import Mathlib.ModelTheory.Encoding /-! # First-Order Substructures + This file defines substructures of first-order structures in a similar manner to the various substructures appearing in the algebra library. ## Main Definitions -* A `FirstOrder.Language.Substructure` is defined so that `L.Substructure M` is the type of all -substructures of the `L`-structure `M`. -* `FirstOrder.Language.Substructure.closure` is defined so that if `s : Set M`, `closure L s` is -the least substructure of `M` containing `s`. -* `FirstOrder.Language.Substructure.comap` is defined so that `s.comap f` is the preimage of the -substructure `s` under the homomorphism `f`, as a substructure. -* `FirstOrder.Language.Substructure.map` is defined so that `s.map f` is the image of the -substructure `s` under the homomorphism `f`, as a substructure. -* `FirstOrder.Language.Hom.range` is defined so that `f.range` is the range of the -homomorphism `f`, as a substructure. -* `FirstOrder.Language.Hom.domRestrict` and `FirstOrder.Language.Hom.codRestrict` restrict -the domain and codomain respectively of first-order homomorphisms to substructures. -* `FirstOrder.Language.Embedding.domRestrict` and `FirstOrder.Language.Embedding.codRestrict` -restrict the domain and codomain respectively of first-order embeddings to substructures. -* `FirstOrder.Language.Substructure.inclusion` is the inclusion embedding between substructures. + +- A `FirstOrder.Language.Substructure` is defined so that `L.Substructure M` is the type of all + substructures of the `L`-structure `M`. +- `FirstOrder.Language.Substructure.closure` is defined so that if `s : Set M`, `closure L s` is + the least substructure of `M` containing `s`. +- `FirstOrder.Language.Substructure.comap` is defined so that `s.comap f` is the preimage of the + substructure `s` under the homomorphism `f`, as a substructure. +- `FirstOrder.Language.Substructure.map` is defined so that `s.map f` is the image of the + substructure `s` under the homomorphism `f`, as a substructure. +- `FirstOrder.Language.Hom.range` is defined so that `f.range` is the range of the + homomorphism `f`, as a substructure. +- `FirstOrder.Language.Hom.domRestrict` and `FirstOrder.Language.Hom.codRestrict` restrict + the domain and codomain respectively of first-order homomorphisms to substructures. +- `FirstOrder.Language.Embedding.domRestrict` and `FirstOrder.Language.Embedding.codRestrict` + restrict the domain and codomain respectively of first-order embeddings to substructures. +- `FirstOrder.Language.Substructure.inclusion` is the inclusion embedding between substructures. ## Main Results -* `L.Substructure M` forms a `CompleteLattice`. +- `L.Substructure M` forms a `CompleteLattice`. -/ diff --git a/Mathlib/ModelTheory/Syntax.lean b/Mathlib/ModelTheory/Syntax.lean index 2c5400fe74973..d77cfb5bd2104 100644 --- a/Mathlib/ModelTheory/Syntax.lean +++ b/Mathlib/ModelTheory/Syntax.lean @@ -9,45 +9,48 @@ import Mathlib.ModelTheory.LanguageMap /-! # Basics on First-Order Syntax + This file defines first-order terms, formulas, sentences, and theories in a style inspired by the [Flypitch project](https://flypitch.github.io/). ## Main Definitions -* A `FirstOrder.Language.Term` is defined so that `L.Term α` is the type of `L`-terms with free + +- A `FirstOrder.Language.Term` is defined so that `L.Term α` is the type of `L`-terms with free variables indexed by `α`. -* A `FirstOrder.Language.Formula` is defined so that `L.Formula α` is the type of `L`-formulas with +- A `FirstOrder.Language.Formula` is defined so that `L.Formula α` is the type of `L`-formulas with free variables indexed by `α`. -* A `FirstOrder.Language.Sentence` is a formula with no free variables. -* A `FirstOrder.Language.Theory` is a set of sentences. -* The variables of terms and formulas can be relabelled with `FirstOrder.Language.Term.relabel`, -`FirstOrder.Language.BoundedFormula.relabel`, and `FirstOrder.Language.Formula.relabel`. -* Given an operation on terms and an operation on relations, +- A `FirstOrder.Language.Sentence` is a formula with no free variables. +- A `FirstOrder.Language.Theory` is a set of sentences. +- The variables of terms and formulas can be relabelled with `FirstOrder.Language.Term.relabel`, + `FirstOrder.Language.BoundedFormula.relabel`, and `FirstOrder.Language.Formula.relabel`. +- Given an operation on terms and an operation on relations, `FirstOrder.Language.BoundedFormula.mapTermRel` gives an operation on formulas. -* `FirstOrder.Language.BoundedFormula.castLE` adds more `Fin`-indexed variables. -* `FirstOrder.Language.BoundedFormula.liftAt` raises the indexes of the `Fin`-indexed variables -above a particular index. -* `FirstOrder.Language.Term.subst` and `FirstOrder.Language.BoundedFormula.subst` substitute -variables with given terms. -* Language maps can act on syntactic objects with functions such as -`FirstOrder.Language.LHom.onFormula`. -* `FirstOrder.Language.Term.constantsVarsEquiv` and -`FirstOrder.Language.BoundedFormula.constantsVarsEquiv` switch terms and formulas between having -constants in the language and having extra variables indexed by the same type. +- `FirstOrder.Language.BoundedFormula.castLE` adds more `Fin`-indexed variables. +- `FirstOrder.Language.BoundedFormula.liftAt` raises the indexes of the `Fin`-indexed variables + above a particular index. +- `FirstOrder.Language.Term.subst` and `FirstOrder.Language.BoundedFormula.subst` substitute + variables with given terms. +- Language maps can act on syntactic objects with functions such as + `FirstOrder.Language.LHom.onFormula`. +- `FirstOrder.Language.Term.constantsVarsEquiv` and + `FirstOrder.Language.BoundedFormula.constantsVarsEquiv` switch terms and formulas between having + constants in the language and having extra variables indexed by the same type. ## Implementation Notes -* Formulas use a modified version of de Bruijn variables. Specifically, a `L.BoundedFormula α n` -is a formula with some variables indexed by a type `α`, which cannot be quantified over, and some -indexed by `Fin n`, which can. For any `φ : L.BoundedFormula α (n + 1)`, we define the formula -`∀' φ : L.BoundedFormula α n` by universally quantifying over the variable indexed by -`n : Fin (n + 1)`. + +- Formulas use a modified version of de Bruijn variables. Specifically, a `L.BoundedFormula α n` + is a formula with some variables indexed by a type `α`, which cannot be quantified over, and some + indexed by `Fin n`, which can. For any `φ : L.BoundedFormula α (n + 1)`, we define the formula + `∀' φ : L.BoundedFormula α n` by universally quantifying over the variable indexed by + `n : Fin (n + 1)`. ## References + For the Flypitch project: - [J. Han, F. van Doorn, *A formal proof of the independence of the continuum hypothesis*] -[flypitch_cpp] + [flypitch_cpp] - [J. Han, F. van Doorn, *A formalization of forcing and the unprovability of -the continuum hypothesis*][flypitch_itp] - + the continuum hypothesis*][flypitch_itp] -/ diff --git a/Mathlib/ModelTheory/Types.lean b/Mathlib/ModelTheory/Types.lean index de6f5b035eb54..b88070ddf1d06 100644 --- a/Mathlib/ModelTheory/Types.lean +++ b/Mathlib/ModelTheory/Types.lean @@ -7,29 +7,33 @@ import Mathlib.ModelTheory.Satisfiability /-! # Type Spaces + This file defines the space of complete types over a first-order theory. (Note that types in model theory are different from types in type theory.) ## Main Definitions -* `FirstOrder.Language.Theory.CompleteType`: + +- `FirstOrder.Language.Theory.CompleteType`: `T.CompleteType α` consists of complete types over the theory `T` with variables `α`. -* `FirstOrder.Language.Theory.typeOf` is the type of a given tuple. -* `FirstOrder.Language.Theory.realizedTypes`: `T.realizedTypes M α` is the set of +- `FirstOrder.Language.Theory.typeOf` is the type of a given tuple. +- `FirstOrder.Language.Theory.realizedTypes`: `T.realizedTypes M α` is the set of types in `T.CompleteType α` that are realized in `M` - that is, the type of some tuple in `M`. ## Main Results -* `FirstOrder.Language.Theory.CompleteType.nonempty_iff`: + +- `FirstOrder.Language.Theory.CompleteType.nonempty_iff`: The space `T.CompleteType α` is nonempty exactly when `T` is satisfiable. -* `FirstOrder.Language.Theory.CompleteType.exists_modelType_is_realized_in`: Every type is realized -in some model. +- `FirstOrder.Language.Theory.CompleteType.exists_modelType_is_realized_in`: Every type is realized + in some model. ## Implementation Notes -* Complete types are implemented as maximal consistent theories in an expanded language. -More frequently they are described as maximal consistent sets of formulas, but this is equivalent. + +- Complete types are implemented as maximal consistent theories in an expanded language. + More frequently they are described as maximal consistent sets of formulas, but this is equivalent. ## TODO -* Connect `T.CompleteType α` to sets of formulas `L.Formula α`. +- Connect `T.CompleteType α` to sets of formulas `L.Formula α`. -/ diff --git a/Mathlib/ModelTheory/Ultraproducts.lean b/Mathlib/ModelTheory/Ultraproducts.lean index 29dfb0f79aa78..efe11fdadc8cd 100644 --- a/Mathlib/ModelTheory/Ultraproducts.lean +++ b/Mathlib/ModelTheory/Ultraproducts.lean @@ -7,19 +7,22 @@ import Mathlib.ModelTheory.Quotients import Mathlib.Order.Filter.Germ.Basic import Mathlib.Order.Filter.Ultrafilter -/-! # Ultraproducts and Łoś's Theorem +/-! +# Ultraproducts and Łoś's Theorem ## Main Definitions -* `FirstOrder.Language.Ultraproduct.Structure` is the ultraproduct structure on `Filter.Product`. + +- `FirstOrder.Language.Ultraproduct.Structure` is the ultraproduct structure on `Filter.Product`. ## Main Results -* Łoś's Theorem: `FirstOrder.Language.Ultraproduct.sentence_realize`. An ultraproduct models a -sentence `φ` if and only if the set of structures in the product that model `φ` is in the -ultrafilter. + +- Łoś's Theorem: `FirstOrder.Language.Ultraproduct.sentence_realize`. An ultraproduct models a + sentence `φ` if and only if the set of structures in the product that model `φ` is in the + ultrafilter. ## Tags -ultraproduct, Los's theorem +ultraproduct, Los's theorem -/ universe u v From 9b1d3da278df41f304e90a85ab4bd30eaa60fccb Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 4 Aug 2024 11:42:43 +0000 Subject: [PATCH 015/359] feat(WithBot): un{bot,top}_{eq,lt}_iff (#15421) this adds analogues to the existing unbot_le_iff lemmas. From the Carleson project. --- Mathlib/Order/WithBot.lean | 40 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 40 insertions(+) diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index 9720f0923febe..364cc4b359466 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -181,6 +181,18 @@ instance instTop [Top α] : Top (WithBot α) where @[simp, norm_cast] lemma coe_eq_top [Top α] {a : α} : (a : WithBot α) = ⊤ ↔ a = ⊤ := coe_eq_coe @[simp, norm_cast] lemma top_eq_coe [Top α] {a : α} : ⊤ = (a : WithBot α) ↔ ⊤ = a := coe_eq_coe +theorem unbot_eq_iff {a : WithBot α} {b : α} (h : a ≠ ⊥) : + a.unbot h = b ↔ a = b := by + induction a + · simpa using h rfl + · simp + +theorem eq_unbot_iff {a : α} {b : WithBot α} (h : b ≠ ⊥) : + a = b.unbot h ↔ a = b := by + induction b + · simpa using h rfl + · simp + section LE variable [LE α] @@ -294,6 +306,18 @@ protected theorem bot_lt_iff_ne_bot : ∀ {x : WithBot α}, ⊥ < x ↔ x ≠ | ⊥ => iff_of_false (WithBot.not_lt_bot _) <| by simp | (x : α) => by simp [bot_lt_coe] +theorem lt_unbot_iff {a : α} {b : WithBot α} (h : b ≠ ⊥) : + a < unbot b h ↔ (a : WithBot α) < b := by + induction b + · simpa [bot_lt_coe] using h rfl + · simp + +theorem unbot_lt_iff {a : WithBot α} (h : a ≠ ⊥) {b : α} : + unbot a h < b ↔ a < (b : WithBot α) := by + induction a + · simpa [bot_lt_coe] using h rfl + · simp + theorem unbot'_lt_iff {a : WithBot α} {b c : α} (h : a = ⊥ → b < c) : a.unbot' b < c ↔ a < c := by induction a @@ -759,6 +783,14 @@ instance instBot [Bot α] : Bot (WithTop α) where @[simp, norm_cast] lemma coe_eq_bot [Bot α] {a : α} : (a : WithTop α) = ⊥ ↔ a = ⊥ := coe_eq_coe @[simp, norm_cast] lemma bot_eq_coe [Bot α] {a : α} : (⊥ : WithTop α) = a ↔ ⊥ = a := coe_eq_coe +theorem untop_eq_iff {a : WithTop α} {b : α} (h : a ≠ ⊤) : + a.untop h = b ↔ a = b := + WithBot.unbot_eq_iff (α := αᵒᵈ) h + +theorem eq_untop_iff {a : α} {b : WithTop α} (h : b ≠ ⊤) : + a = b.untop h ↔ a = b := + WithBot.eq_unbot_iff (α := αᵒᵈ) h + section LE variable [LE α] @@ -879,6 +911,14 @@ theorem lt_ofDual_iff {a : WithBot α} {b : WithTop αᵒᵈ} : theorem ofDual_lt_ofDual_iff {a b : WithTop αᵒᵈ} : WithTop.ofDual a < WithTop.ofDual b ↔ b < a := Iff.rfl +theorem lt_untop_iff {a : α} {b : WithTop α} (h : b ≠ ⊤) : + a < b.untop h ↔ a < b := + WithBot.unbot_lt_iff (α := αᵒᵈ) h + +theorem untop_lt_iff {a : WithTop α} {b : α} (h : a ≠ ⊤) : + a.untop h < b ↔ a < b := + WithBot.lt_unbot_iff (α := αᵒᵈ) h + theorem lt_untop'_iff {a : WithTop α} {b c : α} (h : a = ⊤ → c < b) : c < a.untop' b ↔ c < a := WithBot.unbot'_lt_iff (α := αᵒᵈ) h From b29f5cc9688a130d501e81a0b1a62e9e668b91c1 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 4 Aug 2024 11:42:44 +0000 Subject: [PATCH 016/359] =?UTF-8?q?chore(VitaliFamily):=20rename=20`=CE=B1?= =?UTF-8?q?`=20=E2=86=92=20`X`=20(#15428)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The module docstring uses `X` while the code used `α`. --- .../MeasureTheory/Covering/VitaliFamily.lean | 132 +++++++++--------- 1 file changed, 66 insertions(+), 66 deletions(-) diff --git a/Mathlib/MeasureTheory/Covering/VitaliFamily.lean b/Mathlib/MeasureTheory/Covering/VitaliFamily.lean index 5fa24afca8715..e4c3a72dd39a7 100644 --- a/Mathlib/MeasureTheory/Covering/VitaliFamily.lean +++ b/Mathlib/MeasureTheory/Covering/VitaliFamily.lean @@ -49,7 +49,7 @@ Vitali relations there) open MeasureTheory Metric Set Filter TopologicalSpace MeasureTheory.Measure open scoped Topology -variable {α : Type*} [MetricSpace α] +variable {X : Type*} [MetricSpace X] /-- On a metric space `X` with a measure `μ`, consider for each `x : X` a family of measurable sets with nonempty interiors, called `setsAt x`. This family is a Vitali family if it satisfies the @@ -63,31 +63,31 @@ differentiations of measure that apply in both contexts. -/ -- Porting note(#5171): this linter isn't ported yet. -- @[nolint has_nonempty_instance] -structure VitaliFamily {m : MeasurableSpace α} (μ : Measure α) where +structure VitaliFamily {m : MeasurableSpace X} (μ : Measure X) where /-- Sets of the family "centered" at a given point. -/ - setsAt : α → Set (Set α) + setsAt : X → Set (Set X) /-- All sets of the family are measurable. -/ - measurableSet : ∀ x : α, ∀ s ∈ setsAt x, MeasurableSet s + measurableSet : ∀ x : X, ∀ s ∈ setsAt x, MeasurableSet s /-- All sets of the family have nonempty interior. -/ - nonempty_interior : ∀ x : α, ∀ s ∈ setsAt x, (interior s).Nonempty + nonempty_interior : ∀ x : X, ∀ s ∈ setsAt x, (interior s).Nonempty /-- For any closed ball around `x`, there exists a set of the family contained in this ball. -/ - nontrivial : ∀ (x : α), ∀ ε > (0 : ℝ), ∃ s ∈ setsAt x, s ⊆ closedBall x ε + nontrivial : ∀ (x : X), ∀ ε > (0 : ℝ), ∃ s ∈ setsAt x, s ⊆ closedBall x ε /-- Consider a (possibly non-measurable) set `s`, and for any `x` in `s` a subfamily `f x` of `setsAt x` containing sets of arbitrarily small diameter. Then one can extract a disjoint subfamily covering almost all `s`. -/ - covering : ∀ (s : Set α) (f : α → Set (Set α)), - (∀ x ∈ s, f x ⊆ setsAt x) → (∀ x ∈ s, ∀ ε > (0 : ℝ), ∃ a ∈ f x, a ⊆ closedBall x ε) → - ∃ t : Set (α × Set α), (∀ p ∈ t, p.1 ∈ s) ∧ (t.PairwiseDisjoint fun p ↦ p.2) ∧ + covering : ∀ (s : Set X) (f : X → Set (Set X)), + (∀ x ∈ s, f x ⊆ setsAt x) → (∀ x ∈ s, ∀ ε > (0 : ℝ), ∃ t ∈ f x, t ⊆ closedBall x ε) → + ∃ t : Set (X × Set X), (∀ p ∈ t, p.1 ∈ s) ∧ (t.PairwiseDisjoint fun p ↦ p.2) ∧ (∀ p ∈ t, p.2 ∈ f p.1) ∧ μ (s \ ⋃ p ∈ t, p.2) = 0 namespace VitaliFamily -variable {m0 : MeasurableSpace α} {μ : Measure α} +variable {m0 : MeasurableSpace X} {μ : Measure X} /-- A Vitali family for a measure `μ` is also a Vitali family for any measure absolutely continuous with respect to `μ`. -/ -def mono (v : VitaliFamily μ) (ν : Measure α) (hν : ν ≪ μ) : VitaliFamily ν where +def mono (v : VitaliFamily μ) (ν : Measure X) (hν : ν ≪ μ) : VitaliFamily ν where __ := v covering s f h h' := let ⟨t, ts, disj, mem_f, hμ⟩ := v.covering s f h h' @@ -97,34 +97,34 @@ def mono (v : VitaliFamily μ) (ν : Measure α) (hν : ν ≪ μ) : VitaliFamil every point `x` in `s` belongs to arbitrarily small sets in `v.setsAt x ∩ f x`. This is precisely the subfamilies for which the Vitali family definition ensures that one can extract a disjoint covering of almost all `s`. -/ -def FineSubfamilyOn (v : VitaliFamily μ) (f : α → Set (Set α)) (s : Set α) : Prop := - ∀ x ∈ s, ∀ ε > 0, ∃ a ∈ v.setsAt x ∩ f x, a ⊆ closedBall x ε +def FineSubfamilyOn (v : VitaliFamily μ) (f : X → Set (Set X)) (s : Set X) : Prop := + ∀ x ∈ s, ∀ ε > 0, ∃ t ∈ v.setsAt x ∩ f x, t ⊆ closedBall x ε namespace FineSubfamilyOn -variable {v : VitaliFamily μ} {f : α → Set (Set α)} {s : Set α} (h : v.FineSubfamilyOn f s) +variable {v : VitaliFamily μ} {f : X → Set (Set X)} {s : Set X} (h : v.FineSubfamilyOn f s) theorem exists_disjoint_covering_ae : - ∃ t : Set (α × Set α), - (∀ p : α × Set α, p ∈ t → p.1 ∈ s) ∧ + ∃ t : Set (X × Set X), + (∀ p : X × Set X, p ∈ t → p.1 ∈ s) ∧ (t.PairwiseDisjoint fun p => p.2) ∧ - (∀ p : α × Set α, p ∈ t → p.2 ∈ v.setsAt p.1 ∩ f p.1) ∧ - μ (s \ ⋃ (p : α × Set α) (_ : p ∈ t), p.2) = 0 := + (∀ p : X × Set X, p ∈ t → p.2 ∈ v.setsAt p.1 ∩ f p.1) ∧ + μ (s \ ⋃ (p : X × Set X) (_ : p ∈ t), p.2) = 0 := v.covering s (fun x => v.setsAt x ∩ f x) (fun _ _ => inter_subset_left) h /-- Given `h : v.FineSubfamilyOn f s`, then `h.index` is a set parametrizing a disjoint covering of almost every `s`. -/ -protected def index : Set (α × Set α) := +protected def index : Set (X × Set X) := h.exists_disjoint_covering_ae.choose -- Porting note: Needed to add `(_h : FineSubfamilyOn v f s)` /-- Given `h : v.FineSubfamilyOn f s`, then `h.covering p` is a set in the family, for `p ∈ h.index`, such that these sets form a disjoint covering of almost every `s`. -/ @[nolint unusedArguments] -protected def covering (_h : FineSubfamilyOn v f s) : α × Set α → Set α := +protected def covering (_h : FineSubfamilyOn v f s) : X × Set X → Set X := fun p => p.2 -theorem index_subset : ∀ p : α × Set α, p ∈ h.index → p.1 ∈ s := +theorem index_subset : ∀ p : X × Set X, p ∈ h.index → p.1 ∈ s := h.exists_disjoint_covering_ae.choose_spec.1 theorem covering_disjoint : h.index.PairwiseDisjoint h.covering := @@ -133,24 +133,24 @@ theorem covering_disjoint : h.index.PairwiseDisjoint h.covering := theorem covering_disjoint_subtype : Pairwise (Disjoint on fun x : h.index => h.covering x) := (pairwise_subtype_iff_pairwise_set _ _).2 h.covering_disjoint -theorem covering_mem {p : α × Set α} (hp : p ∈ h.index) : h.covering p ∈ f p.1 := +theorem covering_mem {p : X × Set X} (hp : p ∈ h.index) : h.covering p ∈ f p.1 := (h.exists_disjoint_covering_ae.choose_spec.2.2.1 p hp).2 -theorem covering_mem_family {p : α × Set α} (hp : p ∈ h.index) : h.covering p ∈ v.setsAt p.1 := +theorem covering_mem_family {p : X × Set X} (hp : p ∈ h.index) : h.covering p ∈ v.setsAt p.1 := (h.exists_disjoint_covering_ae.choose_spec.2.2.1 p hp).1 theorem measure_diff_biUnion : μ (s \ ⋃ p ∈ h.index, h.covering p) = 0 := h.exists_disjoint_covering_ae.choose_spec.2.2.2 -theorem index_countable [SecondCountableTopology α] : h.index.Countable := +theorem index_countable [SecondCountableTopology X] : h.index.Countable := h.covering_disjoint.countable_of_nonempty_interior fun _ hx => v.nonempty_interior _ _ (h.covering_mem_family hx) -protected theorem measurableSet_u {p : α × Set α} (hp : p ∈ h.index) : +protected theorem measurableSet_u {p : X × Set X} (hp : p ∈ h.index) : MeasurableSet (h.covering p) := v.measurableSet p.1 _ (h.covering_mem_family hp) -theorem measure_le_tsum_of_absolutelyContinuous [SecondCountableTopology α] {ρ : Measure α} +theorem measure_le_tsum_of_absolutelyContinuous [SecondCountableTopology X] {ρ : Measure X} (hρ : ρ ≪ μ) : ρ s ≤ ∑' p : h.index, ρ (h.covering p) := calc ρ s ≤ ρ ((s \ ⋃ p ∈ h.index, h.covering p) ∪ ⋃ p ∈ h.index, h.covering p) := @@ -161,7 +161,7 @@ theorem measure_le_tsum_of_absolutelyContinuous [SecondCountableTopology α] {ρ rw [hρ h.measure_diff_biUnion, zero_add, measure_biUnion h.index_countable h.covering_disjoint fun x hx => h.measurableSet_u hx] -theorem measure_le_tsum [SecondCountableTopology α] : μ s ≤ ∑' x : h.index, μ (h.covering x) := +theorem measure_le_tsum [SecondCountableTopology X] : μ s ≤ ∑' x : h.index, μ (h.covering x) := h.measure_le_tsum_of_absolutelyContinuous Measure.AbsolutelyContinuous.rfl end FineSubfamilyOn @@ -170,86 +170,86 @@ end FineSubfamilyOn contained in a `δ`-neighborhood on `x`. This does not change the local filter at a point, but it can be convenient to get a nicer global behavior. -/ def enlarge (v : VitaliFamily μ) (δ : ℝ) (δpos : 0 < δ) : VitaliFamily μ where - setsAt x := v.setsAt x ∪ { a | MeasurableSet a ∧ (interior a).Nonempty ∧ ¬a ⊆ closedBall x δ } - measurableSet x a ha := by - cases' ha with ha ha - exacts [v.measurableSet _ _ ha, ha.1] - nonempty_interior x a ha := by - cases' ha with ha ha - exacts [v.nonempty_interior _ _ ha, ha.2.1] + setsAt x := v.setsAt x ∪ {s | MeasurableSet s ∧ (interior s).Nonempty ∧ ¬s ⊆ closedBall x δ} + measurableSet := by + rintro x s (hs | hs) + exacts [v.measurableSet _ _ hs, hs.1] + nonempty_interior := by + rintro x s (hs | hs) + exacts [v.nonempty_interior _ _ hs, hs.2.1] nontrivial := by intro x ε εpos - rcases v.nontrivial x ε εpos with ⟨a, ha, h'a⟩ - exact ⟨a, mem_union_left _ ha, h'a⟩ + rcases v.nontrivial x ε εpos with ⟨s, hs, h's⟩ + exact ⟨s, mem_union_left _ hs, h's⟩ covering := by intro s f fset ffine - let g : α → Set (Set α) := fun x => f x ∩ v.setsAt x - have : ∀ x ∈ s, ∀ ε : ℝ, ε > 0 → ∃ (a : Set α), a ∈ g x ∧ a ⊆ closedBall x ε := by + let g : X → Set (Set X) := fun x => f x ∩ v.setsAt x + have : ∀ x ∈ s, ∀ ε : ℝ, ε > 0 → ∃ t ∈ g x, t ⊆ closedBall x ε := by intro x hx ε εpos - obtain ⟨a, af, ha⟩ : ∃ a ∈ f x, a ⊆ closedBall x (min ε δ) := + obtain ⟨t, tf, ht⟩ : ∃ t ∈ f x, t ⊆ closedBall x (min ε δ) := ffine x hx (min ε δ) (lt_min εpos δpos) - rcases fset x hx af with (h'a | h'a) - · exact ⟨a, ⟨af, h'a⟩, ha.trans (closedBall_subset_closedBall (min_le_left _ _))⟩ - · refine False.elim (h'a.2.2 ?_) - exact ha.trans (closedBall_subset_closedBall (min_le_right _ _)) + rcases fset x hx tf with (h't | h't) + · exact ⟨t, ⟨tf, h't⟩, ht.trans (closedBall_subset_closedBall (min_le_left _ _))⟩ + · refine False.elim (h't.2.2 ?_) + exact ht.trans (closedBall_subset_closedBall (min_le_right _ _)) rcases v.covering s g (fun x _ => inter_subset_right) this with ⟨t, ts, tdisj, tg, μt⟩ exact ⟨t, ts, tdisj, fun p hp => (tg p hp).1, μt⟩ variable (v : VitaliFamily μ) -/-- Given a vitali family `v`, then `v.filterAt x` is the filter on `Set α` made of those families +/-- Given a vitali family `v`, then `v.filterAt x` is the filter on `Set X` made of those families that contain all sets of `v.setsAt x` of a sufficiently small diameter. This filter makes it possible to express limiting behavior when sets in `v.setsAt x` shrink to `x`. -/ -def filterAt (x : α) : Filter (Set α) := (𝓝 x).smallSets ⊓ 𝓟 (v.setsAt x) +def filterAt (x : X) : Filter (Set X) := (𝓝 x).smallSets ⊓ 𝓟 (v.setsAt x) -theorem _root_.Filter.HasBasis.vitaliFamily {ι : Sort*} {p : ι → Prop} {s : ι → Set α} {x : α} +theorem _root_.Filter.HasBasis.vitaliFamily {ι : Sort*} {p : ι → Prop} {s : ι → Set X} {x : X} (h : (𝓝 x).HasBasis p s) : (v.filterAt x).HasBasis p (fun i ↦ {t ∈ v.setsAt x | t ⊆ s i}) := by simpa only [← Set.setOf_inter_eq_sep] using h.smallSets.inf_principal _ -theorem filterAt_basis_closedBall (x : α) : - (v.filterAt x).HasBasis (0 < ·) ({a ∈ v.setsAt x | a ⊆ closedBall x ·}) := +theorem filterAt_basis_closedBall (x : X) : + (v.filterAt x).HasBasis (0 < ·) ({t ∈ v.setsAt x | t ⊆ closedBall x ·}) := nhds_basis_closedBall.vitaliFamily v -theorem mem_filterAt_iff {x : α} {s : Set (Set α)} : - s ∈ v.filterAt x ↔ ∃ ε > (0 : ℝ), ∀ a ∈ v.setsAt x, a ⊆ closedBall x ε → a ∈ s := by +theorem mem_filterAt_iff {x : X} {s : Set (Set X)} : + s ∈ v.filterAt x ↔ ∃ ε > (0 : ℝ), ∀ t ∈ v.setsAt x, t ⊆ closedBall x ε → t ∈ s := by simp only [(v.filterAt_basis_closedBall x).mem_iff, ← and_imp, subset_def, mem_setOf] -instance filterAt_neBot (x : α) : (v.filterAt x).NeBot := +instance filterAt_neBot (x : X) : (v.filterAt x).NeBot := (v.filterAt_basis_closedBall x).neBot_iff.2 <| v.nontrivial _ _ -theorem eventually_filterAt_iff {x : α} {P : Set α → Prop} : - (∀ᶠ a in v.filterAt x, P a) ↔ ∃ ε > (0 : ℝ), ∀ a ∈ v.setsAt x, a ⊆ closedBall x ε → P a := +theorem eventually_filterAt_iff {x : X} {P : Set X → Prop} : + (∀ᶠ t in v.filterAt x, P t) ↔ ∃ ε > (0 : ℝ), ∀ t ∈ v.setsAt x, t ⊆ closedBall x ε → P t := v.mem_filterAt_iff -theorem tendsto_filterAt_iff {ι : Type*} {l : Filter ι} {f : ι → Set α} {x : α} : +theorem tendsto_filterAt_iff {ι : Type*} {l : Filter ι} {f : ι → Set X} {x : X} : Tendsto f l (v.filterAt x) ↔ (∀ᶠ i in l, f i ∈ v.setsAt x) ∧ ∀ ε > (0 : ℝ), ∀ᶠ i in l, f i ⊆ closedBall x ε := by simp only [filterAt, tendsto_inf, nhds_basis_closedBall.smallSets.tendsto_right_iff, tendsto_principal, and_comm, mem_powerset_iff] -theorem eventually_filterAt_mem_setsAt (x : α) : ∀ᶠ a in v.filterAt x, a ∈ v.setsAt x := +theorem eventually_filterAt_mem_setsAt (x : X) : ∀ᶠ t in v.filterAt x, t ∈ v.setsAt x := (v.tendsto_filterAt_iff.mp tendsto_id).1 -theorem eventually_filterAt_subset_closedBall (x : α) {ε : ℝ} (hε : 0 < ε) : - ∀ᶠ a : Set α in v.filterAt x, a ⊆ closedBall x ε := +theorem eventually_filterAt_subset_closedBall (x : X) {ε : ℝ} (hε : 0 < ε) : + ∀ᶠ t : Set X in v.filterAt x, t ⊆ closedBall x ε := (v.tendsto_filterAt_iff.mp tendsto_id).2 ε hε -theorem eventually_filterAt_measurableSet (x : α) : ∀ᶠ a in v.filterAt x, MeasurableSet a := by +theorem eventually_filterAt_measurableSet (x : X) : ∀ᶠ t in v.filterAt x, MeasurableSet t := by filter_upwards [v.eventually_filterAt_mem_setsAt x] with _ ha using v.measurableSet _ _ ha -theorem frequently_filterAt_iff {x : α} {P : Set α → Prop} : - (∃ᶠ a in v.filterAt x, P a) ↔ ∀ ε > (0 : ℝ), ∃ a ∈ v.setsAt x, a ⊆ closedBall x ε ∧ P a := by +theorem frequently_filterAt_iff {x : X} {P : Set X → Prop} : + (∃ᶠ t in v.filterAt x, P t) ↔ ∀ ε > (0 : ℝ), ∃ t ∈ v.setsAt x, t ⊆ closedBall x ε ∧ P t := by simp only [(v.filterAt_basis_closedBall x).frequently_iff, ← and_assoc, subset_def, mem_setOf] -theorem eventually_filterAt_subset_of_nhds {x : α} {o : Set α} (hx : o ∈ 𝓝 x) : - ∀ᶠ a in v.filterAt x, a ⊆ o := +theorem eventually_filterAt_subset_of_nhds {x : X} {o : Set X} (hx : o ∈ 𝓝 x) : + ∀ᶠ t in v.filterAt x, t ⊆ o := (eventually_smallSets_subset.2 hx).filter_mono inf_le_left -theorem fineSubfamilyOn_of_frequently (v : VitaliFamily μ) (f : α → Set (Set α)) (s : Set α) - (h : ∀ x ∈ s, ∃ᶠ a in v.filterAt x, a ∈ f x) : v.FineSubfamilyOn f s := by +theorem fineSubfamilyOn_of_frequently (v : VitaliFamily μ) (f : X → Set (Set X)) (s : Set X) + (h : ∀ x ∈ s, ∃ᶠ t in v.filterAt x, t ∈ f x) : v.FineSubfamilyOn f s := by intro x hx ε εpos - obtain ⟨a, av, ha, af⟩ : ∃ (a : Set α) , a ∈ v.setsAt x ∧ a ⊆ closedBall x ε ∧ a ∈ f x := + obtain ⟨t, tv, ht, tf⟩ : ∃ t ∈ v.setsAt x, t ⊆ closedBall x ε ∧ t ∈ f x := v.frequently_filterAt_iff.1 (h x hx) ε εpos - exact ⟨a, ⟨av, af⟩, ha⟩ + exact ⟨t, ⟨tv, tf⟩, ht⟩ end VitaliFamily From cdba113ff891688d1b19bd9808172b3b9e5d3e53 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 4 Aug 2024 11:42:45 +0000 Subject: [PATCH 017/359] feat: add `tsub_eq_tsub_of_add_eq_add` (#15484) --- Mathlib/Algebra/Order/Sub/Defs.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Algebra/Order/Sub/Defs.lean b/Mathlib/Algebra/Order/Sub/Defs.lean index 45c741d9a8896..eff553b812467 100644 --- a/Mathlib/Algebra/Order/Sub/Defs.lean +++ b/Mathlib/Algebra/Order/Sub/Defs.lean @@ -309,6 +309,12 @@ theorem add_tsub_cancel_right (a b : α) : a + b - b = a := theorem add_tsub_cancel_left (a b : α) : a + b - a = b := Contravariant.AddLECancellable.add_tsub_cancel_left +/-- A more general version of the reverse direction of `sub_eq_sub_iff_add_eq_add` -/ +theorem tsub_eq_tsub_of_add_eq_add (h : a + d = c + b) : a - b = c - d := by + calc a - b = a + d - d - b := by rw [add_tsub_cancel_right] + _ = c + b - b - d := by rw [h, tsub_right_comm] + _ = c - d := by rw [add_tsub_cancel_right] + theorem lt_add_of_tsub_lt_left (h : a - b < c) : a < b + c := Contravariant.AddLECancellable.lt_add_of_tsub_lt_left h From 776ac28d1aa92ebfba597b13a40b7c5f2b85f370 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sun, 4 Aug 2024 12:42:06 +0000 Subject: [PATCH 018/359] feat(RingTheory/Smooth): Correspondence between square-zero lifts and retractions of the conormal sequence. (#14138) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/RingTheory/Ideal/Cotangent.lean | 18 ++ Mathlib/RingTheory/Smooth/Kaehler.lean | 210 ++++++++++++++++++++++++ 3 files changed, 229 insertions(+) create mode 100644 Mathlib/RingTheory/Smooth/Kaehler.lean diff --git a/Mathlib.lean b/Mathlib.lean index 82042139dae44..a2a2b713a4b21 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3912,6 +3912,7 @@ import Mathlib.RingTheory.RootsOfUnity.Lemmas import Mathlib.RingTheory.RootsOfUnity.Minpoly import Mathlib.RingTheory.SimpleModule import Mathlib.RingTheory.Smooth.Basic +import Mathlib.RingTheory.Smooth.Kaehler import Mathlib.RingTheory.Smooth.StandardSmooth import Mathlib.RingTheory.Support import Mathlib.RingTheory.SurjectiveOnStalks diff --git a/Mathlib/RingTheory/Ideal/Cotangent.lean b/Mathlib/RingTheory/Ideal/Cotangent.lean index 2dd5e66fce644..3156cab853769 100644 --- a/Mathlib/RingTheory/Ideal/Cotangent.lean +++ b/Mathlib/RingTheory/Ideal/Cotangent.lean @@ -98,6 +98,16 @@ theorem to_quotient_square_comp_toCotangent : theorem toCotangent_to_quotient_square (x : I) : I.cotangentToQuotientSquare (I.toCotangent x) = (I ^ 2).mkQ x := rfl +lemma Cotangent.smul_eq_zero_of_mem {I : Ideal R} + {x} (hx : x ∈ I) (m : I.Cotangent) : x • m = 0 := by + obtain ⟨m, rfl⟩ := Ideal.toCotangent_surjective _ m + rw [← map_smul, Ideal.toCotangent_eq_zero, pow_two] + exact Ideal.mul_mem_mul hx m.2 + +lemma isTorsionBySet_cotangent : + Module.IsTorsionBySet R I.Cotangent I := + fun m x ↦ m.smul_eq_zero_of_mem x.2 + /-- `I ⧸ I ^ 2` as an ideal of `R ⧸ I ^ 2`. -/ def cotangentIdeal (I : Ideal R) : Ideal (R ⧸ I ^ 2) := Submodule.map (Quotient.mk (I ^ 2)|>.toSemilinearMap) I @@ -163,6 +173,14 @@ theorem _root_.AlgHom.ker_kerSquareLift (f : A →ₐ[R] B) : · intro x hx; obtain ⟨x, rfl⟩ := Ideal.Quotient.mk_surjective x; exact ⟨x, hx, rfl⟩ · rintro _ ⟨x, hx, rfl⟩; exact hx +instance Algebra.kerSquareLift : Algebra (R ⧸ (RingHom.ker (algebraMap R A) ^ 2)) A := + (Algebra.ofId R A).kerSquareLift.toAlgebra + +instance [Algebra A B] [IsScalarTower R A B] : + IsScalarTower R (A ⧸ (RingHom.ker (algebraMap A B) ^ 2)) B := + IsScalarTower.of_algebraMap_eq' + (IsScalarTower.toAlgHom R A B).kerSquareLift.comp_algebraMap.symm + /-- The quotient ring of `I ⧸ I ^ 2` is `R ⧸ I`. -/ def quotCotangent : (R ⧸ I ^ 2) ⧸ I.cotangentIdeal ≃+* R ⧸ I := by refine (Ideal.quotEquivOfEq (Ideal.map_eq_submodule_map _ _).symm).trans ?_ diff --git a/Mathlib/RingTheory/Smooth/Kaehler.lean b/Mathlib/RingTheory/Smooth/Kaehler.lean new file mode 100644 index 0000000000000..1a649e6bd938c --- /dev/null +++ b/Mathlib/RingTheory/Smooth/Kaehler.lean @@ -0,0 +1,210 @@ +/- +Copyright (c) 2024 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import Mathlib.RingTheory.Kaehler.Basic + +/-! +# Relation of smoothness and `Ω[S⁄R]` + +## Main results + +- `retractionKerToTensorEquivSection`: + Given a surjective algebra homomorphism `f : P →ₐ[R] S` with square-zero kernel `I`, + there is a one-to-one correspondence between `P`-linear retractions of `I →ₗ[P] S ⊗[P] Ω[P/R]` + and algebra homomorphism sections of `f`. + +## Future projects + +- Show that relative smooth iff `H¹(L_{S/R}) = 0` and `Ω[S/R]` is projective. +- Show that being smooth is local on stalks. +- Show that being formally smooth is Zariski-local (very hard). + +-/ + +universe u + +open TensorProduct KaehlerDifferential + +open Function (Surjective) + +variable {R P S : Type u} [CommRing R] [CommRing P] [CommRing S] +variable [Algebra R S] [Algebra R P] [Algebra P S] [IsScalarTower R P S] +variable (hf : Surjective (algebraMap P S)) (hf' : (RingHom.ker (algebraMap P S)) ^ 2 = ⊥) + +section ofSection + +-- Suppose we have a section (as alghom) of `P →ₐ[R] S`. +variable (g : S →ₐ[R] P) (hg : (IsScalarTower.toAlgHom R P S).comp g = AlgHom.id R S) + +/-- +Given a surjective algebra homomorphism `f : P →ₐ[R] S` with square-zero kernel `I`, +and a section `g : S →ₐ[R] P` (as an algebra homomorphism), +we get an `R`-derivation `P → I` via `x ↦ x - g (f x)`. +-/ +@[simps] +def derivationOfSectionOfKerSqZero (f : P →ₐ[R] S) (hf' : (RingHom.ker f) ^ 2 = ⊥) (g : S →ₐ[R] P) + (hg : f.comp g = AlgHom.id R S) : Derivation R P (RingHom.ker f) where + toFun x := ⟨x - g (f x), by + simpa [RingHom.mem_ker, sub_eq_zero] using AlgHom.congr_fun hg.symm (f x)⟩ + map_add' x y := by simp only [map_add, AddSubmonoid.mk_add_mk, Subtype.mk.injEq]; ring + map_smul' x y := by + ext + simp only [Algebra.smul_def, _root_.map_mul, ← IsScalarTower.algebraMap_apply, + AlgHom.commutes, RingHom.id_apply, Submodule.coe_smul_of_tower] + ring + map_one_eq_zero' := by simp only [LinearMap.coe_mk, AddHom.coe_mk, _root_.map_one, sub_self, + AddSubmonoid.mk_eq_zero] + leibniz' a b := by + have : (a - g (f a)) * (b - g (f b)) = 0 := by + rw [← Ideal.mem_bot, ← hf', pow_two] + apply Ideal.mul_mem_mul + · simpa [RingHom.mem_ker, sub_eq_zero] using AlgHom.congr_fun hg.symm (f a) + · simpa [RingHom.mem_ker, sub_eq_zero] using AlgHom.congr_fun hg.symm (f b) + ext + rw [← sub_eq_zero] + conv_rhs => rw [← neg_zero, ← this] + simp only [LinearMap.coe_mk, AddHom.coe_mk, _root_.map_mul, SetLike.mk_smul_mk, smul_eq_mul, + mul_sub, AddSubmonoid.mk_add_mk, sub_mul, neg_sub] + ring + +lemma isScalarTower_of_section_of_ker_sqZero : + letI := g.toRingHom.toAlgebra; IsScalarTower P S (RingHom.ker (algebraMap P S)) := by + letI := g.toRingHom.toAlgebra + constructor + intro p s m + ext + show g (p • s) * m = p * (g s * m) + simp only [Algebra.smul_def, _root_.map_mul, mul_assoc, mul_left_comm _ (g s)] + congr 1 + rw [← sub_eq_zero, ← Ideal.mem_bot, ← hf', pow_two, ← sub_mul] + refine Ideal.mul_mem_mul ?_ m.2 + simpa [RingHom.mem_ker, sub_eq_zero] using AlgHom.congr_fun hg (algebraMap P S p) + +/-- +Given a surjective algebra hom `f : P →ₐ[R] S` with square-zero kernel `I`, +and a section `g : S →ₐ[R] P` (as algebra homs), +we get a retraction of the injection `I → S ⊗[P] Ω[P/R]`. +-/ +noncomputable +def retractionOfSectionOfKerSqZero : S ⊗[P] Ω[P⁄R] →ₗ[P] RingHom.ker (algebraMap P S) := + letI := g.toRingHom.toAlgebra + haveI := isScalarTower_of_section_of_ker_sqZero hf' g hg + letI f : _ →ₗ[P] RingHom.ker (algebraMap P S) := (derivationOfSectionOfKerSqZero + (IsScalarTower.toAlgHom R P S) hf' g hg).liftKaehlerDifferential + (f.liftBaseChange S).restrictScalars P + +@[simp] +lemma retractionOfSectionOfKerSqZero_tmul_D (s : S) (t : P) : + retractionOfSectionOfKerSqZero hf' g hg (s ⊗ₜ .D _ _ t) = + g s * t - g s * g (algebraMap _ _ t) := by + letI := g.toRingHom.toAlgebra + haveI := isScalarTower_of_section_of_ker_sqZero hf' g hg + simp only [retractionOfSectionOfKerSqZero, AlgHom.toRingHom_eq_coe, LinearMap.coe_restrictScalars, + LinearMap.liftBaseChange_tmul, SetLike.val_smul_of_tower] + erw [Derivation.liftKaehlerDifferential_comp_D] + exact mul_sub (g s) t (g (algebraMap P S t)) + +lemma retractionOfSectionOfKerSqZero_comp_kerToTensor : + (retractionOfSectionOfKerSqZero hf' g hg).comp (kerToTensor R P S) = LinearMap.id := by + ext x; simp [(RingHom.mem_ker _).mp x.2] + +end ofSection + +section ofRetraction + +variable (l : S ⊗[P] Ω[P⁄R] →ₗ[P] RingHom.ker (algebraMap P S)) +variable (hl : l.comp (kerToTensor R P S) = LinearMap.id) + +-- suppose we have a (set-theoretic) section +variable (σ : S → P) (hσ : ∀ x, algebraMap P S (σ x) = x) + +lemma sectionOfRetractionKerToTensorAux_prop (x y) (h : algebraMap P S x = algebraMap P S y) : + x - l (1 ⊗ₜ .D _ _ x) = y - l (1 ⊗ₜ .D _ _ y) := by + rw [sub_eq_iff_eq_add, sub_add_comm, ← sub_eq_iff_eq_add, ← Submodule.coe_sub, + ← map_sub, ← tmul_sub, ← map_sub] + exact congr_arg Subtype.val (LinearMap.congr_fun hl.symm ⟨x - y, by simp [RingHom.mem_ker, h]⟩) + +/-- +Given a surjective algebra homomorphism `f : P →ₐ[R] S` with square-zero kernel `I`. +Let `σ` be an arbitrary (set-theoretic) section of `f`. +Suppose we have a retraction `l` of the injection `I →ₗ[P] S ⊗[P] Ω[P/R]`, then +`x ↦ σ x - l (1 ⊗ D (σ x))` is an algebra homomorphism and a section to `f`. +-/ +noncomputable +def sectionOfRetractionKerToTensorAux : S →ₐ[R] P where + toFun x := σ x - l (1 ⊗ₜ .D _ _ (σ x)) + map_one' := by simp [sectionOfRetractionKerToTensorAux_prop l hl (σ 1) 1 (by simp [hσ])] + map_mul' a b := by + have (x y) : (l x).1 * (l y).1 = 0 := by + rw [← Ideal.mem_bot, ← hf', pow_two]; exact Ideal.mul_mem_mul (l x).2 (l y).2 + simp only [sectionOfRetractionKerToTensorAux_prop l hl (σ (a * b)) (σ a * σ b) (by simp [hσ]), + Derivation.leibniz, tmul_add, tmul_smul, map_add, map_smul, AddSubmonoid.coe_add, this, + Submodule.coe_toAddSubmonoid, SetLike.val_smul, smul_eq_mul, mul_sub, sub_mul, sub_zero] + ring + map_add' a b := by + simp only [sectionOfRetractionKerToTensorAux_prop l hl (σ (a + b)) (σ a + σ b) (by simp [hσ]), + map_add, tmul_add, AddSubmonoid.coe_add, Submodule.coe_toAddSubmonoid, add_sub_add_comm] + map_zero' := by simp [sectionOfRetractionKerToTensorAux_prop l hl (σ 0) 0 (by simp [hσ])] + commutes' r := by + simp [sectionOfRetractionKerToTensorAux_prop l hl + (σ (algebraMap R S r)) (algebraMap R P r) (by simp [hσ, ← IsScalarTower.algebraMap_apply])] + +lemma sectionOfRetractionKerToTensorAux_algebraMap (x : P) : + sectionOfRetractionKerToTensorAux hf' l hl σ hσ (algebraMap P S x) = x - l (1 ⊗ₜ .D _ _ x) := + sectionOfRetractionKerToTensorAux_prop l hl _ x (by simp [hσ]) + +lemma toAlgHom_comp_sectionOfRetractionKerToTensorAux : + (IsScalarTower.toAlgHom R P S).comp + (sectionOfRetractionKerToTensorAux hf' l hl σ hσ) = AlgHom.id _ _ := by + ext x + obtain ⟨x, rfl⟩ := hf x + simp [sectionOfRetractionKerToTensorAux_algebraMap, (RingHom.mem_ker _).mp] + +/-- +Given a surjective algebra homomorphism `f : P →ₐ[R] S` with square-zero kernel `I`. +Suppose we have a retraction `l` of the injection `I →ₗ[P] S ⊗[P] Ω[P/R]`, then +`x ↦ σ x - l (1 ⊗ D (σ x))` is an algebra homomorphism and a section to `f`, +where `σ` is an arbitrary (set-theoretic) section of `f` +-/ +noncomputable def sectionOfRetractionKerToTensor : S →ₐ[R] P := + sectionOfRetractionKerToTensorAux hf' l hl _ (fun x ↦ (hf x).choose_spec) + +@[simp] +lemma sectionOfRetractionKerToTensor_algebraMap (x : P) : + sectionOfRetractionKerToTensor hf hf' l hl (algebraMap P S x) = x - l (1 ⊗ₜ .D _ _ x) := + sectionOfRetractionKerToTensorAux_algebraMap hf' l hl _ _ x + +@[simp] +lemma toAlgHom_comp_sectionOfRetractionKerToTensor : + (IsScalarTower.toAlgHom R P S).comp + (sectionOfRetractionKerToTensor hf hf' l hl) = AlgHom.id _ _ := + toAlgHom_comp_sectionOfRetractionKerToTensorAux hf _ _ _ _ _ + +end ofRetraction + +/-- +Given a surjective algebra homomorphism `f : P →ₐ[R] S` with square-zero kernel `I`, +there is a one-to-one correspondence between `P`-linear retractions of `I →ₗ[P] S ⊗[P] Ω[P/R]` +and algebra homomorphism sections of `f`. +-/ +noncomputable +def retractionKerToTensorEquivSection : + { l // l ∘ₗ (kerToTensor R P S) = LinearMap.id } ≃ + { g // (IsScalarTower.toAlgHom R P S).comp g = AlgHom.id R S } where + toFun l := ⟨_, toAlgHom_comp_sectionOfRetractionKerToTensor hf hf' _ l.2⟩ + invFun g := ⟨_, retractionOfSectionOfKerSqZero_comp_kerToTensor hf' _ g.2⟩ + left_inv l := by + ext s p + obtain ⟨s, rfl⟩ := hf s + have (x y) : (l.1 x).1 * (l.1 y).1 = 0 := by + rw [← Ideal.mem_bot, ← hf', pow_two]; exact Ideal.mul_mem_mul (l.1 x).2 (l.1 y).2 + simp only [AlgebraTensorModule.curry_apply, + Derivation.coe_comp, LinearMap.coe_comp, LinearMap.coe_restrictScalars, Derivation.coeFn_coe, + Function.comp_apply, curry_apply, retractionOfSectionOfKerSqZero_tmul_D, + sectionOfRetractionKerToTensor_algebraMap, ← mul_sub, sub_sub_cancel] + rw [sub_mul] + simp only [this, Algebra.algebraMap_eq_smul_one, ← smul_tmul', LinearMapClass.map_smul, + SetLike.val_smul, smul_eq_mul, sub_zero] + right_inv g := by ext s; obtain ⟨s, rfl⟩ := hf s; simp From 766dd6223ad28a2405bfc3ed54746bc810ca2841 Mon Sep 17 00:00:00 2001 From: grunweg Date: Sun, 4 Aug 2024 12:42:07 +0000 Subject: [PATCH 019/359] chore: remove fix-comments.py (#15242) This script has served its purpose during the port; now that all align information has been removed from mathlib, it is not useful for other projects either. Projects which wish to use such functionality should check out the v3-eol tag and use the script from that commit. --- scripts/fix-comments.py | 162 ---------------------------------------- 1 file changed, 162 deletions(-) delete mode 100755 scripts/fix-comments.py diff --git a/scripts/fix-comments.py b/scripts/fix-comments.py deleted file mode 100755 index f5a3de4157093..0000000000000 --- a/scripts/fix-comments.py +++ /dev/null @@ -1,162 +0,0 @@ -#!/usr/bin/env python3 - -import sys -import os -from pathlib import Path -import subprocess -import re - -if len(sys.argv) != 2 or not sys.argv[1].endswith('.lean'): - print("usage: fix-comments.py X.lean") - sys.exit(1) - -leanfile = sys.argv[1] - -is_clean = subprocess.run( - ['git', 'status', '--untracked-files=no', '--porcelain'], - capture_output=True, - check=True, - encoding='utf-8').stdout.rstrip() - -if is_clean != "": - print("Certain files tracked by git have uncommitted changes.\n") - os.system("git status --untracked-files=no") - print("\n") - s = input("Type y to continue. ") - if s != 'y': - sys.exit(1) - -root_dir = subprocess.run( - ['git', 'rev-parse', '--show-toplevel'], - capture_output=True, - check=True, - encoding='utf-8').stdout.rstrip() - -align_files = subprocess.run( - ['git', 'grep', '-l', '^#align'], - cwd=root_dir, - capture_output=True, - check=True, - encoding='utf-8') - -name_map = dict() -for f in align_files.stdout.splitlines(): - with open(os.path.join(root_dir, f), encoding="utf-8") as fh: - contents = fh.read() - for p in contents.split(sep='\n#align')[1:]: - n3, n4, *_ = p.split(maxsplit=2) - name_map[n3] = n4 - -def replace_names(s): - # Terrible hack to treat `.` as a word character - # (to match qualified names) - s = s.replace('.', 'Ᾰ') - # re.DOTALL means that `.` can also match a newline. - # `\A` and `\Z` match only at the start/end of the string respectively. - w = re.findall(r'(?:\b|\A).+?(?:\b|\Z)', s, flags=re.DOTALL) - for i in range(len(w)): - w[i] = w[i].replace('Ᾰ', '.') - w[i] = name_map.get(w[i], w[i]) - return ''.join(w) - -def process_backticked_names(s): - w = s.split(sep='`') - for i in range(len(w)): - if i % 2 == 1: - w[i] = replace_names(w[i]) - return '`'.join(w) - -rewritten_contents = '' - -in_block_comment = False -in_line_comment = False -prev_char = None -comment_so_far = None # contains end marker but not begin marker - -def finish_comment(): - global rewritten_contents - global in_block_comment - global in_line_comment - global comment_so_far - if comment_so_far is not None: - rewritten_contents += process_backticked_names(comment_so_far) - in_block_comment = False - in_line_comment = False - comment_so_far = None - -with open(leanfile, encoding="utf-8") as F: - while 1: - char = F.read(1) - if not char: - finish_comment() - break - - if in_block_comment or in_line_comment: - comment_so_far = comment_so_far + char - else: - rewritten_contents += char - - if in_block_comment and prev_char == '-' and char == '/': - finish_comment() - - if in_line_comment and char == '\n': - finish_comment() - - if comment_so_far is None and prev_char == '/' and char == '-': - in_block_comment = True - comment_so_far = '' - - if comment_so_far is None and prev_char == '-' and char == '-': - in_line_comment = True - comment_so_far = '' - - prev_char = char - -def mktree(path, sha, tree=True): - if path == Path('.'): - return sha - if tree: - inp = f"040000 tree {sha}\t{path.name}" - else: - inp = f"100644 blob {sha}\t{path.name}" - tree_sha = subprocess.run( - ['git', 'mktree'], - cwd=root_dir, - input=inp, - capture_output=True, - check=True, - encoding='utf8').stdout.rstrip() - return mktree(path.parent, tree_sha) - -path = Path(subprocess.run( - ['git', 'ls-files', '--full-name', leanfile], - capture_output=True, - check=True, - encoding='utf-8').stdout.rstrip()) - -blob_sha = subprocess.run( - ['git', 'hash-object', '-w', '--stdin'], - input=rewritten_contents, - cwd=root_dir, - capture_output=True, - check=True, - encoding='utf-8').stdout.rstrip() - -tree_sha = mktree(path, blob_sha, tree=False) - -print(f"The script will now interactively suggest changes to {leanfile}.\n") -s = input("Type y to continue. ") -if s != 'y': - sys.exit(1) - -subprocess.run(['git', 'restore', '--patch', '--source=' + tree_sha, '--', leanfile], check=True) - -r = subprocess.run(['git', 'diff', '--quiet', leanfile]) -if r.returncode == 0: - pass -elif r.returncode == 1: # file was changed - print("\nPerhaps you would now like to run:") - print(f"git add {leanfile} && git commit -m 'auto: naming'") -else: - # something went wrong - r.check_returncode() From ba4c368b95071021cb9ec4938c041578c8b7164e Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 4 Aug 2024 12:42:09 +0000 Subject: [PATCH 020/359] feat(ENat): not_lt_zero, coe_lt_top (#15342) from the Carleson project --- Mathlib/Data/ENat/Basic.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index e94cabd57203f..e641b878298b9 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -251,6 +251,14 @@ theorem lt_add_one_iff (hm : n ≠ ⊤) : m < n + 1 ↔ m ≤ n := theorem le_coe_iff {n : ℕ∞} {k : ℕ} : n ≤ ↑k ↔ ∃ (n₀ : ℕ), n = n₀ ∧ n₀ ≤ k := WithTop.le_coe_iff +@[simp] +lemma not_lt_zero (n : ℕ∞) : ¬ n < 0 := by + cases n <;> simp + +@[simp] +lemma coe_lt_top (n : ℕ) : (n : ℕ∞) < ⊤ := + WithTop.coe_lt_top n + @[elab_as_elim] theorem nat_induction {P : ℕ∞ → Prop} (a : ℕ∞) (h0 : P 0) (hsuc : ∀ n : ℕ, P n → P n.succ) (htop : (∀ n : ℕ, P n) → P ⊤) : P a := by From ae960ed3505e3c361100a573b31abd62738f081b Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 4 Aug 2024 12:42:10 +0000 Subject: [PATCH 021/359] feat(ENat): iSup_eq_zero (#15345) from the Carleson project --- Mathlib/Data/ENat/Lattice.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Data/ENat/Lattice.lean b/Mathlib/Data/ENat/Lattice.lean index fe8f5b3fd827a..297e4191742e3 100644 --- a/Mathlib/Data/ENat/Lattice.lean +++ b/Mathlib/Data/ENat/Lattice.lean @@ -71,6 +71,9 @@ lemma sInf_eq_zero : sInf s = 0 ↔ 0 ∈ s := by lemma sSup_eq_zero' : sSup s = 0 ↔ s = ∅ ∨ s = {0} := sSup_eq_bot' +lemma iSup_eq_zero : iSup f = 0 ↔ ∀ i, f i = 0 := + iSup_eq_bot + lemma sSup_eq_top_of_infinite (h : s.Infinite) : sSup s = ⊤ := by apply (sSup_eq_top ..).mpr intro x hx From 0b653510d11af5633ea9db4a07b865c88e24256d Mon Sep 17 00:00:00 2001 From: grunweg Date: Sun, 4 Aug 2024 12:42:11 +0000 Subject: [PATCH 022/359] chore(Topology): remove some uses of open Classical (#15481) --- Mathlib/Topology/Algebra/Algebra.lean | 5 +---- Mathlib/Topology/Algebra/Group/Basic.lean | 2 -- Mathlib/Topology/Algebra/Group/Compact.lean | 7 ------- Mathlib/Topology/Algebra/StarSubalgebra.lean | 6 ------ Mathlib/Topology/Algebra/UniformField.lean | 3 +-- Mathlib/Topology/Algebra/UniformGroup.lean | 2 -- Mathlib/Topology/Algebra/UniformRing.lean | 6 ------ Mathlib/Topology/Clopen.lean | 2 +- Mathlib/Topology/Compactness/Compact.lean | 8 +++++--- .../Topology/Compactness/LocallyCompact.lean | 6 ++++-- .../Topology/Compactness/SigmaCompact.lean | 19 ++++++++++++------- Mathlib/Topology/Connected/Basic.lean | 4 ++-- Mathlib/Topology/Connected/Clopen.lean | 2 +- Mathlib/Topology/Constructions.lean | 7 ++++--- Mathlib/Topology/ExtremallyDisconnected.lean | 1 - .../Topology/FiberBundle/Trivialization.lean | 6 ++++-- Mathlib/Topology/Instances/ENNReal.lean | 6 +++++- Mathlib/Topology/Irreducible.lean | 4 +++- .../Topology/MetricSpace/GromovHausdorff.lean | 12 ++++-------- .../MetricSpace/GromovHausdorffRealized.lean | 9 +-------- Mathlib/Topology/PartitionOfUnity.lean | 16 ++++++++++------ Mathlib/Topology/Separation.lean | 5 ++--- Mathlib/Topology/ShrinkingLemma.lean | 11 ++++++----- Mathlib/Topology/UniformSpace/Cauchy.lean | 7 +------ .../UniformConvergenceTopology.lean | 3 +-- Mathlib/Topology/UnitInterval.lean | 5 +---- 26 files changed, 69 insertions(+), 95 deletions(-) diff --git a/Mathlib/Topology/Algebra/Algebra.lean b/Mathlib/Topology/Algebra/Algebra.lean index b08b2ee2c95c0..250871fe500e9 100644 --- a/Mathlib/Topology/Algebra/Algebra.lean +++ b/Mathlib/Topology/Algebra/Algebra.lean @@ -27,10 +27,7 @@ TODO: add continuous algebra isomorphisms. -/ -open scoped Classical -open Set TopologicalSpace Algebra BigOperators - -open scoped Classical +open Set TopologicalSpace Algebra universe u v w diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 75dd9c5b48682..8a5349a3e7a28 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -31,8 +31,6 @@ groups. topological space, group, topological group -/ - -open scoped Classical open Set Filter TopologicalSpace Function Topology Pointwise MulOpposite universe u v w x diff --git a/Mathlib/Topology/Algebra/Group/Compact.lean b/Mathlib/Topology/Algebra/Group/Compact.lean index 4b991cdbbf986..57b0155e2e6a8 100644 --- a/Mathlib/Topology/Algebra/Group/Compact.lean +++ b/Mathlib/Topology/Algebra/Group/Compact.lean @@ -15,13 +15,6 @@ imports developing either positive compacts or the compact open topology. -/ - -open scoped Classical -open Set Filter TopologicalSpace Function - -open scoped Classical -open Topology Filter Pointwise - universe u v w x variable {α : Type u} {β : Type v} {G : Type w} {H : Type x} diff --git a/Mathlib/Topology/Algebra/StarSubalgebra.lean b/Mathlib/Topology/Algebra/StarSubalgebra.lean index ffedf7f0ce282..c535323f3084f 100644 --- a/Mathlib/Topology/Algebra/StarSubalgebra.lean +++ b/Mathlib/Topology/Algebra/StarSubalgebra.lean @@ -22,12 +22,6 @@ The topological closure of a star subalgebra is still a star subalgebra, which as a star algebra is a topological star algebra. -/ - -open scoped Classical -open Set TopologicalSpace - -open scoped Classical - namespace StarSubalgebra section TopologicalStarAlgebra diff --git a/Mathlib/Topology/Algebra/UniformField.lean b/Mathlib/Topology/Algebra/UniformField.lean index 4bdd67535c6e8..96baeef3010b7 100644 --- a/Mathlib/Topology/Algebra/UniformField.lean +++ b/Mathlib/Topology/Algebra/UniformField.lean @@ -31,10 +31,8 @@ type class and the main results are the instances `UniformSpace.Completion.Field `UniformSpace.Completion.TopologicalDivisionRing`. -/ - noncomputable section -open scoped Classical open uniformity Topology open Set UniformSpace UniformSpace.Completion Filter @@ -88,6 +86,7 @@ theorem continuous_hatInv [CompletableTopField K] {x : hat K} (h : x ≠ 0) : erw [denseInducing_coe.nhds_eq_comap (0 : K), ← Filter.comap_inf, eq_bot] exact comap_bot +open Classical in /- The value of `hat_inv` at zero is not really specified, although it's probably zero. Here we explicitly enforce the `inv_zero` axiom. diff --git a/Mathlib/Topology/Algebra/UniformGroup.lean b/Mathlib/Topology/Algebra/UniformGroup.lean index d9872ca154ef9..bd3442028ba4a 100644 --- a/Mathlib/Topology/Algebra/UniformGroup.lean +++ b/Mathlib/Topology/Algebra/UniformGroup.lean @@ -34,10 +34,8 @@ group naturally induces a uniform structure. the quotient of a Banach space by a subspace is complete. -/ - noncomputable section -open scoped Classical open Uniformity Topology Filter Pointwise section UniformGroup diff --git a/Mathlib/Topology/Algebra/UniformRing.lean b/Mathlib/Topology/Algebra/UniformRing.lean index bda12edeb545c..d5096b53d82a7 100644 --- a/Mathlib/Topology/Algebra/UniformRing.lean +++ b/Mathlib/Topology/Algebra/UniformRing.lean @@ -33,12 +33,6 @@ the main constructions deal with continuous ring morphisms. TODO: Generalise the results here from the concrete `Completion` to any `AbstractCompletion`. -/ - -open scoped Classical -open Set Filter TopologicalSpace AddCommGroup - -open scoped Classical - noncomputable section universe u diff --git a/Mathlib/Topology/Clopen.lean b/Mathlib/Topology/Clopen.lean index 2a18e7931ec4e..c560e10009cd9 100644 --- a/Mathlib/Topology/Clopen.lean +++ b/Mathlib/Topology/Clopen.lean @@ -12,7 +12,7 @@ import Mathlib.Data.Set.BoolIndicator A clopen set is a set that is both closed and open. -/ -open Set Filter Topology TopologicalSpace Classical +open Set Filter Topology TopologicalSpace universe u v diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 1071071d3bfaa..61cae794ebb3b 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -26,7 +26,7 @@ We define the following properties for sets in a topological space: is compact. -/ -open Set Filter Topology TopologicalSpace Classical Function +open Set Filter Topology TopologicalSpace Function universe u v @@ -184,9 +184,10 @@ lemma IsCompact.elim_nhds_subcover_nhdsSet' (hs : IsCompact s) (U : ∀ x ∈ s, exact mem_interior_iff_mem_nhds.1 hy lemma IsCompact.elim_nhds_subcover_nhdsSet (hs : IsCompact s) {U : X → Set X} - (hU : ∀ x ∈ s, U x ∈ 𝓝 x) : ∃ t : Finset X, (∀ x ∈ t, x ∈ s) ∧ (⋃ x ∈ t, U x) ∈ 𝓝ˢ s := + (hU : ∀ x ∈ s, U x ∈ 𝓝 x) : ∃ t : Finset X, (∀ x ∈ t, x ∈ s) ∧ (⋃ x ∈ t, U x) ∈ 𝓝ˢ s := by let ⟨t, ht⟩ := hs.elim_nhds_subcover_nhdsSet' (fun x _ => U x) hU - ⟨t.image (↑), fun x hx => + classical + exact ⟨t.image (↑), fun x hx => let ⟨y, _, hyx⟩ := Finset.mem_image.1 hx hyx ▸ y.2, by rwa [Finset.set_biUnion_finset_image]⟩ @@ -516,6 +517,7 @@ lemma eq_finite_iUnion_of_isTopologicalBasis_of_isCompact_open (b : ι → Set X subst this obtain ⟨t, ht⟩ := hUc.elim_finite_subcover (b ∘ f') (fun i => hb.isOpen (Set.mem_range_self _)) (by rw [e]) + classical refine ⟨t.image f', Set.toFinite _, le_antisymm ?_ ?_⟩ · refine Set.Subset.trans ht ?_ simp only [Set.iUnion_subset_iff] diff --git a/Mathlib/Topology/Compactness/LocallyCompact.lean b/Mathlib/Topology/Compactness/LocallyCompact.lean index 51c7fa91e8a0b..346a860454e8b 100644 --- a/Mathlib/Topology/Compactness/LocallyCompact.lean +++ b/Mathlib/Topology/Compactness/LocallyCompact.lean @@ -12,7 +12,8 @@ We define the following classes of topological spaces: * `LocallyCompactSpace`: for every point `x`, every open neighborhood of `x` contains a compact neighborhood of `x`. The definition is formulated in terms of the neighborhood filter. -/ -open Set Filter Topology TopologicalSpace Classical + +open Set Filter Topology TopologicalSpace variable {X : Type*} {Y : Type*} {ι : Type*} variable [TopologicalSpace X] [TopologicalSpace Y] {s t : Set X} @@ -101,7 +102,8 @@ instance Pi.locallyCompactSpace [∀ i, CompactSpace (X i)] : LocallyCompactSpac refine ⟨s.pi n'', ?_, subset_trans (fun _ => ?_) hsub, ?_⟩ · exact (set_pi_mem_nhds_iff hs _).mpr fun i _ => hn'' i · exact forall₂_imp fun i _ hi' => hsub' i hi' - · rw [← Set.univ_pi_ite] + · classical + rw [← Set.univ_pi_ite] refine isCompact_univ_pi fun i => ?_ by_cases h : i ∈ s · rw [if_pos h] diff --git a/Mathlib/Topology/Compactness/SigmaCompact.lean b/Mathlib/Topology/Compactness/SigmaCompact.lean index 6b9d8a254f801..7518114ce3802 100644 --- a/Mathlib/Topology/Compactness/SigmaCompact.lean +++ b/Mathlib/Topology/Compactness/SigmaCompact.lean @@ -13,7 +13,8 @@ import Mathlib.Topology.Compactness.LocallyCompact of a countable collection of compact subspaces. -/ -open Set Filter Topology TopologicalSpace Classical + +open Set Filter Topology TopologicalSpace universe u v @@ -367,15 +368,18 @@ theorem exists_superset_of_isCompact {s : Set X} (hs : IsCompact s) : ∃ n, s exact mem_iUnion.2 ⟨k + 1, K.subset_interior_succ _ hk⟩ · exact Monotone.directed_le fun _ _ h ↦ interior_mono <| K.subset h +open Classical in /-- The minimal `n` such that `x ∈ K n`. -/ protected noncomputable def find (x : X) : ℕ := Nat.find (K.exists_mem x) -theorem mem_find (x : X) : x ∈ K (K.find x) := - Nat.find_spec (K.exists_mem x) +theorem mem_find (x : X) : x ∈ K (K.find x) := by + classical + exact Nat.find_spec (K.exists_mem x) -theorem mem_iff_find_le {x : X} {n : ℕ} : x ∈ K n ↔ K.find x ≤ n := - ⟨fun h => Nat.find_min' (K.exists_mem x) h, fun h => K.subset h <| K.mem_find x⟩ +theorem mem_iff_find_le {x : X} {n : ℕ} : x ∈ K n ↔ K.find x ≤ n := by + classical + exact ⟨fun h => Nat.find_min' (K.exists_mem x) h, fun h => K.subset h <| K.mem_find x⟩ /-- Prepend the empty set to a compact exhaustion `K n`. -/ def shiftr : CompactExhaustion X where @@ -385,8 +389,9 @@ def shiftr : CompactExhaustion X where iUnion_eq' := iUnion_eq_univ_iff.2 fun x => ⟨K.find x + 1, K.mem_find x⟩ @[simp] -theorem find_shiftr (x : X) : K.shiftr.find x = K.find x + 1 := - Nat.find_comp_succ _ _ (not_mem_empty _) +theorem find_shiftr (x : X) : K.shiftr.find x = K.find x + 1 := by + classical + exact Nat.find_comp_succ _ _ (not_mem_empty _) theorem mem_diff_shiftr_find (x : X) : x ∈ K.shiftr (K.find x + 1) \ K.shiftr (K.find x) := ⟨K.mem_find _, diff --git a/Mathlib/Topology/Connected/Basic.lean b/Mathlib/Topology/Connected/Basic.lean index af3cfad6372d4..8d8d2a542f23d 100644 --- a/Mathlib/Topology/Connected/Basic.lean +++ b/Mathlib/Topology/Connected/Basic.lean @@ -34,9 +34,7 @@ and in particular https://ncatlab.org/nlab/show/too+simple+to+be+simple#relationship_to_biased_definitions. -/ - open Set Function Topology TopologicalSpace Relation -open scoped Classical universe u v @@ -425,6 +423,7 @@ theorem IsConnected.prod [TopologicalSpace β] {s : Set α} {t : Set β} (hs : I theorem isPreconnected_univ_pi [∀ i, TopologicalSpace (π i)] {s : ∀ i, Set (π i)} (hs : ∀ i, IsPreconnected (s i)) : IsPreconnected (pi univ s) := by rintro u v uo vo hsuv ⟨f, hfs, hfu⟩ ⟨g, hgs, hgv⟩ + classical rcases exists_finset_piecewise_mem_of_mem_nhds (uo.mem_nhds hfu) g with ⟨I, hI⟩ induction' I using Finset.induction_on with i I _ ihI · refine ⟨g, hgs, ⟨?_, hgv⟩⟩ @@ -457,6 +456,7 @@ that contains this point. -/ def connectedComponent (x : α) : Set α := ⋃₀ { s : Set α | IsPreconnected s ∧ x ∈ s } +open Classical in /-- Given a set `F` in a topological space `α` and a point `x : α`, the connected component of `x` in `F` is the connected component of `x` in the subtype `F` seen as a set in `α`. This definition does not make sense if `x` is not in `F` so we return the diff --git a/Mathlib/Topology/Connected/Clopen.lean b/Mathlib/Topology/Connected/Clopen.lean index ba60382112ad4..e6f92357c8096 100644 --- a/Mathlib/Topology/Connected/Clopen.lean +++ b/Mathlib/Topology/Connected/Clopen.lean @@ -22,7 +22,6 @@ to clopen sets. -/ open Set Function Topology TopologicalSpace Relation -open scoped Classical universe u v @@ -278,6 +277,7 @@ theorem isConnected_iff_sUnion_disjoint_open {s : Set α} : ∀ U : Finset (Set α), (∀ u v : Set α, u ∈ U → v ∈ U → (s ∩ (u ∩ v)).Nonempty → u = v) → (∀ u ∈ U, IsOpen u) → (s ⊆ ⋃₀ ↑U) → ∃ u ∈ U, s ⊆ u := by rw [IsConnected, isPreconnected_iff_subset_of_disjoint] + classical refine ⟨fun ⟨hne, h⟩ U hU hUo hsU => ?_, fun h => ⟨?_, fun u v hu hv hs hsuv => ?_⟩⟩ · induction U using Finset.induction_on with | empty => exact absurd (by simpa using hsU) hne.not_subset_empty diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 8c43399801e63..51eb2dee9f145 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -30,10 +30,8 @@ product, sum, disjoint union, subspace, quotient space -/ - noncomputable section -open scoped Classical open Topology TopologicalSpace Set Filter Function universe u v @@ -1273,6 +1271,7 @@ theorem isOpen_pi_iff {s : Set (∀ a, π a)} : · exact Subset.trans (pi_mono fun i hi => (eval_image_pi_subset hi).trans (h1 i).choose_spec.1) h2 · rintro ⟨I, t, ⟨h1, h2⟩⟩ + classical refine ⟨I, fun a => ite (a ∈ I) (t a) univ, fun i => ?_, ?_⟩ · by_cases hi : i ∈ I · use t i @@ -1339,7 +1338,8 @@ theorem pi_generateFrom_eq {π : ι → Type*} {g : ∀ a, Set (Set (π a))} : rintro _ ⟨s, i, hi, rfl⟩ letI := fun a => generateFrom (g a) exact isOpen_set_pi i.finite_toSet (fun a ha => GenerateOpen.basic _ (hi a ha)) - · refine le_iInf fun i => coinduced_le_iff_le_induced.1 <| le_generateFrom fun s hs => ?_ + · classical + refine le_iInf fun i => coinduced_le_iff_le_induced.1 <| le_generateFrom fun s hs => ?_ refine GenerateOpen.basic _ ⟨update (fun i => univ) i s, {i}, ?_⟩ simp [hs] @@ -1365,6 +1365,7 @@ theorem pi_generateFrom_eq_finite {π : ι → Type*} {g : ∀ a, Set (Set (π a refine isOpen_iff_forall_mem_open.2 fun f hf => ?_ choose c hcg hfc using fun a => sUnion_eq_univ_iff.1 (hg a) (f a) refine ⟨pi i t ∩ pi ((↑i)ᶜ : Set ι) c, inter_subset_left, ?_, ⟨hf, fun a _ => hfc a⟩⟩ + classical rw [← univ_pi_piecewise] refine GenerateOpen.basic _ ⟨_, fun a => ?_, rfl⟩ by_cases a ∈ i <;> simp [*] diff --git a/Mathlib/Topology/ExtremallyDisconnected.lean b/Mathlib/Topology/ExtremallyDisconnected.lean index 285752c080264..bdba4f48c8dba 100644 --- a/Mathlib/Topology/ExtremallyDisconnected.lean +++ b/Mathlib/Topology/ExtremallyDisconnected.lean @@ -30,7 +30,6 @@ compact Hausdorff spaces. noncomputable section -open scoped Classical open Function Set universe u diff --git a/Mathlib/Topology/FiberBundle/Trivialization.lean b/Mathlib/Topology/FiberBundle/Trivialization.lean index 463d0e8552fbb..cd79bc62c044f 100644 --- a/Mathlib/Topology/FiberBundle/Trivialization.lean +++ b/Mathlib/Topology/FiberBundle/Trivialization.lean @@ -48,8 +48,7 @@ type of linear trivializations is not even particularly well-behaved. -/ open TopologicalSpace Filter Set Bundle Function - -open scoped Topology Classical Bundle +open scoped Topology variable {ι : Type*} {B : Type*} {F : Type*} {E : B → Type*} variable (F) {Z : Type*} [TopologicalSpace B] [TopologicalSpace F] {proj : Z → B} @@ -216,6 +215,7 @@ section Zero variable [∀ x, Zero (E x)] +open Classical in /-- A fiberwise inverse to `e`. This is the function `F → E b` that induces a local inverse `B × F → TotalSpace F E` of `e` on `e.baseSet`. It is defined to be `0` outside `e.baseSet`. -/ protected noncomputable def symm (e : Pretrivialization F (π F E)) (b : B) (y : F) : E b := @@ -666,6 +666,7 @@ theorem frontier_preimage (e : Trivialization F proj) (s : Set B) : rw [← (e.isImage_preimage_prod s).frontier.preimage_eq, frontier_prod_univ_eq, (e.isImage_preimage_prod _).preimage_eq, e.source_eq, preimage_inter] +open Classical in /-- Given two bundle trivializations `e`, `e'` of `proj : Z → B` and a set `s : Set B` such that the base sets of `e` and `e'` intersect `frontier s` on the same set and `e p = e' p` whenever `proj p ∈ e.baseSet ∩ frontier s`, `e.piecewise e' s Hs Heq` is the bundle trivialization over @@ -713,6 +714,7 @@ noncomputable def piecewiseLe [LinearOrder B] [OrderTopology B] (e e' : Triviali · simp [e.coe_fst', e'.coe_fst', *] · simp [coordChange_apply_snd, *] +open Classical in /-- Given two bundle trivializations `e`, `e'` over disjoint sets, `e.disjoint_union e' H` is the bundle trivialization over the union of the base sets that agrees with `e` and `e'` over their base sets. -/ diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 2b2eb2b8c31e2..338cc5c4ba0ef 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -17,7 +17,7 @@ import Mathlib.Topology.Order.T5 noncomputable section open Set Filter Metric Function -open scoped Classical Topology ENNReal NNReal Filter +open scoped Topology ENNReal NNReal variable {α : Type*} {β : Type*} {γ : Type*} @@ -348,6 +348,7 @@ protected theorem Tendsto.mul_const {f : Filter α} {m : α → ℝ≥0∞} {a b theorem tendsto_finset_prod_of_ne_top {ι : Type*} {f : ι → α → ℝ≥0∞} {x : Filter α} {a : ι → ℝ≥0∞} (s : Finset ι) (h : ∀ i ∈ s, Tendsto (f i) x (𝓝 (a i))) (h' : ∀ i ∈ s, a i ≠ ∞) : Tendsto (fun b => ∏ c ∈ s, f c b) x (𝓝 (∏ c ∈ s, a c)) := by + classical induction' s using Finset.induction with a s has IH · simp [tendsto_const_nhds] simp only [Finset.prod_insert has] @@ -920,6 +921,7 @@ theorem tsum_union_le (f : α → ℝ≥0∞) (s t : Set α) : calc ∑' x : ↑(s ∪ t), f x = ∑' x : ⋃ b, cond b s t, f x := tsum_congr_set_coe _ union_eq_iUnion _ ≤ _ := by simpa using tsum_iUnion_le f (cond · s t) +open Classical in theorem tsum_eq_add_tsum_ite {f : β → ℝ≥0∞} (b : β) : ∑' x, f x = f b + ∑' x, ite (x = b) 0 (f x) := tsum_eq_add_tsum_ite' b ENNReal.summable @@ -1064,6 +1066,7 @@ theorem summable_sigma {β : α → Type*} {f : (Σ x, β x) → ℝ≥0} : theorem indicator_summable {f : α → ℝ≥0} (hf : Summable f) (s : Set α) : Summable (s.indicator f) := by + classical refine NNReal.summable_of_le (fun a => le_trans (le_of_eq (s.indicator_apply f a)) ?_) hf split_ifs · exact le_refl (f a) @@ -1109,6 +1112,7 @@ theorem tsum_pos {g : α → ℝ≥0} (hg : Summable g) (i : α) (hi : 0 < g i) rw [← tsum_zero] exact tsum_lt_tsum (fun a => zero_le _) hi hg +open Classical in theorem tsum_eq_add_tsum_ite {f : α → ℝ≥0} (hf : Summable f) (i : α) : ∑' x, f x = f i + ∑' x, ite (x = i) 0 (f x) := by refine tsum_eq_add_tsum_ite' i (NNReal.summable_of_le (fun i' => ?_) hf) diff --git a/Mathlib/Topology/Irreducible.lean b/Mathlib/Topology/Irreducible.lean index c66f3d0b3ecbd..fe622ea65817c 100644 --- a/Mathlib/Topology/Irreducible.lean +++ b/Mathlib/Topology/Irreducible.lean @@ -27,7 +27,7 @@ https://ncatlab.org/nlab/show/too+simple+to+be+simple#relationship_to_biased_def -/ -open Set Classical +open Set variable {X : Type*} {Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {s t : Set X} @@ -220,6 +220,7 @@ theorem isIrreducible_iff_sInter : IsIrreducible s ↔ ∀ (U : Finset (Set X)), (∀ u ∈ U, IsOpen u) → (∀ u ∈ U, (s ∩ u).Nonempty) → (s ∩ ⋂₀ ↑U).Nonempty := by + classical refine ⟨fun h U hu hU => ?_, fun h => ⟨?_, ?_⟩⟩ · induction U using Finset.induction_on with | empty => simpa using h.nonempty @@ -271,6 +272,7 @@ theorem IsPreirreducible.subset_irreducible {S U : Set X} (ht : IsPreirreducible replace ht : IsIrreducible t := ⟨⟨z, h₂ (h₁ hz)⟩, ht⟩ refine ⟨⟨z, h₁ hz⟩, ?_⟩ rintro u v hu hv ⟨x, hx, hx'⟩ ⟨y, hy, hy'⟩ + classical obtain ⟨x, -, hx'⟩ : Set.Nonempty (t ∩ ⋂₀ ↑({U, u, v} : Finset (Set X))) := by refine isIrreducible_iff_sInter.mp ht {U, u, v} ?_ ?_ · simp [*] diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean index d86b7fcd9f9ff..26357a33b0c8c 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean @@ -41,20 +41,16 @@ i.e., it is complete and second countable. We also prove the Gromov compactness noncomputable section -open scoped Classical Topology ENNReal Cardinal +open scoped Topology ENNReal Cardinal +open Set Function TopologicalSpace Filter Metric Quotient Bornology +open BoundedContinuousFunction Nat Int kuratowskiEmbedding +open Sum (inl inr) local notation "ℓ_infty_ℝ" => lp (fun n : ℕ => ℝ) ∞ universe u v w -open scoped Classical -open Set Function TopologicalSpace Filter Metric Quotient Bornology - -open BoundedContinuousFunction Nat Int kuratowskiEmbedding - -open Sum (inl inr) - attribute [local instance] metricSpaceSum namespace GromovHausdorff diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean index 30d66364942da..daa03f9b799fb 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean @@ -32,16 +32,9 @@ space structure on `X ⊕ Y`. The corresponding metric quotient is `OptimalGHCou noncomputable section -open scoped Classical -open Topology NNReal - universe u v w -open scoped Classical -open Set Function TopologicalSpace Filter Metric Quotient - -open BoundedContinuousFunction - +open Topology NNReal Set Function TopologicalSpace Filter Metric Quotient BoundedContinuousFunction open Sum (inl inr) attribute [local instance] metricSpaceSum diff --git a/Mathlib/Topology/PartitionOfUnity.lean b/Mathlib/Topology/PartitionOfUnity.lean index a439497b79088..a98ad48c12803 100644 --- a/Mathlib/Topology/PartitionOfUnity.lean +++ b/Mathlib/Topology/PartitionOfUnity.lean @@ -73,13 +73,9 @@ We use `WellOrderingRel j i` instead of `j < i` in the definition of partition of unity, bump function, Urysohn's lemma, normal space, paracompact space -/ - universe u v -open Function Set Filter - -open scoped Classical -open Topology +open Function Set Filter Topology noncomputable section @@ -333,6 +329,7 @@ theorem nonneg (i : ι) (x : X) : 0 ≤ f i x := theorem le_one (i : ι) (x : X) : f i x ≤ 1 := f.le_one' i x +open Classical in /-- A `BumpCovering` that consists of a single function, uniformly equal to one, defined as an example for `Inhabited` instance. -/ protected def single (i : ι) (s : Set X) : BumpCovering ι X s where @@ -347,8 +344,10 @@ protected def single (i : ι) (s : Set X) : BumpCovering ι X s where le_one' := update_le_iff.2 ⟨le_rfl, fun _ _ _ => zero_le_one⟩ eventuallyEq_one' x _ := ⟨i, by rw [Pi.single_eq_same, ContinuousMap.coe_one]⟩ +open Classical in @[simp] -theorem coe_single (i : ι) (s : Set X) : ⇑(BumpCovering.single i s) = Pi.single i 1 := rfl +theorem coe_single (i : ι) (s : Set X) : ⇑(BumpCovering.single i s) = Pi.single i 1 := by + rfl instance [Inhabited ι] : Inhabited (BumpCovering ι X s) := ⟨BumpCovering.single default s⟩ @@ -453,6 +452,7 @@ theorem toPOUFun_zero_of_zero {i : ι} {x : X} (h : f i x = 0) : f.toPOUFun i x theorem support_toPOUFun_subset (i : ι) : support (f.toPOUFun i) ⊆ support (f i) := fun _ => mt <| f.toPOUFun_zero_of_zero +open Classical in theorem toPOUFun_eq_mul_prod (i : ι) (x : X) (t : Finset ι) (ht : ∀ j, WellOrderingRel j i → f j x ≠ 0 → j ∈ t) : f.toPOUFun i x = f i x * ∏ j ∈ t.filter fun j => WellOrderingRel j i, (1 - f j x) := by @@ -470,6 +470,7 @@ theorem sum_toPOUFun_eq (x : X) : ∑ᶠ i, f.toPOUFun i x = 1 - ∏ᶠ i, (1 - have B : (mulSupport fun i => 1 - f i x) ⊆ s := by rw [hs, mulSupport_one_sub] exact fun i => id + classical letI : LinearOrder ι := linearOrderOfSTO WellOrderingRel rw [finsum_eq_sum_of_support_subset _ A, finprod_eq_prod_of_mulSupport_subset _ B, Finset.prod_one_sub_ordered, sub_sub_cancel] @@ -477,6 +478,7 @@ theorem sum_toPOUFun_eq (x : X) : ∑ᶠ i, f.toPOUFun i x = 1 - ∏ᶠ i, (1 - convert f.toPOUFun_eq_mul_prod _ _ _ fun j _ hj => _ rwa [Finite.mem_toFinset] +open Classical in theorem exists_finset_toPOUFun_eventuallyEq (i : ι) (x : X) : ∃ t : Finset ι, f.toPOUFun i =ᶠ[𝓝 x] f i * ∏ j ∈ t.filter fun j => WellOrderingRel j i, (1 - f j) := by rcases f.locallyFinite x with ⟨U, hU, hf⟩ @@ -519,11 +521,13 @@ def toPartitionOfUnity : PartitionOfUnity ι X s where theorem toPartitionOfUnity_apply (i : ι) (x : X) : f.toPartitionOfUnity i x = f i x * ∏ᶠ (j) (_ : WellOrderingRel j i), (1 - f j x) := rfl +open Classical in theorem toPartitionOfUnity_eq_mul_prod (i : ι) (x : X) (t : Finset ι) (ht : ∀ j, WellOrderingRel j i → f j x ≠ 0 → j ∈ t) : f.toPartitionOfUnity i x = f i x * ∏ j ∈ t.filter fun j => WellOrderingRel j i, (1 - f j x) := f.toPOUFun_eq_mul_prod i x t ht +open Classical in theorem exists_finset_toPartitionOfUnity_eventuallyEq (i : ι) (x : X) : ∃ t : Finset ι, f.toPartitionOfUnity i =ᶠ[𝓝 x] f i * ∏ j ∈ t.filter fun j => WellOrderingRel j i, (1 - f j) := f.exists_finset_toPOUFun_eventuallyEq i x diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index e29cc2436cc61..42a2218df3fec 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -124,10 +124,7 @@ If the space is also Lindelöf: -/ - - open Function Set Filter Topology TopologicalSpace -open scoped Classical universe u v @@ -807,6 +804,7 @@ theorem Dense.diff_singleton [T1Space X] {s : Set X} (hs : Dense s) (x : X) [NeB obtains a dense set. -/ theorem Dense.diff_finset [T1Space X] [∀ x : X, NeBot (𝓝[≠] x)] {s : Set X} (hs : Dense s) (t : Finset X) : Dense (s \ t) := by + classical induction t using Finset.induction_on with | empty => simpa using hs | insert _ ih => @@ -1138,6 +1136,7 @@ theorem IsCompact.binary_compact_cover {K U V : Set X} theorem IsCompact.finite_compact_cover {s : Set X} (hs : IsCompact s) {ι : Type*} (t : Finset ι) (U : ι → Set X) (hU : ∀ i ∈ t, IsOpen (U i)) (hsC : s ⊆ ⋃ i ∈ t, U i) : ∃ K : ι → Set X, (∀ i, IsCompact (K i)) ∧ (∀ i, K i ⊆ U i) ∧ s = ⋃ i ∈ t, K i := by + classical induction' t using Finset.induction with x t hx ih generalizing U s · refine ⟨fun _ => ∅, fun _ => isCompact_empty, fun i => empty_subset _, ?_⟩ simpa only [subset_empty_iff, Finset.not_mem_empty, iUnion_false, iUnion_empty] using hsC diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index 8e4dd9510854c..6124e15d59b41 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -26,11 +26,8 @@ We prove two versions of the lemma: normal space, shrinking lemma -/ - open Set Function -open scoped Classical - noncomputable section variable {ι X : Type*} [TopologicalSpace X] [NormalSpace X] @@ -70,9 +67,11 @@ variable {u : ι → Set X} {s : Set X} instance : CoeFun (PartialRefinement u s) fun _ => ι → Set X := ⟨toFun⟩ -protected theorem subset (v : PartialRefinement u s) (i : ι) : v i ⊆ u i := - if h : i ∈ v.carrier then subset_closure.trans (v.closure_subset h) else (v.apply_eq h).le +protected theorem subset (v : PartialRefinement u s) (i : ι) : v i ⊆ u i := by + classical + exact if h : i ∈ v.carrier then subset_closure.trans (v.closure_subset h) else (v.apply_eq h).le +open Classical in instance : PartialOrder (PartialRefinement u s) where le v₁ v₂ := v₁.carrier ⊆ v₂.carrier ∧ ∀ i ∈ v₁.carrier, v₁ i = v₂ i le_refl v := ⟨Subset.refl _, fun _ _ => rfl⟩ @@ -98,6 +97,7 @@ their carriers. -/ def chainSupCarrier (c : Set (PartialRefinement u s)) : Set ι := ⋃ v ∈ c, carrier v +open Classical in /-- Choice of an element of a nonempty chain of partial refinements. If `i` belongs to one of `carrier v`, `v ∈ c`, then `find c ne i` is one of these partial refinements. -/ def find (c : Set (PartialRefinement u s)) (ne : c.Nonempty) (i : ι) : PartialRefinement u s := @@ -164,6 +164,7 @@ theorem exists_gt (v : PartialRefinement u s) (hs : IsClosed s) (i : ι) (hi : i have C : IsClosed (s ∩ ⋂ (j) (_ : j ≠ i), (v j)ᶜ) := IsClosed.inter hs (isClosed_biInter fun _ _ => isClosed_compl_iff.2 <| v.isOpen _) rcases normal_exists_closure_subset C (v.isOpen i) I with ⟨vi, ovi, hvi, cvi⟩ + classical refine ⟨⟨update v i vi, insert i v.carrier, ?_, ?_, ?_, ?_⟩, ?_, ?_⟩ · intro j rcases eq_or_ne j i with (rfl| hne) <;> simp [*, v.isOpen] diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index c7d5336049b68..eadf9a7d228a9 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -11,14 +11,9 @@ import Mathlib.Topology.UniformSpace.Basic # Theory of Cauchy filters in uniform spaces. Complete uniform spaces. Totally bounded subsets. -/ - universe u v -open scoped Classical -open Filter TopologicalSpace Set UniformSpace Function - -open scoped Classical -open Uniformity Topology Filter +open Filter Function TopologicalSpace Topology Set UniformSpace Uniformity variable {α : Type u} {β : Type v} [uniformSpace : UniformSpace α] diff --git a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean index 9a1597f6bab59..bc5bf5abab8fd 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean @@ -131,10 +131,9 @@ connection API to do most of the work. uniform convergence -/ - noncomputable section -open scoped Classical Topology Uniformity +open scoped Topology Uniformity open Set Filter section TypeAlias diff --git a/Mathlib/Topology/UnitInterval.lean b/Mathlib/Topology/UnitInterval.lean index 250d8ef87e1ff..1a8a9c4a3a525 100644 --- a/Mathlib/Topology/UnitInterval.lean +++ b/Mathlib/Topology/UnitInterval.lean @@ -19,10 +19,7 @@ We provide basic instances, as well as a custom tactic for discharging noncomputable section -open scoped Classical -open Topology Filter - -open Set Int Set.Icc +open Topology Filter Set Int Set.Icc /-! ### The unit interval -/ From 78fb0af360f6f03e25ee06b76581b75e097d9b91 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Sun, 4 Aug 2024 13:29:38 +0000 Subject: [PATCH 023/359] feat(LightCondensed): characterise epimorphisms in light condensed sets and modules (#13495) This PR characterises epimorphisms in light condensed sets and modules as the locally surjective morphisms. Here, the condition of locally surjective is phrased in terms of continuous surjections of light profinite sets. --- Mathlib.lean | 1 + Mathlib/Condensed/Light/Epi.lean | 66 ++++++++++++++++++++++++++++++++ 2 files changed, 67 insertions(+) create mode 100644 Mathlib/Condensed/Light/Epi.lean diff --git a/Mathlib.lean b/Mathlib.lean index a2a2b713a4b21..c18909907b036 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1967,6 +1967,7 @@ import Mathlib.Condensed.Equivalence import Mathlib.Condensed.Explicit import Mathlib.Condensed.Functors import Mathlib.Condensed.Light.Basic +import Mathlib.Condensed.Light.Epi import Mathlib.Condensed.Light.Explicit import Mathlib.Condensed.Light.Functors import Mathlib.Condensed.Light.Module diff --git a/Mathlib/Condensed/Light/Epi.lean b/Mathlib/Condensed/Light/Epi.lean new file mode 100644 index 0000000000000..d1431c31a99fa --- /dev/null +++ b/Mathlib/Condensed/Light/Epi.lean @@ -0,0 +1,66 @@ +/- +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.ConcreteCategory.EpiMono +import Mathlib.CategoryTheory.Sites.EpiMono +import Mathlib.CategoryTheory.Sites.Coherent.LocallySurjective +import Mathlib.Condensed.Light.Module +/-! + +# Epimorphisms of light condensed objects + +This file characterises epimorphisms in light condensed sets and modules as the locally surjective +morphisms. Here, the condition of locally surjective is phrased in terms of continuous surjections +of light profinite sets. +-/ + +universe v u w u' v' + +open CategoryTheory Sheaf Limits ConcreteCategory GrothendieckTopology + +attribute [local instance] ConcreteCategory.hasCoeToSort ConcreteCategory.instFunLike + +namespace LightCondensed + +variable (A : Type u') [Category.{v'} A] [ConcreteCategory.{w} A] + [PreservesFiniteProducts (CategoryTheory.forget A)] + +variable {X Y : LightCondensed.{u} A} (f : X ⟶ Y) + +lemma isLocallySurjective_iff_locallySurjective_on_lightProfinite : IsLocallySurjective f ↔ + ∀ (S : LightProfinite) (y : Y.val.obj ⟨S⟩), + (∃ (S' : LightProfinite) (φ : S' ⟶ S) (_ : Function.Surjective φ) (x : X.val.obj ⟨S'⟩), + f.val.app ⟨S'⟩ x = Y.val.map ⟨φ⟩ y) := by + rw [coherentTopology.isLocallySurjective_iff, + regularTopology.isLocallySurjective_iff] + simp_rw [LightProfinite.effectiveEpi_iff_surjective] + +end LightCondensed + +namespace LightCondSet + +variable {X Y : LightCondSet.{u}} (f : X ⟶ Y) + +lemma epi_iff_locallySurjective_on_lightProfinite : Epi f ↔ + ∀ (S : LightProfinite) (y : Y.val.obj ⟨S⟩), + (∃ (S' : LightProfinite) (φ : S' ⟶ S) (_ : Function.Surjective φ) (x : X.val.obj ⟨S'⟩), + f.val.app ⟨S'⟩ x = Y.val.map ⟨φ⟩ y) := by + rw [← isLocallySurjective_iff_epi'] + exact LightCondensed.isLocallySurjective_iff_locallySurjective_on_lightProfinite _ f + +end LightCondSet + +namespace LightCondMod + +variable (R : Type u) [Ring R] {X Y : LightCondMod.{u} R} (f : X ⟶ Y) + +lemma epi_iff_locallySurjective_on_lightProfinite : Epi f ↔ + ∀ (S : LightProfinite) (y : Y.val.obj ⟨S⟩), + (∃ (S' : LightProfinite) (φ : S' ⟶ S) (_ : Function.Surjective φ) (x : X.val.obj ⟨S'⟩), + f.val.app ⟨S'⟩ x = Y.val.map ⟨φ⟩ y) := by + rw [← isLocallySurjective_iff_epi'] + exact LightCondensed.isLocallySurjective_iff_locallySurjective_on_lightProfinite _ f + +end LightCondMod From 26b8d0d1e09a38db4b9fa4667b8792225bc27d6a Mon Sep 17 00:00:00 2001 From: grunweg Date: Sun, 4 Aug 2024 13:29:40 +0000 Subject: [PATCH 024/359] chore: run fix-comments.py on all of mathlib (#15222) --- Mathlib/Algebra/Algebra/Unitization.lean | 2 +- Mathlib/Algebra/Category/Grp/Adjunctions.lean | 4 +- .../Grp/EquivalenceGroupAddGroup.lean | 8 +-- .../Algebra/Category/ModuleCat/Abelian.lean | 12 ++-- .../ModuleCat/Differentials/Presheaf.lean | 2 +- Mathlib/Algebra/Group/Semiconj/Defs.lean | 2 +- Mathlib/Algebra/Group/Semiconj/Units.lean | 2 +- Mathlib/Algebra/MvPolynomial/Basic.lean | 2 +- Mathlib/Algebra/Vertex/HVertexOperator.lean | 2 +- .../ProjectiveSpectrum/StructureSheaf.lean | 2 +- Mathlib/Analysis/Analytic/Composition.lean | 46 +++++++-------- Mathlib/Analysis/Convolution.lean | 22 +++---- .../Analysis/CstarAlgebra/Exponential.lean | 4 +- .../Analysis/Normed/Algebra/Exponential.lean | 57 ++++++++++-------- .../Normed/Algebra/MatrixExponential.lean | 5 +- .../Normed/Algebra/TrivSqZeroExt.lean | 11 ++-- .../Analysis/Normed/Group/InfiniteSum.lean | 2 +- .../SpecialFunctions/Exponential.lean | 59 ++++++++++--------- Mathlib/CategoryTheory/Abelian/Ext.lean | 5 +- Mathlib/CategoryTheory/Category/RelCat.lean | 6 +- .../CategoryTheory/Groupoid/Subgroupoid.lean | 2 +- Mathlib/CategoryTheory/Linear/Yoneda.lean | 12 ++-- Mathlib/Computability/TMToPartrec.lean | 2 +- Mathlib/Control/Monad/Writer.lean | 4 +- Mathlib/Data/Finset/NAry.lean | 4 +- Mathlib/Data/Finset/Union.lean | 2 +- Mathlib/Data/Int/Interval.lean | 2 +- Mathlib/Data/Nat/Squarefree.lean | 8 +-- Mathlib/Data/Quot.lean | 2 +- Mathlib/Data/Real/Pi/Bounds.lean | 2 +- Mathlib/Data/ZMod/Algebra.lean | 4 +- Mathlib/FieldTheory/Finite/Polynomial.lean | 4 +- .../RingedSpace/LocallyRingedSpace.lean | 12 ++-- .../LocallyRingedSpace/HasColimits.lean | 2 +- .../Geometry/RingedSpace/OpenImmersion.lean | 2 +- Mathlib/Init/Data/Nat/Lemmas.lean | 2 +- Mathlib/Lean/Meta/CongrTheorems.lean | 2 +- .../Eigenspace/Triangularizable.lean | 2 +- Mathlib/LinearAlgebra/Multilinear/Basis.lean | 2 +- .../LinearAlgebra/QuadraticForm/Basic.lean | 2 +- Mathlib/Logic/Equiv/Array.lean | 4 +- Mathlib/Logic/Nontrivial/Defs.lean | 2 +- .../Covering/BesicovitchVectorSpace.lean | 2 +- .../Function/LpSeminorm/Basic.lean | 2 +- .../MeasureTheory/Integral/PeakFunction.lean | 2 +- .../ModularForms/CongruenceSubgroups.lean | 8 +-- Mathlib/Order/Circular.lean | 2 +- Mathlib/Order/Filter/Bases.lean | 2 +- Mathlib/Order/RelIso/Basic.lean | 2 +- .../Kernel/Disintegration/CondCdf.lean | 2 +- .../RepresentationTheory/Action/Basic.lean | 2 +- .../RepresentationTheory/Action/Concrete.lean | 2 +- .../RepresentationTheory/Action/Monoidal.lean | 2 +- Mathlib/RingTheory/Adjoin/Field.lean | 6 +- Mathlib/RingTheory/DedekindDomain/Basic.lean | 2 +- .../DedekindDomain/FiniteAdeleRing.lean | 2 +- Mathlib/RingTheory/Derivation/Basic.lean | 2 +- .../RingTheory/Derivation/ToSquareZero.lean | 2 +- Mathlib/RingTheory/LocalProperties.lean | 6 +- .../RingTheory/Localization/BaseChange.lean | 2 +- Mathlib/RingTheory/Polynomial/Bernstein.lean | 2 +- Mathlib/RingTheory/Trace/Basic.lean | 2 +- Mathlib/RingTheory/Valuation/Basic.lean | 2 +- Mathlib/RingTheory/WittVector/Compare.lean | 8 +-- Mathlib/RingTheory/WittVector/IsPoly.lean | 2 +- Mathlib/Tactic/Conv.lean | 2 +- Mathlib/Testing/SlimCheck/Functions.lean | 8 +-- Mathlib/Testing/SlimCheck/Sampleable.lean | 10 ++-- Mathlib/Topology/Algebra/Ring/Basic.lean | 2 +- Mathlib/Topology/Category/TopCat/Opens.lean | 2 +- .../Topology/FiberBundle/Constructions.lean | 2 +- .../Topology/MetricSpace/GromovHausdorff.lean | 14 ++--- Mathlib/Topology/Sheaves/Forget.lean | 2 +- .../Topology/Sheaves/PresheafOfFunctions.lean | 2 +- 74 files changed, 229 insertions(+), 216 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Unitization.lean b/Mathlib/Algebra/Algebra/Unitization.lean index 6c7d8934847e7..491a87438106e 100644 --- a/Mathlib/Algebra/Algebra/Unitization.lean +++ b/Mathlib/Algebra/Algebra/Unitization.lean @@ -608,7 +608,7 @@ def inrNonUnitalAlgHom (R A : Type*) [CommSemiring R] [NonUnitalSemiring A] [Mod map_add' := inr_add R map_mul' := inr_mul R -/-- The coercion from a non-unital `R`-algebra `A` to its unitization `unitization R A` +/-- The coercion from a non-unital `R`-algebra `A` to its unitization `Unitization R A` realized as a non-unital star algebra homomorphism. -/ @[simps!] def inrNonUnitalStarAlgHom (R A : Type*) [CommSemiring R] [StarAddMonoid R] diff --git a/Mathlib/Algebra/Category/Grp/Adjunctions.lean b/Mathlib/Algebra/Category/Grp/Adjunctions.lean index 2cbb769fcc95c..18557f976c125 100644 --- a/Mathlib/Algebra/Category/Grp/Adjunctions.lean +++ b/Mathlib/Algebra/Category/Grp/Adjunctions.lean @@ -19,7 +19,7 @@ category of abelian groups. with generators `x : X`. * `Grp.free`: constructs the functor associating to a type `X` the free group with generators `x : X`. -* `abelianize`: constructs the functor which associates to a group `G` its abelianization `Gᵃᵇ`. +* `Grp.abelianize`: constructs the functor which associates to a group `G` its abelianization `Gᵃᵇ`. ## Main statements @@ -27,7 +27,7 @@ category of abelian groups. of the forgetful functor from abelian groups to types. * `Grp.adj`: proves that `Grp.free` is the left adjoint of the forgetful functor from groups to types. -* `abelianizeAdj`: proves that `abelianize` is left adjoint to the forgetful functor from +* `abelianizeAdj`: proves that `Grp.abelianize` is left adjoint to the forgetful functor from abelian groups to groups. -/ diff --git a/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean b/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean index 0f8741b4e733e..4c58186eba37a 100644 --- a/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean +++ b/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean @@ -27,7 +27,7 @@ private instance (X : CommGrp) : MulOneClass X.α := X.str.toMulOneClass private instance (X : AddGrp) : AddZeroClass X.α := X.str.toAddZeroClass private instance (X : AddCommGrp) : AddZeroClass X.α := X.str.toAddZeroClass -/-- The functor `Group ⥤ AddGroup` by sending `X ↦ additive X` and `f ↦ f`. +/-- The functor `Group ⥤ AddGroup` by sending `X ↦ Additive X` and `f ↦ f`. -/ @[simps] def toAddGrp : Grp ⥤ AddGrp where @@ -38,7 +38,7 @@ end Grp namespace CommGrp -/-- The functor `CommGroup ⥤ AddCommGroup` by sending `X ↦ additive X` and `f ↦ f`. +/-- The functor `CommGroup ⥤ AddCommGroup` by sending `X ↦ Additive X` and `f ↦ f`. -/ @[simps] def toAddCommGrp : CommGrp ⥤ AddCommGrp where @@ -49,7 +49,7 @@ end CommGrp namespace AddGrp -/-- The functor `AddGroup ⥤ Group` by sending `X ↦ multiplicative Y` and `f ↦ f`. +/-- The functor `AddGroup ⥤ Group` by sending `X ↦ Multiplicative Y` and `f ↦ f`. -/ @[simps] def toGrp : AddGrp ⥤ Grp where @@ -60,7 +60,7 @@ end AddGrp namespace AddCommGrp -/-- The functor `AddCommGroup ⥤ CommGroup` by sending `X ↦ multiplicative Y` and `f ↦ f`. +/-- The functor `AddCommGroup ⥤ CommGroup` by sending `X ↦ Multiplicative Y` and `f ↦ f`. -/ @[simps] def toCommGrp : AddCommGrp ⥤ CommGrp where diff --git a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean index a864d63a882b6..274f4ea89cd2d 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean @@ -37,9 +37,9 @@ def normalMono (hf : Mono f) : NormalMono f where might help you understand what's going on here: ``` calc - M ≃ₗ[R] f.ker.quotient : (submodule.quot_equiv_of_eq_bot _ (ker_eq_bot_of_mono _)).symm - ... ≃ₗ[R] f.range : linear_map.quot_ker_equiv_range f - ... ≃ₗ[R] r.range.mkq.ker : linear_equiv.of_eq _ _ (submodule.ker_mkq _).symm + M ≃ₗ[R] f.ker.quotient : (Submodule.quotEquivOfEqBot _ (ker_eq_bot_of_mono _)).symm + ... ≃ₗ[R] f.range : LinearMap.quotKerEquivRange f + ... ≃ₗ[R] r.range.mkq.ker : LinearEquiv.ofEq _ _ (Submodule.ker_mkQ _).symm ``` -/ IsKernel.isoKernel _ _ (kernelIsLimit _) @@ -57,9 +57,9 @@ def normalEpi (hf : Epi f) : NormalEpi f where /- The following invalid Lean code might help you understand what's going on here: ``` calc f.ker.subtype.range.quotient - ≃ₗ[R] f.ker.quotient : submodule.quot_equiv_of_eq _ _ (submodule.range_subtype _) - ... ≃ₗ[R] f.range : linear_map.quot_ker_equiv_range f - ... ≃ₗ[R] N : linear_equiv.of_top _ (range_eq_top_of_epi _) + ≃ₗ[R] f.ker.quotient : Submodule.quotEquivOfEq _ _ (Submodule.range_subtype _) + ... ≃ₗ[R] f.range : LinearMap.quotKerEquivRange f + ... ≃ₗ[R] N : LinearEquiv.ofTop _ (range_eq_top_of_epi _) ``` -/ IsCokernel.cokernelIso _ _ (cokernelIsColimit _) diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean index 45dec643aa137..1c06b0681c172 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean @@ -208,7 +208,7 @@ noncomputable def derivation' : (relativeDifferentials' φ').Derivation' φ' := Derivation'.mk (fun X ↦ CommRingCat.KaehlerDifferential.D (φ'.app X)) (fun X Y f x ↦ by rw [relativeDifferentials'_map_apply, CommRingCat.KaehlerDifferential.map_d]) -/-- The derivation `derivation' φ'` is universal. -/ +/-- The derivation `Derivation' φ'` is universal. -/ noncomputable def isUniversal' : (derivation' φ').Universal := Derivation'.Universal.mk (fun {M'} d' ↦ Hom.mk'' (fun X ↦ (d'.app X).desc) (fun X Y f ↦ diff --git a/Mathlib/Algebra/Group/Semiconj/Defs.lean b/Mathlib/Algebra/Group/Semiconj/Defs.lean index 41c05e4803a72..f20bc77af2491 100644 --- a/Mathlib/Algebra/Group/Semiconj/Defs.lean +++ b/Mathlib/Algebra/Group/Semiconj/Defs.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -Some proofs and docs came from `algebra/commute` (c) Neil Strickland +Some proofs and docs came from `Algebra/Commute` (c) Neil Strickland -/ import Mathlib.Algebra.Group.Defs import Mathlib.Init.Logic diff --git a/Mathlib/Algebra/Group/Semiconj/Units.lean b/Mathlib/Algebra/Group/Semiconj/Units.lean index 5f00f4ae6bd86..142307a0a9682 100644 --- a/Mathlib/Algebra/Group/Semiconj/Units.lean +++ b/Mathlib/Algebra/Group/Semiconj/Units.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -Some proofs and docs came from `algebra/commute` (c) Neil Strickland +Some proofs and docs came from `Algebra/Commute` (c) Neil Strickland -/ import Mathlib.Algebra.Group.Semiconj.Defs import Mathlib.Algebra.Group.Units diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index 20f20bdbd0d10..4f61b06dd1dd6 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -54,7 +54,7 @@ In the definitions below, we use the following notation: * `map (f : R → S₁) p` : returns the multivariate polynomial obtained from `p` by the change of coefficient semiring corresponding to `f` * `aeval (g : σ → S₁) p` : evaluates the multivariate polynomial obtained from `p` by the change - of coefficient semiring corresponding to `g` (`a` stands for `algebra`) + of coefficient semiring corresponding to `g` (`a` stands for `Algebra`) ## Implementation notes diff --git a/Mathlib/Algebra/Vertex/HVertexOperator.lean b/Mathlib/Algebra/Vertex/HVertexOperator.lean index af95acfde4a24..229a327c1952f 100644 --- a/Mathlib/Algebra/Vertex/HVertexOperator.lean +++ b/Mathlib/Algebra/Vertex/HVertexOperator.lean @@ -8,7 +8,7 @@ import Mathlib.RingTheory.HahnSeries.Multiplication /-! # Vertex operators In this file we introduce heterogeneous vertex operators using Hahn series. When `R = ℂ`, `V = W`, -and `Γ = ℤ`, then this is the usual notion of `meromorphic left-moving 2D field`. The notion we use +and `Γ = ℤ`, then this is the usual notion of "meromorphic left-moving 2D field". The notion we use here allows us to consider composites and scalar-multiply by multivariable Laurent series. ## Definitions * `HVertexOperator` : An `R`-linear map from an `R`-module `V` to `HahnModule Γ W`. diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean index e76145cb2aa2c..747ed8efb6ca5 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/StructureSheaf.lean @@ -9,7 +9,7 @@ import Mathlib.RingTheory.GradedAlgebra.HomogeneousLocalization import Mathlib.Geometry.RingedSpace.LocallyRingedSpace /-! -# The structure sheaf on `projective_spectrum 𝒜`. +# The structure sheaf on `ProjectiveSpectrum 𝒜`. In `Mathlib.AlgebraicGeometry.Topology`, we have given a topology on `ProjectiveSpectrum 𝒜`; in this file we will construct a sheaf on `ProjectiveSpectrum 𝒜`. diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 743f4565a50a5..baf9f2551f7e6 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -25,10 +25,10 @@ To formalize this, we use compositions of an integer `N`, i.e., its decompositio a sum `i₁ + ... + iₙ` of positive integers. Given such a composition `c` and two formal multilinear series `q` and `p`, let `q.comp_along_composition p c` be the above multilinear function. Then the `N`-th coefficient in the power series expansion of `g ∘ f` is the sum of these -terms over all `c : composition N`. +terms over all `c : Composition N`. To complete the proof, we need to show that this power series has a positive radius of convergence. -This follows from the fact that `composition N` has cardinality `2^(N-1)` and estimates on +This follows from the fact that `Composition N` has cardinality `2^(N-1)` and estimates on the norm of `qₙ` and `pₖ`, which give summability. We also need to show that it indeed converges to `g ∘ f`. For this, we note that the composition of partial sums converges to `g ∘ f`, and that it corresponds to a part of the whole sum, on a subset that increases to the whole space. By @@ -49,7 +49,7 @@ The main technical difficulty is to write down things. In particular, we need to `q.comp_along_composition p c` and to show that it is indeed a continuous multilinear function. This requires a whole interface built on the class `Composition`. Once this is set, the main difficulty is to reorder the sums, writing the composition of the partial sums as a sum -over some subset of `Σ n, composition n`. We need to check that the reordering is a bijection, +over some subset of `Σ n, Composition n`. We need to check that the reordering is a bijection, running over difficulties due to the dependent nature of the types under consideration, that are controlled thanks to the interface for `Composition`. @@ -58,8 +58,8 @@ follow from the associativity of composition of analytic functions, as there is the formal multilinear series representing a function (and also, it holds even when the radius of convergence of the series is `0`). Instead, we give a direct proof, which amounts to reordering double sums in a careful way. The change of variables is a canonical (combinatorial) bijection -`Composition.sigmaEquivSigmaPi` between `(Σ (a : composition n), composition a.length)` and -`(Σ (c : composition n), Π (i : fin c.length), composition (c.blocks_fun i))`, and is described +`Composition.sigmaEquivSigmaPi` between `(Σ (a : Composition n), Composition a.length)` and +`(Σ (c : Composition n), Π (i : Fin c.length), Composition (c.blocks_fun i))`, and is described in more details below in the paragraph on associativity. -/ @@ -94,9 +94,9 @@ possible compositions of `n`. /-- Given a formal multilinear series `p`, a composition `c` of `n` and the index `i` of a -block of `c`, we may define a function on `fin n → E` by picking the variables in the `i`-th block +block of `c`, we may define a function on `Fin n → E` by picking the variables in the `i`-th block of `n`, and applying the corresponding coefficient of `p` to these variables. This function is -called `p.apply_composition c v i` for `v : fin n → E` and `i : fin c.length`. -/ +called `p.applyComposition c v i` for `v : Fin n → E` and `i : Fin c.length`. -/ def applyComposition (p : FormalMultilinearSeries 𝕜 E F) {n : ℕ} (c : Composition n) : (Fin n → E) → Fin c.length → F := fun v i => p (c.blocksFun i) (v ∘ c.embedding i) @@ -128,8 +128,8 @@ theorem removeZero_applyComposition (p : FormalMultilinearSeries 𝕜 E F) {n : ext v i simp [applyComposition, zero_lt_one.trans_le (c.one_le_blocksFun i), removeZero_of_pos] -/-- Technical lemma stating how `p.apply_composition` commutes with updating variables. This -will be the key point to show that functions constructed from `apply_composition` retain +/-- Technical lemma stating how `p.applyComposition` commutes with updating variables. This +will be the key point to show that functions constructed from `applyComposition` retain multilinearity. -/ theorem applyComposition_update (p : FormalMultilinearSeries 𝕜 E F) {n : ℕ} (c : Composition n) (j : Fin n) (v : Fin n → E) (z : E) : @@ -507,7 +507,7 @@ theorem le_comp_radius_of_summable (q : FormalMultilinearSeries 𝕜 F G) ### Composing analytic functions Now, we will prove that the composition of the partial sums of `q` and `p` up to order `N` is -given by a sum over some large subset of `Σ n, composition n` of `q.comp_along_composition p`, to +given by a sum over some large subset of `Σ n, Composition n` of `q.comp_along_composition p`, to deduce that the series for `q.comp p` indeed converges to `g ∘ f` when `q` is a power series for `g` and `p` is a power series for `f`. @@ -835,28 +835,28 @@ By definition, ``` (r.comp q).comp p n v = ∑_{i₁ + ... + iₖ = n} (r.comp q)ₖ (p_{i₁} (v₀, ..., v_{i₁ -1}), p_{i₂} (...), ..., p_{iₖ}(...)) -= ∑_{a : composition n} (r.comp q) a.length (apply_composition p a v) += ∑_{a : Composition n} (r.comp q) a.length (applyComposition p a v) ``` decomposing `r.comp q` in the same way, we get ``` (r.comp q).comp p n v -= ∑_{a : composition n} ∑_{b : composition a.length} - r b.length (apply_composition q b (apply_composition p a v)) += ∑_{a : Composition n} ∑_{b : Composition a.length} + r b.length (applyComposition q b (applyComposition p a v)) ``` On the other hand, ``` -r.comp (q.comp p) n v = ∑_{c : composition n} r c.length (apply_composition (q.comp p) c v) +r.comp (q.comp p) n v = ∑_{c : Composition n} r c.length (applyComposition (q.comp p) c v) ``` -Here, `apply_composition (q.comp p) c v` is a vector of length `c.length`, whose `i`-th term is +Here, `applyComposition (q.comp p) c v` is a vector of length `c.length`, whose `i`-th term is given by `(q.comp p) (c.blocks_fun i) (v_l, v_{l+1}, ..., v_{m-1})` where `{l, ..., m-1}` is the `i`-th block in the composition `c`, of length `c.blocks_fun i` by definition. To compute this term, -we expand it as `∑_{dᵢ : composition (c.blocks_fun i)} q dᵢ.length (apply_composition p dᵢ v')`, +we expand it as `∑_{dᵢ : Composition (c.blocks_fun i)} q dᵢ.length (applyComposition p dᵢ v')`, where `v' = (v_l, v_{l+1}, ..., v_{m-1})`. Therefore, we get ``` r.comp (q.comp p) n v = -∑_{c : composition n} ∑_{d₀ : composition (c.blocks_fun 0), - ..., d_{c.length - 1} : composition (c.blocks_fun (c.length - 1))} - r c.length (λ i, q dᵢ.length (apply_composition p dᵢ v'ᵢ)) +∑_{c : Composition n} ∑_{d₀ : Composition (c.blocks_fun 0), + ..., d_{c.length - 1} : Composition (c.blocks_fun (c.length - 1))} + r c.length (λ i, q dᵢ.length (applyComposition p dᵢ v'ᵢ)) ``` To show that these terms coincide, we need to explain how to reindex the sums to put them in bijection (and then the terms we are summing will correspond to each other). Suppose we have a @@ -887,7 +887,7 @@ namespace Composition variable {n : ℕ} -/-- Rewriting equality in the dependent type `Σ (a : composition n), composition a.length)` in +/-- Rewriting equality in the dependent type `Σ (a : Composition n), Composition a.length)` in non-dependent terms with lists, requiring that the blocks coincide. -/ theorem sigma_composition_eq_iff (i j : Σ a : Composition n, Composition a.length) : i = j ↔ i.1.blocks = j.1.blocks ∧ i.2.blocks = j.2.blocks := by @@ -899,7 +899,7 @@ theorem sigma_composition_eq_iff (i j : Σ a : Composition n, Composition a.leng induction H; congr; ext1; exact h' /-- Rewriting equality in the dependent type -`Σ (c : composition n), Π (i : fin c.length), composition (c.blocks_fun i)` in +`Σ (c : Composition n), Π (i : Fin c.length), Composition (c.blocks_fun i)` in non-dependent terms with lists, requiring that the lists of blocks coincide. -/ theorem sigma_pi_composition_eq_iff (u v : Σ c : Composition n, ∀ i : Fin c.length, Composition (c.blocksFun i)) : @@ -1031,8 +1031,8 @@ theorem sizeUpTo_sizeUpTo_add (a : Composition n) (b : Composition a.length) {i simp only [sigmaCompositionAux, add_assoc, add_left_inj, Fin.val_mk] rw [getElem_of_eq (getElem_splitWrtComposition _ _ _ _), getElem_drop', getElem_take _ _ C] -/-- Natural equivalence between `(Σ (a : composition n), composition a.length)` and -`(Σ (c : composition n), Π (i : fin c.length), composition (c.blocks_fun i))`, that shows up as a +/-- Natural equivalence between `(Σ (a : Composition n), Composition a.length)` and +`(Σ (c : Composition n), Π (i : Fin c.length), Composition (c.blocks_fun i))`, that shows up as a change of variables in the proof that composition of formal multilinear series is associative. Consider a composition `a` of `n` and a composition `b` of `a.length`. Then `b` indicates how to diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 8096bad4b1010..44fc258910888 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -49,12 +49,12 @@ This generality has several advantages `smul` to multiply the functions, that would be an asymmetric definition. # Main Definitions -* `convolution f g L μ x = (f ⋆[L, μ] g) x = ∫ t, L (f t) (g (x - t)) ∂μ` is the convolution of - `f` and `g` w.r.t. the continuous bilinear map `L` and measure `μ`. -* `ConvolutionExistsAt f g x L μ` states that the convolution `(f ⋆[L, μ] g) x` is well-defined - (i.e. the integral exists). -* `ConvolutionExists f g L μ` states that the convolution `f ⋆[L, μ] g` is well-defined at each - point. +* `MeasureTheory.convolution f g L μ x = (f ⋆[L, μ] g) x = ∫ t, L (f t) (g (x - t)) ∂μ` + is the convolution of `f` and `g` w.r.t. the continuous bilinear map `L` and measure `μ`. +* `MeasureTheory.ConvolutionExistsAt f g x L μ` states that the convolution `(f ⋆[L, μ] g) x` + is well-defined (i.e. the integral exists). +* `MeasureTheory.ConvolutionExists f g L μ` states that the convolution `f ⋆[L, μ] g` + is well-defined at each point. # Main Results * `HasCompactSupport.hasFDerivAt_convolution_right` and @@ -66,12 +66,12 @@ This generality has several advantages Versions of these statements for functions depending on a parameter are also given. -* `convolution_tendsto_right`: Given a sequence of nonnegative normalized functions whose support - tends to a small neighborhood around `0`, the convolution tends to the right argument. - This is specialized to bump functions in `ContDiffBump.convolution_tendsto_right`. +* `MeasureTheory.convolution_tendsto_right`: Given a sequence of nonnegative normalized functions + whose support tends to a small neighborhood around `0`, the convolution tends to the right + argument. This is specialized to bump functions in `ContDiffBump.convolution_tendsto_right`. # Notation -The following notations are localized in the locale `convolution`: +The following notations are localized in the locale `Convolution`: * `f ⋆[L, μ] g` for the convolution. Note: you have to use parentheses to apply the convolution to an argument: `(f ⋆[L, μ] g) x`. * `f ⋆[L] g := f ⋆[L, volume] g` @@ -854,7 +854,7 @@ theorem integral_convolution [MeasurableAdd₂ G] [MeasurableNeg G] [NormedSpace variable [MeasurableAdd₂ G] [IsAddRightInvariant ν] [MeasurableNeg G] /-- Convolution is associative. This has a weak but inconvenient integrability condition. -See also `convolution_assoc`. -/ +See also `MeasureTheory.convolution_assoc`. -/ theorem convolution_assoc' (hL : ∀ (x : E) (y : E') (z : E''), L₂ (L x y) z = L₃ x (L₄ y z)) {x₀ : G} (hfg : ∀ᵐ y ∂μ, ConvolutionExistsAt f g y L ν) (hgk : ∀ᵐ x ∂ν, ConvolutionExistsAt g k x L₄ μ) diff --git a/Mathlib/Analysis/CstarAlgebra/Exponential.lean b/Mathlib/Analysis/CstarAlgebra/Exponential.lean index d7bfd49513760..52030ce0f5210 100644 --- a/Mathlib/Analysis/CstarAlgebra/Exponential.lean +++ b/Mathlib/Analysis/CstarAlgebra/Exponential.lean @@ -6,8 +6,8 @@ Authors: Jireh Loreaux import Mathlib.Analysis.Normed.Algebra.Exponential /-! # The exponential map from selfadjoint to unitary -In this file, we establish various properties related to the map `fun a ↦ exp ℂ A (I • a)` -between the subtypes `selfAdjoint A` and `unitary A`. +In this file, we establish various properties related to the map +`fun a ↦ NormedSpace.exp ℂ A (I • a)` between the subtypes `selfAdjoint A` and `unitary A`. ## TODO diff --git a/Mathlib/Analysis/Normed/Algebra/Exponential.lean b/Mathlib/Analysis/Normed/Algebra/Exponential.lean index 62c52c3b53ebb..273d045957370 100644 --- a/Mathlib/Analysis/Normed/Algebra/Exponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/Exponential.lean @@ -13,12 +13,12 @@ import Mathlib.Topology.Algebra.Algebra /-! # Exponential in a Banach algebra -In this file, we define `exp 𝕂 : 𝔸 → 𝔸`, the exponential map in a topological algebra `𝔸` over a -field `𝕂`. +In this file, we define `NormedSpace.exp 𝕂 : 𝔸 → 𝔸`, +the exponential map in a topological algebra `𝔸` over a field `𝕂`. While for most interesting results we need `𝔸` to be normed algebra, we do not require this in the -definition in order to make `exp` independent of a particular choice of norm. The definition also -does not require that `𝔸` be complete, but we need to assume it for most results. +definition in order to make `NormedSpace.exp` independent of a particular choice of norm. The +definition also does not require that `𝔸` be complete, but we need to assume it for most results. We then prove some basic results, but we avoid importing derivatives here to minimize dependencies. Results involving derivatives and comparisons with `Real.exp` and `Complex.exp` can be found in @@ -32,24 +32,27 @@ We prove most result for an arbitrary field `𝕂`, and then specialize to `𝕂 - `NormedSpace.exp_add_of_commute_of_mem_ball` : if `𝕂` has characteristic zero, then given two commuting elements `x` and `y` in the disk of convergence, we have - `exp 𝕂 (x+y) = (exp 𝕂 x) * (exp 𝕂 y)` + `NormedSpace.exp 𝕂 (x+y) = (NormedSpace.exp 𝕂 x) * (NormedSpace.exp 𝕂 y)` - `NormedSpace.exp_add_of_mem_ball` : if `𝕂` has characteristic zero and `𝔸` is commutative, then given two elements `x` and `y` in the disk of convergence, we have - `exp 𝕂 (x+y) = (exp 𝕂 x) * (exp 𝕂 y)` + `NormedSpace.exp 𝕂 (x+y) = (NormedSpace.exp 𝕂 x) * (NormedSpace.exp 𝕂 y)` - `NormedSpace.exp_neg_of_mem_ball` : if `𝕂` has characteristic zero and `𝔸` is a division ring, - then given an element `x` in the disk of convergence, we have `exp 𝕂 (-x) = (exp 𝕂 x)⁻¹`. + then given an element `x` in the disk of convergence, + we have `NormedSpace.exp 𝕂 (-x) = (NormedSpace.exp 𝕂 x)⁻¹`. ### `𝕂 = ℝ` or `𝕂 = ℂ` -- `expSeries_radius_eq_top` : the `FormalMultilinearSeries` defining `exp 𝕂` has infinite - radius of convergence +- `expSeries_radius_eq_top` : the `FormalMultilinearSeries` defining `NormedSpace.exp 𝕂` + has infinite radius of convergence - `NormedSpace.exp_add_of_commute` : given two commuting elements `x` and `y`, we have - `exp 𝕂 (x+y) = (exp 𝕂 x) * (exp 𝕂 y)` -- `NormedSpace.exp_add` : if `𝔸` is commutative, then we have `exp 𝕂 (x+y) = (exp 𝕂 x) * (exp 𝕂 y)` - for any `x` and `y` -- `NormedSpace.exp_neg` : if `𝔸` is a division ring, then we have `exp 𝕂 (-x) = (exp 𝕂 x)⁻¹`. -- `exp_sum_of_commute` : the analogous result to `NormedSpace.exp_add_of_commute` for `Finset.sum`. -- `exp_sum` : the analogous result to `NormedSpace.exp_add` for `Finset.sum`. + `NormedSpace.exp 𝕂 (x+y) = (NormedSpace.exp 𝕂 x) * (NormedSpace.exp 𝕂 y)` +- `NormedSpace.exp_add` : if `𝔸` is commutative, then we have + `NormedSpace.exp 𝕂 (x+y) = (NormedSpace.exp 𝕂 x) * (NormedSpace.exp 𝕂 y)` for any `x` and `y` +- `NormedSpace.exp_neg` : if `𝔸` is a division ring, then we have + `NormedSpace.exp 𝕂 (-x) = (NormedSpace.exp 𝕂 x)⁻¹`. +- `NormedSpace.exp_sum_of_commute` : the analogous result to `NormedSpace.exp_add_of_commute` + for `Finset.sum`. +- `NormedSpace.exp_sum` : the analogous result to `NormedSpace.exp_add` for `Finset.sum`. - `NormedSpace.exp_nsmul` : repeated addition in the domain corresponds to repeated multiplication in the codomain. - `NormedSpace.exp_zsmul` : repeated addition in the domain corresponds to @@ -58,7 +61,7 @@ We prove most result for an arbitrary field `𝕂`, and then specialize to `𝕂 ### Other useful compatibility results - `NormedSpace.exp_eq_exp` : if `𝔸` is a normed algebra over two fields `𝕂` and `𝕂'`, - then `exp 𝕂 = exp 𝕂' 𝔸` + then `NormedSpace.exp 𝕂 = NormedSpace.exp 𝕂' 𝔸` ### Notes @@ -95,13 +98,13 @@ section TopologicalAlgebra variable (𝕂 𝔸 : Type*) [Field 𝕂] [Ring 𝔸] [Algebra 𝕂 𝔸] [TopologicalSpace 𝔸] [TopologicalRing 𝔸] /-- `expSeries 𝕂 𝔸` is the `FormalMultilinearSeries` whose `n`-th term is the map -`(xᵢ) : 𝔸ⁿ ↦ (1/n! : 𝕂) • ∏ xᵢ`. Its sum is the exponential map `exp 𝕂 : 𝔸 → 𝔸`. -/ +`(xᵢ) : 𝔸ⁿ ↦ (1/n! : 𝕂) • ∏ xᵢ`. Its sum is the exponential map `NormedSpace.exp 𝕂 : 𝔸 → 𝔸`. -/ def expSeries : FormalMultilinearSeries 𝕂 𝔸 𝔸 := fun n => (n !⁻¹ : 𝕂) • ContinuousMultilinearMap.mkPiAlgebraFin 𝕂 n 𝔸 variable {𝔸} -/-- `exp 𝕂 : 𝔸 → 𝔸` is the exponential map determined by the action of `𝕂` on `𝔸`. +/-- `NormedSpace.exp 𝕂 : 𝔸 → 𝔸` is the exponential map determined by the action of `𝕂` on `𝔸`. It is defined as the sum of the `FormalMultilinearSeries` `expSeries 𝕂 𝔸`. Note that when `𝔸 = Matrix n n 𝕂`, this is the **Matrix Exponential**; see @@ -252,7 +255,8 @@ theorem analyticAt_exp_of_mem_ball (x : 𝔸) (hx : x ∈ EMetric.ball (0 : 𝔸 exact (hasFPowerSeriesOnBall_exp_of_radius_pos h).analyticAt_of_mem hx /-- In a Banach-algebra `𝔸` over a normed field `𝕂` of characteristic zero, if `x` and `y` are -in the disk of convergence and commute, then `exp 𝕂 (x + y) = (exp 𝕂 x) * (exp 𝕂 y)`. -/ +in the disk of convergence and commute, then +`NormedSpace.exp 𝕂 (x + y) = (NormedSpace.exp 𝕂 x) * (NormedSpace.exp 𝕂 y)`. -/ theorem exp_add_of_commute_of_mem_ball [CharZero 𝕂] {x y : 𝔸} (hxy : Commute x y) (hx : x ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius) (hy : y ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius) : exp 𝕂 (x + y) = exp 𝕂 x * exp 𝕂 y := by @@ -271,7 +275,7 @@ theorem exp_add_of_commute_of_mem_ball [CharZero 𝕂] {x y : 𝔸} (hxy : Commu have : (n ! : 𝕂) ≠ 0 := Nat.cast_ne_zero.mpr n.factorial_ne_zero field_simp [this] -/-- `exp 𝕂 x` has explicit two-sided inverse `exp 𝕂 (-x)`. -/ +/-- `NormedSpace.exp 𝕂 x` has explicit two-sided inverse `NormedSpace.exp 𝕂 (-x)`. -/ noncomputable def invertibleExpOfMemBall [CharZero 𝕂] {x : 𝔸} (hx : x ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius) : Invertible (exp 𝕂 x) where invOf := exp 𝕂 (-x) @@ -297,7 +301,7 @@ theorem invOf_exp_of_mem_ball [CharZero 𝕂] {x : 𝔸} ⅟ (exp 𝕂 x) = exp 𝕂 (-x) := by letI := invertibleExpOfMemBall hx; convert (rfl : ⅟ (exp 𝕂 x) = _) -/-- Any continuous ring homomorphism commutes with `exp`. -/ +/-- Any continuous ring homomorphism commutes with `NormedSpace.exp`. -/ theorem map_exp_of_mem_ball {F} [FunLike F 𝔸 𝔹] [RingHomClass F 𝔸 𝔹] (f : F) (hf : Continuous f) (x : 𝔸) (hx : x ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius) : f (exp 𝕂 x) = exp 𝕂 (f x) := by @@ -352,7 +356,8 @@ variable {𝕂 𝔸 : Type*} [NontriviallyNormedField 𝕂] [NormedCommRing 𝔸 [CompleteSpace 𝔸] /-- In a commutative Banach-algebra `𝔸` over a normed field `𝕂` of characteristic zero, -`exp 𝕂 (x+y) = (exp 𝕂 x) * (exp 𝕂 y)` for all `x`, `y` in the disk of convergence. -/ +`NormedSpace.exp 𝕂 (x+y) = (NormedSpace.exp 𝕂 x) * (NormedSpace.exp 𝕂 y)` +for all `x`, `y` in the disk of convergence. -/ theorem exp_add_of_mem_ball [CharZero 𝕂] {x y : 𝔸} (hx : x ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius) (hy : y ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius) : exp 𝕂 (x + y) = exp 𝕂 x * exp 𝕂 y := @@ -431,7 +436,7 @@ theorem exp_analytic (x : 𝔸) : AnalyticAt 𝕂 (exp 𝕂) x := analyticAt_exp_of_mem_ball x ((expSeries_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) /-- In a Banach-algebra `𝔸` over `𝕂 = ℝ` or `𝕂 = ℂ`, if `x` and `y` commute, then -`exp 𝕂 (x+y) = (exp 𝕂 x) * (exp 𝕂 y)`. -/ +`NormedSpace.exp 𝕂 (x+y) = (NormedSpace.exp 𝕂 x) * (NormedSpace.exp 𝕂 y)`. -/ theorem exp_add_of_commute {x y : 𝔸} (hxy : Commute x y) : exp 𝕂 (x + y) = exp 𝕂 x * exp 𝕂 y := exp_add_of_commute_of_mem_ball hxy ((expSeries_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) ((expSeries_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) @@ -440,7 +445,7 @@ section variable (𝕂) -/-- `exp 𝕂 x` has explicit two-sided inverse `exp 𝕂 (-x)`. -/ +/-- `NormedSpace.exp 𝕂 x` has explicit two-sided inverse `NormedSpace.exp 𝕂 (-x)`. -/ noncomputable def invertibleExp (x : 𝔸) : Invertible (exp 𝕂 x) := invertibleExpOfMemBall <| (expSeries_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _ @@ -463,7 +468,7 @@ theorem exp_mem_unitary_of_mem_skewAdjoint [StarRing 𝔸] [ContinuousStar 𝔸] end /-- In a Banach-algebra `𝔸` over `𝕂 = ℝ` or `𝕂 = ℂ`, if a family of elements `f i` mutually -commute then `exp 𝕂 (∑ i, f i) = ∏ i, exp 𝕂 (f i)`. -/ +commute then `NormedSpace.exp 𝕂 (∑ i, f i) = ∏ i, NormedSpace.exp 𝕂 (f i)`. -/ theorem exp_sum_of_commute {ι} (s : Finset ι) (f : ι → 𝔸) (h : (s : Set ι).Pairwise fun i j => Commute (f i) (f j)) : exp 𝕂 (∑ i ∈ s, f i) = @@ -573,7 +578,7 @@ section CommAlgebra variable {𝕂 𝔸 : Type*} [RCLike 𝕂] [NormedCommRing 𝔸] [NormedAlgebra 𝕂 𝔸] [CompleteSpace 𝔸] /-- In a commutative Banach-algebra `𝔸` over `𝕂 = ℝ` or `𝕂 = ℂ`, -`exp 𝕂 (x+y) = (exp 𝕂 x) * (exp 𝕂 y)`. -/ +`NormedSpace.exp 𝕂 (x+y) = (NormedSpace.exp 𝕂 x) * (NormedSpace.exp 𝕂 y)`. -/ theorem exp_add {x y : 𝔸} : exp 𝕂 (x + y) = exp 𝕂 x * exp 𝕂 y := exp_add_of_mem_ball ((expSeries_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) ((expSeries_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) diff --git a/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean b/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean index fb4f9d8aca29b..2f0463ab0aa36 100644 --- a/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/MatrixExponential.lean @@ -13,7 +13,8 @@ import Mathlib.Topology.UniformSpace.Matrix /-! # Lemmas about the matrix exponential -In this file, we provide results about `exp` on `Matrix`s over a topological or normed algebra. +In this file, we provide results about `NormedSpace.exp` on `Matrix`s +over a topological or normed algebra. Note that generic results over all topological spaces such as `NormedSpace.exp_zero` can be used on matrices without issue, so are not repeated here. The topological results specific to matrices are: @@ -51,7 +52,7 @@ results for general rings are instead stated about `Ring.inverse`: ## TODO -* Show that `Matrix.det (exp 𝕂 A) = exp 𝕂 (Matrix.trace A)` +* Show that `Matrix.det (NormedSpace.exp 𝕂 A) = NormedSpace.exp 𝕂 (Matrix.trace A)` ## References diff --git a/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean index 730859031052a..3f1aa98ec5d75 100644 --- a/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean @@ -52,7 +52,7 @@ variable (𝕜 : Type*) {S R M : Type*} local notation "tsze" => TrivSqZeroExt -open NormedSpace -- For `exp`. +open NormedSpace -- For `NormedSpace.exp`. namespace TrivSqZeroExt @@ -86,7 +86,8 @@ theorem snd_expSeries_of_smul_comm Nat.cast_mul, mul_inv_rev, inv_mul_cancel_right₀ ((Nat.cast_ne_zero (R := 𝕜)).mpr <| Nat.succ_ne_zero n)] -/-- If `exp R x.fst` converges to `e` then `(exp R x).snd` converges to `e • x.snd`. -/ +/-- If `NormedSpace.exp R x.fst` converges to `e` +then `(NormedSpace.exp R x).snd` converges to `e • x.snd`. -/ theorem hasSum_snd_expSeries_of_smul_comm (x : tsze R M) (hx : MulOpposite.op x.fst • x.snd = x.fst • x.snd) {e : R} (h : HasSum (fun n => expSeries 𝕜 R n fun _ => x.fst) e) : @@ -98,7 +99,8 @@ theorem hasSum_snd_expSeries_of_smul_comm (x : tsze R M) inv_one, one_smul, snd_one, sub_zero] exact h.smul_const _ -/-- If `exp R x.fst` converges to `e` then `exp R x` converges to `inl e + inr (e • x.snd)`. -/ +/-- If `NormedSpace.exp R x.fst` converges to `e` +then `NormedSpace.exp R x` converges to `inl e + inr (e • x.snd)`. -/ theorem hasSum_expSeries_of_smul_comm (x : tsze R M) (hx : MulOpposite.op x.fst • x.snd = x.fst • x.snd) {e : R} (h : HasSum (fun n => expSeries 𝕜 R n fun _ => x.fst) e) : @@ -312,7 +314,8 @@ variable [BoundedSMul R M] [BoundedSMul Rᵐᵒᵖ M] [SMulCommClass R Rᵐᵒ variable [IsScalarTower 𝕜 R M] [IsScalarTower 𝕜 Rᵐᵒᵖ M] variable [CompleteSpace R] [CompleteSpace M] --- Evidence that we have sufficient instances on `tsze R N` to make `exp_add_of_commute` usable +-- Evidence that we have sufficient instances on `tsze R N` +-- to make `NormedSpace.exp_add_of_commute` usable example (a b : tsze R M) (h : Commute a b) : exp 𝕜 (a + b) = exp 𝕜 a * exp 𝕜 b := exp_add_of_commute h diff --git a/Mathlib/Analysis/Normed/Group/InfiniteSum.lean b/Mathlib/Analysis/Normed/Group/InfiniteSum.lean index 42d46d8ed1b53..fb02a2775c4d7 100644 --- a/Mathlib/Analysis/Normed/Group/InfiniteSum.lean +++ b/Mathlib/Analysis/Normed/Group/InfiniteSum.lean @@ -16,7 +16,7 @@ In a complete (semi)normed group, there exists a finite set `s` such that the sum `∑ i ∈ t, f i` over any finite set `t` disjoint with `s` has norm less than `ε`; -- `summable_of_norm_bounded`, `Summable.of_norm_bounded_eventually`: if `‖f i‖` is bounded above by +- `Summable.of_norm_bounded`, `Summable.of_norm_bounded_eventually`: if `‖f i‖` is bounded above by a summable series `∑' i, g i`, then `∑' i, f i` is summable as well; the same is true if the inequality hold only off some finite set. diff --git a/Mathlib/Analysis/SpecialFunctions/Exponential.lean b/Mathlib/Analysis/SpecialFunctions/Exponential.lean index b41bef6f28d1f..117201399155c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exponential.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exponential.lean @@ -20,33 +20,35 @@ We prove most results for an arbitrary field `𝕂`, and then specialize to ` ### General case -- `hasStrictFDerivAt_exp_zero_of_radius_pos` : `exp 𝕂` has strict Fréchet derivative +- `hasStrictFDerivAt_exp_zero_of_radius_pos` : `NormedSpace.exp 𝕂` has strict Fréchet derivative `1 : 𝔸 →L[𝕂] 𝔸` at zero, as long as it converges on a neighborhood of zero (see also `hasStrictDerivAt_exp_zero_of_radius_pos` for the case `𝔸 = 𝕂`) - `hasStrictFDerivAt_exp_of_lt_radius` : if `𝕂` has characteristic zero and `𝔸` is commutative, - then given a point `x` in the disk of convergence, `exp 𝕂` has strict Fréchet derivative - `exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at x (see also `hasStrictDerivAt_exp_of_lt_radius` for the case - `𝔸 = 𝕂`) -- `hasStrictFDerivAt_exp_smul_const_of_mem_ball`: even when `𝔸` is non-commutative, if we have - an intermediate algebra `𝕊` which is commutative, then the function `(u : 𝕊) ↦ exp 𝕂 (u • x)`, - still has strict Fréchet derivative `exp 𝕂 (t • x) • (1 : 𝕊 →L[𝕂] 𝕊).smulRight x` at `t` if + then given a point `x` in the disk of convergence, `NormedSpace.exp 𝕂` has strict Fréchet + derivative `NormedSpace.exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at x + (see also `hasStrictDerivAt_exp_of_lt_radius` for the case `𝔸 = 𝕂`) +- `hasStrictFDerivAt_exp_smul_const_of_mem_ball`: even when `𝔸` is non-commutative, + if we have an intermediate algebra `𝕊` which is commutative, the function + `(u : 𝕊) ↦ NormedSpace.exp 𝕂 (u • x)`, still has strict Fréchet derivative + `NormedSpace.exp 𝕂 (t • x) • (1 : 𝕊 →L[𝕂] 𝕊).smulRight x` at `t` if `t • x` is in the radius of convergence. ### `𝕂 = ℝ` or `𝕂 = ℂ` -- `hasStrictFDerivAt_exp_zero` : `exp 𝕂` has strict Fréchet derivative `1 : 𝔸 →L[𝕂] 𝔸` at zero - (see also `hasStrictDerivAt_exp_zero` for the case `𝔸 = 𝕂`) -- `hasStrictFDerivAt_exp` : if `𝔸` is commutative, then given any point `x`, `exp 𝕂` has strict - Fréchet derivative `exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at x (see also `hasStrictDerivAt_exp` for the - case `𝔸 = 𝕂`) +- `hasStrictFDerivAt_exp_zero` : `NormedSpace.exp 𝕂` has strict Fréchet derivative `1 : 𝔸 →L[𝕂] 𝔸` + at zero (see also `hasStrictDerivAt_exp_zero` for the case `𝔸 = 𝕂`) +- `hasStrictFDerivAt_exp` : if `𝔸` is commutative, then given any point `x`, `NormedSpace.exp 𝕂` + has strict Fréchet derivative `NormedSpace.exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at x + (see also `hasStrictDerivAt_exp` for the case `𝔸 = 𝕂`) - `hasStrictFDerivAt_exp_smul_const`: even when `𝔸` is non-commutative, if we have - an intermediate algebra `𝕊` which is commutative, then the function `(u : 𝕊) ↦ exp 𝕂 (u • x)` - still has strict Fréchet derivative `exp 𝕂 (t • x) • (1 : 𝔸 →L[𝕂] 𝔸).smulRight x` at `t`. + an intermediate algebra `𝕊` which is commutative, the function + `(u : 𝕊) ↦ NormedSpace.exp 𝕂 (u • x)` still has strict Fréchet derivative + `NormedSpace.exp 𝕂 (t • x) • (1 : 𝔸 →L[𝕂] 𝔸).smulRight x` at `t`. ### Compatibility with `Real.exp` and `Complex.exp` -- `Complex.exp_eq_exp_ℂ` : `Complex.exp = exp ℂ ℂ` -- `Real.exp_eq_exp_ℝ` : `Real.exp = exp ℝ ℝ` +- `Complex.exp_eq_exp_ℂ` : `Complex.exp = NormedSpace.exp ℂ ℂ` +- `Real.exp_eq_exp_ℝ` : `Real.exp = NormedSpace.exp ℝ ℝ` -/ @@ -83,8 +85,8 @@ variable {𝕂 𝔸 : Type*} [NontriviallyNormedField 𝕂] [NormedCommRing 𝔸 [CompleteSpace 𝔸] /-- The exponential map in a commutative Banach algebra `𝔸` over a normed field `𝕂` of -characteristic zero has Fréchet derivative `exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at any point `x` in the -disk of convergence. -/ +characteristic zero has Fréchet derivative `NormedSpace.exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` +at any point `x`in the disk of convergence. -/ theorem hasFDerivAt_exp_of_mem_ball [CharZero 𝕂] {x : 𝔸} (hx : x ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius) : HasFDerivAt (exp 𝕂) (exp 𝕂 x • (1 : 𝔸 →L[𝕂] 𝔸)) x := by @@ -103,8 +105,8 @@ theorem hasFDerivAt_exp_of_mem_ball [CharZero 𝕂] {x : 𝔸} ring /-- The exponential map in a commutative Banach algebra `𝔸` over a normed field `𝕂` of -characteristic zero has strict Fréchet derivative `exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at any point `x` in -the disk of convergence. -/ +characteristic zero has strict Fréchet derivative `NormedSpace.exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` +at any point `x` in the disk of convergence. -/ theorem hasStrictFDerivAt_exp_of_mem_ball [CharZero 𝕂] {x : 𝔸} (hx : x ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius) : HasStrictFDerivAt (exp 𝕂) (exp 𝕂 x • (1 : 𝔸 →L[𝕂] 𝔸)) x := @@ -118,14 +120,14 @@ section deriv variable {𝕂 : Type*} [NontriviallyNormedField 𝕂] [CompleteSpace 𝕂] /-- The exponential map in a complete normed field `𝕂` of characteristic zero has strict derivative -`exp 𝕂 x` at any point `x` in the disk of convergence. -/ +`NormedSpace.exp 𝕂 x` at any point `x` in the disk of convergence. -/ theorem hasStrictDerivAt_exp_of_mem_ball [CharZero 𝕂] {x : 𝕂} (hx : x ∈ EMetric.ball (0 : 𝕂) (expSeries 𝕂 𝕂).radius) : HasStrictDerivAt (exp 𝕂) (exp 𝕂 x) x := by simpa using (hasStrictFDerivAt_exp_of_mem_ball hx).hasStrictDerivAt /-- The exponential map in a complete normed field `𝕂` of characteristic zero has derivative -`exp 𝕂 x` at any point `x` in the disk of convergence. -/ +`NormedSpace.exp 𝕂 x` at any point `x` in the disk of convergence. -/ theorem hasDerivAt_exp_of_mem_ball [CharZero 𝕂] {x : 𝕂} (hx : x ∈ EMetric.ball (0 : 𝕂) (expSeries 𝕂 𝕂).radius) : HasDerivAt (exp 𝕂) (exp 𝕂 x) x := (hasStrictDerivAt_exp_of_mem_ball hx).hasDerivAt @@ -165,12 +167,12 @@ section RCLikeCommAlgebra variable {𝕂 𝔸 : Type*} [RCLike 𝕂] [NormedCommRing 𝔸] [NormedAlgebra 𝕂 𝔸] [CompleteSpace 𝔸] /-- The exponential map in a commutative Banach algebra `𝔸` over `𝕂 = ℝ` or `𝕂 = ℂ` has strict -Fréchet derivative `exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at any point `x`. -/ +Fréchet derivative `NormedSpace.exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at any point `x`. -/ theorem hasStrictFDerivAt_exp {x : 𝔸} : HasStrictFDerivAt (exp 𝕂) (exp 𝕂 x • (1 : 𝔸 →L[𝕂] 𝔸)) x := hasStrictFDerivAt_exp_of_mem_ball ((expSeries_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) /-- The exponential map in a commutative Banach algebra `𝔸` over `𝕂 = ℝ` or `𝕂 = ℂ` has -Fréchet derivative `exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at any point `x`. -/ +Fréchet derivative `NormedSpace.exp 𝕂 x • 1 : 𝔸 →L[𝕂] 𝔸` at any point `x`. -/ theorem hasFDerivAt_exp {x : 𝔸} : HasFDerivAt (exp 𝕂) (exp 𝕂 x • (1 : 𝔸 →L[𝕂] 𝔸)) x := hasStrictFDerivAt_exp.hasFDerivAt @@ -180,12 +182,13 @@ section DerivRCLike variable {𝕂 : Type*} [RCLike 𝕂] -/-- The exponential map in `𝕂 = ℝ` or `𝕂 = ℂ` has strict derivative `exp 𝕂 x` at any point -`x`. -/ +/-- The exponential map in `𝕂 = ℝ` or `𝕂 = ℂ` has strict derivative `NormedSpace.exp 𝕂 x` +at any point `x`. -/ theorem hasStrictDerivAt_exp {x : 𝕂} : HasStrictDerivAt (exp 𝕂) (exp 𝕂 x) x := hasStrictDerivAt_exp_of_mem_ball ((expSeries_radius_eq_top 𝕂 𝕂).symm ▸ edist_lt_top _ _) -/-- The exponential map in `𝕂 = ℝ` or `𝕂 = ℂ` has derivative `exp 𝕂 x` at any point `x`. -/ +/-- The exponential map in `𝕂 = ℝ` or `𝕂 = ℂ` has derivative `NormedSpace.exp 𝕂 x` +at any point `x`. -/ theorem hasDerivAt_exp {x : 𝕂} : HasDerivAt (exp 𝕂) (exp 𝕂 x) x := hasStrictDerivAt_exp.hasDerivAt @@ -391,7 +394,7 @@ section tsum_tprod variable {𝕂 𝔸 : Type*} [RCLike 𝕂] [NormedCommRing 𝔸] [NormedAlgebra 𝕂 𝔸] [CompleteSpace 𝔸] -/-- If `f` has sum `a`, then `exp ∘ f` has product `exp a`. -/ +/-- If `f` has sum `a`, then `NormedSpace.exp ∘ f` has product `NormedSpace.exp a`. -/ lemma HasSum.exp {ι : Type*} {f : ι → 𝔸} {a : 𝔸} (h : HasSum f a) : HasProd (exp 𝕂 ∘ f) (exp 𝕂 a) := Tendsto.congr (fun s ↦ exp_sum s f) <| Tendsto.exp h diff --git a/Mathlib/CategoryTheory/Abelian/Ext.lean b/Mathlib/CategoryTheory/Abelian/Ext.lean index 71cc0990a96ce..63698a865b338 100644 --- a/Mathlib/CategoryTheory/Abelian/Ext.lean +++ b/Mathlib/CategoryTheory/Abelian/Ext.lean @@ -14,7 +14,7 @@ import Mathlib.CategoryTheory.Linear.Yoneda # Ext We define `Ext R C n : Cᵒᵖ ⥤ C ⥤ Module R` for any `R`-linear abelian category `C` -by (left) deriving in the first argument of the bifunctor `(X, Y) ↦ Module.of R (unop X ⟶ Y)`. +by (left) deriving in the first argument of the bifunctor `(X, Y) ↦ ModuleCat.of R (unop X ⟶ Y)`. ## Implementation @@ -32,7 +32,8 @@ open CategoryTheory Limits variable (R : Type*) [Ring R] (C : Type*) [Category C] [Abelian C] [Linear R C] [EnoughProjectives C] -/-- `Ext R C n` is defined by deriving in the first argument of `(X, Y) ↦ Module.of R (unop X ⟶ Y)` +/-- `Ext R C n` is defined by deriving in +the first argument of `(X, Y) ↦ ModuleCat.of R (unop X ⟶ Y)` (which is the second argument of `linearYoneda`). -/ diff --git a/Mathlib/CategoryTheory/Category/RelCat.lean b/Mathlib/CategoryTheory/Category/RelCat.lean index f2474949d304e..b505cec20e646 100644 --- a/Mathlib/CategoryTheory/Category/RelCat.lean +++ b/Mathlib/CategoryTheory/Category/RelCat.lean @@ -118,7 +118,7 @@ theorem rel_iso_iff {X Y : RelCat} (r : X ⟶ Y) : section Opposite open Opposite -/-- The argument-swap isomorphism from `rel` to its opposite. -/ +/-- The argument-swap isomorphism from `RelCat` to its opposite. -/ def opFunctor : RelCat ⥤ RelCatᵒᵖ where obj X := op X map {X Y} r := op (fun y x => r x y) @@ -154,8 +154,8 @@ def unopFunctor : RelCatᵒᵖ ⥤ RelCat where @[simp] theorem unopFunctor_comp_opFunctor_eq : Functor.comp unopFunctor opFunctor = Functor.id _ := rfl -/-- `rel` is self-dual: The map that swaps the argument order of a - relation induces an equivalence between `rel` and its opposite. -/ +/-- `RelCat` is self-dual: The map that swaps the argument order of a + relation induces an equivalence between `RelCat` and its opposite. -/ @[simps] def opEquivalence : Equivalence RelCat RelCatᵒᵖ where functor := opFunctor diff --git a/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean b/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean index 3a47e674116f4..b2efc306cfa42 100644 --- a/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean @@ -33,7 +33,7 @@ Given a type `C` with associated `groupoid C` instance. * `CategoryTheory.Subgroupoid.comap` is the "preimage" map of subgroupoids along a functor. * `CategoryTheory.Subgroupoid.map` is the "image" map of subgroupoids along a functor _injective on objects_. -* `CategoryTheory.Subgroupoid.vertexSubgroup` is the subgroup of the `vertex group` at a given +* `CategoryTheory.Subgroupoid.vertexSubgroup` is the subgroup of the *vertex group* at a given vertex `v`, assuming `v` is contained in the `CategoryTheory.Subgroupoid` (meaning, by definition, that the arrow `𝟙 v` is contained in the subgroupoid). diff --git a/Mathlib/CategoryTheory/Linear/Yoneda.lean b/Mathlib/CategoryTheory/Linear/Yoneda.lean index c75ea9e86c4c0..7022906c6dc3a 100644 --- a/Mathlib/CategoryTheory/Linear/Yoneda.lean +++ b/Mathlib/CategoryTheory/Linear/Yoneda.lean @@ -11,8 +11,8 @@ import Mathlib.CategoryTheory.Preadditive.Yoneda.Basic # The Yoneda embedding for `R`-linear categories The Yoneda embedding for `R`-linear categories `C`, -sends an object `X : C` to the `Module R`-valued presheaf on `C`, -with value on `Y : Cᵒᵖ` given by `Module.of R (unop Y ⟶ X)`. +sends an object `X : C` to the `ModuleCat R`-valued presheaf on `C`, +with value on `Y : Cᵒᵖ` given by `ModuleCat.of R (unop Y ⟶ X)`. TODO: `linearYoneda R C` is `R`-linear. TODO: In fact, `linearYoneda` itself is additive and `R`-linear. @@ -32,8 +32,8 @@ variable (C) -- and similarly in `linearCoyoneda`, otherwise many simp lemmas are not triggered automatically. -- Eventually, doing so allows more proofs to be automatic! /-- The Yoneda embedding for `R`-linear categories `C`, -sending an object `X : C` to the `Module R`-valued presheaf on `C`, -with value on `Y : Cᵒᵖ` given by `Module.of R (unop Y ⟶ X)`. -/ +sending an object `X : C` to the `ModuleCat R`-valued presheaf on `C`, +with value on `Y : Cᵒᵖ` given by `ModuleCat.of R (unop Y ⟶ X)`. -/ @[simps] def linearYoneda : C ⥤ Cᵒᵖ ⥤ ModuleCat R where obj X := @@ -44,8 +44,8 @@ def linearYoneda : C ⥤ Cᵒᵖ ⥤ ModuleCat R where (Linear.rightComp R _ f) } /-- The Yoneda embedding for `R`-linear categories `C`, -sending an object `Y : Cᵒᵖ` to the `Module R`-valued copresheaf on `C`, -with value on `X : C` given by `Module.of R (unop Y ⟶ X)`. -/ +sending an object `Y : Cᵒᵖ` to the `ModuleCat R`-valued copresheaf on `C`, +with value on `X : C` given by `ModuleCat.of R (unop Y ⟶ X)`. -/ @[simps] def linearCoyoneda : Cᵒᵖ ⥤ C ⥤ ModuleCat R where obj Y := diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index d31e42e629f70..3ae1f1aa58149 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -427,7 +427,7 @@ def Cont.eval : Cont → List ℕ →. List ℕ /-- The set of configurations of the machine: * `halt v`: The machine is about to stop and `v : List ℕ` is the result. -* `ret k v`: The machine is about to pass `v : List ℕ` to continuation `k : cont`. +* `ret k v`: The machine is about to pass `v : List ℕ` to continuation `k : Cont`. We don't have a state corresponding to normal evaluation because these are evaluated immediately to a `ret` "in zero steps" using the `stepNormal` function. -/ diff --git a/Mathlib/Control/Monad/Writer.lean b/Mathlib/Control/Monad/Writer.lean index 46479df03cf3d..bd76158292c0c 100644 --- a/Mathlib/Control/Monad/Writer.lean +++ b/Mathlib/Control/Monad/Writer.lean @@ -59,10 +59,10 @@ protected theorem ext {ω : Type u} (x x' : WriterT ω M α) (h : x.run = x'.run variable {ω : Type u} {α β : Type u} [Monad M] -/-- Creates an instance of Monad, with an explicitly given empty and append operation. +/-- Creates an instance of `Monad`, with explicitly given `empty` and `append` operations. Previously, this would have used an instance of `[Monoid ω]` as input. -In practice, however, WriterT is used for logging and creating lists so restricting to +In practice, however, `WriterT` is used for logging and creating lists so restricting to monoids with `Mul` and `One` can make `WriterT` cumbersome to use. This is used to derive instances for both `[EmptyCollection ω] [Append ω]` and `[Monoid ω]`. diff --git a/Mathlib/Data/Finset/NAry.lean b/Mathlib/Data/Finset/NAry.lean index 03a786f2d20c5..74ad2d6a47ad7 100644 --- a/Mathlib/Data/Finset/NAry.lean +++ b/Mathlib/Data/Finset/NAry.lean @@ -425,7 +425,7 @@ theorem image₂_right_identity {f : γ → β → γ} {b : β} (h : ∀ a, f a /-- If each partial application of `f` is injective, and images of `s` under those partial applications are disjoint (but not necessarily distinct!), then the size of `t` divides the size of -`finset.image₂ f s t`. -/ +`Finset.image₂ f s t`. -/ theorem card_dvd_card_image₂_right (hf : ∀ a ∈ s, Injective (f a)) (hs : ((fun a => t.image <| f a) '' s).PairwiseDisjoint id) : t.card ∣ (image₂ f s t).card := by classical @@ -446,7 +446,7 @@ theorem card_dvd_card_image₂_right (hf : ∀ a ∈ s, Injective (f a)) /-- If each partial application of `f` is injective, and images of `t` under those partial applications are disjoint (but not necessarily distinct!), then the size of `s` divides the size of -`finset.image₂ f s t`. -/ +`Finset.image₂ f s t`. -/ theorem card_dvd_card_image₂_left (hf : ∀ b ∈ t, Injective fun a => f a b) (ht : ((fun b => s.image fun a => f a b) '' t).PairwiseDisjoint id) : s.card ∣ (image₂ f s t).card := by rw [← image₂_swap]; exact card_dvd_card_image₂_right hf ht diff --git a/Mathlib/Data/Finset/Union.lean b/Mathlib/Data/Finset/Union.lean index e8fe619aa4671..b54499d3e58d9 100644 --- a/Mathlib/Data/Finset/Union.lean +++ b/Mathlib/Data/Finset/Union.lean @@ -129,7 +129,7 @@ lemma biUnion_insert [DecidableEq α] {a : α} : (insert a s).biUnion t = t a lemma biUnion_congr (hs : s₁ = s₂) (ht : ∀ a ∈ s₁, t₁ a = t₂ a) : s₁.biUnion t₁ = s₂.biUnion t₂ := ext fun x ↦ by - -- Porting note: this entire proof was `simp [or_and_distrib_right, exists_or_distrib]` + -- Porting note: this entire proof was `simp [or_and_right, exists_or]` simp_rw [mem_biUnion] apply exists_congr simp (config := { contextual := true }) only [hs, and_congr_right_iff, ht, implies_true] diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index b156700f0e723..c7ea171ab05d7 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -114,7 +114,7 @@ theorem card_Ioo : (Ioo a b).card = (b - a - 1).toNat := (card_map _).trans <| c theorem card_uIcc : (uIcc a b).card = (b - a).natAbs + 1 := (card_map _).trans <| Int.ofNat.inj <| by - -- Porting note (#11215): TODO: Restore `int.coe_nat_inj` and remove the `change` + -- Porting note (#11215): TODO: Restore `Int.ofNat.inj` and remove the `change` change ((↑) : ℕ → ℤ) _ = ((↑) : ℕ → ℤ) _ rw [card_range, sup_eq_max, inf_eq_min, Int.toNat_of_nonneg (sub_nonneg_of_le <| le_add_one min_le_max), Int.ofNat_add, diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index 5a25a60e4b644..4f5bc89c79f72 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -422,7 +422,7 @@ namespace Tactic namespace NormNum -/-- A predicate representing partial progress in a proof of `squarefree`. -/ +/-- A predicate representing partial progress in a proof of `Squarefree`. -/ def SquarefreeHelper (n k : ℕ) : Prop := 0 < k → (∀ m, Nat.Prime m → m ∣ bit1 n → bit1 k ≤ m) → Squarefree (bit1 n) @@ -507,7 +507,7 @@ theorem not_squarefree_mul (a aa b n : ℕ) (ha : a * a = aa) (hb : aa * b = n) rw [← hb, ← ha] exact fun H => ne_of_gt h₁ (Nat.isUnit_iff.1 <| H _ ⟨_, rfl⟩) -/-- Given `e` a natural numeral and `a : nat` with `a^2 ∣ n`, return `⊢ ¬ squarefree e`. -/ +/-- Given `e` a natural numeral and `a : ℕ` with `a^2 ∣ n`, return `⊢ ¬ Squarefree e`. -/ unsafe def prove_non_squarefree (e : expr) (n a : ℕ) : tactic expr := do let ea := reflect a let eaa := reflect (a * a) @@ -554,7 +554,7 @@ unsafe def prove_squarefree_aux : let p₂ ← prove_squarefree_aux ic en en1 n1 ek' k' pure <| q(squarefreeHelper_2).mk_app [en, ek, ek', ec, p₁, pc, p₀, p₂] -/-- Given `n > 0` a squarefree natural numeral, returns `⊢ squarefree n`. -/ +/-- Given `n > 0` a squarefree natural numeral, returns `⊢ Squarefree n`. -/ unsafe def prove_squarefree (en : expr) (n : ℕ) : tactic expr := match match_numeral en with | match_numeral_result.one => pure q(@squarefree_one ℕ _) @@ -572,7 +572,7 @@ unsafe def prove_squarefree (en : expr) (n : ℕ) : tactic expr := pure <| q(squarefree_bit1).mk_app [en', p] | _ => failed -/-- Evaluates the `squarefree` predicate on naturals. -/ +/-- Evaluates the `Squarefree` predicate on naturals. -/ @[norm_num] unsafe def eval_squarefree : expr → tactic (expr × expr) | q(@Squarefree ℕ $(inst) $(e)) => do diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 11e37fc6687da..ed7af7bc3672e 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -83,7 +83,7 @@ theorem factor_mk_eq {α : Type*} (r s : α → α → Prop) (h : ∀ x y, r x y variable {γ : Sort*} {r : α → α → Prop} {s : β → β → Prop} --- Porting note: used to be an Alias of `quot.lift_beta`. +-- Porting note: used to be an Alias of `Quot.lift_mk`. theorem lift_mk (f : α → γ) (h : ∀ a₁ a₂, r a₁ a₂ → f a₁ = f a₂) (a : α) : Quot.lift f h (Quot.mk r a) = f a := rfl diff --git a/Mathlib/Data/Real/Pi/Bounds.lean b/Mathlib/Data/Real/Pi/Bounds.lean index 6492a70294f01..20aaebc121269 100644 --- a/Mathlib/Data/Real/Pi/Bounds.lean +++ b/Mathlib/Data/Real/Pi/Bounds.lean @@ -8,7 +8,7 @@ import Mathlib.Analysis.SpecialFunctions.Trigonometric.Bounds /-! # Pi -This file contains lemmas which establish bounds on `real.pi`. +This file contains lemmas which establish bounds on `Real.pi`. Notably, these include `pi_gt_sqrtTwoAddSeries` and `pi_lt_sqrtTwoAddSeries`, which bound `π` using series; numerical bounds on `π` such as `pi_gt_314`and `pi_lt_315` (more precise versions are given, too). diff --git a/Mathlib/Data/ZMod/Algebra.lean b/Mathlib/Data/ZMod/Algebra.lean index e8dc26f580687..2f5737b396fd4 100644 --- a/Mathlib/Data/ZMod/Algebra.lean +++ b/Mathlib/Data/ZMod/Algebra.lean @@ -36,8 +36,8 @@ abbrev algebra' (h : m ∣ n) : Algebra (ZMod n) R := end -/-- The `zmod p`-algebra structure on a ring of characteristic `p`. This is not an -instance since it creates a diamond with `algebra.id`. +/-- The `ZMod p`-algebra structure on a ring of characteristic `p`. This is not an +instance since it creates a diamond with `Algebra.id`. See note [reducible non-instances]. -/ abbrev algebra (p : ℕ) [CharP R p] : Algebra (ZMod p) R := algebra' R p dvd_rfl diff --git a/Mathlib/FieldTheory/Finite/Polynomial.lean b/Mathlib/FieldTheory/Finite/Polynomial.lean index 82f17ea5ec684..e6e7fa481079a 100644 --- a/Mathlib/FieldTheory/Finite/Polynomial.lean +++ b/Mathlib/FieldTheory/Finite/Polynomial.lean @@ -154,7 +154,7 @@ universe u variable (σ : Type u) (K : Type u) [Fintype K] --- Porting note: `@[derive [add_comm_group, module K, inhabited]]` done by hand. +-- Porting note: `@[derive [AddCommGroup, Module K, Inhabited]]` done by hand. /-- The submodule of multivariate polynomials whose degree of each variable is strictly less than the cardinality of K. -/ def R [CommRing K] : Type u := @@ -169,7 +169,7 @@ noncomputable instance [CommRing K] : Module K (R σ K) := noncomputable instance [CommRing K] : Inhabited (R σ K) := inferInstanceAs (Inhabited (restrictDegree σ K (Fintype.card K - 1))) -/-- Evaluation in the `mv_polynomial.R` subtype. -/ +/-- Evaluation in the `MvPolynomial.R` subtype. -/ def evalᵢ [CommRing K] : R σ K →ₗ[K] (σ → K) → K := (evalₗ K σ).comp (restrictDegree σ K (Fintype.card K - 1)).subtype diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean index c95a6c1bc0c50..25e4dfacdf7a2 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean @@ -11,7 +11,7 @@ import Mathlib.Geometry.RingedSpace.Stalks We define (bundled) locally ringed spaces (as `SheafedSpace CommRing` along with the fact that the stalks are local rings), and morphisms between these (morphisms in `SheafedSpace` with -`is_local_ring_hom` on the stalk maps). +`IsLocalRingHom` on the stalk maps). -/ -- Explicit universe annotations were used in this file to improve perfomance #12737 @@ -45,7 +45,7 @@ namespace LocallyRingedSpace variable (X : LocallyRingedSpace.{u}) -/-- An alias for `to_SheafedSpace`, where the result type is a `RingedSpace`. +/-- An alias for `toSheafedSpace`, where the result type is a `RingedSpace`. This allows us to use dot-notation for the `RingedSpace` namespace. -/ def toRingedSpace : RingedSpace := @@ -61,8 +61,8 @@ instance : CoeSort LocallyRingedSpace (Type u) := instance (x : X) : LocalRing (X.presheaf.stalk x) := X.localRing x --- PROJECT: how about a typeclass "has_structure_sheaf" to mediate the 𝒪 notation, rather --- than defining it over and over for PresheafedSpace, LRS, Scheme, etc. +-- PROJECT: how about a typeclass "HasStructureSheaf" to mediate the 𝒪 notation, rather +-- than defining it over and over for `PresheafedSpace`, `LocallyRingedSpace`, `Scheme`, etc. /-- The structure sheaf of a locally ringed space. -/ def 𝒪 : Sheaf CommRingCat X.toTopCat := X.sheaf @@ -155,7 +155,7 @@ theorem comp_val_c_app {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ /-- Given two locally ringed spaces `X` and `Y`, an isomorphism between `X` and `Y` as _sheafed_ spaces can be lifted to a morphism `X ⟶ Y` as locally ringed spaces. -See also `iso_of_SheafedSpace_iso`. +See also `isoOfSheafedSpaceIso`. -/ @[simps] def homOfSheafedSpaceHomOfIsIso {X Y : LocallyRingedSpace.{u}} @@ -170,7 +170,7 @@ def homOfSheafedSpaceHomOfIsIso {X Y : LocallyRingedSpace.{u}} /-- Given two locally ringed spaces `X` and `Y`, an isomorphism between `X` and `Y` as _sheafed_ spaces can be lifted to an isomorphism `X ⟶ Y` as locally ringed spaces. -This is related to the property that the functor `forget_to_SheafedSpace` reflects isomorphisms. +This is related to the property that the functor `forgetToSheafedSpace` reflects isomorphisms. In fact, it is slightly stronger as we do not require `f` to come from a morphism between _locally_ ringed spaces. -/ diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index 13b125308e292..a90c40ac98dc7 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -13,7 +13,7 @@ import Mathlib.CategoryTheory.Limits.Constructions.LimitsOfProductsAndEqualizers We construct the explicit coproducts and coequalizers of `LocallyRingedSpace`. It then follows that `LocallyRingedSpace` has all colimits, and -`forget_to_SheafedSpace` preserves them. +`forgetToSheafedSpace` preserves them. -/ diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 5e7ed36998f00..7d86d2963f23e 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -27,7 +27,7 @@ Abbreviations are also provided for `SheafedSpace`, `LocallyRingedSpace` and `Sc contained in an open immersion factors though the open immersion. * `AlgebraicGeometry.PresheafedSpace.IsOpenImmersion.toSheafedSpace`: If `f : X ⟶ Y` is an open immersion of presheafed spaces, and `Y` is a sheafed space, then `X` is also a sheafed - space. The morphism as morphisms of sheafed spaces is given by `to_SheafedSpace_hom`. + space. The morphism as morphisms of sheafed spaces is given by `to_SheafedSpaceHom`. * `AlgebraicGeometry.PresheafedSpace.IsOpenImmersion.toLocallyRingedSpace`: If `f : X ⟶ Y` is an open immersion of presheafed spaces, and `Y` is a locally ringed space, then `X` is also a locally ringed space. The morphism as morphisms of locally ringed spaces is given by diff --git a/Mathlib/Init/Data/Nat/Lemmas.lean b/Mathlib/Init/Data/Nat/Lemmas.lean index 9f846953a0d36..aa856ee56343c 100644 --- a/Mathlib/Init/Data/Nat/Lemmas.lean +++ b/Mathlib/Init/Data/Nat/Lemmas.lean @@ -59,7 +59,7 @@ theorem one_eq_succ_zero : 1 = succ 0 := /-! subtraction -Many lemmas are proven more generally in mathlib `algebra/order/sub` -/ +Many lemmas are proven more generally in mathlib `Algebra/Order/Sub` -/ /-! min -/ diff --git a/Mathlib/Lean/Meta/CongrTheorems.lean b/Mathlib/Lean/Meta/CongrTheorems.lean index b06587aec6ce5..0d86cfc7904ae 100644 --- a/Mathlib/Lean/Meta/CongrTheorems.lean +++ b/Mathlib/Lean/Meta/CongrTheorems.lean @@ -62,7 +62,7 @@ where | .eq | .heq => forallBoundedTelescope type (some 3) fun params' type' => do let #[x, y, eq] := params' | unreachable! - -- See if we can prove `eq` from previous parameters. + -- See if we can prove `Eq` from previous parameters. let g := (← mkFreshExprMVar (← inferType eq)).mvarId! let g ← g.clear eq.fvarId! if (← observing? <| prove g args).isSome then diff --git a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean index c88b774dfa384..a8726197a4438 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Triangularizable.lean @@ -105,7 +105,7 @@ theorem iSup_genEigenspace_eq_top [IsAlgClosed K] [FiniteDimensional K V] (f : E have ih_ER : ⨆ (μ : K) (k : ℕ), f'.genEigenspace μ k = ⊤ := ih (finrank K ER) h_dim_ER f' rfl -- The induction hypothesis gives us a statement about subspaces of `ER`. We can transfer this - -- to a statement about subspaces of `V` via `submodule.subtype`: + -- to a statement about subspaces of `V` via `Submodule.subtype`: have ih_ER' : ⨆ (μ : K) (k : ℕ), (f'.genEigenspace μ k).map ER.subtype = ER := by simp only [(Submodule.map_iSup _ _).symm, ih_ER, Submodule.map_subtype_top ER] -- Moreover, every generalized eigenspace of `f'` is contained in the corresponding generalized diff --git a/Mathlib/LinearAlgebra/Multilinear/Basis.lean b/Mathlib/LinearAlgebra/Multilinear/Basis.lean index 11ac55ea2c898..a88d7fa503423 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basis.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basis.lean @@ -14,7 +14,7 @@ This file proves lemmas about the action of multilinear maps on basis vectors. ## TODO * Refactor the proofs in terms of bases of tensor products, once there is an equivalent of - `Basis.tensorProduct` for `pi_tensor_product`. + `Basis.tensorProduct` for `PiTensorProduct`. -/ diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 1d88c79536876..45d6fc2fec4a7 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -646,7 +646,7 @@ end QuadraticMap 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 `associated` +maps to bilinear maps giving this identification is called the <`associated` quadratic map. -/ diff --git a/Mathlib/Logic/Equiv/Array.lean b/Mathlib/Logic/Equiv/Array.lean index fb24a56ff1db0..cd9c6fbef7d1f 100644 --- a/Mathlib/Logic/Equiv/Array.lean +++ b/Mathlib/Logic/Equiv/Array.lean @@ -28,11 +28,11 @@ introduce the "right" equivalence for `Array` (`arrayEquivList`) and align the i namespace Equiv -- /-- The natural equivalence between length-`n` heterogeneous arrays --- and dependent functions from `fin n`. -/ +-- and dependent functions from `Fin n`. -/ -- def darrayEquivFin {n : ℕ} (α : Fin n → Type*) : DArray n α ≃ ∀ i, α i := -- ⟨DArray.read, DArray.mk, fun ⟨f⟩ => rfl, fun f => rfl⟩ --- /-- The natural equivalence between length-`n` arrays and functions from `fin n`. -/ +-- /-- The natural equivalence between length-`n` arrays and functions from `Fin n`. -/ -- def array'EquivFin (n : ℕ) (α : Type*) : Array' n α ≃ (Fin n → α) := -- darrayEquivFin _ diff --git a/Mathlib/Logic/Nontrivial/Defs.lean b/Mathlib/Logic/Nontrivial/Defs.lean index 3113fcbfc6b71..a0173e706b52d 100644 --- a/Mathlib/Logic/Nontrivial/Defs.lean +++ b/Mathlib/Logic/Nontrivial/Defs.lean @@ -58,7 +58,7 @@ instance : Nontrivial Prop := /-- See Note [lower instance priority] -Note that since this and `nonempty_of_inhabited` are the most "obvious" way to find a nonempty +Note that since this and `instNonemptyOfInhabited` are the most "obvious" way to find a nonempty instance if no direct instance can be found, we give this a higher priority than the usual `100`. -/ instance (priority := 500) Nontrivial.to_nonempty [Nontrivial α] : Nonempty α := diff --git a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean index d8f578a40e4ac..8a716d2787e2c 100644 --- a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean +++ b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean @@ -187,7 +187,7 @@ theorem exists_goodδ : · exact ⟨fun _ => 0, by simp, fun i j _ => by simpa only [norm_zero, sub_nonpos, sub_self]⟩ - -- For `δ > 0`, `F δ` is a function from `fin N` to the ball of radius `2` for which two points + -- For `δ > 0`, `F δ` is a function from `Fin N` to the ball of radius `2` for which two points -- in the image are separated by `1 - δ`. choose! F hF using this -- Choose a converging subsequence when `δ → 0`. diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 4806ca383cba5..fb4fe7466b781 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -22,7 +22,7 @@ and is almost everywhere strongly measurable. * `eLpNorm' f p μ` : `(∫ ‖f a‖^p ∂μ) ^ (1/p)` for `f : α → F` and `p : ℝ`, where `α` is a measurable space and `F` is a normed group. -* `eLpNormEssSup f μ` : seminorm in `ℒ∞`, equal to the essential supremum `ess_sup ‖f‖ μ`. +* `eLpNormEssSup f μ` : seminorm in `ℒ∞`, equal to the essential supremum `essSup ‖f‖ μ`. * `eLpNorm f p μ` : for `p : ℝ≥0∞`, seminorm in `ℒp`, equal to `0` for `p=0`, to `eLpNorm' f p μ` for `0 < p < ∞` and to `eLpNormEssSup f μ` for `p = ∞`. diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index 3568596ae8386..3f601b21d0e48 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -30,7 +30,7 @@ functions are also called approximations of unity, or approximations of identity at `0` and integrable. Note that there are related results about convolution with respect to peak functions in the file -`Analysis.Convolution`, such as `convolution_tendsto_right` there. +`Analysis.Convolution`, such as `MeasureTheory.convolution_tendsto_right` there. -/ open Set Filter MeasureTheory MeasureTheory.Measure TopologicalSpace Metric diff --git a/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean b/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean index fe9b89308d935..1e06aabefba00 100644 --- a/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean +++ b/Mathlib/NumberTheory/ModularForms/CongruenceSubgroups.lean @@ -114,8 +114,8 @@ theorem Gamma0_mem (N : ℕ) (A : SL(2, ℤ)) : A ∈ Gamma0 N ↔ ((↑ₘA 1 0 theorem Gamma0_det (N : ℕ) (A : Gamma0 N) : (A.1.1.det : ZMod N) = 1 := by simp [A.1.property] -/-- The group homomorphism from `Gamma0` to `ZMod N` given by mapping a matrix to its lower -right-hand entry. -/ +/-- The group homomorphism from `CongruenceSubgroup.Gamma0` to `ZMod N` given by +mapping a matrix to its lower right-hand entry. -/ def Gamma0Map (N : ℕ) : Gamma0 N →* ZMod N where toFun g := ((↑ₘg 1 1 : ℤ) : ZMod N) map_one' := by simp @@ -157,8 +157,8 @@ theorem Gamma1_to_Gamma0_mem (N : ℕ) (A : Gamma0 N) : A ∈ Gamma1' N ↔ Int.coe_castRingHom, map_apply] exact ha.2.1 -/-- The congruence subgroup `Gamma1` of `SL(2, ℤ)` consisting of matrices whose bottom -row is congruent to `(0,1)` modulo `N`. -/ +/-- The congruence subgroup `Gamma1` of `SL(2, ℤ)` consisting of matrices +whose bottom row is congruent to `(0,1)` modulo `N`. -/ def Gamma1 (N : ℕ) : Subgroup SL(2, ℤ) := Subgroup.map ((Gamma0 N).subtype.comp (Gamma1' N).subtype) ⊤ diff --git a/Mathlib/Order/Circular.lean b/Mathlib/Order/Circular.lean index ed3a140a2222b..3757b6709e94b 100644 --- a/Mathlib/Order/Circular.lean +++ b/Mathlib/Order/Circular.lean @@ -30,7 +30,7 @@ circular preorder is like a circular partial order, but several points can coexi Note that the relations between `CircularPreorder`, `CircularPartialOrder` and `CircularOrder` are subtler than between `Preorder`, `PartialOrder`, `LinearOrder`. In particular, one cannot -simply extend the `btw` of a `CircularPartialOrder` to make it a `CircularOrder`. +simply extend the `Btw` of a `CircularPartialOrder` to make it a `CircularOrder`. One can translate from usual orders to circular ones by "closing the necklace at infinity". See `LE.toBtw` and `LT.toSBtw`. Going the other way involves "cutting the necklace" or diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index addf9041883d1..dee2d71c9f043 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -72,7 +72,7 @@ We use the latter one because, e.g., `𝓝 x` in an `EMetricSpace` or in a `Metr of this form. The other two can be emulated using `s = id` or `p = fun _ ↦ True`. With this approach sometimes one needs to `simp` the statement provided by the `Filter.HasBasis` -machinery, e.g., `simp only [true_and]` or `simp only [forall_const]` can help with the case +machinery, e.g., `simp only [true_and_iff]` or `simp only [forall_const]` can help with the case `p = fun _ ↦ True`. -/ diff --git a/Mathlib/Order/RelIso/Basic.lean b/Mathlib/Order/RelIso/Basic.lean index 7e7f9b45b82e0..55c86fb983e5c 100644 --- a/Mathlib/Order/RelIso/Basic.lean +++ b/Mathlib/Order/RelIso/Basic.lean @@ -574,7 +574,7 @@ theorem coe_fn_toEquiv (f : r ≃r s) : (f.toEquiv : α → β) = f := rfl /-- The map `coe_fn : (r ≃r s) → (α → β)` is injective. Lean fails to parse -`function.injective (fun e : r ≃r s ↦ (e : α → β))`, so we use a trick to say the same. -/ +`Function.Injective (fun e : r ≃r s ↦ (e : α → β))`, so we use a trick to say the same. -/ theorem coe_fn_injective : Injective fun f : r ≃r s => (f : α → β) := DFunLike.coe_injective diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean index 7e4ff7c97ac0a..763e9f9919b95 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean @@ -14,7 +14,7 @@ Given `ρ : Measure (α × ℝ)`, we define the conditional cumulative distribut measure, then for all `a : α` `condCDF ρ a` is monotone and right-continuous with limit 0 at -∞ and limit 1 at +∞, and such that for all `x : ℝ`, `a ↦ condCDF ρ a x` is measurable. For all `x : ℝ` and measurable set `s`, that function satisfies -`∫⁻ a in s, ennreal.of_real (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x)`. +`∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x)`. `condCDF` is build from the more general tools about kernel CDFs developed in the file `Probability.Kernel.Disintegration.CdfToKernel`. In that file, we build a function diff --git a/Mathlib/RepresentationTheory/Action/Basic.lean b/Mathlib/RepresentationTheory/Action/Basic.lean index 4a05b86f36f2e..f92661d066b17 100644 --- a/Mathlib/RepresentationTheory/Action/Basic.lean +++ b/Mathlib/RepresentationTheory/Action/Basic.lean @@ -17,7 +17,7 @@ The prototypical example is `V = ModuleCat R`, where `Action (ModuleCat R) G` is the category of `R`-linear representations of `G`. We check `Action V G ≌ (singleObj G ⥤ V)`, -and construct the restriction functors `res {G H : Mon} (f : G ⟶ H) : Action V H ⥤ Action V G`. +and construct the restriction functors `res {G H : MonCat} (f : G ⟶ H) : Action V H ⥤ Action V G`. -/ diff --git a/Mathlib/RepresentationTheory/Action/Concrete.lean b/Mathlib/RepresentationTheory/Action/Concrete.lean index 5f335f1230685..1ae786bb201e9 100644 --- a/Mathlib/RepresentationTheory/Action/Concrete.lean +++ b/Mathlib/RepresentationTheory/Action/Concrete.lean @@ -62,7 +62,7 @@ def leftRegular (G : Type u) [Monoid G] : Action (Type u) (MonCat.of G) := def diagonal (G : Type u) [Monoid G] (n : ℕ) : Action (Type u) (MonCat.of G) := Action.ofMulAction G (Fin n → G) -/-- We have `fin 1 → G ≅ G` as `G`-sets, with `G` acting by left multiplication. -/ +/-- We have `Fin 1 → G ≅ G` as `G`-sets, with `G` acting by left multiplication. -/ def diagonalOneIsoLeftRegular (G : Type u) [Monoid G] : diagonal G 1 ≅ leftRegular G := Action.mkIso (Equiv.funUnique _ _).toIso fun _ => rfl diff --git a/Mathlib/RepresentationTheory/Action/Monoidal.lean b/Mathlib/RepresentationTheory/Action/Monoidal.lean index 285ece611cd5c..dc555899d679c 100644 --- a/Mathlib/RepresentationTheory/Action/Monoidal.lean +++ b/Mathlib/RepresentationTheory/Action/Monoidal.lean @@ -265,7 +265,7 @@ end Monoidal open MonoidalCategory -/-- Given `X : Action (Type u) (Mon.of G)` for `G` a group, then `G × X` (with `G` acting as left +/-- Given `X : Action (Type u) (MonCat.of G)` for `G` a group, then `G × X` (with `G` acting as left multiplication on the first factor and by `X.ρ` on the second) is isomorphic as a `G`-set to `G × X` (with `G` acting as left multiplication on the first factor and trivially on the second). The isomorphism is given by `(g, x) ↦ (g, g⁻¹ • x)`. -/ diff --git a/Mathlib/RingTheory/Adjoin/Field.lean b/Mathlib/RingTheory/Adjoin/Field.lean index cc9f0cc8053b7..69adb77259598 100644 --- a/Mathlib/RingTheory/Adjoin/Field.lean +++ b/Mathlib/RingTheory/Adjoin/Field.lean @@ -14,9 +14,9 @@ Some lemmas on the ring generated by adjoining an element to a field. ## Main statements -* `lift_of_splits`: If `K` and `L` are field extensions of `F` and we have `s : Finset K` such that -the minimal polynomial of each `x ∈ s` splits in `L` then `Algebra.adjoin F s` embeds in `L`. - +* `Polynomial.lift_of_splits`: If `K` and `L` are field extensions of `F` +and we have `s : Finset K` such that the minimal polynomial of each `x ∈ s` splits in `L` +then `Algebra.adjoin F s` embeds in `L`. -/ diff --git a/Mathlib/RingTheory/DedekindDomain/Basic.lean b/Mathlib/RingTheory/DedekindDomain/Basic.lean index 25c6e6395c1f1..29a6df2a644da 100644 --- a/Mathlib/RingTheory/DedekindDomain/Basic.lean +++ b/Mathlib/RingTheory/DedekindDomain/Basic.lean @@ -121,7 +121,7 @@ The integral closure condition is independent of the choice of field of fraction use `isDedekindDomain_iff` to prove `IsDedekindDomain` for a given `fraction_map`. This is the default implementation, but there are equivalent definitions, -`is_dedekind_domain_dvr` and `is_dedekind_domain_inv`. +`IsDedekindDomainDvr` and `IsDedekindDomainInv`. TODO: Prove that these are actually equivalent definitions. -/ class IsDedekindDomain diff --git a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean index 7b7ad16bb8ddd..d371725adeeeb 100644 --- a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean +++ b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean @@ -300,7 +300,7 @@ open ProdAdicCompletions.IsFiniteAdele /-- The finite adèle ring of `R` is the restricted product over all maximal ideals `v` of `R` of `adicCompletion`, with respect to `adicCompletionIntegers`. -Note that we make this a `Type` rather than a `Subtype` (e.g., a `subalgebra`) since we wish +Note that we make this a `Type` rather than a `Subtype` (e.g., a `Subalgebra`) since we wish to endow it with a finer topology than that of the subspace topology. -/ def FiniteAdeleRing : Type _ := {x : K_hat R K // x.IsFiniteAdele} diff --git a/Mathlib/RingTheory/Derivation/Basic.lean b/Mathlib/RingTheory/Derivation/Basic.lean index 59fa5928122fd..ba01be750abfd 100644 --- a/Mathlib/RingTheory/Derivation/Basic.lean +++ b/Mathlib/RingTheory/Derivation/Basic.lean @@ -23,7 +23,7 @@ See `RingTheory.Derivation.Lie` for - `derivation.lie_algebra`: The `R`-derivations from `A` to `A` form a lie algebra over `R`. and `RingTheory.Derivation.ToSquareZero` for -- `derivation_to_square_zero_equiv_lift`: The `R`-derivations from `A` into a square-zero ideal `I` +- `derivationToSquareZeroEquivLift`: The `R`-derivations from `A` into a square-zero ideal `I` of `B` corresponds to the lifts `A →ₐ[R] B` of the map `A →ₐ[R] B ⧸ I`. ## Future project diff --git a/Mathlib/RingTheory/Derivation/ToSquareZero.lean b/Mathlib/RingTheory/Derivation/ToSquareZero.lean index 354a9d2ea6570..1fc6b65f92b46 100644 --- a/Mathlib/RingTheory/Derivation/ToSquareZero.lean +++ b/Mathlib/RingTheory/Derivation/ToSquareZero.lean @@ -107,7 +107,7 @@ theorem liftOfDerivationToSquareZero_mk_apply' (d : Derivation R A I) (x : A) : (Ideal.Quotient.mk I) (d x) + (algebraMap A (B ⧸ I)) x = algebraMap A (B ⧸ I) x := by simp only [Ideal.Quotient.eq_zero_iff_mem.mpr (d x).prop, zero_add] -/-- Given a tower of algebras `R → A → B`, and a square-zero `I : ideal B`, +/-- Given a tower of algebras `R → A → B`, and a square-zero `I : Ideal B`, there is a 1-1 correspondence between `R`-derivations from `A` to `I` and lifts `A →ₐ[R] B` of the canonical map `A →ₐ[R] B ⧸ I`. -/ @[simps!] diff --git a/Mathlib/RingTheory/LocalProperties.lean b/Mathlib/RingTheory/LocalProperties.lean index d676783897cb7..95ef9fd34099e 100644 --- a/Mathlib/RingTheory/LocalProperties.lean +++ b/Mathlib/RingTheory/LocalProperties.lean @@ -31,9 +31,9 @@ The following properties are covered: * The triviality of an ideal or an element: `ideal_eq_bot_of_localization`, `eq_zero_of_localization` -* `isReduced` : `localization_isReduced`, `isReduced_of_localization_maximal`. -* `finite`: `localization_finite`, `finite_ofLocalizationSpan` -* `finiteType`: `localization_finiteType`, `finiteType_ofLocalizationSpan` +* `IsReduced` : `localization_isReduced`, `isReduced_of_localization_maximal`. +* `RingHom.finite`: `localization_finite`, `finite_ofLocalizationSpan` +* `RingHom.finiteType`: `localization_finiteType`, `finiteType_ofLocalizationSpan` -/ diff --git a/Mathlib/RingTheory/Localization/BaseChange.lean b/Mathlib/RingTheory/Localization/BaseChange.lean index 3c6078a864e9c..31e5f767447d1 100644 --- a/Mathlib/RingTheory/Localization/BaseChange.lean +++ b/Mathlib/RingTheory/Localization/BaseChange.lean @@ -33,7 +33,7 @@ theorem IsLocalizedModule.isBaseChange [IsLocalizedModule S f] : IsBaseChange A cases h₂ (LinearMap.restrictScalars R g'') h; rfl /-- The map `(f : M →ₗ[R] M')` is a localization of modules iff the map -`(localization S) × M → N, (s, m) ↦ s • f m` is the tensor product (insomuch as it is the universal +`(Localization S) × M → N, (s, m) ↦ s • f m` is the tensor product (insomuch as it is the universal bilinear map). In particular, there is an isomorphism between `LocalizedModule S M` and `(Localization S) ⊗[R] M` given by `m/s ↦ (1/s) ⊗ₜ m`. diff --git a/Mathlib/RingTheory/Polynomial/Bernstein.lean b/Mathlib/RingTheory/Polynomial/Bernstein.lean index c55621cd2dcaf..af0d3936dd937 100644 --- a/Mathlib/RingTheory/Polynomial/Bernstein.lean +++ b/Mathlib/RingTheory/Polynomial/Bernstein.lean @@ -18,7 +18,7 @@ The definition of the Bernstein polynomials bernsteinPolynomial (R : Type*) [CommRing R] (n ν : ℕ) : R[X] := (choose n ν) * X^ν * (1 - X)^(n - ν) ``` -and the fact that for `ν : fin (n+1)` these are linearly independent over `ℚ`. +and the fact that for `ν : Fin (n+1)` these are linearly independent over `ℚ`. We prove the basic identities * `(Finset.range (n + 1)).sum (fun ν ↦ bernsteinPolynomial R n ν) = 1` diff --git a/Mathlib/RingTheory/Trace/Basic.lean b/Mathlib/RingTheory/Trace/Basic.lean index 74c46a749d367..42e25f8be8ce6 100644 --- a/Mathlib/RingTheory/Trace/Basic.lean +++ b/Mathlib/RingTheory/Trace/Basic.lean @@ -248,7 +248,7 @@ theorem trace_eq_sum_embeddings [FiniteDimensional K L] [Algebra.IsSeparable K L let pb := adjoin.powerBasis hx rw [trace_eq_trace_adjoin K x, Algebra.smul_def, RingHom.map_mul, ← adjoin.powerBasis_gen hx, trace_eq_sum_embeddings_gen E pb (IsAlgClosed.splits_codomain _)] - -- Porting note: the following `convert` was `exact`, with `← algebra.smul_def, algebra_map_smul` + -- Porting note: the following `convert` was `exact`, with `← Algebra.smul_def, algebraMap_smul` -- in the previous `rw`. · convert (sum_embeddings_eq_finrank_mul L E pb).symm ext diff --git a/Mathlib/RingTheory/Valuation/Basic.lean b/Mathlib/RingTheory/Valuation/Basic.lean index 2a0d81f181a6d..a793d69a95a4a 100644 --- a/Mathlib/RingTheory/Valuation/Basic.lean +++ b/Mathlib/RingTheory/Valuation/Basic.lean @@ -23,7 +23,7 @@ commutative monoid with zero, that in addition satisfies the following two axiom `Valuation R Γ₀`is the type of valuations `R → Γ₀`, with a coercion to the underlying function. If `v` is a valuation from `R` to `Γ₀` then the induced group -homomorphism `units(R) → Γ₀` is called `unit_map v`. +homomorphism `Units(R) → Γ₀` is called `unit_map v`. The equivalence "relation" `IsEquiv v₁ v₂ : Prop` defined in 1.27 of [wedhorn_adic] is not strictly speaking a relation, because `v₁ : Valuation R Γ₁` and `v₂ : Valuation R Γ₂` might diff --git a/Mathlib/RingTheory/WittVector/Compare.lean b/Mathlib/RingTheory/WittVector/Compare.lean index 187e5b80484e2..4f42b10dee496 100644 --- a/Mathlib/RingTheory/WittVector/Compare.lean +++ b/Mathlib/RingTheory/WittVector/Compare.lean @@ -76,11 +76,11 @@ theorem zmodEquivTrunc_apply {x : ZMod (p ^ n)} : /-- The following diagram commutes: ```text - zmod (p^n) ----------------------------> zmod (p^m) + ZMod (p^n) ----------------------------> ZMod (p^m) | | | | v v -TruncatedWittVector p n (zmod p) ----> TruncatedWittVector p m (zmod p) +TruncatedWittVector p n (ZMod p) ----> TruncatedWittVector p m (ZMod p) ``` Here the vertical arrows are `TruncatedWittVector.zmodEquivTrunc`, the horizontal arrow at the top is `ZMod.castHom`, @@ -104,11 +104,11 @@ theorem commutes_symm' {m : ℕ} (hm : n ≤ m) (x : TruncatedWittVector p m (ZM /-- The following diagram commutes: ```text -TruncatedWittVector p n (zmod p) ----> TruncatedWittVector p m (zmod p) +TruncatedWittVector p n (ZMod p) ----> TruncatedWittVector p m (ZMod p) | | | | v v - zmod (p^n) ----------------------------> zmod (p^m) + ZMod (p^n) ----------------------------> ZMod (p^m) ``` Here the vertical arrows are `(TruncatedWittVector.zmodEquivTrunc p _).symm`, the horizontal arrow at the top is `ZMod.castHom`, diff --git a/Mathlib/RingTheory/WittVector/IsPoly.lean b/Mathlib/RingTheory/WittVector/IsPoly.lean index 21a58f41cb574..4c912841041f4 100644 --- a/Mathlib/RingTheory/WittVector/IsPoly.lean +++ b/Mathlib/RingTheory/WittVector/IsPoly.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.MvPolynomial.Funext import Mathlib.Algebra.Ring.ULift import Mathlib.RingTheory.WittVector.Basic /-! -# The `is_poly` predicate +# The `IsPoly` predicate `WittVector.IsPoly` is a (type-valued) predicate on functions `f : Π R, 𝕎 R → 𝕎 R`. It asserts that there is a family of polynomials `φ : ℕ → MvPolynomial ℕ ℤ`, diff --git a/Mathlib/Tactic/Conv.lean b/Mathlib/Tactic/Conv.lean index 24d0957c4454c..840ffec93b3ce 100644 --- a/Mathlib/Tactic/Conv.lean +++ b/Mathlib/Tactic/Conv.lean @@ -72,7 +72,7 @@ open Elab Tactic /-- The command `#conv tac => e` will run a conv tactic `tac` on `e`, and display the resulting expression (discarding the proof). -For example, `#conv rw [true_and] => True ∧ False` displays `False`. +For example, `#conv rw [true_and_iff] => True ∧ False` displays `False`. There are also shorthand commands for several common conv tactics: * `#whnf e` is short for `#conv whnf => e` diff --git a/Mathlib/Testing/SlimCheck/Functions.lean b/Mathlib/Testing/SlimCheck/Functions.lean index 0175626e66433..1f5e707a062a8 100644 --- a/Mathlib/Testing/SlimCheck/Functions.lean +++ b/Mathlib/Testing/SlimCheck/Functions.lean @@ -83,7 +83,7 @@ def apply [DecidableEq α] : TotalFunction α β → α → β /-- Implementation of `Repr (TotalFunction α β)`. -Creates a string for a given `finmap` and output, `x₀ ↦ y₀, .. xₙ ↦ yₙ` +Creates a string for a given `Finmap` and output, `x₀ ↦ y₀, .. xₙ ↦ yₙ` for each of the entries. The brackets are provided by the calling function. -/ def reprAux [Repr α] [Repr β] (m : List (Σ _ : α, β)) : String := @@ -101,7 +101,7 @@ protected def repr [Repr α] [Repr β] : TotalFunction α β → String instance (α : Type u) (β : Type v) [Repr α] [Repr β] : Repr (TotalFunction α β) where reprPrec f _ := TotalFunction.repr f -/-- Create a `finmap` from a list of pairs. -/ +/-- Create a `Finmap` from a list of pairs. -/ def List.toFinmap' (xs : List (α × β)) : List (Σ _ : α, β) := xs.map Prod.toSigma @@ -110,9 +110,9 @@ section universe ua ub variable [SampleableExt.{_,u} α] [SampleableExt.{_,ub} β] --- Porting note: removed, there is no `sizeof` in the new `Sampleable` +-- Porting note: removed, there is no `SizeOf.sizeOf` in the new `Sampleable` --- /-- Redefine `sizeof` to follow the structure of `sampleable` instances. -/ +-- /-- Redefine `SizeOf.sizeOf` to follow the structure of `sampleable` instances. -/ -- def Total.sizeof : TotalFunction α β → ℕ -- | ⟨m, x⟩ => 1 + @SizeOf.sizeOf _ Sampleable.wf m + SizeOf.sizeOf x diff --git a/Mathlib/Testing/SlimCheck/Sampleable.lean b/Mathlib/Testing/SlimCheck/Sampleable.lean index 5aedc428634f7..6e6da32e42935 100644 --- a/Mathlib/Testing/SlimCheck/Sampleable.lean +++ b/Mathlib/Testing/SlimCheck/Sampleable.lean @@ -18,15 +18,15 @@ controlling the size of those values using the `Gen` monad. This class helps minimize examples by creating smaller versions of given values. -When testing a proposition like `∀ n : ℕ, prime n → n ≤ 100`, +When testing a proposition like `∀ n : ℕ, Prime n → n ≤ 100`, `SlimCheck` requires that `ℕ` have an instance of `SampleableExt` and for -`prime n` to be decidable. `SlimCheck` will then use the instance of +`Prime n` to be decidable. `SlimCheck` will then use the instance of `SampleableExt` to generate small examples of ℕ and progressively increase -in size. For each example `n`, `prime n` is tested. If it is false, +in size. For each example `n`, `Prime n` is tested. If it is false, the example will be rejected (not a test success nor a failure) and -`SlimCheck` will move on to other examples. If `prime n` is true, +`SlimCheck` will move on to other examples. If `Prime n` is true, `n ≤ 100` will be tested. If it is false, `n` is a counter-example of -`∀ n : ℕ, prime n → n ≤ 100` and the test fails. If `n ≤ 100` is true, +`∀ n : ℕ, Prime n → n ≤ 100` and the test fails. If `n ≤ 100` is true, the test passes and `SlimCheck` moves on to trying more examples. This is a port of the Haskell QuickCheck library. diff --git a/Mathlib/Topology/Algebra/Ring/Basic.lean b/Mathlib/Topology/Algebra/Ring/Basic.lean index 551ace3d15c4a..84e24b6c25393 100644 --- a/Mathlib/Topology/Algebra/Ring/Basic.lean +++ b/Mathlib/Topology/Algebra/Ring/Basic.lean @@ -20,7 +20,7 @@ of topological (semi)rings. ## Main Results - `Subring.topologicalClosure`/`Subsemiring.topologicalClosure`: the topological closure of a - `Subring`/`Subsemiring` is itself a `sub(semi)ring`. + `Subring`/`Subsemiring` is itself a `Sub(semi)ring`. - The product of two topological (semi)rings is a topological (semi)ring. - The indexed product of topological (semi)rings is a topological (semi)ring. -/ diff --git a/Mathlib/Topology/Category/TopCat/Opens.lean b/Mathlib/Topology/Category/TopCat/Opens.lean index 0108110139eeb..5bebad78ee1ff 100644 --- a/Mathlib/Topology/Category/TopCat/Opens.lean +++ b/Mathlib/Topology/Category/TopCat/Opens.lean @@ -63,7 +63,7 @@ noncomputable def infLELeft (U V : Opens X) : U ⊓ V ⟶ U := noncomputable def infLERight (U V : Opens X) : U ⊓ V ⟶ V := inf_le_right.hom -/-- The inclusion `U i ⟶ supr U` as a morphism in the category of open sets. +/-- The inclusion `U i ⟶ iSup U` as a morphism in the category of open sets. -/ noncomputable def leSupr {ι : Type*} (U : ι → Opens X) (i : ι) : U i ⟶ iSup U := (le_iSup U i).hom diff --git a/Mathlib/Topology/FiberBundle/Constructions.lean b/Mathlib/Topology/FiberBundle/Constructions.lean index a89ebf85048c2..437065450fbd1 100644 --- a/Mathlib/Topology/FiberBundle/Constructions.lean +++ b/Mathlib/Topology/FiberBundle/Constructions.lean @@ -258,7 +258,7 @@ universe u v w₁ w₂ U variable {B : Type u} (F : Type v) (E : B → Type w₁) {B' : Type w₂} (f : B' → B) instance [∀ x : B, TopologicalSpace (E x)] : ∀ x : B', TopologicalSpace ((f *ᵖ E) x) := by - -- Porting note: Original proof was `delta_instance bundle.pullback` + -- Porting note: Original proof was `delta_instance Bundle.Pullback` intro x rw [Bundle.Pullback] infer_instance diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean index 26357a33b0c8c..c81ccacd3fc2e 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean @@ -639,7 +639,7 @@ instance : SecondCountableTopology GHSpace := by choose N e _ using this -- cardinality of the nice finite subset `s p` of `p.rep`, called `N p` let N := fun p : GHSpace => N p (s p) (hs p).1 - -- equiv from `s p`, a nice finite subset of `p.rep`, to `fin (N p)`, called `E p` + -- equiv from `s p`, a nice finite subset of `p.rep`, to `Fin (N p)`, called `E p` let E := fun p : GHSpace => e p (s p) (hs p).1 -- A function `F` associating to `p : GHSpace` the data of all distances between points -- in the `ε`-dense set `s p`. @@ -650,8 +650,8 @@ instance : SecondCountableTopology GHSpace := by `p` and `q` with `F p = F q` are at distance `≤ δ`. For this, we construct a map `Φ` from `s p ⊆ p.rep` (representing `p`) to `q.rep` (representing `q`) which is almost an isometry on `s p`, and - with image `s q`. For this, we compose the identification of `s p` with `fin (N p)` - and the inverse of the identification of `s q` with `fin (N q)`. Together with + with image `s q`. For this, we compose the identification of `s p` with `Fin (N p)` + and the inverse of the identification of `s q` with `Fin (N q)`. Together with the fact that `N p = N q`, this constructs `Ψ` between `s p` and `s q`, and then composing with the canonical inclusion we get `Φ`. -/ have Npq : N p = N q := (Sigma.mk.inj_iff.1 hpq).1 @@ -694,12 +694,12 @@ instance : SecondCountableTopology GHSpace := by intro x y -- have : dist (Φ x) (Φ y) = dist (Ψ x) (Ψ y) := rfl rw [show dist (Φ x) (Φ y) = dist (Ψ x) (Ψ y) from rfl] - -- introduce `i`, that codes both `x` and `Φ x` in `fin (N p) = fin (N q)` + -- introduce `i`, that codes both `x` and `Φ x` in `Fin (N p) = Fin (N q)` let i : ℕ := E p x have hip : i < N p := ((E p) x).2 have hiq : i < N q := by rwa [Npq] at hip have i' : i = (E q) (Ψ x) := by simp only [Ψ, Equiv.apply_symm_apply, Fin.coe_cast] - -- introduce `j`, that codes both `y` and `Φ y` in `fin (N p) = fin (N q)` + -- introduce `j`, that codes both `y` and `Φ y` in `Fin (N p) = Fin (N q)` let j : ℕ := E p y have hjp : j < N p := ((E p) y).2 have hjq : j < N q := by rwa [Npq] at hjp @@ -846,12 +846,12 @@ theorem totallyBounded {t : Set GHSpace} {C : ℝ} {u : ℕ → ℝ} {K : ℕ intro x y have : dist (Φ x) (Φ y) = dist (Ψ x) (Ψ y) := rfl rw [this] - -- introduce `i`, that codes both `x` and `Φ x` in `fin (N p) = fin (N q)` + -- introduce `i`, that codes both `x` and `Φ x` in `Fin (N p) = Fin (N q)` let i : ℕ := E p x have hip : i < N p := ((E p) x).2 have hiq : i < N q := by rwa [Npq] at hip have i' : i = (E q) (Ψ x) := by simp only [Ψ, Equiv.apply_symm_apply, Fin.coe_cast] - -- introduce `j`, that codes both `y` and `Φ y` in `fin (N p) = fin (N q)` + -- introduce `j`, that codes both `y` and `Φ y` in `Fin (N p) = Fin (N q)` let j : ℕ := E p y have hjp : j < N p := ((E p) y).2 have hjq : j < N q := by rwa [Npq] at hjp diff --git a/Mathlib/Topology/Sheaves/Forget.lean b/Mathlib/Topology/Sheaves/Forget.lean index 57b888ecf2844..6783eda87f149 100644 --- a/Mathlib/Topology/Sheaves/Forget.lean +++ b/Mathlib/Topology/Sheaves/Forget.lean @@ -42,7 +42,7 @@ The important special case is when that preserves limits and reflects isomorphisms. Then to check the sheaf condition it suffices to check it on the underlying sheaf of types. -Another useful example is the forgetful functor `TopCommRing ⥤ Top`. +Another useful example is the forgetful functor `TopCommRingCat ⥤ TopCat`. See . In fact we prove a stronger version with arbitrary target category. diff --git a/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean b/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean index 659df7e047ac3..24ed94958851d 100644 --- a/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean +++ b/Mathlib/Topology/Sheaves/PresheafOfFunctions.lean @@ -154,7 +154,7 @@ values in some topological commutative ring `T`. For example, we could construct the presheaf of continuous complex valued functions of `X` as ``` -presheafToTopCommRing X (TopCommRing.of ℂ) +presheafToTopCommRing X (TopCommRingCat.of ℂ) ``` (this requires `import Topology.Instances.Complex`). -/ From 4c53d762f83f00bef37c22ddb26899fb312e6c61 Mon Sep 17 00:00:00 2001 From: grunweg Date: Sun, 4 Aug 2024 13:29:41 +0000 Subject: [PATCH 025/359] chore(RingTheory/Coprime/Lemmas): remove obsolete porting note (#15336) --- Mathlib/RingTheory/Coprime/Lemmas.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/RingTheory/Coprime/Lemmas.lean b/Mathlib/RingTheory/Coprime/Lemmas.lean index d2d8cbb199ec9..37284d7b39900 100644 --- a/Mathlib/RingTheory/Coprime/Lemmas.lean +++ b/Mathlib/RingTheory/Coprime/Lemmas.lean @@ -195,8 +195,6 @@ theorem IsCoprime.pow_left_iff (hm : 0 < m) : IsCoprime (x ^ m) y ↔ IsCoprime refine ⟨fun h ↦ ?_, IsCoprime.pow_left⟩ rw [← Finset.card_range m, ← Finset.prod_const] at h exact h.of_prod_left 0 (Finset.mem_range.mpr hm) - -- Porting note: I'm not sure why `finset` didn't get corrected automatically to `Finset` - -- by Mathport, nor whether this is an issue theorem IsCoprime.pow_right_iff (hm : 0 < m) : IsCoprime x (y ^ m) ↔ IsCoprime x y := isCoprime_comm.trans <| (IsCoprime.pow_left_iff hm).trans <| isCoprime_comm From 1f03949ae423ae9fb010318e73f8969e44288c92 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Sun, 4 Aug 2024 14:32:22 +0000 Subject: [PATCH 026/359] feat: Real.exists_natCast_add_one_le_pow_of_one_le (#15074) Necessary for arguments in approaching a power from above Co-authored-by: Yakov Pechersky --- Mathlib/Data/Nat/Cast/Order/Ring.lean | 28 +++++++++++++++++++++++++++ Mathlib/Data/Real/Archimedean.lean | 26 +++++++++++++++++++++++++ Mathlib/Data/Real/Basic.lean | 13 +++++++++++++ 3 files changed, 67 insertions(+) diff --git a/Mathlib/Data/Nat/Cast/Order/Ring.lean b/Mathlib/Data/Nat/Cast/Order/Ring.lean index 5088ab573657f..b03a53d8607f4 100644 --- a/Mathlib/Data/Nat/Cast/Order/Ring.lean +++ b/Mathlib/Data/Nat/Cast/Order/Ring.lean @@ -90,4 +90,32 @@ theorem abs_ofNat [LinearOrderedRing α] (n : ℕ) [n.AtLeastTwo] : |(no_index (OfNat.ofNat n : α))| = OfNat.ofNat n := abs_cast n +lemma mul_le_pow {a : ℕ} (ha : a ≠ 1) (b : ℕ) : + a * b ≤ a ^ b := by + induction b generalizing a with + | zero => simp + | succ b hb => + rw [mul_add_one, pow_succ] + rcases a with (_|_|a) + · simp + · simp at ha + · rw [mul_add_one, mul_add_one, add_comm (_ * a), add_assoc _ (_ * a)] + rcases b with (_|b) + · simp [add_assoc, add_comm] + refine add_le_add (hb (by simp)) ?_ + rw [pow_succ'] + refine (le_add_left ?_ ?_).trans' ?_ + exact le_mul_of_one_le_right' (one_le_pow _ _ (by simp)) + +lemma two_mul_sq_add_one_le_two_pow_two_mul (k : ℕ) : 2 * k ^ 2 + 1 ≤ 2 ^ (2 * k) := by + induction k with + | zero => simp + | succ k hk => + rw [add_pow_two, one_pow, mul_one, add_assoc, mul_add, add_right_comm] + refine (add_le_add_right hk _).trans ?_ + rw [mul_add 2 k, pow_add, mul_one, pow_two, ← mul_assoc, mul_two, mul_two, add_assoc] + gcongr + rw [← two_mul, ← pow_succ'] + exact le_add_of_le_right (mul_le_pow (by simp) _) + end Nat diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index d2aab812b081f..db4cfc3bee867 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -335,4 +335,30 @@ theorem iUnion_Iic_rat : ⋃ r : ℚ, Iic (r : ℝ) = univ := by theorem iInter_Iic_rat : ⋂ r : ℚ, Iic (r : ℝ) = ∅ := by exact iInter_Iic_eq_empty_iff.mpr not_bddBelow_coe +/-- Exponentiation is eventually larger than linear growth. -/ +lemma exists_natCast_add_one_lt_pow_of_one_lt {a : ℝ} (ha : 1 < a) : + ∃ m : ℕ, (m + 1 : ℝ) < a ^ m := by + obtain ⟨k, posk, hk⟩ : ∃ k : ℕ, 0 < k ∧ 1 / k + 1 < a := by + contrapose! ha + refine le_of_forall_lt_rat_imp_le ?_ + intro q hq + refine (ha q.den (by positivity)).trans ?_ + rw [← le_sub_iff_add_le, div_le_iff (by positivity), sub_mul, one_mul] + norm_cast at hq ⊢ + rw [← q.num_div_den, one_lt_div (by positivity)] at hq + rw [q.mul_den_eq_num] + norm_cast at hq ⊢ + rw [le_tsub_iff_left hq.le] + exact hq + use 2 * k ^ 2 + refine (pow_lt_pow_left hk (by positivity) (by simp [posk.ne'])).trans_le' ?_ + rcases k.zero_le.eq_or_lt with rfl|kpos + · simp + rw [pow_two, mul_left_comm, pow_mul] + have := mul_add_one_le_add_one_pow (a := 1 / k) (by simp) k + rw [div_mul_cancel₀ _ (by simp [kpos.ne'])] at this + refine (pow_le_pow_left (by positivity) this _).trans' ?_ + rw [mul_left_comm, ← pow_two] + exact_mod_cast Nat.two_mul_sq_add_one_le_two_pow_two_mul _ + end Real diff --git a/Mathlib/Data/Real/Basic.lean b/Mathlib/Data/Real/Basic.lean index 05d9d696eef7d..09ea95dad271d 100644 --- a/Mathlib/Data/Real/Basic.lean +++ b/Mathlib/Data/Real/Basic.lean @@ -573,6 +573,19 @@ theorem mk_near_of_forall_near {f : CauSeq ℚ abs} {x : ℝ} {ε : ℝ} sub_le_comm.1 <| le_mk_of_forall_le <| H.imp fun _ h j ij => sub_le_comm.1 (abs_sub_le_iff.1 <| h j ij).2⟩ +lemma mul_add_one_le_add_one_pow {a : ℝ} (ha : 0 ≤ a) (b : ℕ) : a * b + 1 ≤ (a + 1) ^ b := by + rcases ha.eq_or_lt with rfl|ha' + · simp + clear ha + induction b generalizing a with + | zero => simp + | succ b hb => + rw [Nat.cast_add_one, mul_add, mul_one, add_right_comm, pow_succ, mul_add, mul_one, add_comm] + gcongr + · rw [le_mul_iff_one_le_left ha'] + exact (pow_le_pow_left zero_le_one (by simpa using ha'.le) b).trans' (by simp) + · exact hb ha' + end Real /-- A function `f : R → ℝ≥0` is nonarchimedean if it satisfies the strong triangle inequality From 39adedf1189c6e00e3ab39e5d36aef6d93555bbe Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 5 Aug 2024 03:02:55 +0000 Subject: [PATCH 027/359] chore: backports for leanprover/lean4#4814 (part 16) (#15427) see https://github.com/leanprover-community/mathlib4/pull/15245 Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- .../Algebra/Algebra/NonUnitalSubalgebra.lean | 28 +++++++----- Mathlib/Algebra/Algebra/Quasispectrum.lean | 31 ++++++------- Mathlib/Algebra/Algebra/Unitization.lean | 34 +++++++++------ Mathlib/Algebra/DirectSum/Internal.lean | 2 +- Mathlib/Algebra/DirectSum/Module.lean | 7 +-- .../Algebra/Homology/Embedding/TruncGE.lean | 4 +- Mathlib/Algebra/Homology/HomotopyCofiber.lean | 2 +- Mathlib/Algebra/Lie/Subalgebra.lean | 6 +-- .../Algebra/Module/Submodule/Pointwise.lean | 2 +- Mathlib/Algebra/MonoidAlgebra/Basic.lean | 2 +- Mathlib/Algebra/MonoidAlgebra/Degree.lean | 43 +++++++++---------- Mathlib/Algebra/MvPolynomial/Basic.lean | 2 +- Mathlib/Algebra/Star/NonUnitalSubalgebra.lean | 5 +-- Mathlib/Algebra/TrivSqZeroExt.lean | 18 ++++---- Mathlib/Analysis/Convex/Cone/Basic.lean | 2 +- Mathlib/Analysis/Convex/Mul.lean | 18 ++++---- .../CategoryTheory/Abelian/LeftDerived.lean | 4 +- .../Abelian/ProjectiveResolution.lean | 10 +++-- .../CategoryTheory/Abelian/RightDerived.lean | 4 +- .../CategoryTheory/Galois/Decomposition.lean | 8 ++-- .../CategoryTheory/Galois/GaloisObjects.lean | 4 +- .../Localization/Triangulated.lean | 15 +++++-- .../Combinatorics/SimpleGraph/AdjMatrix.lean | 24 ++++++----- .../Combinatorics/SimpleGraph/DegreeSum.lean | 3 -- .../Combinatorics/SimpleGraph/IncMatrix.lean | 6 +-- .../Combinatorics/SimpleGraph/Matching.lean | 4 +- .../SimpleGraph/StronglyRegular.lean | 8 +++- Mathlib/Combinatorics/SimpleGraph/Turan.lean | 16 ++++--- Mathlib/Data/Matrix/Basic.lean | 4 +- Mathlib/LinearAlgebra/DFinsupp.lean | 27 ++++++------ Mathlib/LinearAlgebra/Finsupp.lean | 14 ++++-- Mathlib/LinearAlgebra/Matrix/DotProduct.lean | 12 +++--- Mathlib/LinearAlgebra/Matrix/ToLin.lean | 2 +- Mathlib/LinearAlgebra/QuotientPi.lean | 20 +++++---- .../LinearAlgebra/TensorProduct/Basic.lean | 6 ++- .../LinearAlgebra/TensorProduct/Basis.lean | 2 +- .../LinearAlgebra/TensorProduct/Tower.lean | 19 +++++--- Mathlib/NumberTheory/MulChar/Basic.lean | 2 +- .../RingTheory/HahnSeries/Multiplication.lean | 6 ++- Mathlib/RingTheory/Valuation/Integers.lean | 2 +- Mathlib/Topology/Compactness/Compact.lean | 12 +++--- .../Connected/TotallyDisconnected.lean | 8 ++-- Mathlib/Topology/DenseEmbedding.lean | 32 +++++++------- Mathlib/Topology/DiscreteSubset.lean | 6 ++- Mathlib/Topology/Order/ProjIcc.lean | 4 +- Mathlib/Topology/Separation.lean | 4 +- Mathlib/Topology/ShrinkingLemma.lean | 7 +-- Mathlib/Topology/UniformSpace/Separation.lean | 4 +- 48 files changed, 286 insertions(+), 219 deletions(-) diff --git a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean index 4359122d1476f..9a7700afdd395 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean @@ -502,8 +502,14 @@ end NonUnitalAlgHom namespace NonUnitalAlgebra variable {F : Type*} (R : Type u) {A : Type v} {B : Type w} -variable [CommSemiring R] -variable [NonUnitalNonAssocSemiring A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] +variable [CommSemiring R] [NonUnitalNonAssocSemiring A] [Module R A] + +@[simp] +lemma span_eq_toSubmodule (s : NonUnitalSubalgebra R A) : + Submodule.span R (s : Set A) = s.toSubmodule := by + simp [SetLike.ext'_iff, Submodule.coe_span_eq_self] + +variable [IsScalarTower R A A] [SMulCommClass R A A] variable [NonUnitalNonAssocSemiring B] [Module R B] [IsScalarTower R B B] [SMulCommClass R B B] variable [FunLike F A B] [NonUnitalAlgHomClass F R A B] @@ -640,11 +646,6 @@ lemma adjoin_eq_span (s : Set A) : (adjoin R s).toSubmodule = span R (Subsemigro show Subsemigroup.closure s ≤ (adjoin R s).toSubsemigroup exact Subsemigroup.closure_le.2 (subset_adjoin R) -@[simp] -lemma span_eq_toSubmodule (s : NonUnitalSubalgebra R A) : - Submodule.span R (s : Set A) = s.toSubmodule := by - simp [SetLike.ext'_iff, Submodule.coe_span_eq_self] - variable (R A) @[simp] @@ -821,13 +822,18 @@ section NonAssoc variable {R : Type u} {A : Type v} {B : Type w} variable [CommSemiring R] -variable [NonUnitalNonAssocSemiring A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] -variable [NonUnitalNonAssocSemiring B] [Module R B] [IsScalarTower R B B] [SMulCommClass R B B] +variable [NonUnitalNonAssocSemiring A] [Module R A] variable (S : NonUnitalSubalgebra R A) +theorem range_val : NonUnitalAlgHom.range (NonUnitalSubalgebraClass.subtype S) = S := + ext <| Set.ext_iff.1 <| (NonUnitalSubalgebraClass.subtype S).coe_range.trans Subtype.range_val + instance subsingleton_of_subsingleton [Subsingleton A] : Subsingleton (NonUnitalSubalgebra R A) := ⟨fun B C => ext fun x => by simp only [Subsingleton.elim x 0, zero_mem B, zero_mem C]⟩ +variable [IsScalarTower R A A] [SMulCommClass R A A] +variable [NonUnitalNonAssocSemiring B] [Module R B] + instance _root_.NonUnitalAlgHom.subsingleton [Subsingleton (NonUnitalSubalgebra R A)] : Subsingleton (A →ₙₐ[R] B) := ⟨fun f g => @@ -836,8 +842,6 @@ instance _root_.NonUnitalAlgHom.subsingleton [Subsingleton (NonUnitalSubalgebra Subsingleton.elim (⊤ : NonUnitalSubalgebra R A) ⊥ ▸ mem_top (mem_bot.mp this).symm ▸ (map_zero f).trans (map_zero g).symm⟩ -theorem range_val : NonUnitalAlgHom.range (NonUnitalSubalgebraClass.subtype S) = S := - ext <| Set.ext_iff.1 <| (NonUnitalSubalgebraClass.subtype S).coe_range.trans Subtype.range_val /-- The map `S → T` when `S` is a non-unital subalgebra contained in the non-unital subalgebra `T`. @@ -898,6 +902,8 @@ theorem mem_prod {S : NonUnitalSubalgebra R A} {S₁ : NonUnitalSubalgebra R B} x ∈ prod S S₁ ↔ x.1 ∈ S ∧ x.2 ∈ S₁ := Set.mem_prod +variable [IsScalarTower R B B] [SMulCommClass R B B] + @[simp] theorem prod_top : (prod ⊤ ⊤ : NonUnitalSubalgebra R (A × B)) = ⊤ := by ext; simp diff --git a/Mathlib/Algebra/Algebra/Quasispectrum.lean b/Mathlib/Algebra/Algebra/Quasispectrum.lean index 20101da6f8cbf..cb01e4b60ed6e 100644 --- a/Mathlib/Algebra/Algebra/Quasispectrum.lean +++ b/Mathlib/Algebra/Algebra/Quasispectrum.lean @@ -236,7 +236,7 @@ lemma isQuasiregular_iff_isUnit' (R : Type*) {A : Type*} [CommSemiring R] [NonUn · exact ⟨(Unitization.unitsFstOne_mulEquiv_quasiregular R) ⟨hx.unit, by simp⟩, by simp⟩ variable (R : Type*) {A : Type*} [CommSemiring R] [NonUnitalRing A] - [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] + [Module R A] /-- If `A` is a non-unital `R`-algebra, the `R`-quasispectrum of `a : A` consists of those `r : R` such that if `r` is invertible (in `R`), then `-(r⁻¹ • a)` is not quasiregular. @@ -291,6 +291,7 @@ lemma mem_quasispectrum_iff {R A : Type*} [Semifield R] [Ring A] simp [quasispectrum_eq_spectrum_union_zero] namespace Unitization +variable [IsScalarTower R A A] [SMulCommClass R A A] lemma isQuasiregular_inr_iff (a : A) : IsQuasiregular (a : Unitization R A) ↔ IsQuasiregular a := by @@ -406,7 +407,7 @@ namespace QuasispectrumRestricts section NonUnital variable {R S A : Type*} [Semifield R] [Field S] [NonUnitalRing A] [Module R A] [Module S A] -variable [IsScalarTower S A A] [SMulCommClass S A A] [Algebra R S] {a : A} {f : S → R} +variable [Algebra R S] {a : A} {f : S → R} protected theorem map_zero (h : QuasispectrumRestricts a f) : f 0 = 0 := by rw [← h.left_inv 0, map_zero (algebraMap R S)] @@ -416,7 +417,12 @@ theorem of_subset_range_algebraMap (hf : f.LeftInverse (algebraMap R S)) rightInvOn := fun s hs => by obtain ⟨r, rfl⟩ := h hs; rw [hf r] left_inv := hf -variable [IsScalarTower R S A] +lemma of_quasispectrum_eq {a b : A} {f : S → R} (ha : QuasispectrumRestricts a f) + (h : quasispectrum S a = quasispectrum S b) : QuasispectrumRestricts b f where + rightInvOn := h ▸ ha.rightInvOn + left_inv := ha.left_inv + +variable [IsScalarTower S A A] [SMulCommClass S A A] [IsScalarTower R S A] theorem algebraMap_image (h : QuasispectrumRestricts a f) : algebraMap R S '' quasispectrum R a = quasispectrum S a := by @@ -436,11 +442,6 @@ theorem subset_preimage (h : QuasispectrumRestricts a f) : quasispectrum S a ⊆ f ⁻¹' quasispectrum R a := h.image ▸ (quasispectrum S a).subset_preimage_image f -lemma of_quasispectrum_eq {a b : A} {f : S → R} (ha : QuasispectrumRestricts a f) - (h : quasispectrum S a = quasispectrum S b) : QuasispectrumRestricts b f where - rightInvOn := h ▸ ha.rightInvOn - left_inv := ha.left_inv - protected lemma comp {R₁ R₂ R₃ A : Type*} [Semifield R₁] [Field R₂] [Field R₃] [NonUnitalRing A] [Module R₁ A] [Module R₂ A] [Module R₃ A] [Algebra R₁ R₂] [Algebra R₂ R₃] [Algebra R₁ R₃] [IsScalarTower R₁ R₂ R₃] [IsScalarTower R₂ R₃ A] [IsScalarTower R₃ A A] @@ -500,6 +501,13 @@ theorem of_subset_range_algebraMap (hf : f.LeftInverse (algebraMap R S)) rw [hf r] left_inv := hf +lemma of_spectrum_eq {a b : A} {f : S → R} (ha : SpectrumRestricts a f) + (h : spectrum S a = spectrum S b) : SpectrumRestricts b f where + rightInvOn := by + rw [quasispectrum_eq_spectrum_union_zero, ← h, ← quasispectrum_eq_spectrum_union_zero] + exact QuasispectrumRestricts.rightInvOn ha + left_inv := ha.left_inv + variable [IsScalarTower R S A] theorem algebraMap_image (h : SpectrumRestricts a f) : @@ -519,13 +527,6 @@ theorem apply_mem (h : SpectrumRestricts a f) {s : S} (hs : s ∈ spectrum S a) theorem subset_preimage (h : SpectrumRestricts a f) : spectrum S a ⊆ f ⁻¹' spectrum R a := h.image ▸ (spectrum S a).subset_preimage_image f -lemma of_spectrum_eq {a b : A} {f : S → R} (ha : SpectrumRestricts a f) - (h : spectrum S a = spectrum S b) : SpectrumRestricts b f where - rightInvOn := by - rw [quasispectrum_eq_spectrum_union_zero, ← h, ← quasispectrum_eq_spectrum_union_zero] - exact QuasispectrumRestricts.rightInvOn ha - left_inv := ha.left_inv - end Unital end SpectrumRestricts diff --git a/Mathlib/Algebra/Algebra/Unitization.lean b/Mathlib/Algebra/Algebra/Unitization.lean index 491a87438106e..886c96c934125 100644 --- a/Mathlib/Algebra/Algebra/Unitization.lean +++ b/Mathlib/Algebra/Algebra/Unitization.lean @@ -710,8 +710,8 @@ end AlgHom section StarAlgHom variable {R A C : Type*} [CommSemiring R] [StarRing R] [NonUnitalSemiring A] [StarRing A] -variable [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] [StarModule R A] -variable [Semiring C] [Algebra R C] [StarRing C] [StarModule R C] +variable [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] +variable [Semiring C] [Algebra R C] [StarRing C] /-- See note [partially-applied ext lemmas] -/ @[ext] @@ -721,6 +721,8 @@ theorem starAlgHom_ext {φ ψ : Unitization R A →⋆ₐ[R] C} φ = ψ := Unitization.algHom_ext'' <| DFunLike.congr_fun h +variable [StarModule R C] + /-- Non-unital star algebra homomorphisms from `A` into a unital star `R`-algebra `C` lift uniquely to `Unitization R A →⋆ₐ[R] C`. This is the universal property of the unitization. -/ @[simps! apply symm_apply apply_apply] @@ -737,37 +739,41 @@ def starLift : (A →⋆ₙₐ[R] C) ≃ (Unitization R A →⋆ₐ[R] C) := -- Note (#6057) : tagging simpNF because linter complains @[simp high, nolint simpNF] -theorem starLift_symm_apply_apply (φ : Unitization R A →ₐ[R] C) (a : A) : - Unitization.lift.symm φ a = φ a := +theorem starLift_symm_apply_apply (φ : Unitization R A →⋆ₐ[R] C) (a : A) : + Unitization.starLift.symm φ a = φ a := rfl end StarAlgHom section StarNormal -variable {R A : Type*} [Semiring R] [AddCommMonoid A] [Mul A] [SMulWithZero R A] +variable {R A : Type*} [Semiring R] variable [StarAddMonoid R] [Star A] {a : A} -@[simp] -lemma isStarNormal_inr : IsStarNormal (a : Unitization R A) ↔ IsStarNormal a := by - simp only [isStarNormal_iff, commute_iff_eq, ← inr_star, ← inr_mul, inr_injective.eq_iff] - -variable (R a) in -instance instIsStarNormal (a : A) [IsStarNormal a] : - IsStarNormal (a : Unitization R A) := - isStarNormal_inr.mpr ‹_› @[simp] lemma isSelfAdjoint_inr : IsSelfAdjoint (a : Unitization R A) ↔ IsSelfAdjoint a := by simp only [isSelfAdjoint_iff, ← inr_star, ← inr_mul, inr_injective.eq_iff] +alias ⟨_root_.IsSelfAdjoint.of_inr, _⟩ := isSelfAdjoint_inr + variable (R) in lemma _root_.IsSelfAdjoint.inr (ha : IsSelfAdjoint a) : IsSelfAdjoint (a : Unitization R A) := isSelfAdjoint_inr.mpr ha -alias ⟨_root_.IsSelfAdjoint.of_inr, _⟩ := isSelfAdjoint_inr +variable [AddCommMonoid A] [Mul A] [SMulWithZero R A] + +@[simp] +lemma isStarNormal_inr : IsStarNormal (a : Unitization R A) ↔ IsStarNormal a := by + simp only [isStarNormal_iff, commute_iff_eq, ← inr_star, ← inr_mul, inr_injective.eq_iff] + alias ⟨_root_.IsStarNormal.of_inr, _⟩ := isStarNormal_inr +variable (R a) in +instance instIsStarNormal (a : A) [IsStarNormal a] : + IsStarNormal (a : Unitization R A) := + isStarNormal_inr.mpr ‹_› + end StarNormal end Unitization diff --git a/Mathlib/Algebra/DirectSum/Internal.lean b/Mathlib/Algebra/DirectSum/Internal.lean index 8ac96aa59b9a8..ac5d09c97b8de 100644 --- a/Mathlib/Algebra/DirectSum/Internal.lean +++ b/Mathlib/Algebra/DirectSum/Internal.lean @@ -283,7 +283,7 @@ instance galgebra [AddMonoid ι] [CommSemiring S] [Semiring R] [Algebra S R] (A smul_def := fun _r ⟨i, _xi⟩ => Sigma.subtype_ext (zero_add i).symm <| Algebra.smul_def _ _ @[simp] -theorem setLike.coe_galgebra_toFun [AddMonoid ι] [CommSemiring S] [Semiring R] [Algebra S R] +theorem setLike.coe_galgebra_toFun {ι} [AddMonoid ι] [CommSemiring S] [Semiring R] [Algebra S R] (A : ι → Submodule S R) [SetLike.GradedMonoid A] (s : S) : (DirectSum.GAlgebra.toFun (A := fun i => A i) s) = (algebraMap S R s : R) := rfl diff --git a/Mathlib/Algebra/DirectSum/Module.lean b/Mathlib/Algebra/DirectSum/Module.lean index 36d41b3ceac64..dd4e97f981a94 100644 --- a/Mathlib/Algebra/DirectSum/Module.lean +++ b/Mathlib/Algebra/DirectSum/Module.lean @@ -33,7 +33,7 @@ open DirectSum section General variable {R : Type u} [Semiring R] -variable {ι : Type v} [dec_ι : DecidableEq ι] +variable {ι : Type v} variable {M : ι → Type w} [∀ i, AddCommMonoid (M i)] [∀ i, Module R (M i)] instance : Module R (⨁ i, M i) := @@ -54,6 +54,7 @@ theorem smul_apply (b : R) (v : ⨁ i, M i) (i : ι) : (b • v) i = b • v i : DFinsupp.smul_apply _ _ _ variable (R ι M) +variable [DecidableEq ι] /-- Create the direct sum given a family `M` of `R` modules indexed over `ι`. -/ def lmk : ∀ s : Finset ι, (∀ i : (↑s : Set ι), M i.val) →ₗ[R] ⨁ i, M i := @@ -144,14 +145,14 @@ def linearEquivFunOnFintype [Fintype ι] : (⨁ i, M i) ≃ₗ[R] ∀ i, M i := variable {ι M} @[simp] -theorem linearEquivFunOnFintype_lof [Fintype ι] [DecidableEq ι] (i : ι) (m : M i) : +theorem linearEquivFunOnFintype_lof [Fintype ι] (i : ι) (m : M i) : (linearEquivFunOnFintype R ι M) (lof R ι M i m) = Pi.single i m := by ext a change (DFinsupp.equivFunOnFintype (lof R ι M i m)) a = _ convert _root_.congr_fun (DFinsupp.equivFunOnFintype_single i m) a @[simp] -theorem linearEquivFunOnFintype_symm_single [Fintype ι] [DecidableEq ι] (i : ι) (m : M i) : +theorem linearEquivFunOnFintype_symm_single [Fintype ι] (i : ι) (m : M i) : (linearEquivFunOnFintype R ι M).symm (Pi.single i m) = lof R ι M i m := by change (DFinsupp.equivFunOnFintype.symm (Pi.single i m)) = _ rw [DFinsupp.equivFunOnFintype_symm_single i m] diff --git a/Mathlib/Algebra/Homology/Embedding/TruncGE.lean b/Mathlib/Algebra/Homology/Embedding/TruncGE.lean index 1cb72b2a91514..3cfeee56bada8 100644 --- a/Mathlib/Algebra/Homology/Embedding/TruncGE.lean +++ b/Mathlib/Algebra/Homology/Embedding/TruncGE.lean @@ -40,7 +40,7 @@ it induces an isomorphism in homology in degrees in the image of `e.f`. open CategoryTheory Limits ZeroObject Category variable {ι ι' : Type*} {c : ComplexShape ι} {c' : ComplexShape ι'} - {C : Type*} [Category C] [HasZeroMorphisms C] [HasZeroObject C] + {C : Type*} [Category C] [HasZeroMorphisms C] namespace HomologicalComplex @@ -128,6 +128,8 @@ lemma truncGE'_d_eq_fromOpcycles {i j : ι} (hij : c.Rel i j) {i' j' : ι'} subst hi' hj' simp [truncGE'XIso, truncGE'XIsoOpcycles] +variable [HasZeroObject C] + /-- The canonical truncation of a homological complex relative to an embedding of complex shapes `e` which satisfies `e.IsTruncGE`. -/ noncomputable def truncGE : HomologicalComplex C c' := (K.truncGE' e).extend e diff --git a/Mathlib/Algebra/Homology/HomotopyCofiber.lean b/Mathlib/Algebra/Homology/HomotopyCofiber.lean index 3ada11873a048..6b4f8c2872560 100644 --- a/Mathlib/Algebra/Homology/HomotopyCofiber.lean +++ b/Mathlib/Algebra/Homology/HomotopyCofiber.lean @@ -339,7 +339,7 @@ lemma eq_desc (f : homotopyCofiber φ ⟶ K) (hc : ∀ j, ∃ i, c.Rel i j) : end -lemma descSigma_ext_iff {K : HomologicalComplex C c} +lemma descSigma_ext_iff {φ : F ⟶ G} {K : HomologicalComplex C c} (x y : Σ (α : G ⟶ K), Homotopy (φ ≫ α) 0) : x = y ↔ x.1 = y.1 ∧ (∀ (i j : ι) (_ : c.Rel j i), x.2.hom i j = y.2.hom i j) := by constructor diff --git a/Mathlib/Algebra/Lie/Subalgebra.lean b/Mathlib/Algebra/Lie/Subalgebra.lean index d3af27be100d6..f5b3887358c7b 100644 --- a/Mathlib/Algebra/Lie/Subalgebra.lean +++ b/Mathlib/Algebra/Lie/Subalgebra.lean @@ -201,7 +201,7 @@ theorem coe_to_submodule : ((L' : Submodule R L) : Set L) = L' := section LieModule variable {M : Type w} [AddCommGroup M] [LieRingModule L M] -variable {N : Type w₁} [AddCommGroup N] [LieRingModule L N] [Module R N] [LieModule R L N] +variable {N : Type w₁} [AddCommGroup N] [LieRingModule L N] [Module R N] /-- Given a Lie algebra `L` containing a Lie subalgebra `L' ⊆ L`, together with a Lie ring module `M` of `L`, we may regard `M` as a Lie ring module of `L'` by restriction. -/ @@ -215,11 +215,11 @@ instance lieRingModule : LieRingModule L' M where theorem coe_bracket_of_module (x : L') (m : M) : ⁅x, m⁆ = ⁅(x : L), m⁆ := rfl -variable [Module R M] [LieModule R L M] +variable [Module R M] /-- Given a Lie algebra `L` containing a Lie subalgebra `L' ⊆ L`, together with a Lie module `M` of `L`, we may regard `M` as a Lie module of `L'` by restriction. -/ -instance lieModule : LieModule R L' M where +instance lieModule [LieModule R L M] : LieModule R L' M where smul_lie t x m := by simp only [coe_bracket_of_module, smul_lie, Submodule.coe_smul_of_tower] lie_smul t x m := by simp only [coe_bracket_of_module, lie_smul] diff --git a/Mathlib/Algebra/Module/Submodule/Pointwise.lean b/Mathlib/Algebra/Module/Submodule/Pointwise.lean index 8dd1930ee2af6..99e50eaa44e68 100644 --- a/Mathlib/Algebra/Module/Submodule/Pointwise.lean +++ b/Mathlib/Algebra/Module/Submodule/Pointwise.lean @@ -299,7 +299,7 @@ then this action actually gives a module structure on submodules of `M` over sub section set_acting_on_submodules variable {S : Type*} [Monoid S] -variable [AddCommMonoid M] [Module R M] [DistribMulAction S M] +variable [DistribMulAction S M] /-- Let `s ⊆ R` be a set and `N ≤ M` be a submodule, then `s • N` is the smallest submodule containing diff --git a/Mathlib/Algebra/MonoidAlgebra/Basic.lean b/Mathlib/Algebra/MonoidAlgebra/Basic.lean index 3176c280a11c6..41dceb8c077b0 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Basic.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Basic.lean @@ -1508,7 +1508,7 @@ theorem of'_eq_of [AddZeroClass G] (a : G) : of' k G a = of k G (.ofAdd a) := rf theorem of_injective [Nontrivial k] [AddZeroClass G] : Function.Injective (of k G) := MonoidAlgebra.of_injective -theorem of'_commute [Semiring k] [AddZeroClass G] {a : G} (h : ∀ a', AddCommute a a') +theorem of'_commute [AddZeroClass G] {a : G} (h : ∀ a', AddCommute a a') (f : AddMonoidAlgebra k G) : Commute (of' k G a) f := MonoidAlgebra.of_commute (G := Multiplicative G) h f diff --git a/Mathlib/Algebra/MonoidAlgebra/Degree.lean b/Mathlib/Algebra/MonoidAlgebra/Degree.lean index b453e2c04b610..e0bf576fc7e96 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Degree.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Degree.lean @@ -251,7 +251,12 @@ theorem apply_eq_zero_of_not_le_supDegree {p : R[A]} {a : A} (hlt : ¬ D a ≤ p contrapose! hlt exact Finset.le_sup (Finsupp.mem_support_iff.2 hlt) -variable [AddZeroClass A] {p q : R[A]} +theorem supDegree_withBot_some_comp {s : AddMonoidAlgebra R A} (hs : s.support.Nonempty) : + supDegree (WithBot.some ∘ D) s = supDegree D s := by + unfold AddMonoidAlgebra.supDegree + rw [← Finset.coe_sup' hs, Finset.sup'_eq_sup] + +variable [AddZeroClass A] {p q : R[A]} @[simp] theorem supDegree_zero : (0 : R[A]).supDegree D = ⊥ := by simp [supDegree] @@ -261,10 +266,10 @@ theorem ne_zero_of_supDegree_ne_bot : p.supDegree D ≠ ⊥ → p ≠ 0 := mt (f theorem ne_zero_of_not_supDegree_le {b : B} (h : ¬ p.supDegree D ≤ b) : p ≠ 0 := ne_zero_of_supDegree_ne_bot (fun he => h <| he ▸ bot_le) -variable [Add B] (hadd : ∀ a1 a2, D (a1 + a2) = D a1 + D a2) +variable [Add B] -theorem supDegree_mul_le [CovariantClass B B (· + ·) (· ≤ ·)] - [CovariantClass B B (Function.swap (· + ·)) (· ≤ ·)] : +theorem supDegree_mul_le (hadd : ∀ a1 a2, D (a1 + a2) = D a1 + D a2) + [CovariantClass B B (· + ·) (· ≤ ·)] [CovariantClass B B (Function.swap (· + ·)) (· ≤ ·)] : (p * q).supDegree D ≤ p.supDegree D + q.supDegree D := sup_support_mul_le (fun {_ _} => (hadd _ _).le) p q @@ -282,10 +287,9 @@ theorem supDegree_prod_le {R A B : Type*} [CommSemiring R] [AddCommMonoid A] [Ad rw [Finset.prod_insert his, Finset.sum_insert his] exact (supDegree_mul_le hadd).trans (add_le_add_left ih _) -variable [CovariantClass B B (· + ·) (· < ·)] [CovariantClass B B (Function.swap (· + ·)) (· < ·)] - -theorem apply_add_of_supDegree_le (hD : D.Injective) {ap aq : A} - (hp : p.supDegree D ≤ D ap) (hq : q.supDegree D ≤ D aq) : +theorem apply_add_of_supDegree_le (hadd : ∀ a1 a2, D (a1 + a2) = D a1 + D a2) + [CovariantClass B B (· + ·) (· < ·)] [CovariantClass B B (Function.swap (· + ·)) (· < ·)] + (hD : D.Injective) {ap aq : A} (hp : p.supDegree D ≤ D ap) (hq : q.supDegree D ≤ D aq) : (p * q) (ap + aq) = p ap * q aq := by classical simp_rw [mul_apply, Finsupp.sum] @@ -303,16 +307,11 @@ theorem apply_add_of_supDegree_le (hD : D.Injective) {ap aq : A} · refine fun h => Finset.sum_eq_zero (fun a _ => ite_eq_right_iff.mpr <| fun _ => ?_) rw [Finsupp.not_mem_support_iff.mp h, zero_mul] -theorem supDegree_withBot_some_comp {s : AddMonoidAlgebra R A} (hs : s.support.Nonempty) : - supDegree (WithBot.some ∘ D) s = supDegree D s := by - unfold AddMonoidAlgebra.supDegree - rw [← Finset.coe_sup' hs, Finset.sup'_eq_sup] - end SupDegree section InfDegree -variable [AddZeroClass A] [SemilatticeInf T] [Add T] [OrderTop T] (D : A → T) +variable [SemilatticeInf T] [OrderTop T] (D : A → T) /-- Let `R` be a semiring, let `A` be an `AddZeroClass`, let `T` be an `OrderTop`, and let `D : A → T` be a "degree" function. @@ -328,19 +327,19 @@ theorem le_infDegree_add (f g : R[A]) : (f.infDegree D) ⊓ (g.infDegree D) ≤ (f + g).infDegree D := le_inf_support_add D f g -variable [CovariantClass T T (· + ·) (· ≤ ·)] [CovariantClass T T (Function.swap (· + ·)) (· ≤ ·)] - (D : AddHom A T) in -theorem le_infDegree_mul (f g : R[A]) : - f.infDegree D + g.infDegree D ≤ (f * g).infDegree D := - le_inf_support_mul (fun {a b : A} => (map_add D a b).ge) _ _ - -variable {D} - +variable {D} in theorem infDegree_withTop_some_comp {s : AddMonoidAlgebra R A} (hs : s.support.Nonempty) : infDegree (WithTop.some ∘ D) s = infDegree D s := by unfold AddMonoidAlgebra.infDegree rw [← Finset.coe_inf' hs, Finset.inf'_eq_inf] +theorem le_infDegree_mul + [AddZeroClass A] [Add T] + [CovariantClass T T (· + ·) (· ≤ ·)] [CovariantClass T T (Function.swap (· + ·)) (· ≤ ·)] + (D : AddHom A T) (f g : R[A]) : + f.infDegree D + g.infDegree D ≤ (f * g).infDegree D := + le_inf_support_mul (fun {a b : A} => (map_add D a b).ge) _ _ + end InfDegree end Degrees diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index 4f61b06dd1dd6..9ff01ac8654c1 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -1101,7 +1101,7 @@ theorem eval_assoc {τ} (f : σ → MvPolynomial τ R) (g : τ → R) (p : MvPol theorem eval₂_id {g : σ → R} (p : MvPolynomial σ R) : eval₂ (RingHom.id _) g p = eval g p := rfl -theorem eval_eval₂ {S τ : Type*} {x : τ → S} [CommSemiring R] [CommSemiring S] +theorem eval_eval₂ {S τ : Type*} {x : τ → S} [CommSemiring S] (f : R →+* MvPolynomial τ S) (g : σ → MvPolynomial τ S) (p : MvPolynomial σ R) : eval x (eval₂ f g p) = eval₂ ((eval x).comp f) (fun s => eval x (g s)) p := by apply induction_on p diff --git a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean index 56a02b6d6dfb7..83a147ec797e3 100644 --- a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean @@ -519,9 +519,7 @@ open scoped Pointwise variable [CommSemiring R] [StarRing R] variable [NonUnitalSemiring A] [StarRing A] [Module R A] -variable [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] -variable [NonUnitalSemiring B] [StarRing B] [Module R B] -variable [IsScalarTower R B B] [SMulCommClass R B B] [StarModule R B] +variable [StarModule R A] /-- The pointwise `star` of a non-unital subalgebra is a non-unital subalgebra. -/ instance instInvolutiveStar : InvolutiveStar (NonUnitalSubalgebra R A) where @@ -552,6 +550,7 @@ theorem star_mono : Monotone (star : NonUnitalSubalgebra R A → NonUnitalSubalg fun _ _ h _ hx => h hx variable (R) +variable [IsScalarTower R A A] [SMulCommClass R A A] /-- The star operation on `NonUnitalSubalgebra` commutes with `NonUnitalAlgebra.adjoin`. -/ theorem star_adjoin_comm (s : Set A) : diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index a8f3e31cf0af9..a265c07f161ca 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -702,7 +702,7 @@ end Inv section DivisionSemiring variable {R : Type u} {M : Type v} -variable [DivisionSemiring R] [AddCommGroup M] [Module Rᵐᵒᵖ M] [Module R M] [SMulCommClass R Rᵐᵒᵖ M] +variable [DivisionSemiring R] [AddCommGroup M] [Module Rᵐᵒᵖ M] [Module R M] protected theorem inv_inl (r : R) : (inl r)⁻¹ = (inl (r⁻¹ : R) : tsze R M) := by @@ -724,18 +724,20 @@ protected theorem inv_zero : (0 : tsze R M)⁻¹ = (0 : tsze R M) := by protected theorem inv_one : (1 : tsze R M)⁻¹ = (1 : tsze R M) := by rw [← inl_one, TrivSqZeroExt.inv_inl, inv_one] -protected theorem mul_inv_cancel {x : tsze R M} (hx : fst x ≠ 0) : x * x⁻¹ = 1 := by - ext - · rw [fst_mul, fst_inv, fst_one, mul_inv_cancel hx] - · rw [snd_mul, snd_inv, snd_one, smul_neg, smul_comm, smul_smul, mul_inv_cancel hx, one_smul, - fst_inv, add_left_neg] - protected theorem inv_mul_cancel {x : tsze R M} (hx : fst x ≠ 0) : x⁻¹ * x = 1 := by ext · rw [fst_mul, fst_inv, inv_mul_cancel hx, fst_one] · rw [snd_mul, snd_inv, snd_one, smul_neg, op_smul_op_smul, inv_mul_cancel hx, op_one, one_smul, fst_inv, add_right_neg] +variable [SMulCommClass R Rᵐᵒᵖ M] + +protected theorem mul_inv_cancel {x : tsze R M} (hx : fst x ≠ 0) : x * x⁻¹ = 1 := by + ext + · rw [fst_mul, fst_inv, fst_one, mul_inv_cancel hx] + · rw [snd_mul, snd_inv, snd_one, smul_neg, smul_comm, smul_smul, mul_inv_cancel hx, one_smul, + fst_inv, add_left_neg] + protected theorem mul_inv_rev (a b : tsze R M) : (a * b)⁻¹ = b⁻¹ * a⁻¹ := by ext @@ -763,7 +765,7 @@ end DivisionSemiring section DivisionRing variable {R : Type u} {M : Type v} -variable [DivisionRing R] [AddCommGroup M] [Module Rᵐᵒᵖ M] [Module R M] [SMulCommClass R Rᵐᵒᵖ M] +variable [DivisionRing R] [AddCommGroup M] [Module Rᵐᵒᵖ M] [Module R M] protected theorem inv_neg {x : tsze R M} : (-x)⁻¹ = -(x⁻¹) := by ext <;> simp [inv_neg] diff --git a/Mathlib/Analysis/Convex/Cone/Basic.lean b/Mathlib/Analysis/Convex/Cone/Basic.lean index 6da2cee9e7342..095dee8252a90 100644 --- a/Mathlib/Analysis/Convex/Cone/Basic.lean +++ b/Mathlib/Analysis/Convex/Cone/Basic.lean @@ -194,7 +194,7 @@ end Module section Maps -variable [AddCommMonoid E] [AddCommMonoid F] [AddCommMonoid G] +variable [AddCommMonoid F] [AddCommMonoid G] variable [Module 𝕜 E] [Module 𝕜 F] [Module 𝕜 G] /-- The image of a convex cone under a `𝕜`-linear map is a convex cone. -/ diff --git a/Mathlib/Analysis/Convex/Mul.lean b/Mathlib/Analysis/Convex/Mul.lean index 118ec5b547b62..634ffe1a33097 100644 --- a/Mathlib/Analysis/Convex/Mul.lean +++ b/Mathlib/Analysis/Convex/Mul.lean @@ -26,7 +26,7 @@ variable {𝕜 E F : Type*} section LinearOrderedCommRing variable [LinearOrderedCommRing 𝕜] [LinearOrderedCommRing E] [LinearOrderedAddCommGroup F] [Module 𝕜 E] [Module 𝕜 F] [Module E F] [IsScalarTower 𝕜 E F] [SMulCommClass 𝕜 E F] - [OrderedSMul 𝕜 E] [OrderedSMul 𝕜 F] [OrderedSMul E F] {s : Set 𝕜} {f : 𝕜 → E} {g : 𝕜 → F} + [OrderedSMul 𝕜 F] [OrderedSMul E F] {s : Set 𝕜} {f : 𝕜 → E} {g : 𝕜 → F} lemma ConvexOn.smul' (hf : ConvexOn 𝕜 s f) (hg : ConvexOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ f x) (hg₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ g x) (hfg : MonovaryOn f g s) : ConvexOn 𝕜 s (f • g) := by @@ -47,8 +47,9 @@ lemma ConvexOn.smul' (hf : ConvexOn 𝕜 s f) (hg : ConvexOn 𝕜 s g) (hf₀ : ← smul_smul_smul_comm b b, smul_eq_mul, smul_eq_mul, smul_eq_mul, smul_eq_mul, mul_comm b, add_comm _ ((b * b) • f y • g y), add_add_add_comm, add_comm ((a * b) • f y • g x)] -lemma ConcaveOn.smul' (hf : ConcaveOn 𝕜 s f) (hg : ConcaveOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ f x) - (hg₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ g x) (hfg : AntivaryOn f g s) : ConcaveOn 𝕜 s (f • g) := by +lemma ConcaveOn.smul' [OrderedSMul 𝕜 E] (hf : ConcaveOn 𝕜 s f) (hg : ConcaveOn 𝕜 s g) + (hf₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ f x) (hg₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ g x) (hfg : AntivaryOn f g s) : + ConcaveOn 𝕜 s (f • g) := by refine ⟨hf.1, fun x hx y hy a b ha hb hab ↦ ?_⟩ dsimp refine (smul_le_smul (hf.2 hx hy ha hb hab) (hg.2 hx hy ha hb hab) @@ -66,8 +67,9 @@ lemma ConcaveOn.smul' (hf : ConcaveOn 𝕜 s f) (hg : ConcaveOn 𝕜 s g) (hf₀ ← smul_smul_smul_comm b b, smul_eq_mul, smul_eq_mul, smul_eq_mul, smul_eq_mul, mul_comm b a, add_comm ((a * b) • f x • g y), add_comm ((a * b) • f x • g y), add_add_add_comm] -lemma ConvexOn.smul'' (hf : ConvexOn 𝕜 s f) (hg : ConvexOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → f x ≤ 0) - (hg₀ : ∀ ⦃x⦄, x ∈ s → g x ≤ 0) (hfg : AntivaryOn f g s) : ConcaveOn 𝕜 s (f • g) := by +lemma ConvexOn.smul'' [OrderedSMul 𝕜 E] (hf : ConvexOn 𝕜 s f) (hg : ConvexOn 𝕜 s g) + (hf₀ : ∀ ⦃x⦄, x ∈ s → f x ≤ 0) (hg₀ : ∀ ⦃x⦄, x ∈ s → g x ≤ 0) (hfg : AntivaryOn f g s) : + ConcaveOn 𝕜 s (f • g) := by rw [← neg_smul_neg] exact hf.neg.smul' hg.neg (fun x hx ↦ neg_nonneg.2 <| hf₀ hx) (fun x hx ↦ neg_nonneg.2 <| hg₀ hx) hfg.neg @@ -84,13 +86,13 @@ lemma ConvexOn.smul_concaveOn (hf : ConvexOn 𝕜 s f) (hg : ConcaveOn 𝕜 s g) rw [← neg_convexOn_iff, ← smul_neg] exact hf.smul' hg.neg hf₀ (fun x hx ↦ neg_nonneg.2 <| hg₀ hx) hfg.neg_right -lemma ConcaveOn.smul_convexOn (hf : ConcaveOn 𝕜 s f) (hg : ConvexOn 𝕜 s g) +lemma ConcaveOn.smul_convexOn [OrderedSMul 𝕜 E] (hf : ConcaveOn 𝕜 s f) (hg : ConvexOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ f x) (hg₀ : ∀ ⦃x⦄, x ∈ s → g x ≤ 0) (hfg : MonovaryOn f g s) : ConvexOn 𝕜 s (f • g) := by rw [← neg_concaveOn_iff, ← smul_neg] exact hf.smul' hg.neg hf₀ (fun x hx ↦ neg_nonneg.2 <| hg₀ hx) hfg.neg_right -lemma ConvexOn.smul_concaveOn' (hf : ConvexOn 𝕜 s f) (hg : ConcaveOn 𝕜 s g) +lemma ConvexOn.smul_concaveOn' [OrderedSMul 𝕜 E] (hf : ConvexOn 𝕜 s f) (hg : ConcaveOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → f x ≤ 0) (hg₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ g x) (hfg : MonovaryOn f g s) : ConvexOn 𝕜 s (f • g) := by rw [← neg_concaveOn_iff, ← smul_neg] @@ -102,7 +104,7 @@ lemma ConcaveOn.smul_convexOn' (hf : ConcaveOn 𝕜 s f) (hg : ConvexOn 𝕜 s g rw [← neg_convexOn_iff, ← smul_neg] exact hf.smul'' hg.neg hf₀ (fun x hx ↦ neg_nonpos.2 <| hg₀ hx) hfg.neg_right -variable [IsScalarTower 𝕜 E E] [SMulCommClass 𝕜 E E] {f g : 𝕜 → E} +variable [OrderedSMul 𝕜 E] [IsScalarTower 𝕜 E E] [SMulCommClass 𝕜 E E] {f g : 𝕜 → E} lemma ConvexOn.mul (hf : ConvexOn 𝕜 s f) (hg : ConvexOn 𝕜 s g) (hf₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ f x) (hg₀ : ∀ ⦃x⦄, x ∈ s → 0 ≤ g x) (hfg : MonovaryOn f g s) : diff --git a/Mathlib/CategoryTheory/Abelian/LeftDerived.lean b/Mathlib/CategoryTheory/Abelian/LeftDerived.lean index e3564c2ddbc3a..be9b3ab46d0c2 100644 --- a/Mathlib/CategoryTheory/Abelian/LeftDerived.lean +++ b/Mathlib/CategoryTheory/Abelian/LeftDerived.lean @@ -249,13 +249,13 @@ noncomputable def fromLeftDerivedZero' {X : C} rw [← F.map_comp, complex_d_comp_π_f_zero, F.map_zero]) @[reassoc (attr := simp)] -lemma pOpcycles_comp_fromLeftDerivedZero' {X : C} +lemma pOpcycles_comp_fromLeftDerivedZero' {C} [Category C] [Abelian C] {X : C} (P : ProjectiveResolution X) (F : C ⥤ D) [F.Additive] : HomologicalComplex.pOpcycles _ _ ≫ P.fromLeftDerivedZero' F = F.map (P.π.f 0) := by simp [fromLeftDerivedZero'] @[reassoc] -lemma fromLeftDerivedZero'_naturality {X Y : C} (f : X ⟶ Y) +lemma fromLeftDerivedZero'_naturality {C} [Category C] [Abelian C] {X Y : C} (f : X ⟶ Y) (P : ProjectiveResolution X) (Q : ProjectiveResolution Y) (φ : P.complex ⟶ Q.complex) (comm : φ.f 0 ≫ Q.π.f 0 = P.π.f 0 ≫ f) (F : C ⥤ D) [F.Additive] : diff --git a/Mathlib/CategoryTheory/Abelian/ProjectiveResolution.lean b/Mathlib/CategoryTheory/Abelian/ProjectiveResolution.lean index 36f6551f668d1..b9dd2157df844 100644 --- a/Mathlib/CategoryTheory/Abelian/ProjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Abelian/ProjectiveResolution.lean @@ -197,9 +197,12 @@ abbrev projectiveResolution (Z : C) [HasZeroObject C] ProjectiveResolution Z := (HasProjectiveResolution.out (Z := Z)).some -variable (C) -variable [Abelian C] [HasProjectiveResolutions C] +variable [Abelian C] + +section +variable [HasProjectiveResolutions C] +variable (C) in /-- Taking projective resolutions is functorial, if considered with target the homotopy category (`ℕ`-indexed chain complexes and chain maps up to homotopy). @@ -215,7 +218,6 @@ def projectiveResolutions : C ⥤ HomotopyCategory C (ComplexShape.down ℕ) whe rw [← (HomotopyCategory.quotient _ _).map_comp] apply HomotopyCategory.eq_of_homotopy apply ProjectiveResolution.liftCompHomotopy -variable {C} /-- If `P : ProjectiveResolution X`, then the chosen `(projectiveResolutions C).obj X` is isomorphic (in the homotopy category) to `P.complex`. -/ @@ -244,6 +246,8 @@ lemma ProjectiveResolution.iso_hom_naturality {X Y : C} (f : X ⟶ Y) rw [← cancel_epi (P.iso).inv, iso_inv_naturality_assoc f P Q φ comm, Iso.inv_hom_id_assoc, Iso.inv_hom_id, comp_id] +end + variable [EnoughProjectives C] theorem exact_d_f {X Y : C} (f : X ⟶ Y) : diff --git a/Mathlib/CategoryTheory/Abelian/RightDerived.lean b/Mathlib/CategoryTheory/Abelian/RightDerived.lean index d4a9e4e767785..2c12f1e56530f 100644 --- a/Mathlib/CategoryTheory/Abelian/RightDerived.lean +++ b/Mathlib/CategoryTheory/Abelian/RightDerived.lean @@ -252,14 +252,14 @@ noncomputable def toRightDerivedZero' {X : C} zero_comp, F.map_zero]) @[reassoc (attr := simp)] -lemma toRightDerivedZero'_comp_iCycles {X : C} +lemma toRightDerivedZero'_comp_iCycles {C} [Category C] [Abelian C] {X : C} (P : InjectiveResolution X) (F : C ⥤ D) [F.Additive] : P.toRightDerivedZero' F ≫ HomologicalComplex.iCycles _ _ = F.map (P.ι.f 0) := by simp [toRightDerivedZero'] @[reassoc] -lemma toRightDerivedZero'_naturality {X Y : C} (f : X ⟶ Y) +lemma toRightDerivedZero'_naturality {C} [Category C] [Abelian C] {X Y : C} (f : X ⟶ Y) (P : InjectiveResolution X) (Q : InjectiveResolution Y) (φ : P.cocomplex ⟶ Q.cocomplex) (comm : P.ι.f 0 ≫ φ.f 0 = f ≫ Q.ι.f 0) (F : C ⥤ D) [F.Additive] : diff --git a/Mathlib/CategoryTheory/Galois/Decomposition.lean b/Mathlib/CategoryTheory/Galois/Decomposition.lean index 278333b9eb0f9..77a9a29932f42 100644 --- a/Mathlib/CategoryTheory/Galois/Decomposition.lean +++ b/Mathlib/CategoryTheory/Galois/Decomposition.lean @@ -206,7 +206,7 @@ private lemma mkSelfProdFib_map_π (t : F.obj X) : F.map (Pi.π _ t) (mkSelfProd simp only [mkSelfProdFib, FintypeCat.inv_hom_id_apply] exact Concrete.productEquiv_symm_apply_π.{w, w, w+1} (fun _ : F.obj X ↦ F.obj X) id t -variable {X} {A : C} [IsConnected A] (u : A ⟶ selfProd F X) [Mono u] +variable {X} {A : C} (u : A ⟶ selfProd F X) (a : F.obj A) (h : F.map u a = mkSelfProdFib F X) {F} /-- For each `x : F.obj X`, this is the composition of `u` with the projection at `x`. -/ @@ -220,6 +220,8 @@ private lemma selfProdProj_fiber (x : F.obj X) : simp only [selfProdProj, selfProd, F.map_comp, FintypeCat.comp_apply, h] rw [mkSelfProdFib_map_π F X x] +variable [IsConnected A] + /-- An element `b : F.obj A` defines a permutation of the fiber of `X` by projecting onto the `F.map u b` factor. -/ private noncomputable def fiberPerm (b : F.obj A) : F.obj X ≃ F.obj X := by @@ -235,7 +237,7 @@ private noncomputable def fiberPerm (b : F.obj A) : F.obj X ≃ F.obj X := by private noncomputable def selfProdPermIncl (b : F.obj A) : A ⟶ selfProd F X := u ≫ (Pi.whiskerEquiv (fiberPerm h b) (fun _ => Iso.refl X)).inv -private instance (b : F.obj A) : Mono (selfProdPermIncl h b) := mono_comp _ _ +private instance [Mono u] (b : F.obj A) : Mono (selfProdPermIncl h b) := mono_comp _ _ /-- Key technical lemma: the twisted inclusion `selfProdPermIncl h b` maps `a` to `F.map u b`. -/ private lemma selfProdTermIncl_fib_eq (b : F.obj A) : @@ -256,7 +258,7 @@ private lemma selfProdTermIncl_fib_eq (b : F.obj A) : `f` is obtained by considering `u` and `selfProdPermIncl h b`. Both are inclusions of `A` into `selfProd F X` mapping `b` respectively `a` to the same element in the fiber of `selfProd F X`. Applying `connected_component_unique` yields the result. -/ -private lemma subobj_selfProd_trans (b : F.obj A) : ∃ (f : A ≅ A), F.map f.hom b = a := by +private lemma subobj_selfProd_trans [Mono u] (b : F.obj A) : ∃ (f : A ≅ A), F.map f.hom b = a := by apply connected_component_unique F b a u (selfProdPermIncl h b) exact selfProdTermIncl_fib_eq h b diff --git a/Mathlib/CategoryTheory/Galois/GaloisObjects.lean b/Mathlib/CategoryTheory/Galois/GaloisObjects.lean index 6b18a5f9f688e..1ceed246f266d 100644 --- a/Mathlib/CategoryTheory/Galois/GaloisObjects.lean +++ b/Mathlib/CategoryTheory/Galois/GaloisObjects.lean @@ -146,8 +146,8 @@ lemma comp_autMap {A B : C} [IsConnected A] [IsGalois B] (f : A ⟶ B) (σ : Aut (exists_autMap f σ).choose_spec.left @[simp] -lemma comp_autMap_apply {A B : C} [IsConnected A] [IsGalois B] (f : A ⟶ B) (σ : Aut A) - (a : F.obj A) : +lemma comp_autMap_apply (F : C ⥤ FintypeCat.{w}) {A B : C} [IsConnected A] [IsGalois B] + (f : A ⟶ B) (σ : Aut A) (a : F.obj A) : F.map (autMap f σ).hom (F.map f a) = F.map f (F.map σ.hom a) := by simpa [-comp_autMap] using congrFun (F.congr_map (comp_autMap f σ)) a diff --git a/Mathlib/CategoryTheory/Localization/Triangulated.lean b/Mathlib/CategoryTheory/Localization/Triangulated.lean index 2f6e5e1a8eac5..d97d95429aa7f 100644 --- a/Mathlib/CategoryTheory/Localization/Triangulated.lean +++ b/Mathlib/CategoryTheory/Localization/Triangulated.lean @@ -118,9 +118,7 @@ namespace Triangulated namespace Localization variable (W : MorphismProperty C) [L.IsLocalization W] - [W.IsCompatibleWithTriangulation] [W.HasLeftCalculusOfFractions] - [Preadditive D] [HasZeroObject D] - [∀ (n : ℤ), (shiftFunctor D n).Additive] [L.Additive] + [W.HasLeftCalculusOfFractions] lemma distinguished_cocone_triangle {X Y : D} (f : X ⟶ Y) : ∃ (Z : D) (g : Y ⟶ Z) (h : Z ⟶ X⟦(1 : ℤ)⟧), @@ -137,6 +135,9 @@ lemma distinguished_cocone_triangle {X Y : D} (f : X ⟶ Y) : simp only [assoc, id_comp, ← Functor.map_comp, ← Arrow.comp_left, e.hom_inv_id, Arrow.id_left, Functor.mapArrow_obj_left, Functor.map_id, comp_id] +section +variable [W.IsCompatibleWithTriangulation] + lemma complete_distinguished_triangle_morphism (T₁ T₂ : Triangle D) (hT₁ : T₁ ∈ L.essImageDistTriang) (hT₂ : T₂ ∈ L.essImageDistTriang) (a : T₁.obj₁ ⟶ T₂.obj₁) (b : T₁.obj₂ ⟶ T₂.obj₂) (fac : T₁.mor₁ ≫ b = a ≫ T₂.mor₁) : @@ -178,6 +179,8 @@ lemma complete_distinguished_triangle_morphism (T₁ T₂ : Triangle D) simp only [Functor.map_comp, reassoc_of% hγ, MorphismProperty.LeftFraction.map_comp_map_s_assoc] +variable [HasZeroObject D] [Preadditive D] [∀ (n : ℤ), (shiftFunctor D n).Additive] [L.Additive] + /-- The pretriangulated structure on the localized category. -/ def pretriangulated : Pretriangulated D where distinguishedTriangles := L.essImageDistTriang @@ -193,11 +196,17 @@ instance isTriangulated_functor : letI : Pretriangulated D := pretriangulated L W ⟨fun T hT => ⟨T, Iso.refl _, hT⟩⟩ +end + +variable [HasZeroObject D] [Preadditive D] [∀ (n : ℤ), (shiftFunctor D n).Additive] + lemma isTriangulated [Pretriangulated D] [L.IsTriangulated] [IsTriangulated C] : IsTriangulated D := by have := essSurj_mapComposableArrows L W 2 exact isTriangulated_of_essSurj_mapComposableArrows_two L +variable [W.IsCompatibleWithTriangulation] + instance (n : ℤ) : (shiftFunctor (W.Localization) n).Additive := by rw [Localization.functor_additive_iff W.Q W] exact Functor.additive_of_iso (W.Q.commShiftIso n) diff --git a/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean b/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean index a7f3ee7a0592e..338e9167ee4a1 100644 --- a/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean +++ b/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean @@ -172,7 +172,19 @@ theorem toGraph_adjMatrix_eq [MulZeroOneClass α] [Nontrivial α] : simp only [IsAdjMatrix.toGraph_adj, adjMatrix_apply, ite_eq_left_iff, zero_ne_one] apply Classical.not_not -variable {α} [Fintype V] +variable {α} + +/-- The sum of the identity, the adjacency matrix, and its complement is the all-ones matrix. -/ +theorem one_add_adjMatrix_add_compl_adjMatrix_eq_allOnes [DecidableEq V] [DecidableEq α] + [NonAssocSemiring α] : 1 + G.adjMatrix α + (G.adjMatrix α).compl = Matrix.of fun _ _ ↦ 1 := by + ext i j + unfold Matrix.compl + rw [of_apply, add_apply, adjMatrix_apply, add_apply, adjMatrix_apply, one_apply] + by_cases h : G.Adj i j + · aesop + · split_ifs <;> simp_all + +variable [Fintype V] @[simp] theorem adjMatrix_dotProduct [NonAssocSemiring α] (v : V) (vec : V → α) : @@ -246,16 +258,6 @@ theorem adjMatrix_pow_apply_eq_card_walk [DecidableEq V] [Semiring α] (n : ℕ) cases hp simp at hxy -/-- The sum of the identity, the adjacency matrix, and its complement is the all-ones matrix. -/ -theorem one_add_adjMatrix_add_compl_adjMatrix_eq_allOnes [DecidableEq V] [DecidableEq α] - [NonAssocSemiring α] : 1 + G.adjMatrix α + (G.adjMatrix α).compl = Matrix.of fun _ _ ↦ 1 := by - ext i j - unfold Matrix.compl - rw [of_apply, add_apply, adjMatrix_apply, add_apply, adjMatrix_apply, one_apply] - by_cases h : G.Adj i j - · aesop - · split_ifs <;> simp_all - theorem dotProduct_mulVec_adjMatrix [NonAssocSemiring α] (x y : V → α) : x ⬝ᵥ (G.adjMatrix α).mulVec y = ∑ i : V, ∑ j : V, if G.Adj i j then x i * y j else 0 := by simp only [dotProduct, mulVec, adjMatrix_apply, ite_mul, one_mul, zero_mul, mul_sum, mul_ite, diff --git a/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean b/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean index 83865467fc3a9..5364ee41b6234 100644 --- a/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean +++ b/Mathlib/Combinatorics/SimpleGraph/DegreeSum.lean @@ -48,9 +48,6 @@ section DegreeSum variable [Fintype V] [DecidableRel G.Adj] --- Porting note: Changed to `Fintype (Sym2 V)` to match Combinatorics.SimpleGraph.Basic -variable [Fintype (Sym2 V)] - theorem dart_fst_fiber [DecidableEq V] (v : V) : (univ.filter fun d : G.Dart => d.fst = v) = univ.image (G.dartOfNeighborSet v) := by ext d diff --git a/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean b/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean index f8a15c8481bdd..ec09c234563e3 100644 --- a/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean +++ b/Mathlib/Combinatorics/SimpleGraph/IncMatrix.lean @@ -105,13 +105,13 @@ end MulZeroOneClass section NonAssocSemiring -variable [Fintype (Sym2 α)] [NonAssocSemiring R] {a b : α} {e : Sym2 α} +variable [NonAssocSemiring R] {a b : α} {e : Sym2 α} -theorem sum_incMatrix_apply [Fintype (neighborSet G a)] : +theorem sum_incMatrix_apply [Fintype (Sym2 α)] [Fintype (neighborSet G a)] : ∑ e, G.incMatrix R a e = G.degree a := by classical simp [incMatrix_apply', sum_boole, Set.filter_mem_univ_eq_toFinset] -theorem incMatrix_mul_transpose_diag [Fintype (neighborSet G a)] : +theorem incMatrix_mul_transpose_diag [Fintype (Sym2 α)] [Fintype (neighborSet G a)] : (G.incMatrix R * (G.incMatrix R)ᵀ) a a = G.degree a := by classical rw [← sum_incMatrix_apply] diff --git a/Mathlib/Combinatorics/SimpleGraph/Matching.lean b/Mathlib/Combinatorics/SimpleGraph/Matching.lean index b5c253cf3d920..2b7ea6ab503ed 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Matching.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Matching.lean @@ -192,9 +192,9 @@ namespace ConnectedComponent section Finite -variable [Fintype V] [DecidableRel G.Adj] +variable [Fintype V] -lemma even_card_of_isPerfectMatching [DecidableEq V] +lemma even_card_of_isPerfectMatching [DecidableEq V] [DecidableRel G.Adj] (c : ConnectedComponent G) (hM : M.IsPerfectMatching) : Even (Fintype.card c.supp) := by classical simpa using (hM.induce_connectedComponent_isMatching c).even_card diff --git a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean index 5563ad99d43b5..529afbfeaee47 100644 --- a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean +++ b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean @@ -35,7 +35,7 @@ universe u namespace SimpleGraph -variable {V : Type u} [Fintype V] [DecidableEq V] +variable {V : Type u} [Fintype V] variable (G : SimpleGraph V) [DecidableRel G.Adj] /-- A graph is strongly regular with parameters `n k ℓ μ` if @@ -63,6 +63,8 @@ theorem bot_strongly_regular : (⊥ : SimpleGraph V).IsSRGWith (Fintype.card V) ext simp [mem_commonNeighbors] +variable [DecidableEq V] + /-- Complete graphs are strongly regular. Note that `μ` can take any value for complete graphs, since there are no distinct pairs of non-adjacent vertices. -/ theorem IsSRGWith.top : @@ -158,7 +160,9 @@ theorem IsSRGWith.compl (h : G.IsSRGWith n k ℓ μ) : /-- The parameters of a strongly regular graph with at least one vertex satisfy `k * (k - ℓ - 1) = (n - k - 1) * μ`. -/ -theorem IsSRGWith.param_eq (h : G.IsSRGWith n k ℓ μ) (hn : 0 < n) : +theorem IsSRGWith.param_eq + {V : Type u} [Fintype V] (G : SimpleGraph V) [DecidableRel G.Adj] + (h : G.IsSRGWith n k ℓ μ) (hn : 0 < n) : k * (k - ℓ - 1) = (n - k - 1) * μ := by letI := Classical.decEq V rw [← h.card, Fintype.card_pos_iff] at hn diff --git a/Mathlib/Combinatorics/SimpleGraph/Turan.lean b/Mathlib/Combinatorics/SimpleGraph/Turan.lean index 097481f2c90cb..d97bad0845370 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Turan.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Turan.lean @@ -187,7 +187,9 @@ instance : DecidableRel h.setoid.r := /-- The finpartition derived from `h.setoid`. -/ def finpartition [DecidableEq V] : Finpartition (univ : Finset V) := Finpartition.ofSetoid h.setoid -lemma not_adj_iff_part_eq [DecidableEq V] : +variable [DecidableEq V] + +lemma not_adj_iff_part_eq : ¬G.Adj s t ↔ h.finpartition.part s = h.finpartition.part t := by change h.setoid.r s t ↔ _ rw [← Finpartition.mem_part_ofSetoid_iff_rel] @@ -195,7 +197,7 @@ lemma not_adj_iff_part_eq [DecidableEq V] : change t ∈ fp.part s ↔ fp.part s = fp.part t rw [fp.mem_part_iff_part_eq_part (mem_univ t) (mem_univ s), eq_comm] -lemma degree_eq_card_sub_part_card [DecidableEq V] : +lemma degree_eq_card_sub_part_card : G.degree s = Fintype.card V - (h.finpartition.part s).card := calc _ = (univ.filter (G.Adj s)).card := by @@ -208,7 +210,7 @@ lemma degree_eq_card_sub_part_card [DecidableEq V] : simp [setoid] /-- The parts of a Turán-maximal graph form an equipartition. -/ -theorem isEquipartition [DecidableEq V] : h.finpartition.IsEquipartition := by +theorem isEquipartition : h.finpartition.IsEquipartition := by set fp := h.finpartition by_contra hn rw [Finpartition.not_isEquipartition] at hn @@ -228,7 +230,7 @@ theorem isEquipartition [DecidableEq V] : h.finpartition.IsEquipartition := by have : large.card ≤ Fintype.card V := by simpa using card_le_card large.subset_univ omega -lemma card_parts_le [DecidableEq V] : h.finpartition.parts.card ≤ r := by +lemma card_parts_le : h.finpartition.parts.card ≤ r := by by_contra! l obtain ⟨z, -, hz⟩ := h.finpartition.exists_subset_part_bijOn have ncf : ¬G.CliqueFree z.card := by @@ -241,7 +243,7 @@ lemma card_parts_le [DecidableEq V] : h.finpartition.parts.card ≤ r := by /-- There are `min n r` parts in a graph on `n` vertices satisfying `G.IsTuranMaximal r`. `min` handles the `n < r` case, when `G` is complete but still `r + 1`-cliquefree for having insufficiently many vertices. -/ -theorem card_parts [DecidableEq V] : h.finpartition.parts.card = min (Fintype.card V) r := by +theorem card_parts : h.finpartition.parts.card = min (Fintype.card V) r := by set fp := h.finpartition apply le_antisymm (le_min fp.card_parts_le_card h.card_parts_le) by_contra! l @@ -281,6 +283,8 @@ theorem nonempty_iso_turanGraph (h : G.IsTuranMaximal r) : end IsTuranMaximal +variable [DecidableEq V] + /-- **Turán's theorem**, reverse direction. Any graph isomorphic to `turanGraph n r` is itself Turán-maximal if `0 < r`. -/ @@ -292,7 +296,7 @@ theorem isTuranMaximal_of_iso (f : G ≃g turanGraph n r) (hr : 0 < r) : G.IsTur fun H _ cf ↦ (f.symm.comp g).card_edgeFinset_eq ▸ j.2 H cf /-- Turán-maximality with `0 < r` transfers across graph isomorphisms. -/ -theorem IsTuranMaximal.iso {W : Type*} [Fintype W] {H : SimpleGraph W} +theorem IsTuranMaximal.iso {W : Type*} [DecidableEq W] [Fintype W] {H : SimpleGraph W} [DecidableRel H.Adj] (h : G.IsTuranMaximal r) (f : G ≃g H) (hr : 0 < r) : H.IsTuranMaximal r := isTuranMaximal_of_iso (h.nonempty_iso_turanGraph.some.comp f.symm) hr diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index 4731af2e85614..3ad3b1225f011 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -470,7 +470,7 @@ theorem diagonal_map [Zero α] [Zero β] {f : α → β} (h : f 0 = 0) {d : n simp only [diagonal_apply, map_apply] split_ifs <;> simp [h] -protected theorem map_natCast [DecidableEq n] [AddMonoidWithOne α] [AddMonoidWithOne β] +protected theorem map_natCast [AddMonoidWithOne α] [AddMonoidWithOne β] {f : α → β} (h : f 0 = 0) (d : ℕ) : (d : Matrix n n α).map f = diagonal (fun _ => f d) := diagonal_map h @@ -481,7 +481,7 @@ protected theorem map_ofNat [AddMonoidWithOne α] [AddMonoidWithOne β] (no_index (OfNat.ofNat d) : Matrix n n α).map f = diagonal (fun _ => f (OfNat.ofNat d)) := diagonal_map h -protected theorem map_intCast [DecidableEq n] [AddGroupWithOne α] [AddGroupWithOne β] +protected theorem map_intCast [AddGroupWithOne α] [AddGroupWithOne β] {f : α → β} (h : f 0 = 0) (d : ℤ) : (d : Matrix n n α).map f = diagonal (fun _ => f d) := diagonal_map h diff --git a/Mathlib/LinearAlgebra/DFinsupp.lean b/Mathlib/LinearAlgebra/DFinsupp.lean index b909ff03f9f8d..c8e7dc428808d 100644 --- a/Mathlib/LinearAlgebra/DFinsupp.lean +++ b/Mathlib/LinearAlgebra/DFinsupp.lean @@ -41,7 +41,6 @@ variable [Semiring R] [∀ i, AddCommMonoid (M i)] [∀ i, Module R (M i)] variable [AddCommMonoid N] [Module R N] section DecidableEq - variable [DecidableEq ι] /-- `DFinsupp.mk` as a `LinearMap`. -/ @@ -69,12 +68,6 @@ theorem lhom_ext' ⦃φ ψ : (Π₀ i, M i) →ₗ[R] N⦄ (h : ∀ i, φ.comp ( φ = ψ := lhom_ext fun i => LinearMap.congr_fun (h i) -/-- Interpret `fun (f : Π₀ i, M i) ↦ f i` as a linear map. -/ -def lapply (i : ι) : (Π₀ i, M i) →ₗ[R] M i where - toFun f := f i - map_add' f g := add_apply f g i - map_smul' c f := smul_apply c f i - -- This lemma has always been bad, but the linter only noticed after lean4#2644. @[simp, nolint simpNF] theorem lmk_apply (s : Finset ι) (x) : (lmk s : _ →ₗ[R] Π₀ i, M i) x = mk s x := @@ -84,6 +77,14 @@ theorem lmk_apply (s : Finset ι) (x) : (lmk s : _ →ₗ[R] Π₀ i, M i) x = m theorem lsingle_apply (i : ι) (x : M i) : (lsingle i : (M i) →ₗ[R] _) x = single i x := rfl +end DecidableEq + +/-- Interpret `fun (f : Π₀ i, M i) ↦ f i` as a linear map. -/ +def lapply (i : ι) : (Π₀ i, M i) →ₗ[R] M i where + toFun f := f i + map_add' f g := add_apply f g i + map_smul' c f := smul_apply c f i + @[simp] theorem lapply_apply (i : ι) (f : Π₀ i, M i) : (lapply i : (Π₀ i, M i) →ₗ[R] _) f = f i := rfl @@ -108,7 +109,7 @@ instance moduleOfLinearMap [Semiring S] [Module S N] [SMulCommClass R S N] : DFinsupp.module variable (S) - +variable [DecidableEq ι] instance {R : Type*} {S : Type*} [Semiring R] [Semiring S] (σ : R →+* S) {σ' : S →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] (M : Type*) (M₂ : Type*) @@ -160,8 +161,6 @@ theorem lsum_single [Semiring S] [Module S N] [SMulCommClass R S N] (F : ∀ i, end Lsum -end DecidableEq - /-! ### Bundled versions of `DFinsupp.mapRange` The names should match the equivalent bundled `Finsupp.mapRange` definitions. @@ -231,7 +230,7 @@ end mapRange section CoprodMap -variable [DecidableEq ι] [∀ x : N, Decidable (x ≠ 0)] +variable [DecidableEq ι] /-- Given a family of linear maps `f i : M i →ₗ[R] N`, we can form a linear map `(Π₀ i, M i) →ₗ[R] N` which sends `x : Π₀ i, M i` to the sum over `i` of `f i` applied to `x i`. @@ -240,7 +239,7 @@ See also `LinearMap.coprod` for the binary product version. -/ def coprodMap (f : ∀ i : ι, M i →ₗ[R] N) : (Π₀ i, M i) →ₗ[R] N := (DFinsupp.lsum ℕ fun _ : ι => LinearMap.id) ∘ₗ DFinsupp.mapRange.linearMap f -theorem coprodMap_apply (f : ∀ i : ι, M i →ₗ[R] N) (x : Π₀ i, M i) : +theorem coprodMap_apply [∀ x : N, Decidable (x ≠ 0)] (f : ∀ i : ι, M i →ₗ[R] N) (x : Π₀ i, M i) : coprodMap f x = DFinsupp.sum (mapRange (fun i => f i) (fun _ => LinearMap.map_zero _) x) fun _ => id := @@ -498,7 +497,7 @@ theorem independent_iff_dfinsupp_sumAddHom_injective (p : ι → AddSubgroup N) forms a linearly independent family. See also `CompleteLattice.Independent.linearIndependent'`. -/ -theorem Independent.linearIndependent [NoZeroSMulDivisors R N] (p : ι → Submodule R N) +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) : LinearIndependent R v := by let _ := Classical.decEq ι @@ -517,7 +516,7 @@ theorem Independent.linearIndependent [NoZeroSMulDivisors R N] (p : ι → Submo 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} +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 := let _ := Classical.decEq ι ⟨fun hv => hv.linearIndependent _ (fun i => Submodule.mem_span_singleton_self <| v i) h_ne_zero, diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index 8b398a6d36163..e03004b27ef2f 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -120,7 +120,7 @@ theorem LinearEquiv.finsuppUnique_apply (f : α →₀ M) : variable {α} @[simp] -theorem LinearEquiv.finsuppUnique_symm_apply [Unique α] (m : M) : +theorem LinearEquiv.finsuppUnique_symm_apply (m : M) : (LinearEquiv.finsuppUnique R M α).symm m = Finsupp.single default m := by ext; simp [LinearEquiv.finsuppUnique, Equiv.funUnique, single, Pi.single, equivFunOnFinite, Function.update] @@ -1306,8 +1306,10 @@ end LinearMap namespace Submodule -variable {S : Type*} [Semiring S] [Module R S] [SMulCommClass R R S] [SMulCommClass R S S] - [IsScalarTower R S S] +variable {S : Type*} [Semiring S] [Module R S] [SMulCommClass R R S] + +section +variable [SMulCommClass R S S] /-- If `M` and `N` are submodules of an `R`-algebra `S`, `m : ι → M` is a family of elements, then there is an `R`-linear map from `ι →₀ N` to `S` which maps `{ n_i }` to the sum of `m_i * n_i`. @@ -1323,6 +1325,10 @@ theorem mulLeftMap_apply_single {M N : Submodule R S} {ι : Type*} (m : ι → M mulLeftMap N m (Finsupp.single i n) = (m i).1 * n.1 := by simp [mulLeftMap] +end + +variable [IsScalarTower R S S] + /-- If `M` and `N` are submodules of an `R`-algebra `S`, `n : ι → N` is a family of elements, then there is an `R`-linear map from `ι →₀ M` to `S` which maps `{ m_i }` to the sum of `m_i * n_i`. This is used in the definition of linearly disjointness. -/ @@ -1337,7 +1343,7 @@ theorem mulRightMap_apply_single {M N : Submodule R S} {ι : Type*} (n : ι → mulRightMap M n (Finsupp.single i m) = m.1 * (n i).1 := by simp [mulRightMap] -theorem mulLeftMap_eq_mulRightMap_of_commute +theorem mulLeftMap_eq_mulRightMap_of_commute [SMulCommClass R S S] {M : Submodule R S} (N : Submodule R S) {ι : Type*} (m : ι → M) (hc : ∀ (i : ι) (n : N), Commute (m i).1 n.1) : mulLeftMap N m = mulRightMap N m := by ext i n; simp [(hc i n).eq] diff --git a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean index fbe5cc10e5f76..2d3e7dd8b3acf 100644 --- a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean +++ b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean @@ -105,34 +105,34 @@ theorem dotProduct_self_star_eq_zero {v : n → R} : dotProduct v (star v) = 0 by simp [Function.funext_iff, mul_eq_zero] @[simp] -lemma conjTranspose_mul_self_eq_zero {A : Matrix m n R} : Aᴴ * A = 0 ↔ A = 0 := +lemma conjTranspose_mul_self_eq_zero {n} {A : Matrix m n R} : Aᴴ * A = 0 ↔ A = 0 := ⟨fun h => Matrix.ext fun i j => (congr_fun <| dotProduct_star_self_eq_zero.1 <| Matrix.ext_iff.2 h j j) i, fun h => h ▸ Matrix.mul_zero _⟩ @[simp] -lemma self_mul_conjTranspose_eq_zero {A : Matrix m n R} : A * Aᴴ = 0 ↔ A = 0 := +lemma self_mul_conjTranspose_eq_zero {m} {A : Matrix m n R} : A * Aᴴ = 0 ↔ A = 0 := ⟨fun h => Matrix.ext fun i j => (congr_fun <| dotProduct_self_star_eq_zero.1 <| Matrix.ext_iff.2 h i i) j, fun h => h ▸ Matrix.zero_mul _⟩ -lemma conjTranspose_mul_self_mul_eq_zero (A : Matrix m n R) (B : Matrix n p R) : +lemma conjTranspose_mul_self_mul_eq_zero {p} (A : Matrix m n R) (B : Matrix n p R) : (Aᴴ * A) * B = 0 ↔ A * B = 0 := by refine ⟨fun h => ?_, fun h => by simp only [Matrix.mul_assoc, h, Matrix.mul_zero]⟩ apply_fun (Bᴴ * ·) at h rwa [Matrix.mul_zero, Matrix.mul_assoc, ← Matrix.mul_assoc, ← conjTranspose_mul, conjTranspose_mul_self_eq_zero] at h -lemma self_mul_conjTranspose_mul_eq_zero (A : Matrix m n R) (B : Matrix m p R) : +lemma self_mul_conjTranspose_mul_eq_zero {p} (A : Matrix m n R) (B : Matrix m p R) : (A * Aᴴ) * B = 0 ↔ Aᴴ * B = 0 := by simpa only [conjTranspose_conjTranspose] using conjTranspose_mul_self_mul_eq_zero Aᴴ _ -lemma mul_self_mul_conjTranspose_eq_zero (A : Matrix m n R) (B : Matrix p m R) : +lemma mul_self_mul_conjTranspose_eq_zero {p} (A : Matrix m n R) (B : Matrix p m R) : B * (A * Aᴴ) = 0 ↔ B * A = 0 := by rw [← conjTranspose_eq_zero, conjTranspose_mul, conjTranspose_mul, conjTranspose_conjTranspose, self_mul_conjTranspose_mul_eq_zero, ← conjTranspose_mul, conjTranspose_eq_zero] -lemma mul_conjTranspose_mul_self_eq_zero (A : Matrix m n R) (B : Matrix p n R) : +lemma mul_conjTranspose_mul_self_eq_zero {p} (A : Matrix m n R) (B : Matrix p n R) : B * (Aᴴ * A) = 0 ↔ B * Aᴴ = 0 := by simpa only [conjTranspose_conjTranspose] using mul_self_mul_conjTranspose_eq_zero Aᴴ _ diff --git a/Mathlib/LinearAlgebra/Matrix/ToLin.lean b/Mathlib/LinearAlgebra/Matrix/ToLin.lean index 4ae77f01f8732..fcce5c87672f1 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLin.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLin.lean @@ -933,7 +933,7 @@ variable {R M M₁ M₂ ι ι₁ ι₂ : Type*} [CommSemiring R] variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] variable [Module R M] [Module R M₁] [Module R M₂] variable [Fintype ι] [Fintype ι₁] [Fintype ι₂] -variable [DecidableEq ι] [DecidableEq ι₁] [DecidableEq ι₂] +variable [DecidableEq ι] [DecidableEq ι₁] variable (b : Basis ι R M) (b₁ : Basis ι₁ R M₁) (b₂ : Basis ι₂ R M₂) /-- The standard basis of the space linear maps between two modules diff --git a/Mathlib/LinearAlgebra/QuotientPi.lean b/Mathlib/LinearAlgebra/QuotientPi.lean index 4099632f4a7f3..b3f8f3b83c3c1 100644 --- a/Mathlib/LinearAlgebra/QuotientPi.lean +++ b/Mathlib/LinearAlgebra/QuotientPi.lean @@ -72,12 +72,22 @@ theorem quotientPiLift_mk (p : ∀ i, Submodule R (Ms i)) (f : ∀ i, Ms i → -- Porting note (#11083): split up the definition to avoid timeouts. Still slow. namespace quotientPi_aux -variable [Fintype ι] [DecidableEq ι] (p : ∀ i, Submodule R (Ms i)) +variable (p : ∀ i, Submodule R (Ms i)) @[simp] def toFun : ((∀ i, Ms i) ⧸ pi Set.univ p) → ∀ i, Ms i ⧸ p i := quotientPiLift p (fun i => (p i).mkQ) fun i => (ker_mkQ (p i)).ge +theorem map_add (x y : ((i : ι) → Ms i) ⧸ pi Set.univ p) : + toFun p (x + y) = toFun p x + toFun p y := + LinearMap.map_add (quotientPiLift p (fun i => (p i).mkQ) fun i => (ker_mkQ (p i)).ge) x y + +theorem map_smul (r : R) (x : ((i : ι) → Ms i) ⧸ pi Set.univ p) : + toFun p (r • x) = (RingHom.id R r) • toFun p x := + LinearMap.map_smul (quotientPiLift p (fun i => (p i).mkQ) fun i => (ker_mkQ (p i)).ge) r x + +variable [Fintype ι] [DecidableEq ι] + @[simp] def invFun : (∀ i, Ms i ⧸ p i) → (∀ i, Ms i) ⧸ pi Set.univ p := piQuotientLift p (pi Set.univ p) single fun _ => le_comap_single_pi p @@ -100,14 +110,6 @@ theorem right_inv : Function.RightInverse (invFun p) (toFun p) := by rw [Pi.single_eq_same, Pi.single_eq_same] · rw [Pi.single_eq_of_ne (Ne.symm hij), Pi.single_eq_of_ne (Ne.symm hij), Quotient.mk_zero] -theorem map_add (x y : ((i : ι) → Ms i) ⧸ pi Set.univ p) : - toFun p (x + y) = toFun p x + toFun p y := - LinearMap.map_add (quotientPiLift p (fun i => (p i).mkQ) fun i => (ker_mkQ (p i)).ge) x y - -theorem map_smul (r : R) (x : ((i : ι) → Ms i) ⧸ pi Set.univ p) : - toFun p (r • x) = (RingHom.id R r) • toFun p x := - LinearMap.map_smul (quotientPiLift p (fun i => (p i).mkQ) fun i => (ker_mkQ (p i)).ge) r x - end quotientPi_aux open quotientPi_aux in diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index fd7d15eb5777f..393c10b371f81 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -48,7 +48,7 @@ variable {R'' : Type*} [Semiring R''] variable {M : Type*} {N : Type*} {P : Type*} {Q : Type*} {S : Type*} {T : Type*} variable [AddCommMonoid M] [AddCommMonoid N] [AddCommMonoid P] variable [AddCommMonoid Q] [AddCommMonoid S] [AddCommMonoid T] -variable [Module R M] [Module R N] [Module R P] [Module R Q] [Module R S] [Module R T] +variable [Module R M] [Module R N] [Module R Q] [Module R S] [Module R T] variable [DistribMulAction R' M] variable [Module R'' M] variable (M N) @@ -478,6 +478,8 @@ theorem exists_eq_tmul_of_forall (x : TensorProduct R M N) end Module +variable [Module R P] + section UMP variable {M N} @@ -1046,6 +1048,8 @@ end TensorProduct open scoped TensorProduct +variable [Module R P] + namespace LinearMap variable {N} diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basis.lean b/Mathlib/LinearAlgebra/TensorProduct/Basis.lean index ff1d35fb30692..7b3a50a9cf6ea 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basis.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basis.lean @@ -49,7 +49,7 @@ theorem Basis.tensorProduct_repr_tmul_apply (b : Basis ι S M) (c : Basis κ R N (Basis.tensorProduct b c).repr (m ⊗ₜ n) (i, j) = c.repr n j • b.repr m i := by simp [Basis.tensorProduct, mul_comm] -variable (S) +variable (S : Type*) [Semiring S] [Algebra R S] /-- The lift of an `R`-basis of `M` to an `S`-basis of the base change `S ⊗[R] M`. -/ noncomputable diff --git a/Mathlib/LinearAlgebra/TensorProduct/Tower.lean b/Mathlib/LinearAlgebra/TensorProduct/Tower.lean index b0028cf05f079..da36cbf10939d 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Tower.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Tower.lean @@ -59,11 +59,11 @@ open Algebra (lsmul) section Semiring variable [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] -variable [AddCommMonoid M] [Module R M] [Module A M] [Module B M] -variable [IsScalarTower R A M] [IsScalarTower R B M] [SMulCommClass A B M] +variable [AddCommMonoid M] [Module R M] [Module A M] +variable [IsScalarTower R A M] variable [AddCommMonoid N] [Module R N] -variable [AddCommMonoid P] [Module R P] [Module A P] [Module B P] -variable [IsScalarTower R A P] [IsScalarTower R B P] [SMulCommClass A B P] +variable [AddCommMonoid P] [Module R P] [Module A P] +variable [IsScalarTower R A P] variable [AddCommMonoid Q] [Module R Q] variable [AddCommMonoid P'] [Module R P'] [Module A P'] [Module B P'] variable [IsScalarTower R A P'] [IsScalarTower R B P'] [SMulCommClass A B P'] @@ -129,6 +129,9 @@ theorem lift_tmul (f : M →ₗ[A] N →ₗ[R] P) (x : M) (y : N) : lift f (x variable (R A B M N P Q) +section +variable [Module B P] [IsScalarTower R B P] [SMulCommClass A B P] + /-- Heterobasic version of `TensorProduct.uncurry`: Linearly constructing a linear map `M ⊗[R] N →[A] P` given a bilinear map `M →[A] N →[R] P` @@ -292,6 +295,8 @@ variable {M} in @[simp] theorem rid_symm_apply (m : M) : (AlgebraTensorModule.rid R A M).symm m = m ⊗ₜ 1 := rfl +end + end Semiring section CommSemiring @@ -300,14 +305,14 @@ variable [CommSemiring R] [CommSemiring A] [Semiring B] [Algebra R A] [Algebra R variable [AddCommMonoid M] [Module R M] [Module A M] [Module B M] variable [IsScalarTower R A M] [IsScalarTower R B M] [SMulCommClass A B M] variable [AddCommMonoid N] [Module R N] -variable [AddCommMonoid P] [Module R P] [Module A P] [Module B P] -variable [IsScalarTower R A P] [IsScalarTower R B P] [SMulCommClass A B P] +variable [AddCommMonoid P] [Module A P] variable [AddCommMonoid Q] [Module R Q] variable (R A B M N P Q) attribute [local ext high] TensorProduct.ext section assoc +variable [Module R P] [IsScalarTower R A P] variable [Algebra A B] [IsScalarTower A B M] /-- Heterobasic version of `TensorProduct.assoc`: @@ -362,6 +367,7 @@ theorem cancelBaseChange_symm_tmul (m : M) (n : N) : end cancelBaseChange section leftComm +variable [Module R P] [IsScalarTower R A P] /-- Heterobasic version of `TensorProduct.leftComm` -/ def leftComm : M ⊗[A] (P ⊗[R] Q) ≃ₗ[A] P ⊗[A] (M ⊗[R] Q) := @@ -416,6 +422,7 @@ theorem rightComm_symm_tmul (m : M) (p : P) (q : Q) : end rightComm section tensorTensorTensorComm +variable [Module R P] [IsScalarTower R A P] /-- Heterobasic version of `tensorTensorTensorComm`. -/ def tensorTensorTensorComm : diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index c86cefab9d8c1..9113b004f434c 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -208,7 +208,7 @@ theorem equivToUnitHom_symm_coe (f : Rˣ →* R'ˣ) (a : Rˣ) : equivToUnitHom.s ofUnitHom_coe f a @[simp] -lemma coe_toMonoidHom [CommMonoid R] (χ : MulChar R R') +lemma coe_toMonoidHom (χ : MulChar R R') (x : R) : χ.toMonoidHom x = χ x := rfl /-! diff --git a/Mathlib/RingTheory/HahnSeries/Multiplication.lean b/Mathlib/RingTheory/HahnSeries/Multiplication.lean index 1757fb871037f..5436740e4da6f 100644 --- a/Mathlib/RingTheory/HahnSeries/Multiplication.lean +++ b/Mathlib/RingTheory/HahnSeries/Multiplication.lean @@ -156,7 +156,7 @@ end SMul section SMulZeroClass -variable [PartialOrder Γ] [PartialOrder Γ'] [VAdd Γ Γ'] [IsOrderedCancelVAdd Γ Γ'] [Zero R] +variable [PartialOrder Γ] [PartialOrder Γ'] [VAdd Γ Γ'] [IsOrderedCancelVAdd Γ Γ'] [AddCommMonoid V] instance instBaseSMulZeroClass [SMulZeroClass R V] : @@ -165,9 +165,11 @@ instance instBaseSMulZeroClass [SMulZeroClass R V] : @[simp] theorem of_smul [SMulZeroClass R V] (r : R) (x : HahnSeries Γ V) : (of R) (r • x) = r • (of R) x := rfl -@[simp] theorem of_symm_smul [PartialOrder Γ] [SMulZeroClass R V] (r : R) (x : HahnModule Γ R V) : +@[simp] theorem of_symm_smul [SMulZeroClass R V] (r : R) (x : HahnModule Γ R V) : (of R).symm (r • x) = r • (of R).symm x := rfl +variable [Zero R] + instance instSMulZeroClass [SMulZeroClass R V] : SMulZeroClass (HahnSeries Γ R) (HahnModule Γ' R V) where smul_zero x := by diff --git a/Mathlib/RingTheory/Valuation/Integers.lean b/Mathlib/RingTheory/Valuation/Integers.lean index e727743126fed..d14885dd76c84 100644 --- a/Mathlib/RingTheory/Valuation/Integers.lean +++ b/Mathlib/RingTheory/Valuation/Integers.lean @@ -60,7 +60,7 @@ theorem integer.integers : v.Integers v.integer := namespace Integers -variable {v O} [CommRing O] [Algebra O R] +variable {v O} theorem one_of_isUnit' {x : O} (hx : IsUnit x) (H : ∀ x, v (algebraMap O R x) ≤ 1) : v (algebraMap O R x) = 1 := diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index 61cae794ebb3b..e8f56fd2a5354 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -148,7 +148,7 @@ lemma IsCompact.le_nhds_of_unique_clusterPt (hs : IsCompact s) {l : Filter X} {y /-- If values of `f : Y → X` belong to a compact set `s` eventually along a filter `l` and `y` is a unique `MapClusterPt` for `f` along `l` in `s`, then `f` tends to `𝓝 y` along `l`. -/ -lemma IsCompact.tendsto_nhds_of_unique_mapClusterPt {l : Filter Y} {y : X} {f : Y → X} +lemma IsCompact.tendsto_nhds_of_unique_mapClusterPt {Y} {l : Filter Y} {y : X} {f : Y → X} (hs : IsCompact s) (hmem : ∀ᶠ x in l, f x ∈ s) (h : ∀ x ∈ s, MapClusterPt x l f → x = y) : Tendsto f l (𝓝 y) := hs.le_nhds_of_unique_clusterPt (mem_map.2 hmem) h @@ -365,7 +365,7 @@ theorem isCompact_iff_finite_subfamily_closed : /-- If `s : Set (X × Y)` belongs to `𝓝 x ×ˢ l` for all `x` from a compact set `K`, then it belongs to `(𝓝ˢ K) ×ˢ l`, i.e., there exist an open `U ⊇ K` and `t ∈ l` such that `U ×ˢ t ⊆ s`. -/ -theorem IsCompact.mem_nhdsSet_prod_of_forall {K : Set X} {l : Filter Y} {s : Set (X × Y)} +theorem IsCompact.mem_nhdsSet_prod_of_forall {K : Set X} {Y} {l : Filter Y} {s : Set (X × Y)} (hK : IsCompact K) (hs : ∀ x ∈ K, s ∈ 𝓝 x ×ˢ l) : s ∈ (𝓝ˢ K) ×ˢ l := by refine hK.induction_on (by simp) (fun t t' ht hs ↦ ?_) (fun t t' ht ht' ↦ ?_) fun x hx ↦ ?_ · exact prod_mono (nhdsSet_mono ht) le_rfl hs @@ -375,19 +375,19 @@ theorem IsCompact.mem_nhdsSet_prod_of_forall {K : Set X} {l : Filter Y} {s : Set refine ⟨u, nhdsWithin_le_nhds (huo.mem_nhds hx), mem_of_superset ?_ hs⟩ exact prod_mem_prod (huo.mem_nhdsSet.2 Subset.rfl) hv -theorem IsCompact.nhdsSet_prod_eq_biSup {K : Set X} (hK : IsCompact K) (l : Filter Y) : +theorem IsCompact.nhdsSet_prod_eq_biSup {K : Set X} (hK : IsCompact K) {Y} (l : Filter Y) : (𝓝ˢ K) ×ˢ l = ⨆ x ∈ K, 𝓝 x ×ˢ l := le_antisymm (fun s hs ↦ hK.mem_nhdsSet_prod_of_forall <| by simpa using hs) (iSup₂_le fun x hx ↦ prod_mono (nhds_le_nhdsSet hx) le_rfl) -theorem IsCompact.prod_nhdsSet_eq_biSup {K : Set Y} (hK : IsCompact K) (l : Filter X) : +theorem IsCompact.prod_nhdsSet_eq_biSup {K : Set Y} (hK : IsCompact K) {X} (l : Filter X) : l ×ˢ (𝓝ˢ K) = ⨆ y ∈ K, l ×ˢ 𝓝 y := by simp only [prod_comm (f := l), hK.nhdsSet_prod_eq_biSup, map_iSup] /-- If `s : Set (X × Y)` belongs to `l ×ˢ 𝓝 y` for all `y` from a compact set `K`, then it belongs to `l ×ˢ (𝓝ˢ K)`, i.e., there exist `t ∈ l` and an open `U ⊇ K` such that `t ×ˢ U ⊆ s`. -/ -theorem IsCompact.mem_prod_nhdsSet_of_forall {K : Set Y} {l : Filter X} {s : Set (X × Y)} +theorem IsCompact.mem_prod_nhdsSet_of_forall {K : Set Y} {X} {l : Filter X} {s : Set (X × Y)} (hK : IsCompact K) (hs : ∀ y ∈ K, s ∈ l ×ˢ 𝓝 y) : s ∈ l ×ˢ 𝓝ˢ K := (hK.prod_nhdsSet_eq_biSup l).symm ▸ by simpa using hs @@ -763,7 +763,7 @@ lemma le_nhds_of_unique_clusterPt [CompactSpace X] {l : Filter X} {y : X} /-- If `y` is a unique `MapClusterPt` for `f` along `l` and the codomain of `f` is a compact space, then `f` tends to `𝓝 y` along `l`. -/ -lemma tendsto_nhds_of_unique_mapClusterPt [CompactSpace X] {l : Filter Y} {y : X} {f : Y → X} +lemma tendsto_nhds_of_unique_mapClusterPt [CompactSpace X] {Y} {l : Filter Y} {y : X} {f : Y → X} (h : ∀ x, MapClusterPt x l f → x = y) : Tendsto f l (𝓝 y) := le_nhds_of_unique_clusterPt h diff --git a/Mathlib/Topology/Connected/TotallyDisconnected.lean b/Mathlib/Topology/Connected/TotallyDisconnected.lean index b01fb5f23d02b..0da1004c037f4 100644 --- a/Mathlib/Topology/Connected/TotallyDisconnected.lean +++ b/Mathlib/Topology/Connected/TotallyDisconnected.lean @@ -295,8 +295,8 @@ theorem PreconnectedSpace.constant {Y : Type*} [TopologicalSpace Y] [DiscreteTop /-- Refinement of `IsPreconnected.constant` only assuming the map factors through a discrete subset of the target. -/ theorem IsPreconnected.constant_of_mapsTo {S : Set α} (hS : IsPreconnected S) - {T : Set β} [DiscreteTopology T] {f : α → β} (hc : ContinuousOn f S) (hTm : MapsTo f S T) - {x y : α} (hx : x ∈ S) (hy : y ∈ S) : f x = f y := by + {β} [TopologicalSpace β] {T : Set β} [DiscreteTopology T] {f : α → β} (hc : ContinuousOn f S) + (hTm : MapsTo f S T) {x y : α} (hx : x ∈ S) (hy : y ∈ S) : f x = f y := by let F : S → T := hTm.restrict f S T suffices F ⟨x, hx⟩ = F ⟨y, hy⟩ by rwa [← Subtype.coe_inj] at this exact (isPreconnected_iff_preconnectedSpace.mp hS).constant (hc.restrict_mapsTo _) @@ -304,8 +304,8 @@ theorem IsPreconnected.constant_of_mapsTo {S : Set α} (hS : IsPreconnected S) /-- A version of `IsPreconnected.constant_of_mapsTo` that assumes that the codomain is nonempty and proves that `f` is equal to `const α y` on `S` for some `y ∈ T`. -/ theorem IsPreconnected.eqOn_const_of_mapsTo {S : Set α} (hS : IsPreconnected S) - {T : Set β} [DiscreteTopology T] {f : α → β} (hc : ContinuousOn f S) (hTm : MapsTo f S T) - (hne : T.Nonempty) : ∃ y ∈ T, EqOn f (const α y) S := by + {β} [TopologicalSpace β] {T : Set β} [DiscreteTopology T] {f : α → β} (hc : ContinuousOn f S) + (hTm : MapsTo f S T) (hne : T.Nonempty) : ∃ y ∈ T, EqOn f (const α y) S := by rcases S.eq_empty_or_nonempty with (rfl | ⟨x, hx⟩) · exact hne.imp fun _ hy => ⟨hy, eqOn_empty _ _⟩ · exact ⟨f x, hTm hx, fun x' hx' => hS.constant_of_mapsTo hc hTm hx' hx⟩ diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index e1d90fe1c0361..b6754581c6062 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.lean @@ -40,7 +40,7 @@ structure DenseInducing [TopologicalSpace α] [TopologicalSpace β] (i : α → namespace DenseInducing variable [TopologicalSpace α] [TopologicalSpace β] -variable {i : α → β} (di : DenseInducing i) +variable {i : α → β} theorem nhds_eq_comap (di : DenseInducing i) : ∀ a : α, 𝓝 a = comap i (𝓝 <| i a) := di.toInducing.nhds_eq_comap @@ -48,7 +48,7 @@ theorem nhds_eq_comap (di : DenseInducing i) : ∀ a : α, 𝓝 a = comap i ( protected theorem continuous (di : DenseInducing i) : Continuous i := di.toInducing.continuous -theorem closure_range : closure (range i) = univ := +theorem closure_range (di : DenseInducing i) : closure (range i) = univ := di.dense.closure_range protected theorem preconnectedSpace [PreconnectedSpace α] (di : DenseInducing i) : @@ -90,7 +90,7 @@ protected theorem prod [TopologicalSpace γ] [TopologicalSpace δ] {e₁ : α open TopologicalSpace /-- If the domain of a `DenseInducing` map is a separable space, then so is the codomain. -/ -protected theorem separableSpace [SeparableSpace α] : SeparableSpace β := +protected theorem separableSpace [SeparableSpace α] (di : DenseInducing i) : SeparableSpace β := di.dense.separableSpace di.continuous variable [TopologicalSpace δ] {f : γ → α} {g : γ → δ} {h : δ → β} @@ -127,20 +127,21 @@ variable [TopologicalSpace γ] def extend (di : DenseInducing i) (f : α → γ) (b : β) : γ := @limUnder _ _ _ ⟨f (di.dense.some b)⟩ (comap i (𝓝 b)) f -theorem extend_eq_of_tendsto [T2Space γ] {b : β} {c : γ} {f : α → γ} +theorem extend_eq_of_tendsto [T2Space γ] (di : DenseInducing i) {b : β} {c : γ} {f : α → γ} (hf : Tendsto f (comap i (𝓝 b)) (𝓝 c)) : di.extend f b = c := haveI := di.comap_nhds_neBot hf.limUnder_eq -theorem extend_eq_at [T2Space γ] {f : α → γ} {a : α} (hf : ContinuousAt f a) : - di.extend f (i a) = f a := +theorem extend_eq_at [T2Space γ] (di : DenseInducing i) {f : α → γ} {a : α} + (hf : ContinuousAt f a) : di.extend f (i a) = f a := extend_eq_of_tendsto _ <| di.nhds_eq_comap a ▸ hf -theorem extend_eq_at' [T2Space γ] {f : α → γ} {a : α} (c : γ) (hf : Tendsto f (𝓝 a) (𝓝 c)) : - di.extend f (i a) = f a := +theorem extend_eq_at' [T2Space γ] (di : DenseInducing i) {f : α → γ} {a : α} (c : γ) + (hf : Tendsto f (𝓝 a) (𝓝 c)) : di.extend f (i a) = f a := di.extend_eq_at (continuousAt_of_tendsto_nhds hf) -theorem extend_eq [T2Space γ] {f : α → γ} (hf : Continuous f) (a : α) : di.extend f (i a) = f a := +theorem extend_eq [T2Space γ] (di : DenseInducing i) {f : α → γ} (hf : Continuous f) (a : α) : + di.extend f (i a) = f a := di.extend_eq_at hf.continuousAt /-- Variation of `extend_eq` where we ask that `f` has a limit along `comap i (𝓝 b)` for each @@ -216,17 +217,17 @@ namespace DenseEmbedding open TopologicalSpace variable [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] [TopologicalSpace δ] -variable {e : α → β} (de : DenseEmbedding e) +variable {e : α → β} -theorem inj_iff {x y} : e x = e y ↔ x = y := +theorem inj_iff (de : DenseEmbedding e) {x y} : e x = e y ↔ x = y := de.inj.eq_iff -theorem to_embedding : Embedding e := +theorem to_embedding (de : DenseEmbedding e) : Embedding e := { induced := de.induced inj := de.inj } /-- If the domain of a `DenseEmbedding` is a separable space, then so is its codomain. -/ -protected theorem separableSpace [SeparableSpace α] : SeparableSpace β := +protected theorem separableSpace [SeparableSpace α] (de : DenseEmbedding e) : SeparableSpace β := de.toDenseInducing.separableSpace /-- The product of two dense embeddings is a dense embedding. -/ @@ -241,7 +242,8 @@ def subtypeEmb {α : Type*} (p : α → Prop) (e : α → β) (x : { x // p x }) { x // x ∈ closure (e '' { x | p x }) } := ⟨e x, subset_closure <| mem_image_of_mem e x.prop⟩ -protected theorem subtype (p : α → Prop) : DenseEmbedding (subtypeEmb p e) := +protected theorem subtype (de : DenseEmbedding e) (p : α → Prop) : + DenseEmbedding (subtypeEmb p e) := { dense := dense_iff_closure_eq.2 <| by ext ⟨x, hx⟩ @@ -253,7 +255,7 @@ protected theorem subtype (p : α → Prop) : DenseEmbedding (subtypeEmb p e) := simp [subtypeEmb, nhds_subtype_eq_comap, de.toInducing.nhds_eq_comap, comap_comap, (· ∘ ·)] } -theorem dense_image {s : Set α} : Dense (e '' s) ↔ Dense s := +theorem dense_image (de : DenseEmbedding e) {s : Set α} : Dense (e '' s) ↔ Dense s := de.toDenseInducing.dense_image end DenseEmbedding diff --git a/Mathlib/Topology/DiscreteSubset.lean b/Mathlib/Topology/DiscreteSubset.lean index b5a31660a9a69..4122752067f57 100644 --- a/Mathlib/Topology/DiscreteSubset.lean +++ b/Mathlib/Topology/DiscreteSubset.lean @@ -37,7 +37,7 @@ formalise this as `Filter.codiscrete`. open Set Filter Function Topology -variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {f : X → Y} +variable {X Y : Type*} [TopologicalSpace Y] {f : X → Y} section cofinite_cocompact @@ -47,6 +47,8 @@ lemma tendsto_cofinite_cocompact_iff : refine forall₂_congr (fun K _ ↦ ?_) simp only [mem_compl_iff, eventually_cofinite, not_not, preimage] +variable [TopologicalSpace X] + lemma Continuous.discrete_of_tendsto_cofinite_cocompact [T1Space X] [WeaklyLocallyCompactSpace Y] (hf' : Continuous f) (hf : Tendsto f cofinite (cocompact _)) : DiscreteTopology X := by @@ -78,6 +80,8 @@ end cofinite_cocompact section codiscrete_filter +variable [TopologicalSpace X] + /-- Criterion for a subset `S ⊆ X` to be closed and discrete in terms of the punctured neighbourhood filter at an arbitrary point of `X`. (Compare `discreteTopology_subtype_iff`.) -/ theorem isClosed_and_discrete_iff {S : Set X} : diff --git a/Mathlib/Topology/Order/ProjIcc.lean b/Mathlib/Topology/Order/ProjIcc.lean index 36aa2f133e03a..993c017bb989f 100644 --- a/Mathlib/Topology/Order/ProjIcc.lean +++ b/Mathlib/Topology/Order/ProjIcc.lean @@ -16,14 +16,14 @@ to show that `Set.IccExtend h f` is continuous if and only if `f` is continuous. open Set Filter Topology -variable {α β γ : Type*} [LinearOrder α] [TopologicalSpace γ] {a b c : α} {h : a ≤ b} +variable {α β γ : Type*} [LinearOrder α] {a b c : α} {h : a ≤ b} protected theorem Filter.Tendsto.IccExtend (f : γ → Icc a b → β) {la : Filter α} {lb : Filter β} {lc : Filter γ} (hf : Tendsto (↿f) (lc ×ˢ la.map (projIcc a b h)) lb) : Tendsto (↿(IccExtend h ∘ f)) (lc ×ˢ la) lb := hf.comp <| tendsto_id.prod_map tendsto_map -variable [TopologicalSpace α] [OrderTopology α] [TopologicalSpace β] +variable [TopologicalSpace α] [OrderTopology α] [TopologicalSpace β] [TopologicalSpace γ] @[continuity] theorem continuous_projIcc : Continuous (projIcc a b h) := diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index 42a2218df3fec..af4b880e2169a 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -653,7 +653,7 @@ instance (priority := 100) [T1Space X] : R0Space X where instance : T1Space (CofiniteTopology X) := t1Space_iff_continuous_cofinite_of.mpr continuous_id -theorem t1Space_antitone : Antitone (@T1Space X) := fun a _ h _ => +theorem t1Space_antitone {X} : Antitone (@T1Space X) := fun a _ h _ => @T1Space.mk _ a fun x => (T1Space.t1 x).mono h theorem continuousWithinAt_update_of_ne [T1Space X] [DecidableEq X] [TopologicalSpace Y] {f : X → Y} @@ -827,7 +827,7 @@ theorem eq_of_tendsto_nhds [TopologicalSpace Y] [T1Space Y] {f : X → Y} {x : X have fact₂ : Tendsto f (pure x) (𝓝 y) := h.comp (tendsto_id'.2 <| pure_le_nhds x) fact₂ fact₁ (Eq.refl <| f x) -theorem Filter.Tendsto.eventually_ne [TopologicalSpace Y] [T1Space Y] {g : X → Y} +theorem Filter.Tendsto.eventually_ne {X} [TopologicalSpace Y] [T1Space Y] {g : X → Y} {l : Filter X} {b₁ b₂ : Y} (hg : Tendsto g l (𝓝 b₁)) (hb : b₁ ≠ b₂) : ∀ᶠ z in l, g z ≠ b₂ := hg.eventually (isOpen_compl_singleton.eventually_mem hb) diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index 6124e15d59b41..2d988c29214b8 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -30,7 +30,7 @@ open Set Function noncomputable section -variable {ι X : Type*} [TopologicalSpace X] [NormalSpace X] +variable {ι X : Type*} [TopologicalSpace X] namespace ShrinkingLemma @@ -154,7 +154,8 @@ theorem le_chainSup {c : Set (PartialRefinement u s)} (hc : IsChain (· ≤ ·) /-- If `s` is a closed set, `v` is a partial refinement, and `i` is an index such that `i ∉ v.carrier`, then there exists a partial refinement that is strictly greater than `v`. -/ -theorem exists_gt (v : PartialRefinement u s) (hs : IsClosed s) (i : ι) (hi : i ∉ v.carrier) : +theorem exists_gt [NormalSpace X] (v : PartialRefinement u s) (hs : IsClosed s) (i : ι) + (hi : i ∉ v.carrier) : ∃ v' : PartialRefinement u s, v < v' := by have I : (s ∩ ⋂ (j) (_ : j ≠ i), (v j)ᶜ) ⊆ v i := by simp only [subset_def, mem_inter_iff, mem_iInter, and_imp] @@ -193,7 +194,7 @@ end ShrinkingLemma open ShrinkingLemma -variable {u : ι → Set X} {s : Set 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/Separation.lean b/Mathlib/Topology/UniformSpace/Separation.lean index a0c604d0f8d4c..4513d502f08bc 100644 --- a/Mathlib/Topology/UniformSpace/Separation.lean +++ b/Mathlib/Topology/UniformSpace/Separation.lean @@ -165,8 +165,8 @@ theorem eq_of_forall_symmetric {α : Type*} [UniformSpace α] [T0Space α] {x y theorem eq_of_clusterPt_uniformity [T0Space α] {x y : α} (h : ClusterPt (x, y) (𝓤 α)) : x = y := (inseparable_iff_clusterPt_uniformity.2 h).eq -theorem Filter.Tendsto.inseparable_iff_uniformity {l : Filter β} [NeBot l] {f g : β → α} {a b : α} - (ha : Tendsto f l (𝓝 a)) (hb : Tendsto g l (𝓝 b)) : +theorem Filter.Tendsto.inseparable_iff_uniformity {β} {l : Filter β} [NeBot l] {f g : β → α} + {a b : α} (ha : Tendsto f l (𝓝 a)) (hb : Tendsto g l (𝓝 b)) : Inseparable a b ↔ Tendsto (fun x ↦ (f x, g x)) l (𝓤 α) := by refine ⟨fun h ↦ (ha.prod_mk_nhds hb).mono_right h.nhds_le_uniformity, fun h ↦ ?_⟩ rw [inseparable_iff_clusterPt_uniformity] From 94a546478327fc441cc4f0a7f07a6046abd79cc9 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 5 Aug 2024 05:22:30 +0000 Subject: [PATCH 028/359] chore: fix comments in Mathlib/Analysis/Analytic/Composition.lean (#15498) --- Mathlib/Analysis/Analytic/Composition.lean | 82 +++++++++++----------- 1 file changed, 41 insertions(+), 41 deletions(-) diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index baf9f2551f7e6..3026497f16084 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -23,7 +23,7 @@ Then `g ∘ f` is obtained by summing all these multilinear functions. To formalize this, we use compositions of an integer `N`, i.e., its decompositions into a sum `i₁ + ... + iₙ` of positive integers. Given such a composition `c` and two formal -multilinear series `q` and `p`, let `q.comp_along_composition p c` be the above multilinear +multilinear series `q` and `p`, let `q.compAlongComposition p c` be the above multilinear function. Then the `N`-th coefficient in the power series expansion of `g ∘ f` is the sum of these terms over all `c : Composition N`. @@ -46,7 +46,7 @@ summability of the norms, this implies the overall convergence. ## Implementation details The main technical difficulty is to write down things. In particular, we need to define precisely -`q.comp_along_composition p c` and to show that it is indeed a continuous multilinear +`q.compAlongComposition p c` and to show that it is indeed a continuous multilinear function. This requires a whole interface built on the class `Composition`. Once this is set, the main difficulty is to reorder the sums, writing the composition of the partial sums as a sum over some subset of `Σ n, Composition n`. We need to check that the reordering is a bijection, @@ -59,7 +59,7 @@ the formal multilinear series representing a function (and also, it holds even w convergence of the series is `0`). Instead, we give a direct proof, which amounts to reordering double sums in a careful way. The change of variables is a canonical (combinatorial) bijection `Composition.sigmaEquivSigmaPi` between `(Σ (a : Composition n), Composition a.length)` and -`(Σ (c : Composition n), Π (i : Fin c.length), Composition (c.blocks_fun i))`, and is described +`(Σ (c : Composition n), Π (i : Fin c.length), Composition (c.blocksFun i))`, and is described in more details below in the paragraph on associativity. -/ @@ -173,7 +173,7 @@ variable [TopologicalAddGroup F] [ContinuousConstSMul 𝕜 F] /-- Given a formal multilinear series `p`, a composition `c` of `n` and a continuous multilinear map `f` in `c.length` variables, one may form a continuous multilinear map in `n` variables by applying the right coefficient of `p` to each block of the composition, and then applying `f` to -the resulting vector. It is called `f.comp_along_composition p c`. -/ +the resulting vector. It is called `f.compAlongComposition p c`. -/ def compAlongComposition {n : ℕ} (p : FormalMultilinearSeries 𝕜 E F) (c : Composition n) (f : ContinuousMultilinearMap 𝕜 (fun _i : Fin c.length => F) G) : ContinuousMultilinearMap 𝕜 (fun _i : Fin n => E) G where @@ -205,7 +205,7 @@ variable [TopologicalAddGroup G] [ContinuousConstSMul 𝕜 G] /-- Given two formal multilinear series `q` and `p` and a composition `c` of `n`, one may form a continuous multilinear map in `n` variables by applying the right coefficient of `p` to each block of the composition, and then applying `q c.length` to the resulting vector. It is -called `q.comp_along_composition p c`. -/ +called `q.compAlongComposition p c`. -/ def compAlongComposition {n : ℕ} (q : FormalMultilinearSeries 𝕜 F G) (p : FormalMultilinearSeries 𝕜 E F) (c : Composition n) : ContinuousMultilinearMap 𝕜 (fun _i : Fin n => E) G := @@ -218,7 +218,7 @@ theorem compAlongComposition_apply {n : ℕ} (q : FormalMultilinearSeries 𝕜 F rfl /-- Formal composition of two formal multilinear series. The `n`-th coefficient in the composition -is defined to be the sum of `q.comp_along_composition p c` over all compositions of +is defined to be the sum of `q.compAlongComposition p c` over all compositions of `n`. In other words, this term (as a multilinear function applied to `v_0, ..., v_{n-1}`) is `∑'_{k} ∑'_{i₁ + ... + iₖ = n} qₖ (p_{i_1} (...), ..., p_{i_k} (...))`, where one puts all variables `v_0, ..., v_{n-1}` in increasing order in the dots. @@ -287,7 +287,7 @@ variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 namespace FormalMultilinearSeries -/-- The norm of `f.comp_along_composition p c` is controlled by the product of +/-- The norm of `f.compAlongComposition p c` is controlled by the product of the norms of the relevant bits of `f` and `p`. -/ theorem compAlongComposition_bound {n : ℕ} (p : FormalMultilinearSeries 𝕜 E F) (c : Composition n) (f : ContinuousMultilinearMap 𝕜 (fun _i : Fin c.length => F) G) (v : Fin n → E) : @@ -306,7 +306,7 @@ theorem compAlongComposition_bound {n : ℕ} (p : FormalMultilinearSeries 𝕜 E rw [← c.blocksFinEquiv.prod_comp, ← Finset.univ_sigma_univ, Finset.prod_sigma] congr -/-- The norm of `q.comp_along_composition p c` is controlled by the product of +/-- The norm of `q.compAlongComposition p c` is controlled by the product of the norms of the relevant bits of `q` and `p`. -/ theorem compAlongComposition_norm {n : ℕ} (q : FormalMultilinearSeries 𝕜 F G) (p : FormalMultilinearSeries 𝕜 E F) (c : Composition n) : @@ -432,7 +432,7 @@ theorem comp_summable_nnreal (q : FormalMultilinearSeries 𝕜 F G) (p : FormalM ∃ r > (0 : ℝ≥0), Summable fun i : Σ n, Composition n => ‖q.compAlongComposition p i.2‖₊ * r ^ i.1 := by /- This follows from the fact that the growth rate of `‖qₙ‖` and `‖pₙ‖` is at most geometric, - giving a geometric bound on each `‖q.comp_along_composition p op‖`, together with the + giving a geometric bound on each `‖q.compAlongComposition p op‖`, together with the fact that there are `2^(n-1)` compositions of `n`, giving at most a geometric loss. -/ rcases ENNReal.lt_iff_exists_nnreal_btwn.1 (lt_min zero_lt_one hq) with ⟨rq, rq_pos, hrq⟩ rcases ENNReal.lt_iff_exists_nnreal_btwn.1 (lt_min zero_lt_one hp) with ⟨rp, rp_pos, hrp⟩ @@ -507,19 +507,19 @@ theorem le_comp_radius_of_summable (q : FormalMultilinearSeries 𝕜 F G) ### Composing analytic functions Now, we will prove that the composition of the partial sums of `q` and `p` up to order `N` is -given by a sum over some large subset of `Σ n, Composition n` of `q.comp_along_composition p`, to +given by a sum over some large subset of `Σ n, Composition n` of `q.compAlongComposition p`, to deduce that the series for `q.comp p` indeed converges to `g ∘ f` when `q` is a power series for `g` and `p` is a power series for `f`. This proof is a big reindexing argument of a sum. Since it is a bit involved, we define first -the source of the change of variables (`comp_partial_source`), its target -(`comp_partial_target`) and the change of variables itself (`comp_change_of_variables`) before -giving the main statement in `comp_partial_sum`. -/ +the source of the change of variables (`compPartialSumSource`), its target +(`compPartialSumTarget`) and the change of variables itself (`compChangeOfVariables`) before +giving the main statement in `comp_partialSum`. -/ /-- Source set in the change of variables to compute the composition of partial sums of formal power series. -See also `comp_partial_sum`. -/ +See also `comp_partialSum`. -/ def compPartialSumSource (m M N : ℕ) : Finset (Σ n, Fin n → ℕ) := Finset.sigma (Finset.Ico m M) (fun n : ℕ => Fintype.piFinset fun _i : Fin n => Finset.Ico 1 N : _) @@ -578,7 +578,7 @@ theorem compPartialSumTargetSet_image_compPartialSumSource (m M N : ℕ) /-- Target set in the change of variables to compute the composition of partial sums of formal power series, here given a a finset. -See also `comp_partial_sum`. -/ +See also `comp_partialSum`. -/ def compPartialSumTarget (m M N : ℕ) : Finset (Σ n, Composition n) := Set.Finite.toFinset <| ((Finset.finite_toSet _).dependent_image _).subset <| @@ -590,9 +590,9 @@ theorem mem_compPartialSumTarget_iff {m M N : ℕ} {a : Σ n, Composition n} : m ≤ a.2.length ∧ a.2.length < M ∧ ∀ j : Fin a.2.length, a.2.blocksFun j < N := by simp [compPartialSumTarget, compPartialSumTargetSet] -/-- `comp_change_of_variables m M N` is a bijection between `comp_partial_sum_source m M N` -and `comp_partial_sum_target m M N`, yielding equal sums for functions that correspond to each -other under the bijection. As `comp_change_of_variables m M N` is a dependent function, stating +/-- `compChangeOfVariables m M N` is a bijection between `compPartialSumSource m M N` +and `compPartialSumTarget m M N`, yielding equal sums for functions that correspond to each +other under the bijection. As `compChangeOfVariables m M N` is a dependent function, stating that it is a bijection is not directly possible, but the consequence on sums can be stated more easily. -/ theorem compChangeOfVariables_sum {α : Type*} [AddCommMonoid α] (m M N : ℕ) @@ -600,9 +600,9 @@ theorem compChangeOfVariables_sum {α : Type*} [AddCommMonoid α] (m M N : ℕ) (h : ∀ (e) (he : e ∈ compPartialSumSource m M N), f e = g (compChangeOfVariables m M N e he)) : ∑ e ∈ compPartialSumSource m M N, f e = ∑ e ∈ compPartialSumTarget m M N, g e := by apply Finset.sum_bij (compChangeOfVariables m M N) - -- We should show that the correspondance we have set up is indeed a bijection + -- We should show that the correspondence we have set up is indeed a bijection -- between the index sets of the two sums. - -- 1 - show that the image belongs to `comp_partial_sum_target m N N` + -- 1 - show that the image belongs to `compPartialSumTarget m N N` · rintro ⟨k, blocks_fun⟩ H rw [mem_compPartialSumSource_iff] at H -- Porting note: added @@ -629,7 +629,7 @@ theorem compChangeOfVariables_sum {α : Type*} [AddCommMonoid α] (m M N : ℕ) · intro i hi apply compPartialSumTargetSet_image_compPartialSumSource m M N i simpa [compPartialSumTarget] using hi - -- 4 - show that the composition gives the `comp_along_composition` application + -- 4 - show that the composition gives the `compAlongComposition` application · rintro ⟨k, blocks_fun⟩ H rw [h] @@ -652,8 +652,8 @@ theorem compPartialSumTarget_tendsto_atTop : simp only [Finset.mem_image_of_mem, Finset.mem_coe, Finset.mem_univ] /-- Composing the partial sums of two multilinear series coincides with the sum over all -compositions in `comp_partial_sum_target 0 N N`. This is precisely the motivation for the -definition of `comp_partial_sum_target`. -/ +compositions in `compPartialSumTarget 0 N N`. This is precisely the motivation for the +definition of `compPartialSumTarget`. -/ theorem comp_partialSum (q : FormalMultilinearSeries 𝕜 F G) (p : FormalMultilinearSeries 𝕜 E F) (N : ℕ) (z : E) : q.partialSum N (∑ i ∈ Finset.Ico 1 N, p i fun _j => z) = @@ -666,9 +666,9 @@ theorem comp_partialSum (q : FormalMultilinearSeries 𝕜 F G) (p : FormalMultil ∑ i ∈ compPartialSumTarget 0 N N, q.compAlongComposition p i.2 fun _j => z by simpa only [FormalMultilinearSeries.partialSum, ContinuousMultilinearMap.map_sum_finset] using H -- rewrite the first sum as a big sum over a sigma type, in the finset - -- `comp_partial_sum_target 0 N N` + -- `compPartialSumTarget 0 N N` rw [Finset.range_eq_Ico, Finset.sum_sigma'] - -- use `comp_change_of_variables_sum`, saying that this change of variables respects sums + -- use `compChangeOfVariables_sum`, saying that this change of variables respects sums apply compChangeOfVariables_sum 0 N N rintro ⟨k, blocks_fun⟩ H apply congr _ (compChangeOfVariables_length 0 N N H).symm @@ -722,7 +722,7 @@ theorem HasFPowerSeriesAt.comp {g : F → G} {f : E → F} {q : FormalMultilinea we will write `q.comp p` applied to `y` as a big sum over all compositions. Since the sum is summable, to get its convergence it suffices to get the convergence along some increasing sequence of sets. - We will use the sequence of sets `comp_partial_sum_target 0 n n`, + We will use the sequence of sets `compPartialSumTarget 0 n n`, along which the sum is exactly the composition of the partial sums of `q` and `p`, by design. To show that it converges to `g (f (x + y))`, pointwise convergence would not be enough, but we have uniform convergence to save the day. -/ @@ -757,7 +757,7 @@ theorem HasFPowerSeriesAt.comp {g : F → G} {f : E → F} {q : FormalMultilinea rw [← EMetric.isOpen_ball.nhdsWithin_eq B₂] at A convert Hg.tendstoLocallyUniformlyOn.tendsto_comp B₁.continuousWithinAt B₂ A simp only [add_sub_cancel] - -- Third step: the sum over all compositions in `comp_partial_sum_target 0 n n` converges to + -- Third step: the sum over all compositions in `compPartialSumTarget 0 n n` converges to -- `g (f (x + y))`. As this sum is exactly the composition of the partial sum, this is a direct -- consequence of the second step have C : @@ -848,14 +848,14 @@ On the other hand, r.comp (q.comp p) n v = ∑_{c : Composition n} r c.length (applyComposition (q.comp p) c v) ``` Here, `applyComposition (q.comp p) c v` is a vector of length `c.length`, whose `i`-th term is -given by `(q.comp p) (c.blocks_fun i) (v_l, v_{l+1}, ..., v_{m-1})` where `{l, ..., m-1}` is the -`i`-th block in the composition `c`, of length `c.blocks_fun i` by definition. To compute this term, -we expand it as `∑_{dᵢ : Composition (c.blocks_fun i)} q dᵢ.length (applyComposition p dᵢ v')`, +given by `(q.comp p) (c.blocksFun i) (v_l, v_{l+1}, ..., v_{m-1})` where `{l, ..., m-1}` is the +`i`-th block in the composition `c`, of length `c.blocksFun i` by definition. To compute this term, +we expand it as `∑_{dᵢ : Composition (c.blocksFun i)} q dᵢ.length (applyComposition p dᵢ v')`, where `v' = (v_l, v_{l+1}, ..., v_{m-1})`. Therefore, we get ``` r.comp (q.comp p) n v = -∑_{c : Composition n} ∑_{d₀ : Composition (c.blocks_fun 0), - ..., d_{c.length - 1} : Composition (c.blocks_fun (c.length - 1))} +∑_{c : Composition n} ∑_{d₀ : Composition (c.blocksFun 0), + ..., d_{c.length - 1} : Composition (c.blocksFun (c.length - 1))} r c.length (λ i, q dᵢ.length (applyComposition p dᵢ v'ᵢ)) ``` To show that these terms coincide, we need to explain how to reindex the sums to put them in @@ -875,7 +875,7 @@ made of two blocks of length `4` and `9`, i.e., `c = [4, 9]`. But one can also r the new first block was initially made of two blocks of size `2`, so `d₀ = [2, 2]`, and the new second block was initially made of three blocks of size `3`, `4` and `2`, so `d₁ = [3, 4, 2]`. -This equivalence is called `Composition.sigma_equiv_sigma_pi n` below. +This equivalence is called `Composition.sigmaEquivSigmaPi n` below. We start with preliminary results on compositions, of a very specialized nature, then define the equivalence `Composition.sigmaEquivSigmaPi n`, and we deduce finally the associativity of @@ -899,7 +899,7 @@ theorem sigma_composition_eq_iff (i j : Σ a : Composition n, Composition a.leng induction H; congr; ext1; exact h' /-- Rewriting equality in the dependent type -`Σ (c : Composition n), Π (i : Fin c.length), Composition (c.blocks_fun i)` in +`Σ (c : Composition n), Π (i : Fin c.length), Composition (c.blocksFun i)` in non-dependent terms with lists, requiring that the lists of blocks coincide. -/ theorem sigma_pi_composition_eq_iff (u v : Σ c : Composition n, ∀ i : Fin c.length, Composition (c.blocksFun i)) : @@ -987,10 +987,10 @@ set_option linter.deprecated false in Consider a composition `a` of `n` and a composition `b` of `a.length`. Grouping together some blocks of `a` according to `b` as in `a.gather b`, one can compute the total size of the blocks -of `a` up to an index `size_up_to b i + j` (where the `j` corresponds to a set of blocks of `a` +of `a` up to an index `sizeUpTo b i + j` (where the `j` corresponds to a set of blocks of `a` that do not fill a whole block of `a.gather b`). The first part corresponds to a sum of blocks in `a.gather b`, and the second one to a sum of blocks in the next block of -`sigma_composition_aux a b`. This is the content of this lemma. -/ +`sigmaCompositionAux a b`. This is the content of this lemma. -/ theorem sizeUpTo_sizeUpTo_add (a : Composition n) (b : Composition a.length) {i j : ℕ} (hi : i < b.length) (hj : j < blocksFun b ⟨i, hi⟩) : sizeUpTo a (sizeUpTo b i + j) = @@ -1032,7 +1032,7 @@ theorem sizeUpTo_sizeUpTo_add (a : Composition n) (b : Composition a.length) {i rw [getElem_of_eq (getElem_splitWrtComposition _ _ _ _), getElem_drop', getElem_take _ _ C] /-- Natural equivalence between `(Σ (a : Composition n), Composition a.length)` and -`(Σ (c : Composition n), Π (i : Fin c.length), Composition (c.blocks_fun i))`, that shows up as a +`(Σ (c : Composition n), Π (i : Fin c.length), Composition (c.blocksFun i))`, that shows up as a change of variables in the proof that composition of formal multilinear series is associative. Consider a composition `a` of `n` and a composition `b` of `a.length`. Then `b` indicates how to @@ -1062,7 +1062,7 @@ def sigmaEquivSigmaPi (n : ℕ) : exact Composition.length_pos_of_pos _ (Composition.blocks_pos' _ _ _) blocks_sum := by dsimp only [Composition.length]; simp [sum_ofFn] }⟩ left_inv := by - -- the fact that we have a left inverse is essentially `join_split_wrt_composition`, + -- the fact that we have a left inverse is essentially `join_splitWrtComposition`, -- but we need to massage it to take care of the dependent setting. rintro ⟨a, b⟩ rw [sigma_composition_eq_iff] @@ -1139,14 +1139,14 @@ theorem comp_assoc (r : FormalMultilinearSeries 𝕜 G H) (q : FormalMultilinear -- check that the `r` components are the same. Based on `Composition.length_gather` apply r.congr (Composition.length_gather a b).symm intro i hi1 hi2 - -- check that the `q` components are the same. Based on `length_sigma_composition_aux` + -- check that the `q` components are the same. Based on `length_sigmaCompositionAux` apply q.congr (length_sigmaCompositionAux a b _).symm intro j hj1 hj2 - -- check that the `p` components are the same. Based on `blocks_fun_sigma_composition_aux` + -- check that the `p` components are the same. Based on `blocksFun_sigmaCompositionAux` apply p.congr (blocksFun_sigmaCompositionAux a b _ _).symm intro k hk1 hk2 -- finally, check that the coordinates of `v` one is using are the same. Based on - -- `size_up_to_size_up_to_add`. + -- `sizeUpTo_sizeUpTo_add`. refine congr_arg v (Fin.ext ?_) dsimp [Composition.embedding] rw [sizeUpTo_sizeUpTo_add _ _ hi1 hj1, add_assoc] From d56e389960165aa122e7c97c098d67e4def09470 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 5 Aug 2024 05:35:00 +0000 Subject: [PATCH 029/359] =?UTF-8?q?chore(*):=20`Cdf`=20=E2=86=92=20`CDF`?= =?UTF-8?q?=20(#15496)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib.lean | 6 +++--- Mathlib/Probability/{Cdf.lean => CDF.lean} | 2 +- Mathlib/Probability/Distributions/Exponential.lean | 2 +- Mathlib/Probability/Distributions/Gamma.lean | 2 +- .../Disintegration/{CdfToKernel.lean => CDFToKernel.lean} | 0 .../Kernel/Disintegration/{CondCdf.lean => CondCDF.lean} | 4 ++-- .../Probability/Kernel/Disintegration/StandardBorel.lean | 8 ++++---- 7 files changed, 12 insertions(+), 12 deletions(-) rename Mathlib/Probability/{Cdf.lean => CDF.lean} (98%) rename Mathlib/Probability/Kernel/Disintegration/{CdfToKernel.lean => CDFToKernel.lean} (100%) rename Mathlib/Probability/Kernel/Disintegration/{CondCdf.lean => CondCDF.lean} (99%) diff --git a/Mathlib.lean b/Mathlib.lean index c18909907b036..29748850f8ccb 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3603,7 +3603,7 @@ import Mathlib.Order.WithBot import Mathlib.Order.Zorn import Mathlib.Order.ZornAtoms import Mathlib.Probability.BorelCantelli -import Mathlib.Probability.Cdf +import Mathlib.Probability.CDF import Mathlib.Probability.CondCount import Mathlib.Probability.ConditionalExpectation import Mathlib.Probability.ConditionalProbability @@ -3625,8 +3625,8 @@ import Mathlib.Probability.Kernel.Composition import Mathlib.Probability.Kernel.CondDistrib import Mathlib.Probability.Kernel.Condexp import Mathlib.Probability.Kernel.Disintegration.Basic -import Mathlib.Probability.Kernel.Disintegration.CdfToKernel -import Mathlib.Probability.Kernel.Disintegration.CondCdf +import Mathlib.Probability.Kernel.Disintegration.CDFToKernel +import Mathlib.Probability.Kernel.Disintegration.CondCDF import Mathlib.Probability.Kernel.Disintegration.Density import Mathlib.Probability.Kernel.Disintegration.Integral import Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes diff --git a/Mathlib/Probability/Cdf.lean b/Mathlib/Probability/CDF.lean similarity index 98% rename from Mathlib/Probability/Cdf.lean rename to Mathlib/Probability/CDF.lean index 6f64b28d7f57a..f17cc568c2d9e 100644 --- a/Mathlib/Probability/Cdf.lean +++ b/Mathlib/Probability/CDF.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ -import Mathlib.Probability.Kernel.Disintegration.CondCdf +import Mathlib.Probability.Kernel.Disintegration.CondCDF /-! # Cumulative distribution function of a real probability measure diff --git a/Mathlib/Probability/Distributions/Exponential.lean b/Mathlib/Probability/Distributions/Exponential.lean index 98b6a9710016b..ad5ce86b273a2 100644 --- a/Mathlib/Probability/Distributions/Exponential.lean +++ b/Mathlib/Probability/Distributions/Exponential.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Claus Clausen, Patrick Massot -/ import Mathlib.Probability.Notation -import Mathlib.Probability.Cdf +import Mathlib.Probability.CDF import Mathlib.Probability.Distributions.Gamma /-! # Exponential distributions over ℝ diff --git a/Mathlib/Probability/Distributions/Gamma.lean b/Mathlib/Probability/Distributions/Gamma.lean index 48d3df5b8482f..0548d04bfd3c0 100644 --- a/Mathlib/Probability/Distributions/Gamma.lean +++ b/Mathlib/Probability/Distributions/Gamma.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Josha Dekker -/ import Mathlib.Probability.Notation -import Mathlib.Probability.Cdf +import Mathlib.Probability.CDF import Mathlib.Analysis.SpecialFunctions.Gamma.Basic /-! # Gamma distributions over ℝ diff --git a/Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean similarity index 100% rename from Mathlib/Probability/Kernel/Disintegration/CdfToKernel.lean rename to Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean similarity index 99% rename from Mathlib/Probability/Kernel/Disintegration/CondCdf.lean rename to Mathlib/Probability/Kernel/Disintegration/CondCDF.lean index 763e9f9919b95..0659092ade23a 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCdf.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.MeasureTheory.Decomposition.RadonNikodym -import Mathlib.Probability.Kernel.Disintegration.CdfToKernel +import Mathlib.Probability.Kernel.Disintegration.CDFToKernel /-! # Conditional cumulative distribution function @@ -17,7 +17,7 @@ and limit 1 at +∞, and such that for all `x : ℝ`, `a ↦ condCDF ρ a x` is `∫⁻ a in s, ENNReal.ofReal (condCDF ρ a x) ∂ρ.fst = ρ (s ×ˢ Iic x)`. `condCDF` is build from the more general tools about kernel CDFs developed in the file -`Probability.Kernel.Disintegration.CdfToKernel`. In that file, we build a function +`Probability.Kernel.Disintegration.CDFToKernel`. In that file, we build a function `α × β → StieltjesFunction` (which is `α × β → ℝ → ℝ` with additional properties) from a function `α × β → ℚ → ℝ`. The restriction to `ℚ` allows to prove some properties like measurability more easily. Here we apply that construction to the case `β = Unit` and then drop `β` to build diff --git a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean index 0a3aec0c493b1..1659fdaec44bc 100644 --- a/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/StandardBorel.lean @@ -5,9 +5,9 @@ Authors: Rémy Degenne -/ import Mathlib.Probability.Kernel.MeasureCompProd import Mathlib.Probability.Kernel.Disintegration.Basic -import Mathlib.Probability.Kernel.Disintegration.CondCdf +import Mathlib.Probability.Kernel.Disintegration.CondCDF import Mathlib.Probability.Kernel.Disintegration.Density -import Mathlib.Probability.Kernel.Disintegration.CdfToKernel +import Mathlib.Probability.Kernel.Disintegration.CDFToKernel import Mathlib.MeasureTheory.Constructions.Polish.EmbeddingReal /-! @@ -30,14 +30,14 @@ For `κ : Kernel α (β × ℝ)`, the construction of the conditional kernel pro * Extend that function to `(α × β) → StieltjesFunction`. See the file `MeasurableStieltjes.lean`. * Finally obtain from the measurable Stieltjes function a measure on `ℝ` for each element of `α × β` in a measurable way: we have obtained a `Kernel (α × β) ℝ`. - See the file `CdfToKernel.lean` for that step. + See the file `CDFToKernel.lean` for that step. The first step (building the measurable function on `ℚ`) is done differently depending on whether `α` is countable or not. * If `α` is countable, we can provide for each `a : α` a function `f : β → ℚ → ℝ` and proceed as above to obtain a `Kernel β ℝ`. Since `α` is countable, measurability is not an issue and we can put those together into a `Kernel (α × β) ℝ`. The construction of that `f` is done in - the `CondCdf.lean` file. + the `CondCDF.lean` file. * If `α` is not countable, we can't proceed separately for each `a : α` and have to build a function `f : α × β → ℚ → ℝ` which is measurable on the product. We are able to do so if `β` has a countably generated σ-algebra (this is the case in particular for standard Borel spaces). From 8c4038b2507733e36fe8af2a3a92e662efb9fc0f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 5 Aug 2024 06:34:16 +0000 Subject: [PATCH 030/359] chore: backports for leanprover/lean4#4814 (part 2) (#15246) See #15245 --- Mathlib/Algebra/Ring/InjSurj.lean | 16 +++++++++++++--- Mathlib/Data/FunLike/Equiv.lean | 2 +- Mathlib/Data/List/Basic.lean | 10 ++++++---- Mathlib/Data/Nat/Find.lean | 3 +-- 4 files changed, 21 insertions(+), 10 deletions(-) diff --git a/Mathlib/Algebra/Ring/InjSurj.lean b/Mathlib/Algebra/Ring/InjSurj.lean index 8f46d988a67dc..659f7585b8136 100644 --- a/Mathlib/Algebra/Ring/InjSurj.lean +++ b/Mathlib/Algebra/Ring/InjSurj.lean @@ -11,12 +11,13 @@ import Mathlib.Algebra.GroupWithZero.InjSurj # Pulling back rings along injective maps, and pushing them forward along surjective maps -/ -variable {α β : Type*} [Zero β] [One β] [Add β] [Mul β] [Neg β] [Sub β] [SMul ℕ β] [SMul ℤ β] - [Pow β ℕ] [NatCast β] [IntCast β] +variable {α β : Type*} namespace Function.Injective variable (f : β → α) (hf : Injective f) +variable [Add β] [Mul β] + /-- Pullback a `LeftDistribClass` instance along an injective function. -/ theorem leftDistribClass [Mul α] [Add α] [LeftDistribClass α] (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) : LeftDistribClass β where @@ -27,6 +28,9 @@ theorem rightDistribClass [Mul α] [Add α] [RightDistribClass α] (add : ∀ x (mul : ∀ x y, f (x * y) = f x * f y) : RightDistribClass β where right_distrib x y z := hf <| by simp only [*, right_distrib] +variable [Zero β] [One β] [Neg β] [Sub β] [SMul ℕ β] [SMul ℤ β] + [Pow β ℕ] [NatCast β] [IntCast β] + /-- Pullback a `Distrib` instance along an injective function. -/ -- See note [reducible non-instances] protected abbrev distrib [Distrib α] (add : ∀ x y, f (x + y) = f x + f y) @@ -37,7 +41,8 @@ protected abbrev distrib [Distrib α] (add : ∀ x y, f (x + y) = f x + f y) /-- A type endowed with `-` and `*` has distributive negation, if it admits an injective map that preserves `-` and `*` to a type which has distributive negation. -/ -- -- See note [reducible non-instances] -protected abbrev hasDistribNeg [Mul α] [HasDistribNeg α] (neg : ∀ a, f (-a) = -f a) +protected abbrev hasDistribNeg (f : β → α) (hf : Injective f) [Mul α] [HasDistribNeg α] + (neg : ∀ a, f (-a) = -f a) (mul : ∀ a b, f (a * b) = f a * f b) : HasDistribNeg β := { hf.involutiveNeg _ neg, ‹Mul β› with neg_mul := fun x y => hf <| by erw [neg, mul, neg, neg_mul, mul], @@ -191,6 +196,8 @@ end Function.Injective namespace Function.Surjective variable (f : α → β) (hf : Surjective f) +variable [Add β] [Mul β] + /-- Pushforward a `LeftDistribClass` instance along a surjective function. -/ theorem leftDistribClass [Mul α] [Add α] [LeftDistribClass α] (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) : LeftDistribClass β where @@ -208,6 +215,9 @@ protected abbrev distrib [Distrib α] (add : ∀ x y, f (x + y) = f x + f y) __ := hf.leftDistribClass f add mul __ := hf.rightDistribClass f add mul +variable [Zero β] [One β] [Neg β] [Sub β] [SMul ℕ β] [SMul ℤ β] + [Pow β ℕ] [NatCast β] [IntCast β] + /-- A type endowed with `-` and `*` has distributive negation, if it admits a surjective map that preserves `-` and `*` from a type which has distributive negation. -/ -- See note [reducible non-instances] diff --git a/Mathlib/Data/FunLike/Equiv.lean b/Mathlib/Data/FunLike/Equiv.lean index 6841944ac6568..f99cdf26ac4fd 100644 --- a/Mathlib/Data/FunLike/Equiv.lean +++ b/Mathlib/Data/FunLike/Equiv.lean @@ -149,7 +149,7 @@ class EquivLike (E : Sort*) (α β : outParam (Sort*)) where namespace EquivLike -variable {E F α β γ : Sort*} [iE : EquivLike E α β] [iF : EquivLike F β γ] +variable {E F α β γ : Sort*} [EquivLike E α β] [EquivLike F β γ] theorem inv_injective : Function.Injective (EquivLike.inv : E → β → α) := fun e g h ↦ coe_injective' e g ((right_inv e).eq_rightInverse (h.symm ▸ left_inv g)) h diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 533b04237de27..bd2aa377378ce 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1581,17 +1581,17 @@ end FoldlEqFoldlr' section FoldlEqFoldlr' variable {f : α → β → β} -variable (hf : ∀ a b c, f a (f b c) = f b (f a c)) -theorem foldr_eq_of_comm' : ∀ a b l, foldr f a (b :: l) = foldr f (f b a) l +theorem foldr_eq_of_comm' (hf : ∀ a b c, f a (f b c) = f b (f a c)) : + ∀ a b l, foldr f a (b :: l) = foldr f (f b a) l | a, b, [] => rfl - | a, b, c :: l => by rw [foldr, foldr, foldr, hf, ← foldr_eq_of_comm' ..]; rfl + | a, b, c :: l => by rw [foldr, foldr, foldr, hf, ← foldr_eq_of_comm' hf ..]; rfl end FoldlEqFoldlr' section -variable {op : α → α → α} [ha : Std.Associative op] [hc : Std.Commutative op] +variable {op : α → α → α} [ha : Std.Associative op] /-- Notation for `op a b`. -/ local notation a " ⋆ " b => op a b @@ -1612,6 +1612,8 @@ theorem foldl_op_eq_op_foldr_assoc : | a :: l, a₁, a₂ => by simp only [foldl_cons, foldr_cons, foldl_assoc, ha.assoc]; rw [foldl_op_eq_op_foldr_assoc] +variable [hc : Std.Commutative op] + theorem foldl_assoc_comm_cons {l : List α} {a₁ a₂} : ((a₁ :: l) <*> a₂) = a₁ ⋆ l <*> a₂ := by rw [foldl_cons, hc.comm, foldl_assoc] diff --git a/Mathlib/Data/Nat/Find.lean b/Mathlib/Data/Nat/Find.lean index 6b77f83b04b27..2f06a77d4afe8 100644 --- a/Mathlib/Data/Nat/Find.lean +++ b/Mathlib/Data/Nat/Find.lean @@ -71,8 +71,6 @@ protected theorem find_min : ∀ {m : ℕ}, m < Nat.find H → ¬p m := protected theorem find_min' {m : ℕ} (h : p m) : Nat.find H ≤ m := Nat.le_of_not_lt fun l => Nat.find_min H l h -variable [DecidablePred p] [DecidablePred q] - lemma find_eq_iff (h : ∃ n : ℕ, p n) : Nat.find h = m ↔ p m ∧ ∀ n < m, ¬ p n := by constructor · rintro rfl @@ -95,6 +93,7 @@ lemma find_eq_iff (h : ∃ n : ℕ, p n) : Nat.find h = m ↔ p m ∧ ∀ n < m, @[simp] lemma find_eq_zero (h : ∃ n : ℕ, p n) : Nat.find h = 0 ↔ p 0 := by simp [find_eq_iff] +variable [DecidablePred q] in lemma find_mono (h : ∀ n, q n → p n) {hp : ∃ n, p n} {hq : ∃ n, q n} : Nat.find hp ≤ Nat.find hq := Nat.find_min' _ (h _ (Nat.find_spec hq)) From f74b424c9e20470f2173f546dd1bbec3b9cb6c60 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Mon, 5 Aug 2024 08:52:09 +0000 Subject: [PATCH 031/359] feat(ConditionallyCompleteLattice): ciSup_image, ciSup_or', and other iSup-like lemmas (#15271) Co-authored-by: Yakov Pechersky --- .../ConditionallyCompleteLattice/Basic.lean | 98 ++++++++++++++++++- 1 file changed, 97 insertions(+), 1 deletion(-) diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index 7f821542b2ab7..0bd74776170ed 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -861,6 +861,94 @@ lemma sup_eq_top_of_top_mem [OrderTop α] (h : ⊤ ∈ s) : sSup s = ⊤ := lemma inf_eq_bot_of_bot_mem [OrderBot α] (h : ⊥ ∈ s) : sInf s = ⊥ := bot_unique <| csInf_le (OrderBot.bddBelow s) h +theorem ciSup_subtype [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : Subtype p → α} + (hf : BddAbove (Set.range f)) (hf' : sSup ∅ ≤ iSup f) : + iSup f = ⨆ (i) (h : p i), f ⟨i, h⟩ := by + classical + refine le_antisymm (ciSup_le ?_) ?_ + · intro ⟨i, h⟩ + have : f ⟨i, h⟩ = (fun i : ι ↦ ⨆ (h : p i), f ⟨i, h⟩) i := by simp [h] + rw [this] + refine le_ciSup (f := (fun i : ι ↦ ⨆ (h : p i), f ⟨i, h⟩)) ?_ i + simp_rw [ciSup_eq_ite] + refine (hf.union (bddAbove_singleton (a := sSup ∅))).mono ?_ + intro + simp only [Set.mem_range, Set.union_singleton, Set.mem_insert_iff, Subtype.exists, + forall_exists_index] + intro b hb + split_ifs at hb + · exact Or.inr ⟨_, _, hb⟩ + · simp_all + · refine ciSup_le fun i ↦ ?_ + simp_rw [ciSup_eq_ite] + split_ifs + · exact le_ciSup hf ?_ + · exact hf' + +theorem ciInf_subtype [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : Subtype p → α} + (hf : BddBelow (Set.range f)) (hf' : iInf f ≤ sInf ∅) : + iInf f = ⨅ (i) (h : p i), f ⟨i, h⟩ := + ciSup_subtype (α := αᵒᵈ) hf hf' + +theorem ciSup_subtype' [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : ∀ i, p i → α} + (hf : BddAbove (Set.range (fun i : Subtype p ↦ f i i.prop))) + (hf' : sSup ∅ ≤ ⨆ (i : Subtype p), f i i.prop) : + ⨆ (i) (h), f i h = ⨆ x : Subtype p, f x x.property := + (ciSup_subtype (f := fun x => f x.val x.property) hf hf').symm + +theorem ciInf_subtype' [Nonempty ι] {p : ι → Prop} [Nonempty (Subtype p)] {f : ∀ i, p i → α} + (hf : BddBelow (Set.range (fun i : Subtype p ↦ f i i.prop))) + (hf' : ⨅ (i : Subtype p), f i i.prop ≤ sInf ∅) : + ⨅ (i) (h), f i h = ⨅ x : Subtype p, f x x.property := + (ciInf_subtype (f := fun x => f x.val x.property) hf hf').symm + +theorem ciSup_subtype'' {ι} [Nonempty ι] {s : Set ι} (hs : s.Nonempty) {f : ι → α} + (hf : BddAbove (Set.range fun i : s ↦ f i)) (hf' : sSup ∅ ≤ ⨆ i : s, f i) : + ⨆ i : s, f i = ⨆ (t : ι) (_ : t ∈ s), f t := + haveI : Nonempty s := Set.Nonempty.to_subtype hs + ciSup_subtype hf hf' + +theorem ciInf_subtype'' {ι} [Nonempty ι] {s : Set ι} (hs : s.Nonempty) {f : ι → α} + (hf : BddBelow (Set.range fun i : s ↦ f i)) (hf' : ⨅ i : s, f i ≤ sInf ∅) : + ⨅ i : s, f i = ⨅ (t : ι) (_ : t ∈ s), f t := + haveI : Nonempty s := Set.Nonempty.to_subtype hs + ciInf_subtype hf hf' + +theorem csSup_image [Nonempty β] {s : Set β} (hs : s.Nonempty) {f : β → α} + (hf : BddAbove (Set.range fun i : s ↦ f i)) (hf' : sSup ∅ ≤ ⨆ i : s, f i) : + sSup (f '' s) = ⨆ a ∈ s, f a := by + rw [← ciSup_subtype'' hs hf hf', iSup, Set.image_eq_range] + +theorem csInf_image [Nonempty β] {s : Set β} (hs : s.Nonempty) {f : β → α} + (hf : BddBelow (Set.range fun i : s ↦ f i)) (hf' : ⨅ i : s, f i ≤ sInf ∅) : + sInf (f '' s) = ⨅ a ∈ s, f a := + csSup_image (α := αᵒᵈ) hs hf hf' + +lemma ciSup_image {α ι ι' : Type*} [ConditionallyCompleteLattice α] [Nonempty ι] [Nonempty ι'] + {s : Set ι} (hs : s.Nonempty) {f : ι → ι'} {g : ι' → α} + (hf : BddAbove (Set.range fun i : s ↦ g (f i))) (hg' : sSup ∅ ≤ ⨆ i : s, g (f i)) : + ⨆ i ∈ (f '' s), g i = ⨆ x ∈ s, g (f x) := by + have hg : BddAbove (Set.range fun i : f '' s ↦ g i) := by + simpa [bddAbove_def] using hf + have hf' : sSup ∅ ≤ ⨆ i : f '' s, g i := by + refine hg'.trans ?_ + have : Nonempty s := Set.Nonempty.to_subtype hs + refine ciSup_le ?_ + intro ⟨i, h⟩ + obtain ⟨t, ht⟩ : ∃ t : f '' s, g t = g (f (Subtype.mk i h)) := by + have : f i ∈ f '' s := Set.mem_image_of_mem _ h + exact ⟨⟨f i, this⟩, by simp [this]⟩ + rw [← ht] + refine le_ciSup_set ?_ t.prop + simpa [bddAbove_def] using hf + rw [← csSup_image (by simpa using hs) hg hf', ← csSup_image hs hf hg', ← Set.image_comp, comp_def] + +lemma ciInf_image {α ι ι' : Type*} [ConditionallyCompleteLattice α] [Nonempty ι] [Nonempty ι'] + {s : Set ι} (hs : s.Nonempty) {f : ι → ι'} {g : ι' → α} + (hf : BddBelow (Set.range fun i : s ↦ g (f i))) (hg' : ⨅ i : s, g (f i) ≤ sInf ∅) : + ⨅ i ∈ (f '' s), g i = ⨅ x ∈ s, g (f x) := + ciSup_image (α := αᵒᵈ) hs hf hg' + end ConditionallyCompleteLattice instance Pi.conditionallyCompleteLattice {ι : Type*} {α : ι → Type*} @@ -992,7 +1080,7 @@ theorem cbiSup_eq_of_not_forall {p : ι → Prop} {f : Subtype p → α} (hp : · simp [hi] · apply sup_le · rcases isEmpty_or_nonempty (Subtype p) with hp|hp - · simp only [iSup_of_empty'] + · rw [iSup_of_empty'] convert le_ciSup B i₀ simp [hi₀] · apply ciSup_le @@ -1123,6 +1211,14 @@ theorem ciSup_mono' {ι'} {f : ι → α} {g : ι' → α} (hg : BddAbove (range theorem csInf_le_csInf' {s t : Set α} (h₁ : t.Nonempty) (h₂ : t ⊆ s) : sInf s ≤ sInf t := csInf_le_csInf (OrderBot.bddBelow s) h₁ h₂ +lemma ciSup_or' (p q : Prop) (f : p ∨ q → α) : + ⨆ (h : p ∨ q), f h = (⨆ h : p, f (.inl h)) ⊔ ⨆ h : q, f (.inr h) := by + by_cases hp : p <;> by_cases hq : q + · simp [hp, hq] + · simp [hp, hq] + · simp [hp, hq] + · simp [hp, hq] + end ConditionallyCompleteLinearOrderBot namespace WithTop From 145e6ba702be03cbd05635f29bca8f2c95622da7 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 5 Aug 2024 08:52:10 +0000 Subject: [PATCH 032/359] chore(CstarAlgebra/Module): drop `DecidableEq` assumptions (#15489) Also reuse `map_sum` instead of proving by induction. --- Mathlib/Analysis/CstarAlgebra/Module.lean | 15 ++++++--------- 1 file changed, 6 insertions(+), 9 deletions(-) diff --git a/Mathlib/Analysis/CstarAlgebra/Module.lean b/Mathlib/Analysis/CstarAlgebra/Module.lean index 4051e9d595f1b..8244b2422d6d3 100644 --- a/Mathlib/Analysis/CstarAlgebra/Module.lean +++ b/Mathlib/Analysis/CstarAlgebra/Module.lean @@ -117,17 +117,14 @@ lemma innerₛₗ_apply {x y : E} : innerₛₗ x y = ⟪x, y⟫ := rfl simp [← innerₛₗ_apply] @[simp] -lemma inner_sum_right {ι : Type*} [DecidableEq ι] {s : Finset ι} {x : E} {y : ι → E} : - ⟪x, ∑ i ∈ s, y i⟫ = ∑ i ∈ s, ⟪x, y i⟫ := by - induction s using Finset.induction_on - case empty => simp - case insert a t a_notmem_t hbase => - simp_rw [Finset.sum_insert a_notmem_t] - simp only [inner_add_right, hbase] +lemma inner_sum_right {ι : Type*} {s : Finset ι} {x : E} {y : ι → E} : + ⟪x, ∑ i ∈ s, y i⟫ = ∑ i ∈ s, ⟪x, y i⟫ := + map_sum (innerₛₗ x) .. @[simp] -lemma inner_sum_left {ι : Type*} [DecidableEq ι] {s : Finset ι} {x : ι → E} {y : E} : - ⟪∑ i ∈ s, x i, y⟫ = ∑ i ∈ s, ⟪x i, y⟫ := by rw [← star_inner y]; simp +lemma inner_sum_left {ι : Type*} {s : Finset ι} {x : ι → E} {y : E} : + ⟪∑ i ∈ s, x i, y⟫ = ∑ i ∈ s, ⟪x i, y⟫ := + map_sum (innerₛₗ.flip y) .. @[simp] lemma isSelfAdjoint_inner_self {x : E} : IsSelfAdjoint ⟪x, x⟫ := star_inner _ _ From 5025874dc5f9f8dd1598190e60ef20dda7b42566 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 5 Aug 2024 08:52:12 +0000 Subject: [PATCH 033/359] chore: move toolchain to v4.11.0-rc1 (#15513) Co-authored-by: Johan Commelin --- Archive/MiuLanguage/DecisionNec.lean | 8 +- Archive/MiuLanguage/DecisionSuf.lean | 6 +- Archive/Wiedijk100Theorems/BallotProblem.lean | 9 +- Archive/Wiedijk100Theorems/Partition.lean | 4 +- LongestPole/Main.lean | 1 - Mathlib.lean | 1 - Mathlib/Algebra/Algebra/Equiv.lean | 3 - Mathlib/Algebra/Algebra/Hom.lean | 3 - Mathlib/Algebra/Algebra/NonUnitalHom.lean | 5 - Mathlib/Algebra/BigOperators/Fin.lean | 2 +- Mathlib/Algebra/BigOperators/Group/List.lean | 25 +- .../Algebra/Category/BialgebraCat/Basic.lean | 6 +- .../Algebra/Category/CoalgebraCat/Basic.lean | 6 +- .../Category/HopfAlgebraCat/Basic.lean | 6 +- Mathlib/Algebra/Category/ModuleCat/Basic.lean | 8 +- Mathlib/Algebra/Category/ModuleCat/Sheaf.lean | 2 +- Mathlib/Algebra/CharP/Defs.lean | 2 +- Mathlib/Algebra/CubicDiscriminant.lean | 2 +- Mathlib/Algebra/DirectSum/Module.lean | 3 +- Mathlib/Algebra/Free.lean | 4 +- Mathlib/Algebra/Group/Action/Defs.lean | 2 +- Mathlib/Algebra/Group/Equiv/Basic.lean | 4 - Mathlib/Algebra/Group/Units.lean | 4 - .../Algebra/GroupWithZero/Units/Basic.lean | 1 + .../Algebra/Homology/DerivedCategory/Ext.lean | 3 - .../Homology/HomotopyCategory/HomComplex.lean | 3 - .../Algebra/Homology/ShortComplex/Basic.lean | 2 +- Mathlib/Algebra/Lie/Basic.lean | 10 - Mathlib/Algebra/Lie/DirectSum.lean | 3 +- Mathlib/Algebra/Lie/Weights/Basic.lean | 4 +- Mathlib/Algebra/Module/Defs.lean | 2 +- Mathlib/Algebra/Module/Equiv/Defs.lean | 3 - Mathlib/Algebra/Module/Injective.lean | 2 +- Mathlib/Algebra/Module/LinearMap/Defs.lean | 6 - Mathlib/Algebra/MvPolynomial/Basic.lean | 3 - .../Order/BigOperators/Group/List.lean | 2 +- Mathlib/Algebra/Order/Interval/Basic.lean | 8 +- .../Algebra/Order/Monoid/Unbundled/Basic.lean | 1 + Mathlib/Algebra/Quaternion.lean | 26 +- Mathlib/Algebra/Ring/Equiv.lean | 3 - Mathlib/Algebra/Ring/Ext.lean | 90 ------- Mathlib/Algebra/Ring/Hom/Defs.lean | 6 - Mathlib/Algebra/Ring/Subring/Basic.lean | 2 +- Mathlib/Algebra/Star/StarAlgHom.lean | 3 - Mathlib/Algebra/Star/StarRingHom.lean | 3 - Mathlib/AlgebraicGeometry/Limits.lean | 4 +- .../PrimeSpectrum/Basic.lean | 9 +- .../ProjectiveSpectrum/Scheme.lean | 4 +- Mathlib/AlgebraicGeometry/Scheme.lean | 4 +- Mathlib/AlgebraicGeometry/Spec.lean | 4 +- .../FundamentalGroupoid/InducedMaps.lean | 10 +- .../AlgebraicTopology/SimplicialObject.lean | 4 +- Mathlib/AlgebraicTopology/SimplicialSet.lean | 2 +- Mathlib/Analysis/Analytic/Composition.lean | 2 +- .../BoxIntegral/Partition/Filter.lean | 2 +- .../Calculus/FormalMultilinearSeries.lean | 3 - .../Complex/UpperHalfPlane/Basic.lean | 1 - .../Convex/SimplicialComplex/Basic.lean | 2 +- .../WeakOperatorTopology.lean | 3 - Mathlib/Analysis/Normed/Group/Hom.lean | 3 - Mathlib/Analysis/Normed/Lp/lpSpace.lean | 3 - .../Normed/Operator/WeakOperatorTopology.lean | 4 - Mathlib/Analysis/NormedSpace/ENorm.lean | 3 - .../CategoryTheory/Bicategory/Coherence.lean | 12 +- .../Bicategory/LocallyDiscrete.lean | 8 +- Mathlib/CategoryTheory/Category/Basic.lean | 2 +- .../CategoryTheory/Category/Bipointed.lean | 12 +- .../CategoryTheory/Category/PartialFun.lean | 8 +- Mathlib/CategoryTheory/Category/Pointed.lean | 8 +- Mathlib/CategoryTheory/Category/TwoP.lean | 12 +- Mathlib/CategoryTheory/Comma/Arrow.lean | 2 +- Mathlib/CategoryTheory/Comma/Basic.lean | 2 +- Mathlib/CategoryTheory/Comma/Presheaf.lean | 2 +- .../CategoryTheory/Comma/StructuredArrow.lean | 8 +- .../CategoryTheory/ConnectedComponents.lean | 2 +- Mathlib/CategoryTheory/Dialectica/Basic.lean | 2 +- .../CategoryTheory/DifferentialObject.lean | 2 +- .../CategoryTheory/Endofunctor/Algebra.lean | 4 +- Mathlib/CategoryTheory/Filtered/Small.lean | 5 +- Mathlib/CategoryTheory/Functor/Category.lean | 2 +- .../Galois/Prorepresentability.lean | 2 +- .../CategoryTheory/GradedObject/Monoidal.lean | 8 +- Mathlib/CategoryTheory/Grothendieck.lean | 2 +- .../CategoryTheory/GuitartExact/Basic.lean | 2 +- .../CategoryTheory/Idempotents/Karoubi.lean | 2 +- Mathlib/CategoryTheory/IsConnected.lean | 4 +- .../Limits/Preserves/Ulift.lean | 2 +- .../Limits/Shapes/Biproducts.lean | 2 +- .../CategoryTheory/Limits/Shapes/Images.lean | 6 +- .../CategoryTheory/Limits/Shapes/Types.lean | 1 + Mathlib/CategoryTheory/Limits/Types.lean | 3 - Mathlib/CategoryTheory/Limits/VanKampen.lean | 2 +- .../Localization/Resolution.lean | 4 +- Mathlib/CategoryTheory/Monad/Algebra.lean | 4 +- Mathlib/CategoryTheory/Monad/Basic.lean | 16 +- Mathlib/CategoryTheory/Monad/Coequalizer.lean | 2 +- Mathlib/CategoryTheory/Monad/Equalizer.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Bimod.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Bimon_.lean | 3 +- .../Monoidal/Braided/Basic.lean | 4 +- Mathlib/CategoryTheory/Monoidal/CommMon_.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Comon_.lean | 2 +- .../Monoidal/Free/Coherence.lean | 6 +- Mathlib/CategoryTheory/Monoidal/Mod_.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 2 +- .../Monoidal/NaturalTransformation.lean | 2 +- .../CategoryTheory/Monoidal/Rigid/Basic.lean | 19 +- .../MorphismProperty/IsInvertedBy.lean | 2 +- Mathlib/CategoryTheory/PathCategory.lean | 4 +- Mathlib/CategoryTheory/Quotient.lean | 2 +- .../CategoryTheory/Shift/SingleFunctors.lean | 2 +- .../Sites/ConcreteSheafification.lean | 6 +- .../CategoryTheory/Sites/CoverLifting.lean | 4 +- Mathlib/CategoryTheory/Sites/Coverage.lean | 2 +- .../CategoryTheory/Sites/DenseSubsite.lean | 2 +- Mathlib/CategoryTheory/Sites/Limits.lean | 2 +- Mathlib/CategoryTheory/Sites/Pretopology.lean | 2 +- Mathlib/CategoryTheory/Sites/Sheaf.lean | 10 +- .../CategoryTheory/Sites/SheafOfTypes.lean | 8 +- Mathlib/CategoryTheory/Sites/Sieves.lean | 3 - Mathlib/CategoryTheory/Sites/Subsheaf.lean | 10 +- Mathlib/CategoryTheory/Sites/Types.lean | 5 +- Mathlib/CategoryTheory/Sites/Whiskering.lean | 4 +- Mathlib/CategoryTheory/Square.lean | 2 +- Mathlib/CategoryTheory/Subobject/Basic.lean | 2 +- Mathlib/CategoryTheory/Thin.lean | 2 +- .../CategoryTheory/Triangulated/Basic.lean | 2 +- Mathlib/CategoryTheory/Yoneda.lean | 4 +- Mathlib/Combinatorics/Colex.lean | 4 +- .../Enumerative/Composition.lean | 4 +- .../Combinatorics/Enumerative/Partition.lean | 8 +- Mathlib/Combinatorics/Quiver/Basic.lean | 2 +- .../Combinatorics/SimpleGraph/Acyclic.lean | 9 +- Mathlib/Combinatorics/SimpleGraph/Basic.lean | 2 +- .../SimpleGraph/Connectivity/Subgraph.lean | 8 +- .../SimpleGraph/Hamiltonian.lean | 23 +- Mathlib/Combinatorics/SimpleGraph/Maps.lean | 6 +- Mathlib/Combinatorics/SimpleGraph/Path.lean | 2 +- .../Combinatorics/SimpleGraph/Subgraph.lean | 8 +- Mathlib/Combinatorics/SimpleGraph/Walk.lean | 127 ++++----- Mathlib/Computability/Primrec.lean | 2 +- Mathlib/Control/Traversable/Basic.lean | 4 - Mathlib/Data/Complex/Basic.lean | 3 - Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean | 2 +- Mathlib/Data/Finmap.lean | 3 - Mathlib/Data/Finset/Basic.lean | 4 +- Mathlib/Data/Finsupp/Defs.lean | 3 - Mathlib/Data/Fintype/Fin.lean | 4 +- Mathlib/Data/List/AList.lean | 3 - Mathlib/Data/List/Basic.lean | 248 +++--------------- Mathlib/Data/List/Chain.lean | 5 +- Mathlib/Data/List/Count.lean | 21 +- Mathlib/Data/List/Cycle.lean | 8 +- Mathlib/Data/List/Destutter.lean | 2 +- Mathlib/Data/List/DropRight.lean | 12 +- Mathlib/Data/List/EditDistance/Defs.lean | 2 +- Mathlib/Data/List/Enum.lean | 109 +------- Mathlib/Data/List/Forall2.lean | 4 +- Mathlib/Data/List/GetD.lean | 1 - Mathlib/Data/List/Indexes.lean | 1 + Mathlib/Data/List/Intervals.lean | 2 +- Mathlib/Data/List/Iterate.lean | 3 +- Mathlib/Data/List/Join.lean | 33 +-- Mathlib/Data/List/Lattice.lean | 12 +- Mathlib/Data/List/NatAntidiagonal.lean | 1 - Mathlib/Data/List/Nodup.lean | 51 +--- Mathlib/Data/List/OfFn.lean | 7 +- Mathlib/Data/List/Perm.lean | 13 +- Mathlib/Data/List/Range.lean | 47 ---- Mathlib/Data/List/Rotate.lean | 5 +- Mathlib/Data/List/Sigma.lean | 2 +- Mathlib/Data/List/Sort.lean | 5 +- Mathlib/Data/List/Sublists.lean | 6 +- Mathlib/Data/List/Sym.lean | 2 +- Mathlib/Data/List/TFAE.lean | 4 +- Mathlib/Data/List/Zip.lean | 145 +--------- Mathlib/Data/Multiset/Basic.lean | 30 +-- Mathlib/Data/Multiset/FinsetOps.lean | 1 + Mathlib/Data/Multiset/Nodup.lean | 4 +- Mathlib/Data/Multiset/Powerset.lean | 1 + Mathlib/Data/NNRat/Defs.lean | 3 - Mathlib/Data/NNReal/Basic.lean | 3 - Mathlib/Data/Nat/Defs.lean | 11 +- Mathlib/Data/Nat/Nth.lean | 4 +- Mathlib/Data/Option/Basic.lean | 14 +- Mathlib/Data/PEquiv.lean | 3 - Mathlib/Data/Prod/Basic.lean | 3 - Mathlib/Data/Seq/Seq.lean | 52 ++-- Mathlib/Data/Set/Basic.lean | 3 - Mathlib/Data/Set/Image.lean | 1 + Mathlib/Data/Set/List.lean | 4 +- Mathlib/Data/Set/Pairwise/Basic.lean | 1 + Mathlib/Data/Sigma/Basic.lean | 20 +- Mathlib/Data/Sigma/Order.lean | 2 +- Mathlib/Data/String/Defs.lean | 2 +- Mathlib/Data/String/Lemmas.lean | 1 - Mathlib/Data/Subtype.lean | 20 +- Mathlib/Data/Sym/Basic.lean | 2 +- Mathlib/Data/ULift.lean | 3 - Mathlib/Data/Vector/Zip.lean | 1 - .../RotationNumber/TranslationNumber.lean | 3 - Mathlib/Geometry/Euclidean/Angle/Sphere.lean | 2 +- Mathlib/Geometry/Euclidean/Circumcenter.lean | 2 +- Mathlib/Geometry/Euclidean/Sphere/Basic.lean | 2 +- Mathlib/Geometry/Manifold/MFDeriv/Basic.lean | 2 +- .../RingedSpace/LocallyRingedSpace.lean | 16 +- .../LocallyRingedSpace/HasColimits.lean | 8 +- .../Geometry/RingedSpace/OpenImmersion.lean | 16 +- .../Geometry/RingedSpace/PresheafedSpace.lean | 4 +- .../PresheafedSpace/HasColimits.lean | 6 +- .../Geometry/RingedSpace/SheafedSpace.lean | 4 +- Mathlib/GroupTheory/Congruence/Basic.lean | 5 - Mathlib/GroupTheory/CoprodI.lean | 12 +- Mathlib/GroupTheory/Coxeter/Inversion.lean | 10 +- Mathlib/GroupTheory/GroupAction/Hom.lean | 12 - Mathlib/GroupTheory/HNNExtension.lean | 4 +- .../GroupTheory/MonoidLocalization/Basic.lean | 4 - Mathlib/GroupTheory/PushoutI.lean | 7 +- Mathlib/GroupTheory/SemidirectProduct.lean | 8 +- Mathlib/GroupTheory/Sylow.lean | 3 - Mathlib/Init/Data/Nat/GCD.lean | 32 --- Mathlib/Lean/Name.lean | 11 +- .../LinearAlgebra/AffineSpace/AffineMap.lean | 5 +- .../AffineSpace/AffineSubspace.lean | 3 +- .../AffineSpace/ContinuousAffineEquiv.lean | 3 - .../AffineSpace/Independent.lean | 3 +- Mathlib/LinearAlgebra/Alternating/Basic.lean | 3 - Mathlib/LinearAlgebra/BilinearForm/Basic.lean | 2 - .../LinearAlgebra/CliffordAlgebra/Even.lean | 4 +- .../CliffordAlgebra/EvenEquiv.lean | 2 +- Mathlib/LinearAlgebra/LinearPMap.lean | 2 +- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 3 - Mathlib/LinearAlgebra/Pi.lean | 3 - .../LinearAlgebra/QuadraticForm/Basic.lean | 3 - .../QuadraticForm/QuadraticModuleCat.lean | 12 +- Mathlib/LinearAlgebra/TensorPower.lean | 4 +- Mathlib/Logic/Basic.lean | 22 +- Mathlib/Logic/Embedding/Basic.lean | 4 - Mathlib/Logic/Equiv/Basic.lean | 2 +- Mathlib/Logic/Equiv/Defs.lean | 4 - Mathlib/Logic/Nontrivial/Basic.lean | 1 - Mathlib/Logic/Relation.lean | 18 +- Mathlib/Mathport/Syntax.lean | 1 - .../MeasureTheory/Decomposition/Jordan.lean | 2 +- Mathlib/MeasureTheory/Function/AEEqFun.lean | 3 - .../ConditionalExpectation/AEMeasurable.lean | 8 +- Mathlib/MeasureTheory/Function/Jacobian.lean | 9 +- Mathlib/MeasureTheory/Function/LpSpace.lean | 3 - .../MeasureTheory/MeasurableSpace/Defs.lean | 4 - Mathlib/MeasureTheory/Measure/AddContent.lean | 3 - .../Measure/MeasureSpaceDef.lean | 3 - Mathlib/ModelTheory/Basic.lean | 9 - Mathlib/ModelTheory/ElementaryMaps.lean | 3 - Mathlib/ModelTheory/Encoding.lean | 8 + Mathlib/NumberTheory/ArithmeticFunction.lean | 3 - Mathlib/NumberTheory/ModularForms/Basic.lean | 2 +- Mathlib/NumberTheory/MulChar/Basic.lean | 5 - Mathlib/NumberTheory/NumberField/Basic.lean | 3 - Mathlib/NumberTheory/Pell.lean | 4 +- Mathlib/NumberTheory/Zsqrtd/Basic.lean | 10 +- Mathlib/Order/Filter/Basic.lean | 3 - Mathlib/Order/RelClasses.lean | 2 +- Mathlib/Order/RelIso/Basic.lean | 9 - Mathlib/Order/RelSeries.lean | 2 +- .../Probability/ConditionalExpectation.lean | 2 +- Mathlib/Probability/Kernel/Basic.lean | 2 - .../ProbabilityMassFunction/Basic.lean | 3 - .../RepresentationTheory/Action/Basic.lean | 4 +- Mathlib/RepresentationTheory/Rep.lean | 8 +- Mathlib/RingTheory/AdicCompletion/Basic.lean | 6 - Mathlib/RingTheory/Bialgebra/Equiv.lean | 3 - Mathlib/RingTheory/Bialgebra/Hom.lean | 3 - Mathlib/RingTheory/Coalgebra/Equiv.lean | 3 - Mathlib/RingTheory/Coalgebra/Hom.lean | 3 - .../DedekindDomain/Factorization.lean | 4 +- Mathlib/RingTheory/DedekindDomain/Ideal.lean | 2 +- Mathlib/RingTheory/Filtration.lean | 3 +- Mathlib/RingTheory/HahnSeries/Basic.lean | 6 +- .../RingTheory/HahnSeries/Multiplication.lean | 4 - Mathlib/RingTheory/IsTensorProduct.lean | 10 +- Mathlib/RingTheory/MaximalSpectrum.lean | 2 +- Mathlib/RingTheory/MvPowerSeries/Basic.lean | 6 +- Mathlib/RingTheory/PowerSeries/Basic.lean | 3 +- Mathlib/RingTheory/PrimeSpectrum.lean | 4 +- Mathlib/RingTheory/Valuation/Basic.lean | 8 - Mathlib/RingTheory/WittVector/Defs.lean | 3 - Mathlib/RingTheory/WittVector/Truncated.lean | 3 - Mathlib/SetTheory/ZFC/Basic.lean | 6 - Mathlib/Tactic/Abel.lean | 2 +- Mathlib/Tactic/Attr/Core.lean | 1 - Mathlib/Tactic/Linter/Lint.lean | 1 - Mathlib/Tactic/Linter/UnusedTactic.lean | 1 - Mathlib/Tactic/TFAE.lean | 2 +- Mathlib/Tactic/Widget/Calc.lean | 2 + Mathlib/Tactic/Widget/Conv.lean | 1 + Mathlib/Testing/SlimCheck/Functions.lean | 12 +- Mathlib/Testing/SlimCheck/Testable.lean | 2 +- Mathlib/Topology/Algebra/Algebra.lean | 4 +- .../Topology/Algebra/ContinuousAffineMap.lean | 3 - .../Algebra/Module/Alternating/Basic.lean | 3 - Mathlib/Topology/Algebra/Module/Basic.lean | 6 - .../Algebra/Module/Multilinear/Basic.lean | 3 - Mathlib/Topology/Basic.lean | 2 +- Mathlib/Topology/Bornology/Basic.lean | 6 +- .../Topology/Category/TopCat/OpenNhds.lean | 2 +- .../ContinuousFunction/StoneWeierstrass.lean | 4 +- Mathlib/Topology/LocallyConstant/Basic.lean | 2 - Mathlib/Topology/MetricSpace/Dilation.lean | 3 - Mathlib/Topology/Metrizable/Uniformity.lean | 8 +- Mathlib/Topology/PartialHomeomorph.lean | 6 - Mathlib/Topology/Sheaves/Sheaf.lean | 2 +- .../Sheaves/SheafCondition/OpensLeCover.lean | 2 +- Mathlib/Topology/Sheaves/Skyscraper.lean | 10 +- Mathlib/Topology/ShrinkingLemma.lean | 2 +- Mathlib/Topology/UniformSpace/Basic.lean | 2 +- Mathlib/Util/AssertNoSorry.lean | 13 +- Mathlib/Util/LongNames.lean | 3 +- lake-manifest.json | 10 +- lakefile.lean | 10 +- lean-toolchain | 2 +- scripts/check-yaml.lean | 1 - scripts/noshake.json | 2 +- test/DeriveToExpr.lean | 2 +- test/TCSynth.lean | 2 +- test/Use.lean | 6 +- test/Variable.lean | 2 +- test/eval_elab.lean | 2 +- test/jacobiSym.lean | 2 +- 328 files changed, 675 insertions(+), 1773 deletions(-) delete mode 100644 Mathlib/Init/Data/Nat/GCD.lean diff --git a/Archive/MiuLanguage/DecisionNec.lean b/Archive/MiuLanguage/DecisionNec.lean index ce395bd21c5b7..ba3cc0bc34669 100644 --- a/Archive/MiuLanguage/DecisionNec.lean +++ b/Archive/MiuLanguage/DecisionNec.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Gihan Marasingha -/ import Archive.MiuLanguage.Basic +import Mathlib.Data.List.Basic import Mathlib.Data.List.Count import Mathlib.Data.Nat.ModEq import Mathlib.Tactic.Ring @@ -71,9 +72,10 @@ theorem count_equiv_one_or_two_mod3_of_derivable (en : Miustr) : any_goals apply mod3_eq_1_or_mod3_eq_2 h_ih -- Porting note: `simp_rw [count_append]` usually doesn't work · left; rw [count_append, count_append]; rfl - · right; simp_rw [count_append, count_cons, if_false, two_mul]; simp + · right; simp_rw [count_append, count_cons, beq_iff_eq, ite_false, add_zero, two_mul] · left; rw [count_append, count_append, count_append] - simp_rw [count_cons_self, count_nil, count_cons, ite_false, add_right_comm, add_mod_right] + simp_rw [count_cons_self, count_nil, count_cons, beq_iff_eq, ite_false, add_right_comm, + add_mod_right] simp · left; rw [count_append, count_append, count_append] simp only [ne_eq, not_false_eq_true, count_cons_of_ne, count_nil, add_zero] @@ -127,7 +129,7 @@ theorem goodm_of_rule1 (xs : Miustr) (h₁ : Derivable (xs ++ ↑[I])) (h₂ : G · change ¬M ∈ tail (xs ++ ↑([I] ++ [U])) rw [← append_assoc, tail_append_singleton_of_ne_nil] · simp_rw [mem_append, mem_singleton, or_false]; exact nmtail - · exact append_ne_nil_of_ne_nil_left _ _ this + · exact append_ne_nil_of_left_ne_nil this _ theorem goodm_of_rule2 (xs : Miustr) (_ : Derivable (M :: xs)) (h₂ : Goodm (M :: xs)) : Goodm (↑(M :: xs) ++ xs) := by diff --git a/Archive/MiuLanguage/DecisionSuf.lean b/Archive/MiuLanguage/DecisionSuf.lean index 4e460e9939ba2..10eaa8e26942b 100644 --- a/Archive/MiuLanguage/DecisionSuf.lean +++ b/Archive/MiuLanguage/DecisionSuf.lean @@ -251,7 +251,7 @@ theorem count_I_eq_length_of_count_U_zero_and_neg_mem {ys : Miustr} (hu : count · -- case `x = M` gives a contradiction. exfalso; exact hm (mem_cons_self M xs) · -- case `x = I` - rw [count_cons, if_pos rfl, length, succ_inj'] + rw [count_cons, beq_self_eq_true, if_pos rfl, length, succ_inj'] apply hxs · simpa only [count] · rw [mem_cons, not_or] at hm; exact hm.2 @@ -307,7 +307,9 @@ theorem ind_hyp_suf (k : ℕ) (ys : Miustr) (hu : count U ys = succ k) (hdec : D use as, bs refine ⟨rfl, ?_, ?_, ?_⟩ · -- Porting note: `simp_rw [count_append]` didn't work - rw [count_append] at hu; simp_rw [count_cons, if_true, add_succ, succ_inj'] at hu + rw [count_append] at hu + simp_rw [count_cons, beq_self_eq_true, if_true, add_succ, beq_iff_eq, reduceIte, add_zero, + succ_inj'] at hu rwa [count_append, count_append] · apply And.intro rfl rw [cons_append, cons_append] diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index b037dd8fab914..8b45732a49ddb 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -115,18 +115,19 @@ theorem counted_succ_succ (p q : ℕ) : obtain ⟨hl₀, hl₁, hl₂⟩ := hl obtain hlast | hlast := hl₂ (l.head hlnil) (List.head_mem hlnil) · refine Or.inl ⟨l.tail, ⟨?_, ?_, ?_⟩, ?_⟩ - · rw [List.count_tail l 1 hlnil, hl₀, hlast, if_pos rfl, Nat.add_sub_cancel] + · rw [List.count_tail l 1 hlnil, hl₀, hlast, beq_self_eq_true, if_pos rfl, Nat.add_sub_cancel] · rw [List.count_tail l (-1) hlnil, hl₁, hlast, if_neg (by decide), Nat.sub_zero] · exact fun x hx => hl₂ x (List.mem_of_mem_tail hx) · rw [← hlast, List.head_cons_tail] · refine Or.inr ⟨l.tail, ⟨?_, ?_, ?_⟩, ?_⟩ · rw [List.count_tail l 1 hlnil, hl₀, hlast, if_neg (by decide), Nat.sub_zero] - · rw [List.count_tail l (-1) hlnil, hl₁, hlast, if_pos rfl, Nat.add_sub_cancel] + · rw [List.count_tail l (-1) hlnil, hl₁, hlast, beq_self_eq_true, if_pos rfl, + Nat.add_sub_cancel] · exact fun x hx => hl₂ x (List.mem_of_mem_tail hx) · rw [← hlast, List.head_cons_tail] · rintro (⟨t, ⟨ht₀, ht₁, ht₂⟩, rfl⟩ | ⟨t, ⟨ht₀, ht₁, ht₂⟩, rfl⟩) · refine ⟨?_, ?_, ?_⟩ - · rw [List.count_cons, if_pos rfl, ht₀] + · rw [List.count_cons, beq_self_eq_true, if_pos rfl, ht₀] · rw [List.count_cons, if_neg, ht₁] norm_num · rintro x (_ | _) @@ -134,7 +135,7 @@ theorem counted_succ_succ (p q : ℕ) : · refine ⟨?_, ?_, ?_⟩ · rw [List.count_cons, if_neg, ht₀] norm_num - · rw [List.count_cons, if_pos rfl, ht₁] + · rw [List.count_cons, beq_self_eq_true, if_pos rfl, ht₁] · rintro x (_ | _) exacts [Or.inr rfl, ht₂ x (by tauto)] diff --git a/Archive/Wiedijk100Theorems/Partition.lean b/Archive/Wiedijk100Theorems/Partition.lean index 0179465f55758..38232baf5bdab 100644 --- a/Archive/Wiedijk100Theorems/Partition.lean +++ b/Archive/Wiedijk100Theorems/Partition.lean @@ -235,7 +235,7 @@ theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) rw [← hw₂] exact dvd_mul_left _ _ · intro i - simp_rw [Multiset.count_sum', Multiset.count_replicate, sum_ite_eq] + simp_rw [Multiset.count_sum', Multiset.count_replicate, sum_ite_eq'] split_ifs with h · rcases hf₄ i h with ⟨w, hw₁, hw₂⟩ rwa [← hw₂, Nat.mul_div_cancel _ (hs i h)] @@ -245,7 +245,7 @@ theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) rcases hi with ⟨j, hj₁, hj₂⟩ rwa [Multiset.eq_of_mem_replicate hj₂] · ext i - simp_rw [Multiset.count_sum', Multiset.count_replicate, sum_ite_eq] + simp_rw [Multiset.count_sum', Multiset.count_replicate, sum_ite_eq'] simp only [ne_eq, Multiset.mem_toFinset, not_not, smul_eq_mul, ite_mul, zero_mul, Finsupp.coe_mk] split_ifs with h diff --git a/LongestPole/Main.lean b/LongestPole/Main.lean index 086e6a6d9ac21..1ad6f2a7b33cd 100644 --- a/LongestPole/Main.lean +++ b/LongestPole/Main.lean @@ -6,7 +6,6 @@ Authors: Scott Morrison import ImportGraph import Mathlib.Data.String.Defs import Mathlib.Util.FormatTable -import Batteries.Lean.Util.Path import Cli import LongestPole.SpeedCenterJson diff --git a/Mathlib.lean b/Mathlib.lean index 29748850f8ccb..2af47ca3954ff 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2783,7 +2783,6 @@ import Mathlib.Init.Data.Int.Order import Mathlib.Init.Data.List.Basic import Mathlib.Init.Data.List.Instances import Mathlib.Init.Data.List.Lemmas -import Mathlib.Init.Data.Nat.GCD import Mathlib.Init.Data.Nat.Lemmas import Mathlib.Init.Data.Quot import Mathlib.Init.Data.Sigma.Basic diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index f85d2346ffdc7..244dfed080662 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -132,9 +132,6 @@ protected theorem congr_arg {f : A₁ ≃ₐ[R] A₂} {x x' : A₁} : x = x' → protected theorem congr_fun {f g : A₁ ≃ₐ[R] A₂} (h : f = g) (x : A₁) : f x = g x := DFunLike.congr_fun h x -protected theorem ext_iff {f g : A₁ ≃ₐ[R] A₂} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - theorem coe_fun_injective : @Function.Injective (A₁ ≃ₐ[R] A₂) (A₁ → A₂) fun e => (e : A₁ → A₂) := DFunLike.coe_injective diff --git a/Mathlib/Algebra/Algebra/Hom.lean b/Mathlib/Algebra/Algebra/Hom.lean index 3d2a23fa99344..e82539f2a9c02 100644 --- a/Mathlib/Algebra/Algebra/Hom.lean +++ b/Mathlib/Algebra/Algebra/Hom.lean @@ -190,9 +190,6 @@ protected theorem congr_arg (φ : A →ₐ[R] B) {x y : A} (h : x = y) : φ x = theorem ext {φ₁ φ₂ : A →ₐ[R] B} (H : ∀ x, φ₁ x = φ₂ x) : φ₁ = φ₂ := DFunLike.ext _ _ H -theorem ext_iff {φ₁ φ₂ : A →ₐ[R] B} : φ₁ = φ₂ ↔ ∀ x, φ₁ x = φ₂ x := - DFunLike.ext_iff - @[simp] theorem mk_coe {f : A →ₐ[R] B} (h₁ h₂ h₃ h₄ h₅) : (⟨⟨⟨⟨f, h₁⟩, h₂⟩, h₃, h₄⟩, h₅⟩ : A →ₐ[R] B) = f := ext fun _ => rfl diff --git a/Mathlib/Algebra/Algebra/NonUnitalHom.lean b/Mathlib/Algebra/Algebra/NonUnitalHom.lean index 3cf00e9e5c710..5df753173fa61 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalHom.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalHom.lean @@ -198,11 +198,6 @@ instance : NonUnitalAlgSemiHomClass (A →ₛₙₐ[φ] B) φ A B where theorem ext {f g : A →ₛₙₐ[φ] B} (h : ∀ x, f x = g x) : f = g := coe_injective <| funext h -theorem ext_iff {f g : A →ₛₙₐ[φ] B} : f = g ↔ ∀ x, f x = g x := - ⟨by - rintro rfl x - rfl, ext⟩ - theorem congr_fun {f g : A →ₛₙₐ[φ] B} (h : f = g) (x : A) : f x = g x := h ▸ rfl diff --git a/Mathlib/Algebra/BigOperators/Fin.lean b/Mathlib/Algebra/BigOperators/Fin.lean index a8d0c959a9e3c..e8acaf6380069 100644 --- a/Mathlib/Algebra/BigOperators/Fin.lean +++ b/Mathlib/Algebra/BigOperators/Fin.lean @@ -406,7 +406,7 @@ theorem prod_take_ofFn {n : ℕ} (f : Fin n → α) (i : ℕ) : · have A : (ofFn f).take i = (ofFn f).take i.succ := by rw [← length_ofFn f] at h have : length (ofFn f) ≤ i := not_lt.mp h - rw [take_all_of_le this, take_all_of_le (le_trans this (Nat.le_succ _))] + rw [take_of_length_le this, take_of_length_le (le_trans this (Nat.le_succ _))] have B : ∀ j : Fin n, ((j : ℕ) < i.succ) = ((j : ℕ) < i) := by intro j have : (j : ℕ) < i := lt_of_lt_of_le j.2 (not_lt.mp h) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 57ab5807f612f..750c88af849a9 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -297,9 +297,10 @@ lemma prod_map_eq_pow_single [DecidableEq α] {l : List α} (a : α) (f : α → · rw [map_nil, prod_nil, count_nil, _root_.pow_zero] · specialize h a fun a' ha' hfa' => hf a' ha' (mem_cons_of_mem _ hfa') rw [List.map_cons, List.prod_cons, count_cons, h] + simp only [beq_iff_eq] split_ifs with ha' · rw [ha', _root_.pow_succ'] - · rw [hf a' (Ne.symm ha') (List.mem_cons_self a' as), one_mul, add_zero] + · rw [hf a' ha' (List.mem_cons_self a' as), one_mul, add_zero] @[to_additive] lemma prod_eq_pow_single [DecidableEq M] (a : M) (h : ∀ a', a' ≠ a → a' ∈ l → a' = 1) : @@ -335,7 +336,7 @@ lemma prod_map_erase [DecidableEq α] (f : α → M) {a} : | b :: l, h => by obtain rfl | ⟨ne, h⟩ := List.eq_or_ne_mem_of_mem h · simp only [map, erase_cons_head, prod_cons] - · simp only [map, erase_cons_tail _ (not_beq_of_ne ne.symm), prod_cons, prod_map_erase _ h, + · simp only [map, erase_cons_tail (not_beq_of_ne ne.symm), prod_cons, prod_map_erase _ h, mul_left_comm (f a) (f b)] @[to_additive] lemma Perm.prod_eq (h : Perm l₁ l₂) : prod l₁ = prod l₂ := h.fold_op_eq @@ -463,7 +464,7 @@ theorem prod_set' (L : List G) (n : ℕ) (a : G) : split_ifs with hn · rw [mul_comm _ a, mul_assoc a, prod_drop_succ L n hn, mul_comm _ (drop n L).prod, ← mul_assoc (take n L).prod, prod_take_mul_prod_drop, mul_comm a, mul_assoc] - · simp only [take_all_of_le (le_of_not_lt hn), prod_nil, mul_one, + · simp only [take_of_length_le (le_of_not_lt hn), prod_nil, mul_one, drop_eq_nil_of_le ((le_of_not_lt hn).trans n.le_succ)] @[to_additive] @@ -478,8 +479,8 @@ lemma prod_map_ite_eq {A : Type*} [DecidableEq A] (l : List A) (f g : A → G) ( clear ih by_cases hx : x = a · simp only [hx, ite_true, div_pow, pow_add, pow_one, div_eq_mul_inv, mul_assoc, mul_comm, - mul_left_comm, mul_inv_cancel_left] - · simp only [hx, ite_false, ne_comm.mp hx, add_zero, mul_assoc, mul_comm (g x) _] + mul_left_comm, mul_inv_cancel_left, beq_self_eq_true] + · simp only [hx, ite_false, ne_comm.mp hx, add_zero, mul_assoc, mul_comm (g x) _, beq_iff_eq] end CommGroup @@ -578,13 +579,14 @@ theorem sum_map_count_dedup_filter_eq_countP (p : α → Bool) (l : List α) : match p a with | true => simp only [List.map_cons, List.sum_cons, List.count_eq_zero.2 ha, zero_add] | false => simp only - · by_cases hp : p a - · refine _root_.trans (sum_map_eq_nsmul_single a _ fun _ h _ => by simp [h]) ?_ + · simp only [beq_iff_eq] + by_cases hp : p a + · refine _root_.trans (sum_map_eq_nsmul_single a _ fun _ h _ => by simp [h.symm]) ?_ simp [hp, count_dedup] · refine _root_.trans (List.sum_eq_zero fun n hn => ?_) (by simp [hp]) obtain ⟨a', ha'⟩ := List.mem_map.1 hn split_ifs at ha' with ha - · simp only [ha, mem_filter, mem_dedup, find?, mem_cons, true_or, hp, + · simp only [ha.symm, mem_filter, mem_dedup, find?, mem_cons, true_or, hp, and_false, false_and] at ha' · exact ha'.2.symm @@ -630,10 +632,6 @@ lemma ranges_join (l : List ℕ) : l.ranges.join = range l.sum := by simp [range lemma mem_mem_ranges_iff_lt_sum (l : List ℕ) {n : ℕ} : (∃ s ∈ l.ranges, n ∈ s) ↔ n < l.sum := by simp [mem_mem_ranges_iff_lt_natSum] -@[simp] -theorem length_join (L : List (List α)) : length (join L) = sum (map length L) := by - induction L <;> [rfl; simp only [*, join, map, sum_cons, length_append]] - lemma countP_join (p : α → Bool) : ∀ L : List (List α), countP p L.join = (L.map (countP p)).sum | [] => rfl | a :: l => by rw [join, countP_append, map_cons, sum_cons, countP_join _ l] @@ -643,7 +641,8 @@ lemma count_join [BEq α] (L : List (List α)) (a : α) : L.join.count a = (L.ma @[simp] theorem length_bind (l : List α) (f : α → List β) : - length (List.bind l f) = sum (map (length ∘ f) l) := by rw [List.bind, length_join, map_map] + length (List.bind l f) = sum (map (length ∘ f) l) := by + rw [List.bind, length_join, map_map, Nat.sum_eq_listSum] lemma countP_bind (p : β → Bool) (l : List α) (f : α → List β) : countP p (l.bind f) = sum (map (countP p ∘ f) l) := by rw [List.bind, countP_join, map_map] diff --git a/Mathlib/Algebra/Category/BialgebraCat/Basic.lean b/Mathlib/Algebra/Category/BialgebraCat/Basic.lean index 40a6ab5c49b18..3cf9cd5c6fcbe 100644 --- a/Mathlib/Algebra/Category/BialgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/BialgebraCat/Basic.lean @@ -76,7 +76,7 @@ instance category : Category (BialgebraCat.{v} R) where @[ext] lemma hom_ext {X Y : BialgebraCat.{v} R} (f g : X ⟶ Y) (h : f.toBialgHom = g.toBialgHom) : f = g := - Hom.ext _ _ h + Hom.ext h /-- Typecheck a `BialgHom` as a morphism in `BialgebraCat R`. -/ abbrev ofHom {X Y : Type v} [Ring X] [Ring Y] @@ -145,8 +145,8 @@ variable [Bialgebra R X] [Bialgebra R Y] [Bialgebra R Z] def toBialgebraCatIso (e : X ≃ₐc[R] Y) : BialgebraCat.of R X ≅ BialgebraCat.of R Y where hom := BialgebraCat.ofHom e inv := BialgebraCat.ofHom e.symm - hom_inv_id := Hom.ext _ _ <| DFunLike.ext _ _ e.left_inv - inv_hom_id := Hom.ext _ _ <| DFunLike.ext _ _ e.right_inv + hom_inv_id := Hom.ext <| DFunLike.ext _ _ e.left_inv + inv_hom_id := Hom.ext <| DFunLike.ext _ _ e.right_inv @[simp] theorem toBialgebraCatIso_refl : toBialgebraCatIso (BialgEquiv.refl R X) = .refl _ := rfl diff --git a/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean b/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean index 9591fa4f9270b..2e3227b072f57 100644 --- a/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/CoalgebraCat/Basic.lean @@ -79,7 +79,7 @@ instance category : Category (CoalgebraCat.{v} R) where @[ext] lemma hom_ext {M N : CoalgebraCat.{v} R} (f g : M ⟶ N) (h : f.toCoalgHom = g.toCoalgHom) : f = g := - Hom.ext _ _ h + Hom.ext h /-- Typecheck a `CoalgHom` as a morphism in `CoalgebraCat R`. -/ abbrev ofHom {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] @@ -133,8 +133,8 @@ variable [Coalgebra R X] [Coalgebra R Y] [Coalgebra R Z] def toCoalgebraCatIso (e : X ≃ₗc[R] Y) : CoalgebraCat.of R X ≅ CoalgebraCat.of R Y where hom := CoalgebraCat.ofHom e inv := CoalgebraCat.ofHom e.symm - hom_inv_id := Hom.ext _ _ <| DFunLike.ext _ _ e.left_inv - inv_hom_id := Hom.ext _ _ <| DFunLike.ext _ _ e.right_inv + hom_inv_id := Hom.ext <| DFunLike.ext _ _ e.left_inv + inv_hom_id := Hom.ext <| DFunLike.ext _ _ e.right_inv @[simp] theorem toCoalgebraCatIso_refl : toCoalgebraCatIso (CoalgEquiv.refl R X) = .refl _ := diff --git a/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean b/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean index 3c0086d68fe47..b690a7f5f0bb5 100644 --- a/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean @@ -75,7 +75,7 @@ instance category : Category (HopfAlgebraCat.{v} R) where @[ext] lemma hom_ext {X Y : HopfAlgebraCat.{v} R} (f g : X ⟶ Y) (h : f.toBialgHom = g.toBialgHom) : f = g := - Hom.ext _ _ h + Hom.ext h /-- Typecheck a `BialgHom` as a morphism in `HopfAlgebraCat R`. -/ abbrev ofHom {X Y : Type v} [Ring X] [Ring Y] @@ -129,8 +129,8 @@ variable [HopfAlgebra R X] [HopfAlgebra R Y] [HopfAlgebra R Z] def toHopfAlgebraCatIso (e : X ≃ₐc[R] Y) : HopfAlgebraCat.of R X ≅ HopfAlgebraCat.of R Y where hom := HopfAlgebraCat.ofHom e inv := HopfAlgebraCat.ofHom e.symm - hom_inv_id := Hom.ext _ _ <| DFunLike.ext _ _ e.left_inv - inv_hom_id := Hom.ext _ _ <| DFunLike.ext _ _ e.right_inv + hom_inv_id := Hom.ext <| DFunLike.ext _ _ e.left_inv + inv_hom_id := Hom.ext <| DFunLike.ext _ _ e.right_inv @[simp] theorem toHopfAlgebraCatIso_refl : toHopfAlgebraCatIso (BialgEquiv.refl R X) = .refl _ := diff --git a/Mathlib/Algebra/Category/ModuleCat/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Basic.lean index 767dade2ddf80..48187e0d55415 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Basic.lean @@ -353,10 +353,10 @@ def smulNatTrans : R →+* End (forget₂ (ModuleCat R) AddCommGrp) where toFun r := { app := fun M => M.smul r naturality := fun _ _ _ => smul_naturality _ r } - map_one' := NatTrans.ext _ _ (by aesop_cat) - map_zero' := NatTrans.ext _ _ (by aesop_cat) - map_mul' _ _ := NatTrans.ext _ _ (by aesop_cat) - map_add' _ _ := NatTrans.ext _ _ (by aesop_cat) + map_one' := NatTrans.ext (by aesop_cat) + map_zero' := NatTrans.ext (by aesop_cat) + map_mul' _ _ := NatTrans.ext (by aesop_cat) + map_add' _ _ := NatTrans.ext (by aesop_cat) variable {R} diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean index 37ebf65e896ab..fb6b24bc7f604 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean @@ -55,7 +55,7 @@ instance : Category (SheafOfModules.{v} R) where @[ext] lemma hom_ext {X Y : SheafOfModules.{v} R} {f g : X ⟶ Y} (h : f.val = g.val) : f = g := - Hom.ext _ _ h + Hom.ext h @[simp] lemma id_val (X : SheafOfModules.{v} R) : Hom.val (𝟙 X) = 𝟙 X.val := rfl diff --git a/Mathlib/Algebra/CharP/Defs.lean b/Mathlib/Algebra/CharP/Defs.lean index fafb687f61bd3..b9c236c16d197 100644 --- a/Mathlib/Algebra/CharP/Defs.lean +++ b/Mathlib/Algebra/CharP/Defs.lean @@ -366,7 +366,7 @@ instance Prod.charZero_of_right [CharZero S] : CharZero (R × S) where end Prod instance ULift.charP [AddMonoidWithOne R] (p : ℕ) [CharP R p] : CharP (ULift R) p where - cast_eq_zero_iff' n := Iff.trans (ULift.ext_iff _ _) <| CharP.cast_eq_zero_iff R p n + cast_eq_zero_iff' n := Iff.trans ULift.ext_iff <| CharP.cast_eq_zero_iff R p n instance MulOpposite.charP [AddMonoidWithOne R] (p : ℕ) [CharP R p] : CharP Rᵐᵒᵖ p where cast_eq_zero_iff' n := MulOpposite.unop_inj.symm.trans <| CharP.cast_eq_zero_iff R p n diff --git a/Mathlib/Algebra/CubicDiscriminant.lean b/Mathlib/Algebra/CubicDiscriminant.lean index 0c033b4661a31..b200ea3ac1f44 100644 --- a/Mathlib/Algebra/CubicDiscriminant.lean +++ b/Mathlib/Algebra/CubicDiscriminant.lean @@ -114,7 +114,7 @@ theorem c_of_eq (h : P.toPoly = Q.toPoly) : P.c = Q.c := by rw [← coeff_eq_c, theorem d_of_eq (h : P.toPoly = Q.toPoly) : P.d = Q.d := by rw [← coeff_eq_d, h, coeff_eq_d] theorem toPoly_injective (P Q : Cubic R) : P.toPoly = Q.toPoly ↔ P = Q := - ⟨fun h ↦ Cubic.ext P Q (a_of_eq h) (b_of_eq h) (c_of_eq h) (d_of_eq h), congr_arg toPoly⟩ + ⟨fun h ↦ Cubic.ext (a_of_eq h) (b_of_eq h) (c_of_eq h) (d_of_eq h), congr_arg toPoly⟩ theorem of_a_eq_zero (ha : P.a = 0) : P.toPoly = C P.b * X ^ 2 + C P.c * X + C P.d := by rw [toPoly, ha, C_0, zero_mul, zero_add] diff --git a/Mathlib/Algebra/DirectSum/Module.lean b/Mathlib/Algebra/DirectSum/Module.lean index dd4e97f981a94..8321707132911 100644 --- a/Mathlib/Algebra/DirectSum/Module.lean +++ b/Mathlib/Algebra/DirectSum/Module.lean @@ -178,7 +178,8 @@ variable {ι M} theorem apply_eq_component (f : ⨁ i, M i) (i : ι) : f i = component R ι M i f := rfl -@[ext] +-- Note(kmill): `@[ext]` cannot prove `ext_iff` because `R` is not determined by `f` or `g`. +@[ext (iff := false)] theorem ext {f g : ⨁ i, M i} (h : ∀ i, component R ι M i f = component R ι M i g) : f = g := DFinsupp.ext h diff --git a/Mathlib/Algebra/Free.lean b/Mathlib/Algebra/Free.lean index 8fe99af17d57f..68fa9ba945b69 100644 --- a/Mathlib/Algebra/Free.lean +++ b/Mathlib/Algebra/Free.lean @@ -427,7 +427,7 @@ variable {α : Type u} @[to_additive] instance : Semigroup (FreeSemigroup α) where mul L1 L2 := ⟨L1.1, L1.2 ++ L2.1 :: L2.2⟩ - mul_assoc _L1 _L2 _L3 := FreeSemigroup.ext _ _ rfl <| List.append_assoc _ _ _ + mul_assoc _L1 _L2 _L3 := FreeSemigroup.ext rfl <| List.append_assoc _ _ _ @[to_additive (attr := simp)] theorem head_mul (x y : FreeSemigroup α) : (x * y).1 = x.1 := rfl @@ -632,7 +632,7 @@ end Category @[to_additive] instance [DecidableEq α] : DecidableEq (FreeSemigroup α) := - fun _ _ ↦ decidable_of_iff' _ (FreeSemigroup.ext_iff _ _) + fun _ _ ↦ decidable_of_iff' _ FreeSemigroup.ext_iff end FreeSemigroup diff --git a/Mathlib/Algebra/Group/Action/Defs.lean b/Mathlib/Algebra/Group/Action/Defs.lean index d40bf5a4e7186..113c045227580 100644 --- a/Mathlib/Algebra/Group/Action/Defs.lean +++ b/Mathlib/Algebra/Group/Action/Defs.lean @@ -620,7 +620,7 @@ def monoidHomEquivMulActionIsScalarTower (M N) [Monoid M] [Monoid N] : toFun f := ⟨MulAction.compHom N f, SMul.comp.isScalarTower _⟩ invFun := fun ⟨_, _⟩ ↦ MonoidHom.smulOneHom left_inv f := MonoidHom.ext fun m ↦ mul_one (f m) - right_inv := fun ⟨_, _⟩ ↦ Subtype.ext <| MulAction.ext _ _ <| funext₂ <| smul_one_smul N + right_inv := fun ⟨_, _⟩ ↦ Subtype.ext <| MulAction.ext <| funext₂ <| smul_one_smul N end CompatibleScalar diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 115b71cdc60ad..0a4884f19e160 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -389,10 +389,6 @@ same underlying function. -/ theorem ext {f g : MulEquiv M N} (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h -@[to_additive] -theorem ext_iff {f g : MulEquiv M N} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - @[to_additive (attr := simp)] theorem mk_coe (e : M ≃* N) (e' h₁ h₂ h₃) : (⟨⟨e, e', h₁, h₂⟩, h₃⟩ : M ≃* N) = e := ext fun _ => rfl diff --git a/Mathlib/Algebra/Group/Units.lean b/Mathlib/Algebra/Group/Units.lean index 97c59ccbf7d9c..3286fc67b219d 100644 --- a/Mathlib/Algebra/Group/Units.lean +++ b/Mathlib/Algebra/Group/Units.lean @@ -133,10 +133,6 @@ theorem ext : Function.Injective (val : αˣ → α) theorem eq_iff {a b : αˣ} : (a : α) = b ↔ a = b := ext.eq_iff -@[to_additive] -protected theorem ext_iff {a b : αˣ} : a = b ↔ (a : α) = b := - eq_iff.symm - /-- Units have decidable equality if the base `Monoid` has decidable equality. -/ @[to_additive "Additive units have decidable equality if the base `AddMonoid` has deciable equality."] diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index 2dda38c3fdafd..57475724a359a 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -10,6 +10,7 @@ import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Nontriviality import Mathlib.Tactic.Spread import Mathlib.Util.AssertExists +import Mathlib.Data.Subtype /-! # Lemmas about units in a `MonoidWithZero` or a `GroupWithZero`. diff --git a/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean b/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean index 6d7136744887d..6245ca808b6f0 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean @@ -129,9 +129,6 @@ lemma comp_hom {a b : ℕ} (α : Ext X Y a) (β : Ext Y Z b) {c : ℕ} (h : a + lemma ext {n : ℕ} {α β : Ext X Y n} (h : α.hom = β.hom) : α = β := homEquiv.injective h -protected lemma ext_iff {n : ℕ} {α β : Ext X Y n} : α = β ↔ α.hom = β.hom := - ⟨fun h ↦ by rw [h], ext⟩ - end /-- The canonical map `(X ⟶ Y) → Ext X Y 0`. -/ diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean index 0afeefd8f4fbe..b00732bca4501 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean @@ -604,9 +604,6 @@ instance : Coe (Cocycle F G n) (Cochain F G n) where lemma ext (z₁ z₂ : Cocycle F G n) (h : (z₁ : Cochain F G n) = z₂) : z₁ = z₂ := Subtype.ext h -protected lemma ext_iff (z₁ z₂ : Cocycle F G n) : z₁ = z₂ ↔ (z₁ : Cochain F G n) = z₂ := - Subtype.ext_iff - instance : SMul R (Cocycle F G n) where smul r z := ⟨r • z.1, by have hz := z.2 diff --git a/Mathlib/Algebra/Homology/ShortComplex/Basic.lean b/Mathlib/Algebra/Homology/ShortComplex/Basic.lean index 7ef3f0fb65219..3d8810eab8fcd 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Basic.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Basic.lean @@ -87,7 +87,7 @@ instance : Category (ShortComplex C) where @[ext] lemma hom_ext (f g : S₁ ⟶ S₂) (h₁ : f.τ₁ = g.τ₁) (h₂ : f.τ₂ = g.τ₂) (h₃ : f.τ₃ = g.τ₃) : f = g := - Hom.ext _ _ h₁ h₂ h₃ + Hom.ext h₁ h₂ h₃ /-- A constructor for morphisms in `ShortComplex C` when the commutativity conditions are not obvious. -/ diff --git a/Mathlib/Algebra/Lie/Basic.lean b/Mathlib/Algebra/Lie/Basic.lean index 298909452954e..c8a7cd3805a7b 100644 --- a/Mathlib/Algebra/Lie/Basic.lean +++ b/Mathlib/Algebra/Lie/Basic.lean @@ -352,11 +352,6 @@ theorem coe_injective : @Function.Injective (L₁ →ₗ⁅R⁆ L₂) (L₁ → theorem ext {f g : L₁ →ₗ⁅R⁆ L₂} (h : ∀ x, f x = g x) : f = g := coe_injective <| funext h -theorem ext_iff {f g : L₁ →ₗ⁅R⁆ L₂} : f = g ↔ ∀ x, f x = g x := - ⟨by - rintro rfl x - rfl, ext⟩ - theorem congr_fun {f g : L₁ →ₗ⁅R⁆ L₂} (h : f = g) (x : L₁) : f x = g x := h ▸ rfl @@ -707,11 +702,6 @@ theorem coe_injective : @Function.Injective (M →ₗ⁅R,L⁆ N) (M → N) (↑ theorem ext {f g : M →ₗ⁅R,L⁆ N} (h : ∀ m, f m = g m) : f = g := coe_injective <| funext h -theorem ext_iff {f g : M →ₗ⁅R,L⁆ N} : f = g ↔ ∀ m, f m = g m := - ⟨by - rintro rfl m - rfl, ext⟩ - theorem congr_fun {f g : M →ₗ⁅R,L⁆ N} (h : f = g) (x : M) : f x = g x := h ▸ rfl diff --git a/Mathlib/Algebra/Lie/DirectSum.lean b/Mathlib/Algebra/Lie/DirectSum.lean index b7f6f7dc4c37b..7d4a6efc56417 100644 --- a/Mathlib/Algebra/Lie/DirectSum.lean +++ b/Mathlib/Algebra/Lie/DirectSum.lean @@ -173,7 +173,8 @@ def lieAlgebraComponent (j : ι) : (⨁ i, L i) →ₗ⁅R⁆ L j := toFun := component R ι L j map_lie' := fun {x y} => by simp [component, lapply] } -@[ext] +-- Note(kmill): `ext` cannot generate an iff theorem here since `x` and `y` do not determine `R`. +@[ext (iff := false)] theorem lieAlgebra_ext {x y : ⨁ i, L i} (h : ∀ i, lieAlgebraComponent R ι L i x = lieAlgebraComponent R ι L i y) : x = y := DFinsupp.ext h diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index 84032cda8f80f..cfab261a0e7f0 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -208,7 +208,7 @@ variable {M} @[ext] lemma ext {χ₁ χ₂ : Weight R L M} (h : ∀ x, χ₁ x = χ₂ x) : χ₁ = χ₂ := by cases' χ₁ with f₁ _; cases' χ₂ with f₂ _; aesop -protected lemma ext_iff {χ₁ χ₂ : Weight R L M} : χ₁ = χ₂ ↔ (χ₁ : L → R) = χ₂ := by aesop +lemma ext_iff' {χ₁ χ₂ : Weight R L M} : (χ₁ : L → R) = χ₂ ↔ χ₁ = χ₂ := by aesop lemma exists_ne_zero (χ : Weight R L M) : ∃ x ∈ weightSpace M χ, x ≠ 0 := by @@ -236,7 +236,7 @@ def IsZero (χ : Weight R L M) := (χ : L → R) = 0 @[simp] lemma coe_eq_zero_iff (χ : Weight R L M) : (χ : L → R) = 0 ↔ χ.IsZero := Iff.rfl lemma isZero_iff_eq_zero [Nontrivial (weightSpace M (0 : L → R))] {χ : Weight R L M} : - χ.IsZero ↔ χ = 0 := Weight.ext_iff (χ₂ := 0).symm + χ.IsZero ↔ χ = 0 := Weight.ext_iff' (χ₂ := 0) lemma isZero_zero [Nontrivial (weightSpace M (0 : L → R))] : IsZero (0 : Weight R L M) := rfl diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index f1acf37008a84..22eea21dc9935 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -315,7 +315,7 @@ def ringHomEquivModuleIsScalarTower [Semiring R] [Semiring S] : toFun f := ⟨Module.compHom S f, SMul.comp.isScalarTower _⟩ invFun := fun ⟨_, _⟩ ↦ RingHom.smulOneHom left_inv f := RingHom.ext fun r ↦ mul_one (f r) - right_inv := fun ⟨_, _⟩ ↦ Subtype.ext <| Module.ext _ _ <| funext₂ <| smul_one_smul S + right_inv := fun ⟨_, _⟩ ↦ Subtype.ext <| Module.ext <| funext₂ <| smul_one_smul S section AddCommMonoid diff --git a/Mathlib/Algebra/Module/Equiv/Defs.lean b/Mathlib/Algebra/Module/Equiv/Defs.lean index 28a26d6e92b83..1513cc839f617 100644 --- a/Mathlib/Algebra/Module/Equiv/Defs.lean +++ b/Mathlib/Algebra/Module/Equiv/Defs.lean @@ -227,9 +227,6 @@ variable {e e'} theorem ext (h : ∀ x, e x = e' x) : e = e' := DFunLike.ext _ _ h -theorem ext_iff : e = e' ↔ ∀ x, e x = e' x := - DFunLike.ext_iff - protected theorem congr_arg {x x'} : x = x' → e x = e x' := DFunLike.congr_arg e diff --git a/Mathlib/Algebra/Module/Injective.lean b/Mathlib/Algebra/Module/Injective.lean index f734312c3e26c..fd2e6fec70ebc 100644 --- a/Mathlib/Algebra/Module/Injective.lean +++ b/Mathlib/Algebra/Module/Injective.lean @@ -98,7 +98,7 @@ section Ext variable {i f} -@[ext] +@[ext (iff := false)] theorem ExtensionOf.ext {a b : ExtensionOf i f} (domain_eq : a.domain = b.domain) (to_fun_eq : ∀ ⦃x : a.domain⦄ ⦃y : b.domain⦄, (x : N) = y → a.toLinearPMap x = b.toLinearPMap y) : diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index 86a600850421c..1de087baccedf 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -322,9 +322,6 @@ protected theorem congr_arg {x x' : M} : x = x' → f x = f x' := protected theorem congr_fun (h : f = g) (x : M) : f x = g x := DFunLike.congr_fun h x -protected theorem ext_iff : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - @[simp] theorem mk_coe (f : M →ₛₗ[σ] M₃) (h) : (LinearMap.mk f h : M →ₛₗ[σ] M₃) = f := ext fun _ ↦ rfl @@ -456,9 +453,6 @@ theorem toAddMonoidHom_injective : theorem ext_ring {f g : R →ₛₗ[σ] M₃} (h : f 1 = g 1) : f = g := ext fun x ↦ by rw [← mul_one x, ← smul_eq_mul, f.map_smulₛₗ, g.map_smulₛₗ, h] -theorem ext_ring_iff {σ : R →+* R} {f g : R →ₛₗ[σ] M} : f = g ↔ f 1 = g 1 := - ⟨fun h ↦ h ▸ rfl, ext_ring⟩ - @[ext high] theorem ext_ring_op {σ : Rᵐᵒᵖ →+* S} {f g : R →ₛₗ[σ] M₃} (h : f (1 : R) = g (1 : R)) : f = g := diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index 9ff01ac8654c1..90a2b3d0d0db8 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -529,9 +529,6 @@ theorem support_mul [DecidableEq σ] (p q : MvPolynomial σ R) : theorem ext (p q : MvPolynomial σ R) : (∀ m, coeff m p = coeff m q) → p = q := Finsupp.ext -protected theorem ext_iff (p q : MvPolynomial σ R) : p = q ↔ ∀ m, coeff m p = coeff m q := - ⟨fun h m => by rw [h], ext p q⟩ - @[simp] theorem coeff_add (m : σ →₀ ℕ) (p q : MvPolynomial σ R) : coeff m (p + q) = coeff m p + coeff m q := add_apply p q m diff --git a/Mathlib/Algebra/Order/BigOperators/Group/List.lean b/Mathlib/Algebra/Order/BigOperators/Group/List.lean index 920c30bdab385..7872ed429ca63 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/List.lean @@ -172,7 +172,7 @@ variable [CanonicallyOrderedCommMonoid M] {l : List M} cases' lt_or_le n L.length with h h · rw [prod_take_succ _ _ h] exact le_self_mul - · simp [take_all_of_le h, take_all_of_le (le_trans h (Nat.le_succ _))] + · simp [take_of_length_le h, take_of_length_le (le_trans h (Nat.le_succ _))] end CanonicallyOrderedCommMonoid end List diff --git a/Mathlib/Algebra/Order/Interval/Basic.lean b/Mathlib/Algebra/Order/Interval/Basic.lean index c048f5474096b..2103932e456c4 100644 --- a/Mathlib/Algebra/Order/Interval/Basic.lean +++ b/Mathlib/Algebra/Order/Interval/Basic.lean @@ -453,11 +453,11 @@ instance subtractionCommMonoid {α : Type u} [OrderedAddCommGroup α] : neg := Neg.neg sub := Sub.sub sub_eq_add_neg := fun s t => by - refine NonemptyInterval.ext _ _ (Prod.ext ?_ ?_) <;> + refine NonemptyInterval.ext (Prod.ext ?_ ?_) <;> exact sub_eq_add_neg _ _ neg_neg := fun s => by apply NonemptyInterval.ext; exact neg_neg _ neg_add_rev := fun s t => by - refine NonemptyInterval.ext _ _ (Prod.ext ?_ ?_) <;> + refine NonemptyInterval.ext (Prod.ext ?_ ?_) <;> exact neg_add_rev _ _ neg_eq_of_add := fun s t h => by obtain ⟨a, b, rfl, rfl, hab⟩ := NonemptyInterval.add_eq_zero_iff.1 h @@ -471,11 +471,11 @@ instance divisionCommMonoid : DivisionCommMonoid (NonemptyInterval α) := inv := Inv.inv div := (· / ·) div_eq_mul_inv := fun s t => by - refine NonemptyInterval.ext _ _ (Prod.ext ?_ ?_) <;> + refine NonemptyInterval.ext (Prod.ext ?_ ?_) <;> exact div_eq_mul_inv _ _ inv_inv := fun s => by apply NonemptyInterval.ext; exact inv_inv _ mul_inv_rev := fun s t => by - refine NonemptyInterval.ext _ _ (Prod.ext ?_ ?_) <;> + refine NonemptyInterval.ext (Prod.ext ?_ ?_) <;> exact mul_inv_rev _ _ inv_eq_of_mul := fun s t h => by obtain ⟨a, b, rfl, rfl, hab⟩ := NonemptyInterval.mul_eq_one_iff.1 h diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean index 983d0219d8a27..53a7cbf0f18a3 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Order.Monoid.Unbundled.Defs import Mathlib.Data.Ordering.Basic import Mathlib.Order.MinMax import Mathlib.Tactic.Contrapose +import Mathlib.Tactic.Use /-! # Ordered monoids diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index a0479723d185f..df41bba69c9e6 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -188,7 +188,7 @@ instance : Add ℍ[R,c₁,c₂] := @[simp] theorem add_imK : (a + b).imK = a.imK + b.imK := rfl @[simp] theorem add_im : (a + b).im = a.im + b.im := - QuaternionAlgebra.ext _ _ (zero_add _).symm rfl rfl rfl + QuaternionAlgebra.ext (zero_add _).symm rfl rfl rfl @[simp] theorem mk_add_mk (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) : @@ -210,7 +210,7 @@ instance : Neg ℍ[R,c₁,c₂] := ⟨fun a => ⟨-a.1, -a.2, -a.3, -a.4⟩⟩ @[simp] theorem neg_imK : (-a).imK = -a.imK := rfl @[simp] theorem neg_im : (-a).im = -a.im := - QuaternionAlgebra.ext _ _ neg_zero.symm rfl rfl rfl + QuaternionAlgebra.ext neg_zero.symm rfl rfl rfl @[simp] theorem neg_mk (a₁ a₂ a₃ a₄ : R) : -(mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) = ⟨-a₁, -a₂, -a₃, -a₄⟩ := @@ -231,7 +231,7 @@ instance : Sub ℍ[R,c₁,c₂] := @[simp] theorem sub_imK : (a - b).imK = a.imK - b.imK := rfl @[simp] theorem sub_im : (a - b).im = a.im - b.im := - QuaternionAlgebra.ext _ _ (sub_zero _).symm rfl rfl rfl + QuaternionAlgebra.ext (sub_zero _).symm rfl rfl rfl @[simp] theorem mk_sub_mk (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) : @@ -244,15 +244,15 @@ theorem coe_im : (x : ℍ[R,c₁,c₂]).im = 0 := @[simp] theorem re_add_im : ↑a.re + a.im = a := - QuaternionAlgebra.ext _ _ (add_zero _) (zero_add _) (zero_add _) (zero_add _) + QuaternionAlgebra.ext (add_zero _) (zero_add _) (zero_add _) (zero_add _) @[simp] theorem sub_self_im : a - a.im = a.re := - QuaternionAlgebra.ext _ _ (sub_zero _) (sub_self _) (sub_self _) (sub_self _) + QuaternionAlgebra.ext (sub_zero _) (sub_self _) (sub_self _) (sub_self _) @[simp] theorem sub_self_re : a - a.re = a.im := - QuaternionAlgebra.ext _ _ (sub_self _) (sub_zero _) (sub_zero _) (sub_zero _) + QuaternionAlgebra.ext (sub_self _) (sub_zero _) (sub_zero _) (sub_zero _) /-- Multiplication is given by @@ -312,7 +312,7 @@ instance [SMulCommClass S T R] : SMulCommClass S T ℍ[R,c₁,c₂] where @[simp] theorem smul_imK : (s • a).imK = s • a.imK := rfl @[simp] theorem smul_im {S} [SMulZeroClass S R] (s : S) : (s • a).im = s • a.im := - QuaternionAlgebra.ext _ _ (smul_zero s).symm rfl rfl rfl + QuaternionAlgebra.ext (smul_zero s).symm rfl rfl rfl @[simp] theorem smul_mk (re im_i im_j im_k : R) : @@ -324,7 +324,7 @@ end @[simp, norm_cast] theorem coe_smul [SMulZeroClass S R] (s : S) (r : R) : (↑(s • r) : ℍ[R,c₁,c₂]) = s • (r : ℍ[R,c₁,c₂]) := - QuaternionAlgebra.ext _ _ rfl (smul_zero s).symm (smul_zero s).symm (smul_zero s).symm + QuaternionAlgebra.ext rfl (smul_zero s).symm (smul_zero s).symm (smul_zero s).symm instance : AddCommGroup ℍ[R,c₁,c₂] := (equivProd c₁ c₂).injective.addCommGroup _ rfl (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) @@ -591,7 +591,7 @@ theorem imK_star : (star a).imK = -a.imK := @[simp] theorem im_star : (star a).im = -a.im := - QuaternionAlgebra.ext _ _ neg_zero.symm rfl rfl rfl + QuaternionAlgebra.ext neg_zero.symm rfl rfl rfl @[simp] theorem star_mk (a₁ a₂ a₃ a₄ : R) : star (mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) = ⟨a₁, -a₂, -a₃, -a₄⟩ := @@ -626,7 +626,7 @@ theorem star_coe : star (x : ℍ[R,c₁,c₂]) = x := by ext <;> simp @[simp] theorem star_smul [Monoid S] [DistribMulAction S R] (s : S) (a : ℍ[R,c₁,c₂]) : star (s • a) = s • star a := - QuaternionAlgebra.ext _ _ rfl (smul_neg _ _).symm (smul_neg _ _).symm (smul_neg _ _).symm + QuaternionAlgebra.ext rfl (smul_neg _ _).symm (smul_neg _ _).symm (smul_neg _ _).symm theorem eq_re_of_eq_coe {a : ℍ[R,c₁,c₂]} {x : R} (h : a = x) : a = a.re := by rw [h, coe_re] @@ -730,11 +730,7 @@ instance : IsStarNormal a := inferInstanceAs <| IsStarNormal (R := ℍ[R,-1,-1]) @[ext] theorem ext : a.re = b.re → a.imI = b.imI → a.imJ = b.imJ → a.imK = b.imK → a = b := - QuaternionAlgebra.ext a b - -theorem ext_iff {a b : ℍ[R]} : - a = b ↔ a.re = b.re ∧ a.imI = b.imI ∧ a.imJ = b.imJ ∧ a.imK = b.imK := - QuaternionAlgebra.ext_iff a b + QuaternionAlgebra.ext /-- The imaginary part of a quaternion. -/ nonrec def im (x : ℍ[R]) : ℍ[R] := x.im diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index 08ccaf24f2d1a..92eabfb25d020 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -177,9 +177,6 @@ protected theorem congr_arg {f : R ≃+* S} {x x' : R} : x = x' → f x = f x' : protected theorem congr_fun {f g : R ≃+* S} (h : f = g) (x : R) : f x = g x := DFunLike.congr_fun h x -protected theorem ext_iff {f g : R ≃+* S} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - @[simp] theorem toAddEquiv_eq_coe (f : R ≃+* S) : f.toAddEquiv = ↑f := rfl diff --git a/Mathlib/Algebra/Ring/Ext.lean b/Mathlib/Algebra/Ring/Ext.lean index 7186891d5c6c5..3d7a340d5bff8 100644 --- a/Mathlib/Algebra/Ring/Ext.lean +++ b/Mathlib/Algebra/Ring/Ext.lean @@ -49,12 +49,6 @@ namespace Distrib -- Prove equality of parts using function extensionality. congr -theorem ext_iff {inst₁ inst₂ : Distrib R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end Distrib /-! ### NonUnitalNonAssocSemiring -/ @@ -76,12 +70,6 @@ theorem toDistrib_injective : Function.Injective (@toDistrib R) := by · exact congrArg (·.toAdd.add x y) h · exact congrArg (·.toMul.mul x y) h -theorem ext_iff {inst₁ inst₂ : NonUnitalNonAssocSemiring R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end NonUnitalNonAssocSemiring /-! ### NonUnitalSemiring -/ @@ -98,12 +86,6 @@ theorem toNonUnitalNonAssocSemiring_injective : toNonUnitalNonAssocSemiring_injective <| NonUnitalNonAssocSemiring.ext h_add h_mul -theorem ext_iff {inst₁ inst₂ : NonUnitalSemiring R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end NonUnitalSemiring /-! ### NonAssocSemiring and its ancestors @@ -173,12 +155,6 @@ theorem toNonUnitalNonAssocSemiring_injective : intro _ _ _ ext <;> congr -theorem ext_iff {inst₁ inst₂ : NonAssocSemiring R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end NonAssocSemiring /-! ### NonUnitalNonAssocRing -/ @@ -200,12 +176,6 @@ theorem toNonUnitalNonAssocSemiring_injective : · exact congrArg (·.toAdd.add x y) h · exact congrArg (·.toMul.mul x y) h -theorem ext_iff {inst₁ inst₂ : NonUnitalNonAssocRing R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end NonUnitalNonAssocRing /-! ### NonUnitalRing -/ @@ -233,12 +203,6 @@ theorem toNonUnitalNonAssocring_injective : intro _ _ _ ext <;> congr -theorem ext_iff {inst₁ inst₂ : NonUnitalRing R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end NonUnitalRing /-! ### NonAssocRing and its ancestors @@ -303,12 +267,6 @@ theorem toNonUnitalNonAssocring_injective : intro _ _ _ ext <;> congr -theorem ext_iff {inst₁ inst₂ : NonAssocRing R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end NonAssocRing /-! ### Semiring -/ @@ -343,12 +301,6 @@ theorem toNonAssocSemiring_injective : · exact congrArg (·.toAdd.add x y) h · exact congrArg (·.toMul.mul x y) h -theorem ext_iff {inst₁ inst₂ : Semiring R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end Semiring /-! ### Ring -/ @@ -391,12 +343,6 @@ theorem toSemiring_injective : · exact congrArg (·.toAdd.add x y) h · exact congrArg (·.toMul.mul x y) h -theorem ext_iff {inst₁ inst₂ : Ring R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext ·)⟩ - end Ring /-! ### NonUnitalNonAssocCommSemiring -/ @@ -413,12 +359,6 @@ theorem toNonUnitalNonAssocSemiring_injective : toNonUnitalNonAssocSemiring_injective <| NonUnitalNonAssocSemiring.ext h_add h_mul -theorem ext_iff {inst₁ inst₂ : NonUnitalNonAssocCommSemiring R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end NonUnitalNonAssocCommSemiring /-! ### NonUnitalCommSemiring -/ @@ -435,12 +375,6 @@ theorem toNonUnitalSemiring_injective : toNonUnitalSemiring_injective <| NonUnitalSemiring.ext h_add h_mul -theorem ext_iff {inst₁ inst₂ : NonUnitalCommSemiring R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end NonUnitalCommSemiring -- At present, there is no `NonAssocCommSemiring` in Mathlib. @@ -459,12 +393,6 @@ theorem toNonUnitalNonAssocRing_injective : toNonUnitalNonAssocRing_injective <| NonUnitalNonAssocRing.ext h_add h_mul -theorem ext_iff {inst₁ inst₂ : NonUnitalNonAssocCommRing R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end NonUnitalNonAssocCommRing /-! ### NonUnitalCommRing -/ @@ -481,12 +409,6 @@ theorem toNonUnitalRing_injective : toNonUnitalRing_injective <| NonUnitalRing.ext h_add h_mul -theorem ext_iff {inst₁ inst₂ : NonUnitalCommRing R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end NonUnitalCommRing -- At present, there is no `NonAssocCommRing` in Mathlib. @@ -505,12 +427,6 @@ theorem toSemiring_injective : toSemiring_injective <| Semiring.ext h_add h_mul -theorem ext_iff {inst₁ inst₂ : CommSemiring R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext · ·)⟩ - end CommSemiring /-! ### CommRing -/ @@ -525,10 +441,4 @@ theorem toRing_injective : Function.Injective (@toRing R) := by inst₁ = inst₂ := toRing_injective <| Ring.ext h_add h_mul -theorem ext_iff {inst₁ inst₂ : CommRing R} : - inst₁ = inst₂ ↔ - (local_hAdd[R, inst₁] = local_hAdd[R, inst₂]) ∧ - (local_hMul[R, inst₁] = local_hMul[R, inst₂]) := - ⟨by rintro rfl; constructor <;> rfl, And.elim (ext ·)⟩ - end CommRing diff --git a/Mathlib/Algebra/Ring/Hom/Defs.lean b/Mathlib/Algebra/Ring/Hom/Defs.lean index 6fa837a5ae1bd..a440563a5296e 100644 --- a/Mathlib/Algebra/Ring/Hom/Defs.lean +++ b/Mathlib/Algebra/Ring/Hom/Defs.lean @@ -155,9 +155,6 @@ variable (f : α →ₙ+* β) {x y : α} theorem ext ⦃f g : α →ₙ+* β⦄ : (∀ x, f x = g x) → f = g := DFunLike.ext _ _ -protected theorem ext_iff {f g : α →ₙ+* β} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - @[simp] theorem mk_coe (f : α →ₙ+* β) (h₁ h₂ h₃) : NonUnitalRingHom.mk (MulHom.mk f h₁) h₂ h₃ = f := ext fun _ => rfl @@ -455,9 +452,6 @@ theorem coe_inj ⦃f g : α →+* β⦄ (h : (f : α → β) = g) : f = g := theorem ext ⦃f g : α →+* β⦄ : (∀ x, f x = g x) → f = g := DFunLike.ext _ _ -protected theorem ext_iff {f g : α →+* β} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - @[simp] theorem mk_coe (f : α →+* β) (h₁ h₂ h₃ h₄) : RingHom.mk ⟨⟨f, h₁⟩, h₂⟩ h₃ h₄ = f := ext fun _ => rfl diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index 1dca6c73aa202..c4878d4541ca9 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -850,7 +850,7 @@ theorem exists_list_of_mem_closure {s : Set R} {x : R} (h : x ∈ closure s) : ⟨l ++ m, fun t ht => (List.mem_append.1 ht).elim (hl1 t) (hm1 t), by simp [hl2, hm2]⟩) fun x ⟨L, hL⟩ => ⟨L.map (List.cons (-1)), - List.forall_mem_map_iff.2 fun j hj => List.forall_mem_cons.2 ⟨Or.inr rfl, hL.1 j hj⟩, + List.forall_mem_map.2 fun j hj => List.forall_mem_cons.2 ⟨Or.inr rfl, hL.1 j hj⟩, hL.2 ▸ List.recOn L (by simp) (by simp (config := { contextual := true }) [List.map_cons, add_comm])⟩ diff --git a/Mathlib/Algebra/Star/StarAlgHom.lean b/Mathlib/Algebra/Star/StarAlgHom.lean index e641a3f99fec4..b8748b3c4d8b7 100644 --- a/Mathlib/Algebra/Star/StarAlgHom.lean +++ b/Mathlib/Algebra/Star/StarAlgHom.lean @@ -765,9 +765,6 @@ theorem toRingEquiv_eq_coe (e : A ≃⋆ₐ[R] B) : e.toRingEquiv = e := theorem ext {f g : A ≃⋆ₐ[R] B} (h : ∀ a, f a = g a) : f = g := DFunLike.ext f g h -theorem ext_iff {f g : A ≃⋆ₐ[R] B} : f = g ↔ ∀ a, f a = g a := - DFunLike.ext_iff - /-- The identity map is a star algebra isomorphism. -/ @[refl] def refl : A ≃⋆ₐ[R] A := diff --git a/Mathlib/Algebra/Star/StarRingHom.lean b/Mathlib/Algebra/Star/StarRingHom.lean index cf7dd54269046..cdf64f59320dc 100644 --- a/Mathlib/Algebra/Star/StarRingHom.lean +++ b/Mathlib/Algebra/Star/StarRingHom.lean @@ -319,9 +319,6 @@ theorem toRingEquiv_eq_coe (e : A ≃⋆+* B) : e.toRingEquiv = e := theorem ext {f g : A ≃⋆+* B} (h : ∀ a, f a = g a) : f = g := DFunLike.ext f g h -theorem ext_iff {f g : A ≃⋆+* B} : f = g ↔ ∀ a, f a = g a := - DFunLike.ext_iff - /-- The identity map as a star ring isomorphism. -/ @[refl] def refl : A ≃⋆+* A := diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index 7034f5fe1de27..903e50c8aefa6 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -58,8 +58,8 @@ def Scheme.emptyTo (X : Scheme.{u}) : ∅ ⟶ X := @[ext] theorem Scheme.empty_ext {X : Scheme.{u}} (f g : ∅ ⟶ X) : f = g := -- Porting note (#11041): `ext` regression - LocallyRingedSpace.Hom.ext _ _ <| PresheafedSpace.ext _ _ (by ext a; exact PEmpty.elim a) <| - NatTrans.ext _ _ <| funext fun a => by aesop_cat + LocallyRingedSpace.Hom.ext <| PresheafedSpace.ext _ _ (by ext a; exact PEmpty.elim a) <| + NatTrans.ext <| funext fun a => by aesop_cat theorem Scheme.eq_emptyTo {X : Scheme.{u}} (f : ∅ ⟶ X) : f = Scheme.emptyTo X := Scheme.empty_ext f (Scheme.emptyTo X) diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 07a434ddae9cb..b0d4fb892fa7a 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -99,7 +99,7 @@ theorem isClosed_singleton_iff_isMaximal (x : PrimeSpectrum R) : constructor <;> intro H · rcases x.asIdeal.exists_le_maximal x.2.1 with ⟨m, hm, hxm⟩ exact (congr_arg asIdeal (@H ⟨m, hm.isPrime⟩ hxm)) ▸ hm - · exact fun p hp ↦ PrimeSpectrum.ext _ _ (H.eq_of_le p.2.1 hp).symm + · exact fun p hp ↦ PrimeSpectrum.ext (H.eq_of_le p.2.1 hp).symm theorem isRadical_vanishingIdeal (s : Set (PrimeSpectrum R)) : (vanishingIdeal s).IsRadical := by rw [← vanishingIdeal_closure, ← zeroLocus_vanishingIdeal_eq_closure, @@ -248,9 +248,8 @@ theorem preimage_comap_zeroLocus (s : Set R) : comap f ⁻¹' zeroLocus s = zero theorem comap_injective_of_surjective (f : R →+* S) (hf : Function.Surjective f) : Function.Injective (comap f) := fun x y h => - PrimeSpectrum.ext _ _ - (Ideal.comap_injective_of_surjective f hf - (congr_arg PrimeSpectrum.asIdeal h : (comap f x).asIdeal = (comap f y).asIdeal)) + PrimeSpectrum.ext (Ideal.comap_injective_of_surjective f hf + (congr_arg PrimeSpectrum.asIdeal h : (comap f x).asIdeal = (comap f y).asIdeal)) variable (S) @@ -590,7 +589,7 @@ def primeSpectrum_unique_of_localization_at_minimal (h : I ∈ minimalPrimes R) default := ⟨LocalRing.maximalIdeal (Localization I.primeCompl), (LocalRing.maximalIdeal.isMaximal _).isPrime⟩ - uniq x := PrimeSpectrum.ext _ _ (Localization.AtPrime.prime_unique_of_minimal h x.asIdeal) + uniq x := PrimeSpectrum.ext (Localization.AtPrime.prime_unique_of_minimal h x.asIdeal) end LocalizationAtMinimal diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 8fd09f5b58301..3ed524534c561 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -497,7 +497,7 @@ section fromSpecToSpec lemma fromSpec_toSpec {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) (x : Proj.T| pbo f) : FromSpec.toFun f_deg hm (toSpec 𝒜 f x) = x := by - refine Subtype.ext <| ProjectiveSpectrum.ext _ _ <| HomogeneousIdeal.ext' ?_ + refine Subtype.ext <| ProjectiveSpectrum.ext <| HomogeneousIdeal.ext' ?_ intros i z hzi refine (FromSpec.mem_carrier_iff_of_mem f_deg hm _ _ hzi).trans ?_ exact (ToSpec.mk_mem_carrier _ _).trans (x.1.2.pow_mem_iff_mem m hm) @@ -643,7 +643,7 @@ lemma toSpec_base_apply_eq_comap {f} (x : Proj| pbo f) : lemma toSpec_base_apply_eq {f} (x : Proj| pbo f) : (toSpec 𝒜 f).1.base x = ProjIsoSpecTopComponent.toSpec 𝒜 f x := - toSpec_base_apply_eq_comap 𝒜 x |>.trans <| PrimeSpectrum.ext _ _ <| Ideal.ext fun z => + toSpec_base_apply_eq_comap 𝒜 x |>.trans <| PrimeSpectrum.ext <| Ideal.ext fun z => show ¬ IsUnit _ ↔ z ∈ ProjIsoSpecTopComponent.ToSpec.carrier _ by obtain ⟨z, rfl⟩ := z.mk_surjective rw [← HomogeneousLocalization.isUnit_iff_isUnit_val, diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index 3fc4941e1c97a..93d4ae918f332 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -152,11 +152,11 @@ to `X.presheaf.stalk x` for any `x : X`. -/ def stalkMap (x : X) : Y.presheaf.stalk (f.val.base x) ⟶ X.presheaf.stalk x := f.val.stalkMap x -@[ext] +@[ext (iff := false)] protected lemma ext {f g : X ⟶ Y} (h_base : f.1.base = g.1.base) (h_app : ∀ U, f.app U ≫ X.presheaf.map (eqToHom congr((Opens.map $h_base.symm).obj U)).op = g.app U) : f = g := - LocallyRingedSpace.Hom.ext _ _ <| SheafedSpace.ext h_base + LocallyRingedSpace.Hom.ext <| SheafedSpace.ext _ _ h_base (TopCat.Presheaf.ext fun U ↦ by simpa using h_app U) lemma preimage_iSup {ι} (U : ι → Opens Y) : f ⁻¹ᵁ iSup U = ⨆ i, f ⁻¹ᵁ U i := diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index 78f816940acb1..a18add389dda1 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -247,13 +247,13 @@ def Spec.locallyRingedSpaceMap {R S : CommRingCat.{u}} (f : R ⟶ S) : @[simp] theorem Spec.locallyRingedSpaceMap_id (R : CommRingCat.{u}) : Spec.locallyRingedSpaceMap (𝟙 R) = 𝟙 (Spec.locallyRingedSpaceObj R) := - LocallyRingedSpace.Hom.ext _ _ <| by + LocallyRingedSpace.Hom.ext <| by rw [Spec.locallyRingedSpaceMap_val, Spec.sheafedSpaceMap_id]; rfl theorem Spec.locallyRingedSpaceMap_comp {R S T : CommRingCat.{u}} (f : R ⟶ S) (g : S ⟶ T) : Spec.locallyRingedSpaceMap (f ≫ g) = Spec.locallyRingedSpaceMap g ≫ Spec.locallyRingedSpaceMap f := - LocallyRingedSpace.Hom.ext _ _ <| by + LocallyRingedSpace.Hom.ext <| by rw [Spec.locallyRingedSpaceMap_val, Spec.sheafedSpaceMap_comp]; rfl /-- Spec, as a contravariant functor from commutative rings to locally ringed spaces. diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean index 85ef266e17287..3745a6fd9c752 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean @@ -74,11 +74,11 @@ section Casts /-- Abbreviation for `eqToHom` that accepts points in a topological space -/ abbrev hcast {X : TopCat} {x₀ x₁ : X} (hx : x₀ = x₁) : fromTop x₀ ⟶ fromTop x₁ := - eqToHom <| FundamentalGroupoid.ext _ _ hx + eqToHom <| FundamentalGroupoid.ext hx @[simp] theorem hcast_def {X : TopCat} {x₀ x₁ : X} (hx₀ : x₀ = x₁) : - hcast hx₀ = eqToHom (FundamentalGroupoid.ext _ _ hx₀) := + hcast hx₀ = eqToHom (FundamentalGroupoid.ext hx₀) := rfl variable {X₁ X₂ Y : TopCat.{u}} {f : C(X₁, Y)} {g : C(X₂, Y)} {x₀ x₁ : X₁} {x₂ x₃ : X₂} @@ -97,8 +97,8 @@ theorem eq_path_of_eq_image : (πₘ f).map ⟦p⟧ = hcast (start_path hfg) ≫ (πₘ g).map ⟦q⟧ ≫ hcast (end_path hfg).symm := by rw [Functor.conj_eqToHom_iff_heq ((πₘ f).map ⟦p⟧) ((πₘ g).map ⟦q⟧) - (FundamentalGroupoid.ext _ _ <| start_path hfg) - (FundamentalGroupoid.ext _ _ <| end_path hfg)] + (FundamentalGroupoid.ext <| start_path hfg) + (FundamentalGroupoid.ext <| end_path hfg)] exact heq_path_of_eq_image hfg end Casts @@ -178,7 +178,7 @@ theorem evalAt_eq (x : X) : ⟦H.evalAt x⟧ = hcast (H.apply_zero x).symm ≫ hcast (H.apply_one x).symm.symm := by dsimp only [prodToProdTopI, uhpath01, hcast] refine (@Functor.conj_eqToHom_iff_heq (πₓ Y) _ _ _ _ _ _ _ _ - (FundamentalGroupoid.ext _ _ <| H.apply_one x).symm).mpr ?_ + (FundamentalGroupoid.ext <| H.apply_one x).symm).mpr ?_ simp only [id_eq_path_refl, prodToProdTop_map, Path.Homotopic.prod_lift, map_eq, ← Path.Homotopic.map_lift] apply Path.Homotopic.hpath_hext; intro; rfl diff --git a/Mathlib/AlgebraicTopology/SimplicialObject.lean b/Mathlib/AlgebraicTopology/SimplicialObject.lean index cf8b6f4083165..ac55809d6b622 100644 --- a/Mathlib/AlgebraicTopology/SimplicialObject.lean +++ b/Mathlib/AlgebraicTopology/SimplicialObject.lean @@ -74,7 +74,7 @@ variable {C} @[ext] lemma hom_ext {X Y : SimplicialObject C} (f g : X ⟶ Y) (h : ∀ (n : SimplexCategoryᵒᵖ), f.app n = g.app n) : f = g := - NatTrans.ext _ _ (by ext; apply h) + NatTrans.ext (by ext; apply h) variable (X : SimplicialObject C) @@ -407,7 +407,7 @@ variable {C} @[ext] lemma hom_ext {X Y : CosimplicialObject C} (f g : X ⟶ Y) (h : ∀ (n : SimplexCategory), f.app n = g.app n) : f = g := - NatTrans.ext _ _ (by ext; apply h) + NatTrans.ext (by ext; apply h) variable (X : CosimplicialObject C) diff --git a/Mathlib/AlgebraicTopology/SimplicialSet.lean b/Mathlib/AlgebraicTopology/SimplicialSet.lean index e0ac42b5f0967..e20ccf4e006ff 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet.lean @@ -339,7 +339,7 @@ instance Truncated.hasColimits {n : ℕ} : HasColimits (Truncated n) := by @[ext] lemma Truncated.hom_ext {n : ℕ} {X Y : Truncated n} {f g : X ⟶ Y} (w : ∀ n, f.app n = g.app n) : f = g := - NatTrans.ext _ _ (funext w) + NatTrans.ext (funext w) /-- The skeleton functor on simplicial sets. -/ def sk (n : ℕ) : SSet ⥤ SSet.Truncated n := diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 3026497f16084..2ece61963849f 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -934,7 +934,7 @@ the first two blocks of `a` and its last three blocks, giving `a.gather b = [11, def gather (a : Composition n) (b : Composition a.length) : Composition n where blocks := (a.blocks.splitWrtComposition b).map sum blocks_pos := by - rw [forall_mem_map_iff] + rw [forall_mem_map] intro j hj suffices H : ∀ i ∈ j, 1 ≤ i by calc 0 < j.length := length_pos_of_mem_splitWrtComposition hj diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean index 0727c837e9c46..1489ce4929c4c 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean @@ -230,7 +230,7 @@ instance : DecidableRel ((· ≤ ·) : IntegrationParams → IntegrationParams fun _ _ => And.decidable instance : DecidableEq IntegrationParams := - fun x y => decidable_of_iff _ (IntegrationParams.ext_iff x y).symm + fun _ _ => decidable_of_iff _ IntegrationParams.ext_iff.symm /-- The `BoxIntegral.IntegrationParams` corresponding to the Riemann integral. In the corresponding filter, we require that the diameters of all boxes `J` of a tagged partition are diff --git a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean index 6d6c7ac4cd4ba..a4e5c528314f8 100644 --- a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean +++ b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean @@ -79,9 +79,6 @@ theorem neg_apply (f : FormalMultilinearSeries 𝕜 E F) (n : ℕ) : (-f) n = - protected theorem ext {p q : FormalMultilinearSeries 𝕜 E F} (h : ∀ n, p n = q n) : p = q := funext h -protected theorem ext_iff {p q : FormalMultilinearSeries 𝕜 E F} : p = q ↔ ∀ n, p n = q n := - Function.funext_iff - protected theorem ne_iff {p q : FormalMultilinearSeries 𝕜 E F} : p ≠ q ↔ ∃ n, p n ≠ q n := Function.ne_iff diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 8968cba486234..146016e97eba6 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -62,7 +62,6 @@ instance : Inhabited ℍ := @[ext] theorem ext {a b : ℍ} (h : (a : ℂ) = b) : a = b := Subtype.eq h -protected theorem ext_iff {a b : ℍ} : a = b ↔ (a : ℂ) = b := Subtype.coe_inj.symm @[simp, norm_cast] theorem ext_iff' {a b : ℍ} : (a : ℂ) = b ↔ a = b := UpperHalfPlane.ext_iff.symm instance canLift : CanLift ℂ ℍ ((↑) : ℍ → ℂ) fun z => 0 < z.im := diff --git a/Mathlib/Analysis/Convex/SimplicialComplex/Basic.lean b/Mathlib/Analysis/Convex/SimplicialComplex/Basic.lean index 1ff6671d44e17..b1f12fdb8b073 100644 --- a/Mathlib/Analysis/Convex/SimplicialComplex/Basic.lean +++ b/Mathlib/Analysis/Convex/SimplicialComplex/Basic.lean @@ -211,7 +211,7 @@ instance : Inf (SimplicialComplex 𝕜 E) := inter_subset_convexHull := fun hs ht => K.inter_subset_convexHull hs.1 ht.1 }⟩ instance : SemilatticeInf (SimplicialComplex 𝕜 E) := - { PartialOrder.lift faces SimplicialComplex.ext with + { PartialOrder.lift faces (fun _ _ => SimplicialComplex.ext) with inf := (· ⊓ ·) inf_le_left := fun _ _ _ hs => hs.1 inf_le_right := fun _ _ _ hs => hs.2 diff --git a/Mathlib/Analysis/InnerProductSpace/WeakOperatorTopology.lean b/Mathlib/Analysis/InnerProductSpace/WeakOperatorTopology.lean index 60c0bdf736203..34f154ad37514 100644 --- a/Mathlib/Analysis/InnerProductSpace/WeakOperatorTopology.lean +++ b/Mathlib/Analysis/InnerProductSpace/WeakOperatorTopology.lean @@ -27,9 +27,6 @@ lemma ext_inner {A B : E →WOT[𝕜] F} (h : ∀ x y, ⟪y, A x⟫_𝕜 = ⟪y, rw [ext_iff] exact fun x => ext_inner_left 𝕜 fun y => h x y -lemma ext_inner_iff {A B : E →WOT[𝕜] F} : A = B ↔ ∀ x y, ⟪y, A x⟫_𝕜 = ⟪y, B x⟫_𝕜 := - ⟨fun h _ _ => by simp [h], ext_inner⟩ - open Filter in /-- The defining property of the weak operator topology: a function `f` tends to `A : E →WOT[𝕜] F` along filter `l` iff `⟪y, (f a) x⟫` tends to `⟪y, A x⟫` along the same filter. -/ diff --git a/Mathlib/Analysis/Normed/Group/Hom.lean b/Mathlib/Analysis/Normed/Group/Hom.lean index 5257de7368fd5..50127f79f7499 100644 --- a/Mathlib/Analysis/Normed/Group/Hom.lean +++ b/Mathlib/Analysis/Normed/Group/Hom.lean @@ -104,9 +104,6 @@ theorem coe_inj_iff : f = g ↔ (f : V₁ → V₂) = g := theorem ext (H : ∀ x, f x = g x) : f = g := coe_inj <| funext H -protected theorem ext_iff : f = g ↔ ∀ x, f x = g x := - ⟨by rintro rfl x; rfl, ext⟩ - variable (f g) @[simp] diff --git a/Mathlib/Analysis/Normed/Lp/lpSpace.lean b/Mathlib/Analysis/Normed/Lp/lpSpace.lean index ef586747ea963..58943db12217a 100644 --- a/Mathlib/Analysis/Normed/Lp/lpSpace.lean +++ b/Mathlib/Analysis/Normed/Lp/lpSpace.lean @@ -311,9 +311,6 @@ instance coeFun : CoeFun (lp E p) fun _ => ∀ i, E i := theorem ext {f g : lp E p} (h : (f : ∀ i, E i) = g) : f = g := Subtype.ext h -protected theorem ext_iff {f g : lp E p} : f = g ↔ (f : ∀ i, E i) = g := - Subtype.ext_iff - theorem eq_zero' [IsEmpty α] (f : lp E p) : f = 0 := Subsingleton.elim f 0 diff --git a/Mathlib/Analysis/Normed/Operator/WeakOperatorTopology.lean b/Mathlib/Analysis/Normed/Operator/WeakOperatorTopology.lean index 8dee97fb59fcd..b7a5ea7136815 100644 --- a/Mathlib/Analysis/Normed/Operator/WeakOperatorTopology.lean +++ b/Mathlib/Analysis/Normed/Operator/WeakOperatorTopology.lean @@ -118,10 +118,6 @@ lemma ext_dual {A B : E →WOT[𝕜] F} (h : ∀ x (y : F⋆), y (A x) = y (B x) specialize h x rwa [← NormedSpace.eq_iff_forall_dual_eq 𝕜] at h -protected lemma ext_dual_iff {A B : E →WOT[𝕜] F} : - A = B ↔ ∀ x (y : F⋆), y (A x) = y (B x) := - ⟨fun h _ _ => by simp [h], ext_dual⟩ - @[simp] lemma zero_apply (x : E) : (0 : E →WOT[𝕜] F) x = 0 := by simp only [DFunLike.coe]; rfl unseal ContinuousLinearMapWOT in diff --git a/Mathlib/Analysis/NormedSpace/ENorm.lean b/Mathlib/Analysis/NormedSpace/ENorm.lean index 9c087c57474a6..fccadf993c631 100644 --- a/Mathlib/Analysis/NormedSpace/ENorm.lean +++ b/Mathlib/Analysis/NormedSpace/ENorm.lean @@ -64,9 +64,6 @@ theorem coeFn_injective : Function.Injective ((↑) : ENorm 𝕜 V → V → ℝ theorem ext {e₁ e₂ : ENorm 𝕜 V} (h : ∀ x, e₁ x = e₂ x) : e₁ = e₂ := coeFn_injective <| funext h -theorem ext_iff {e₁ e₂ : ENorm 𝕜 V} : e₁ = e₂ ↔ ∀ x, e₁ x = e₂ x := - ⟨fun h _ => h ▸ rfl, ext⟩ - @[simp, norm_cast] theorem coe_inj {e₁ e₂ : ENorm 𝕜 V} : (e₁ : V → ℝ≥0∞) = e₂ ↔ e₁ = e₂ := coeFn_injective.eq_iff diff --git a/Mathlib/CategoryTheory/Bicategory/Coherence.lean b/Mathlib/CategoryTheory/Bicategory/Coherence.lean index 423ebfb121f69..acb01f1bd07bb 100644 --- a/Mathlib/CategoryTheory/Bicategory/Coherence.lean +++ b/Mathlib/CategoryTheory/Bicategory/Coherence.lean @@ -86,9 +86,9 @@ theorem preinclusion_obj (a : B) : (preinclusion B).obj ⟨a⟩ = a := @[simp] theorem preinclusion_map₂ {a b : B} (f g : Discrete (Path.{v + 1} a b)) (η : f ⟶ g) : - (preinclusion B).map₂ η = eqToHom (congr_arg _ (Discrete.ext _ _ (Discrete.eq_of_hom η))) := by + (preinclusion B).map₂ η = eqToHom (congr_arg _ (Discrete.ext (Discrete.eq_of_hom η))) := by rcases η with ⟨⟨⟩⟩ - cases Discrete.ext _ _ (by assumption) + cases Discrete.ext (by assumption) convert (inclusionPath a b).map_id _ /-- The normalization of the composition of `p : Path a b` and `f : Hom b c`. @@ -151,7 +151,7 @@ theorem normalizeAux_congr {a b c : B} (p : Path a b) {f g : Hom b c} (η : f theorem normalize_naturality {a b c : B} (p : Path a b) {f g : Hom b c} (η : f ⟶ g) : (preinclusion B).map ⟨p⟩ ◁ η ≫ (normalizeIso p g).hom = (normalizeIso p f).hom ≫ - (preinclusion B).map₂ (eqToHom (Discrete.ext _ _ (normalizeAux_congr p η))) := by + (preinclusion B).map₂ (eqToHom (Discrete.ext (normalizeAux_congr p η))) := by rcases η with ⟨η'⟩; clear η induction η' with | id => simp @@ -186,9 +186,9 @@ def normalize (B : Type u) [Quiver.{v + 1} B] : Pseudofunctor (FreeBicategory B) (LocallyDiscrete (Paths B)) where obj a := ⟨a⟩ map f := ⟨normalizeAux nil f⟩ - map₂ η := eqToHom <| Discrete.ext _ _ <| normalizeAux_congr nil η - mapId a := eqToIso <| Discrete.ext _ _ rfl - mapComp f g := eqToIso <| Discrete.ext _ _ <| normalizeAux_nil_comp f g + map₂ η := eqToHom <| Discrete.ext <| normalizeAux_congr nil η + mapId a := eqToIso <| Discrete.ext rfl + mapComp f g := eqToIso <| Discrete.ext <| normalizeAux_nil_comp f g /-- Auxiliary definition for `normalizeEquiv`. -/ def normalizeUnitIso (a b : FreeBicategory B) : diff --git a/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean b/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean index eb3188205ae76..4bce6d2322466 100644 --- a/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean +++ b/Mathlib/CategoryTheory/Bicategory/LocallyDiscrete.lean @@ -82,7 +82,7 @@ instance subsingleton2Hom {a b : LocallyDiscrete C} (f g : a ⟶ b) : Subsinglet /-- Extract the equation from a 2-morphism in a locally discrete 2-category. -/ theorem eq_of_hom {X Y : LocallyDiscrete C} {f g : X ⟶ Y} (η : f ⟶ g) : f = g := - Discrete.ext _ _ η.1.1 + Discrete.ext η.1.1 end LocallyDiscrete @@ -102,9 +102,9 @@ instance locallyDiscreteBicategory : Bicategory (LocallyDiscrete C) where /-- A locally discrete bicategory is strict. -/ instance locallyDiscreteBicategory.strict : Strict (LocallyDiscrete C) where - id_comp f := Discrete.ext _ _ (Category.id_comp _) - comp_id f := Discrete.ext _ _ (Category.comp_id _) - assoc f g h := Discrete.ext _ _ (Category.assoc _ _ _) + id_comp f := Discrete.ext (Category.id_comp _) + comp_id f := Discrete.ext (Category.comp_id _) + assoc f g h := Discrete.ext (Category.assoc _ _ _) variable {I : Type u₁} [Category.{v₁} I] {B : Type u₂} [Bicategory.{w₂, v₂} B] [Strict B] diff --git a/Mathlib/CategoryTheory/Category/Basic.lean b/Mathlib/CategoryTheory/Category/Basic.lean index c5280badf2b12..7ec4a03391b45 100644 --- a/Mathlib/CategoryTheory/Category/Basic.lean +++ b/Mathlib/CategoryTheory/Category/Basic.lean @@ -103,7 +103,7 @@ open Lean Meta Elab.Tactic in @[tactic sorryIfSorry, inherit_doc sorryIfSorry] def evalSorryIfSorry : Tactic := fun _ => do let goalType ← getMainTarget if goalType.hasSorry then - closeMainGoal (← mkSorry goalType true) + closeMainGoal `sorry_if_sorry (← mkSorry goalType true) else throwError "The goal does not contain `sorry`" diff --git a/Mathlib/CategoryTheory/Category/Bipointed.lean b/Mathlib/CategoryTheory/Category/Bipointed.lean index b7aa0001afb3f..ec8c5cb5122de 100644 --- a/Mathlib/CategoryTheory/Category/Bipointed.lean +++ b/Mathlib/CategoryTheory/Category/Bipointed.lean @@ -147,15 +147,15 @@ def pointedToBipointed : Pointed.{u} ⥤ Bipointed where def pointedToBipointedFst : Pointed.{u} ⥤ Bipointed where obj X := ⟨Option X, X.point, none⟩ map f := ⟨Option.map f.toFun, congr_arg _ f.map_point, rfl⟩ - map_id _ := Bipointed.Hom.ext _ _ Option.map_id - map_comp f g := Bipointed.Hom.ext _ _ (Option.map_comp_map f.1 g.1).symm + map_id _ := Bipointed.Hom.ext Option.map_id + map_comp f g := Bipointed.Hom.ext (Option.map_comp_map f.1 g.1).symm /-- The functor from `Pointed` to `Bipointed` which adds a first point. -/ def pointedToBipointedSnd : Pointed.{u} ⥤ Bipointed where obj X := ⟨Option X, none, X.point⟩ map f := ⟨Option.map f.toFun, rfl, congr_arg _ f.map_point⟩ - map_id _ := Bipointed.Hom.ext _ _ Option.map_id - map_comp f g := Bipointed.Hom.ext _ _ (Option.map_comp_map f.1 g.1).symm + map_id _ := Bipointed.Hom.ext Option.map_id + map_comp f g := Bipointed.Hom.ext (Option.map_comp_map f.1 g.1).symm @[simp] theorem pointedToBipointedFst_comp_swap : @@ -197,7 +197,7 @@ def pointedToBipointedFstBipointedToPointedFstAdjunction : cases x · exact f.map_snd.symm · rfl - right_inv := fun f => Pointed.Hom.ext _ _ rfl } + right_inv := fun f => Pointed.Hom.ext rfl } homEquiv_naturality_left_symm := fun f g => by apply Bipointed.Hom.ext funext x @@ -217,7 +217,7 @@ def pointedToBipointedSndBipointedToPointedSndAdjunction : cases x · exact f.map_fst.symm · rfl - right_inv := fun f => Pointed.Hom.ext _ _ rfl } + right_inv := fun f => Pointed.Hom.ext rfl } homEquiv_naturality_left_symm := fun f g => by apply Bipointed.Hom.ext funext x diff --git a/Mathlib/CategoryTheory/Category/PartialFun.lean b/Mathlib/CategoryTheory/Category/PartialFun.lean index eedda582c92ac..b0c1ab026a13c 100644 --- a/Mathlib/CategoryTheory/Category/PartialFun.lean +++ b/Mathlib/CategoryTheory/Category/PartialFun.lean @@ -109,10 +109,10 @@ noncomputable def partialFunToPointed : PartialFun ⥤ Pointed := by exact { obj := fun X => ⟨Option X, none⟩ map := fun f => ⟨Option.elim' none fun a => (f a).toOption, rfl⟩ - map_id := fun X => Pointed.Hom.ext _ _ <| funext fun o => Option.recOn o rfl fun a => (by + map_id := fun X => Pointed.Hom.ext <| funext fun o => Option.recOn o rfl fun a => (by dsimp [CategoryStruct.id] convert Part.some_toOption a) - map_comp := fun f g => Pointed.Hom.ext _ _ <| funext fun o => Option.recOn o rfl fun a => by + map_comp := fun f g => Pointed.Hom.ext <| funext fun o => Option.recOn o rfl fun a => by dsimp [CategoryStruct.comp] rw [Part.bind_toOption g (f a), Option.elim'_eq_elim] } @@ -148,7 +148,7 @@ noncomputable def partialFunEquivPointed : PartialFun.{u} ≌ Pointed := exact ⟨b, ⟨ha, h.symm⟩, rfl⟩) $ NatIso.ofComponents (fun X ↦ Pointed.Iso.mk (by classical exact Equiv.optionSubtypeNe X.point) (by rfl)) - fun {X Y} f ↦ Pointed.Hom.ext _ _ <| funext fun a ↦ by + fun {X Y} f ↦ Pointed.Hom.ext <| funext fun a ↦ by obtain _ | ⟨a, ha⟩ := a · exact f.map_point.symm simp_all [Option.casesOn'_eq_elim, Part.elim_toOption] @@ -165,7 +165,7 @@ noncomputable def typeToPartialFunIsoPartialFunToPointed : hom_inv_id := rfl inv_hom_id := rfl }) fun f => - Pointed.Hom.ext _ _ <| + Pointed.Hom.ext <| funext fun a => Option.recOn a rfl fun a => by convert Part.some_toOption _ simpa using (Part.get_eq_iff_mem (by trivial)).mp rfl diff --git a/Mathlib/CategoryTheory/Category/Pointed.lean b/Mathlib/CategoryTheory/Category/Pointed.lean index 1f0697c06fd0b..a1f720281f992 100644 --- a/Mathlib/CategoryTheory/Category/Pointed.lean +++ b/Mathlib/CategoryTheory/Category/Pointed.lean @@ -99,8 +99,8 @@ between them. -/ def Iso.mk {α β : Pointed} (e : α ≃ β) (he : e α.point = β.point) : α ≅ β where hom := ⟨e, he⟩ inv := ⟨e.symm, e.symm_apply_eq.2 he.symm⟩ - hom_inv_id := Pointed.Hom.ext _ _ e.symm_comp_self - inv_hom_id := Pointed.Hom.ext _ _ e.self_comp_symm + hom_inv_id := Pointed.Hom.ext e.symm_comp_self + inv_hom_id := Pointed.Hom.ext e.self_comp_symm end Pointed @@ -109,8 +109,8 @@ end Pointed def typeToPointed : Type u ⥤ Pointed.{u} where obj X := ⟨Option X, none⟩ map f := ⟨Option.map f, rfl⟩ - map_id _ := Pointed.Hom.ext _ _ Option.map_id - map_comp _ _ := Pointed.Hom.ext _ _ (Option.map_comp_map _ _).symm + map_id _ := Pointed.Hom.ext Option.map_id + map_comp _ _ := Pointed.Hom.ext (Option.map_comp_map _ _).symm /-- `typeToPointed` is the free functor. -/ def typeToPointedForgetAdjunction : typeToPointed ⊣ forget Pointed := diff --git a/Mathlib/CategoryTheory/Category/TwoP.lean b/Mathlib/CategoryTheory/Category/TwoP.lean index 615d115eac84a..a866c977c18f0 100644 --- a/Mathlib/CategoryTheory/Category/TwoP.lean +++ b/Mathlib/CategoryTheory/Category/TwoP.lean @@ -102,16 +102,16 @@ theorem TwoP_swap_comp_forget_to_Bipointed : noncomputable def pointedToTwoPFst : Pointed.{u} ⥤ TwoP where obj X := ⟨Option X, ⟨X.point, none⟩, some_ne_none _⟩ map f := ⟨Option.map f.toFun, congr_arg _ f.map_point, rfl⟩ - map_id _ := Bipointed.Hom.ext _ _ Option.map_id - map_comp f g := Bipointed.Hom.ext _ _ (Option.map_comp_map f.1 g.1).symm + map_id _ := Bipointed.Hom.ext Option.map_id + map_comp f g := Bipointed.Hom.ext (Option.map_comp_map f.1 g.1).symm /-- The functor from `Pointed` to `TwoP` which adds a first point. -/ @[simps] noncomputable def pointedToTwoPSnd : Pointed.{u} ⥤ TwoP where obj X := ⟨Option X, ⟨none, X.point⟩, (some_ne_none _).symm⟩ map f := ⟨Option.map f.toFun, rfl, congr_arg _ f.map_point⟩ - map_id _ := Bipointed.Hom.ext _ _ Option.map_id - map_comp f g := Bipointed.Hom.ext _ _ (Option.map_comp_map f.1 g.1).symm + map_id _ := Bipointed.Hom.ext Option.map_id + map_comp f g := Bipointed.Hom.ext (Option.map_comp_map f.1 g.1).symm @[simp] theorem pointedToTwoPFst_comp_swap : pointedToTwoPFst ⋙ TwoP.swap = pointedToTwoPSnd := @@ -144,7 +144,7 @@ noncomputable def pointedToTwoPFstForgetCompBipointedToPointedFstAdjunction : cases x · exact f.map_snd.symm · rfl - right_inv := fun f => Pointed.Hom.ext _ _ rfl } + right_inv := fun f => Pointed.Hom.ext rfl } homEquiv_naturality_left_symm := fun f g => by apply Bipointed.Hom.ext funext x @@ -163,7 +163,7 @@ noncomputable def pointedToTwoPSndForgetCompBipointedToPointedSndAdjunction : cases x · exact f.map_fst.symm · rfl - right_inv := fun f => Pointed.Hom.ext _ _ rfl } + right_inv := fun f => Pointed.Hom.ext rfl } homEquiv_naturality_left_symm := fun f g => by apply Bipointed.Hom.ext funext x diff --git a/Mathlib/CategoryTheory/Comma/Arrow.lean b/Mathlib/CategoryTheory/Comma/Arrow.lean index 5347f50362668..1f396ae0ab27c 100644 --- a/Mathlib/CategoryTheory/Comma/Arrow.lean +++ b/Mathlib/CategoryTheory/Comma/Arrow.lean @@ -48,7 +48,7 @@ namespace Arrow @[ext] lemma hom_ext {X Y : Arrow T} (f g : X ⟶ Y) (h₁ : f.left = g.left) (h₂ : f.right = g.right) : f = g := - CommaMorphism.ext _ _ h₁ h₂ + CommaMorphism.ext h₁ h₂ @[simp] theorem id_left (f : Arrow T) : CommaMorphism.left (𝟙 f) = 𝟙 f.left := diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index 9fb8bf3f77d2c..8754946c6e5c6 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -107,7 +107,7 @@ variable {X Y Z : Comma L R} {f : X ⟶ Y} {g : Y ⟶ Z} -- was not triggered automatically @[ext] lemma hom_ext (f g : X ⟶ Y) (h₁ : f.left = g.left) (h₂ : f.right = g.right) : f = g := - CommaMorphism.ext _ _ h₁ h₂ + CommaMorphism.ext h₁ h₂ @[simp] theorem id_left : (𝟙 X : CommaMorphism X X).left = 𝟙 X.left := diff --git a/Mathlib/CategoryTheory/Comma/Presheaf.lean b/Mathlib/CategoryTheory/Comma/Presheaf.lean index fa19b40aa6f62..671b3ef28ec5f 100644 --- a/Mathlib/CategoryTheory/Comma/Presheaf.lean +++ b/Mathlib/CategoryTheory/Comma/Presheaf.lean @@ -286,7 +286,7 @@ lemma mk_snd (s : yoneda.obj X ⟶ A) (x : F.obj (op (CostructuredArrow.mk s))) (mk s x).snd = F.map (eqToHom <| by rw [YonedaCollection.mk_fst]) x := rfl -@[ext] +@[ext (iff := false)] lemma ext {p q : YonedaCollection F X} (h : p.fst = q.fst) (h' : F.map (eqToHom <| by rw [h]) q.snd = p.snd) : p = q := by rcases p with ⟨p, p'⟩ diff --git a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean index 4e4ae5e39d4b6..e59aa9a52ccbd 100644 --- a/Mathlib/CategoryTheory/Comma/StructuredArrow.lean +++ b/Mathlib/CategoryTheory/Comma/StructuredArrow.lean @@ -55,7 +55,7 @@ variable {S S' S'' : D} {Y Y' Y'' : C} {T T' : C ⥤ D} -- was not triggered automatically @[ext] lemma hom_ext {X Y : StructuredArrow S T} (f g : X ⟶ Y) (h : f.right = g.right) : f = g := - CommaMorphism.ext _ _ (Subsingleton.elim _ _) h + CommaMorphism.ext (Subsingleton.elim _ _) h @[simp] theorem hom_eq_iff {X Y : StructuredArrow S T} (f g : X ⟶ Y) : f = g ↔ f.right = g.right := @@ -168,7 +168,7 @@ picks up on it. Either way simp solves these. -/ attribute [-simp, nolint simpNF] isoMk_hom_left_down_down isoMk_inv_left_down_down theorem ext {A B : StructuredArrow S T} (f g : A ⟶ B) : f.right = g.right → f = g := - CommaMorphism.ext _ _ (Subsingleton.elim _ _) + CommaMorphism.ext (Subsingleton.elim _ _) theorem ext_iff {A B : StructuredArrow S T} (f g : A ⟶ B) : f = g ↔ f.right = g.right := ⟨fun h => h ▸ rfl, ext f g⟩ @@ -404,7 +404,7 @@ variable {T T' T'' : D} {Y Y' Y'' : C} {S S' : C ⥤ D} -- was not triggered automatically @[ext] lemma hom_ext {X Y : CostructuredArrow S T} (f g : X ⟶ Y) (h : f.left = g.left) : f = g := - CommaMorphism.ext _ _ h (Subsingleton.elim _ _) + CommaMorphism.ext h (Subsingleton.elim _ _) @[simp] theorem hom_eq_iff {X Y : CostructuredArrow S T} (f g : X ⟶ Y) : f = g ↔ f.left = g.left := @@ -513,7 +513,7 @@ picks up on it. Either way simp solves these. -/ attribute [-simp, nolint simpNF] isoMk_hom_right_down_down isoMk_inv_right_down_down theorem ext {A B : CostructuredArrow S T} (f g : A ⟶ B) (h : f.left = g.left) : f = g := - CommaMorphism.ext _ _ h (Subsingleton.elim _ _) + CommaMorphism.ext h (Subsingleton.elim _ _) theorem ext_iff {A B : CostructuredArrow S T} (f g : A ⟶ B) : f = g ↔ f.left = g.left := ⟨fun h => h ▸ rfl, ext f g⟩ diff --git a/Mathlib/CategoryTheory/ConnectedComponents.lean b/Mathlib/CategoryTheory/ConnectedComponents.lean index 28cf64f7165f9..4f198a18289bb 100644 --- a/Mathlib/CategoryTheory/ConnectedComponents.lean +++ b/Mathlib/CategoryTheory/ConnectedComponents.lean @@ -90,7 +90,7 @@ instance (j : ConnectedComponents J) : IsConnected (Component j) := by · refine @List.chain_pmap_of_chain _ _ _ _ _ f (fun x y _ _ h => ?_) _ _ hl₁ h₁₂ _ exact zag_of_zag_obj (Component.ι _) h · erw [List.getLast_pmap _ f (j₁ :: l) (by simpa [h₁₂] using hf) (List.cons_ne_nil _ _)] - exact FullSubcategory.ext _ _ hl₂ + exact FullSubcategory.ext hl₂ /-- The disjoint union of `J`s connected components, written explicitly as a sigma-type with the category structure. diff --git a/Mathlib/CategoryTheory/Dialectica/Basic.lean b/Mathlib/CategoryTheory/Dialectica/Basic.lean index 6fd658bb09a74..37bab00c84ac7 100644 --- a/Mathlib/CategoryTheory/Dialectica/Basic.lean +++ b/Mathlib/CategoryTheory/Dialectica/Basic.lean @@ -105,7 +105,7 @@ instance : Category (Dial C) where ext <;> simp @[ext] theorem hom_ext {X Y : Dial C} {x y : X ⟶ Y} (hf : x.f = y.f) (hF : x.F = y.F) : x = y := - Hom.ext x y hf hF + Hom.ext hf hF /-- An isomorphism in `Dial C` can be induced by isomorphisms on the source and target, diff --git a/Mathlib/CategoryTheory/DifferentialObject.lean b/Mathlib/CategoryTheory/DifferentialObject.lean index 041bf460419cb..ffd8546c41068 100644 --- a/Mathlib/CategoryTheory/DifferentialObject.lean +++ b/Mathlib/CategoryTheory/DifferentialObject.lean @@ -79,7 +79,7 @@ instance categoryOfDifferentialObjects : Category (DifferentialObject S C) where -- Porting note: added @[ext] theorem ext {A B : DifferentialObject S C} {f g : A ⟶ B} (w : f.f = g.f := by aesop_cat) : f = g := - Hom.ext _ _ w + Hom.ext w @[simp] theorem id_f (X : DifferentialObject S C) : (𝟙 X : X ⟶ X).f = 𝟙 X.obj := rfl diff --git a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean index c8223f4b89e1b..26791079e378f 100644 --- a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean +++ b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean @@ -86,7 +86,7 @@ instance (F : C ⥤ C) : CategoryStruct (Algebra F) where @[ext] lemma ext {A B : Algebra F} {f g : A ⟶ B} (w : f.f = g.f := by aesop_cat) : f = g := - Hom.ext _ _ w + Hom.ext w @[simp] theorem id_eq_id : Algebra.Hom.id A = 𝟙 A := @@ -279,7 +279,7 @@ instance (F : C ⥤ C) : CategoryStruct (Coalgebra F) where @[ext] lemma ext {A B : Coalgebra F} {f g : A ⟶ B} (w : f.f = g.f := by aesop_cat) : f = g := - Hom.ext _ _ w + Hom.ext w @[simp] theorem id_eq_id : Coalgebra.Hom.id V = 𝟙 V := diff --git a/Mathlib/CategoryTheory/Filtered/Small.lean b/Mathlib/CategoryTheory/Filtered/Small.lean index 824efc7668bf3..7948241296129 100644 --- a/Mathlib/CategoryTheory/Filtered/Small.lean +++ b/Mathlib/CategoryTheory/Filtered/Small.lean @@ -97,7 +97,7 @@ end FilteredClosureSmall theorem small_fullSubcategory_filteredClosure : Small.{max v w} (FullSubcategory (FilteredClosure f)) := by refine small_of_injective_of_exists (FilteredClosureSmall.abstractFilteredClosureRealization f) - FullSubcategory.ext ?_ + (fun _ _ => FullSubcategory.ext) ?_ rintro ⟨j, h⟩ induction h with | base x => exact ⟨⟨0, ⟨x⟩⟩, rfl⟩ @@ -232,7 +232,8 @@ end CofilteredClosureSmall theorem small_fullSubcategory_cofilteredClosure : Small.{max v w} (FullSubcategory (CofilteredClosure f)) := by refine small_of_injective_of_exists - (CofilteredClosureSmall.abstractCofilteredClosureRealization f) FullSubcategory.ext ?_ + (CofilteredClosureSmall.abstractCofilteredClosureRealization f) + (fun _ _ => FullSubcategory.ext) ?_ rintro ⟨j, h⟩ induction h with | base x => exact ⟨⟨0, ⟨x⟩⟩, rfl⟩ diff --git a/Mathlib/CategoryTheory/Functor/Category.lean b/Mathlib/CategoryTheory/Functor/Category.lean index de5df863f1f5d..654f3f72d8875 100644 --- a/Mathlib/CategoryTheory/Functor/Category.lean +++ b/Mathlib/CategoryTheory/Functor/Category.lean @@ -53,7 +53,7 @@ namespace NatTrans -- written in terms of `F ⟶ G` rather than `NatTrans F G`, -- or `ext` will not retrieve it from the cache. @[ext] -theorem ext' {α β : F ⟶ G} (w : α.app = β.app) : α = β := NatTrans.ext _ _ w +theorem ext' {α β : F ⟶ G} (w : α.app = β.app) : α = β := NatTrans.ext w @[simp] theorem vcomp_eq_comp (α : F ⟶ G) (β : G ⟶ H) : vcomp α β = α ≫ β := rfl diff --git a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean index 7124c23ab3b9f..cee99b90c06b8 100644 --- a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean +++ b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean @@ -92,7 +92,7 @@ variable {F} @[ext] lemma hom_ext {A B : PointedGaloisObject F} {f g : A ⟶ B} (h : f.val = g.val) : f = g := - Hom.ext f g h + Hom.ext h @[simp] lemma id_val (A : PointedGaloisObject F) : 𝟙 A = 𝟙 A.obj := diff --git a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean index 3fe6da4db7748..6875157be3ba4 100644 --- a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean +++ b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean @@ -201,7 +201,7 @@ lemma ιTensorObj₃_tensorHom (f₁ : X₁ ⟶ Y₁) (f₂ : X₂ ⟶ Y₂) (f ← id_tensorHom, ← id_tensorHom, ← MonoidalCategory.tensor_comp_assoc, ι_tensorHom, ← MonoidalCategory.tensor_comp_assoc, id_comp, comp_id] -@[ext] +@[ext (iff := false)] lemma tensorObj₃_ext {j : I} {A : C} (f g : tensorObj X₁ (tensorObj X₂ X₃) j ⟶ A) [H : HasGoodTensorTensor₂₃ X₁ X₂ X₃] (h : ∀ (i₁ i₂ i₃ : I) (hi : i₁ + i₂ + i₃ = j), @@ -245,7 +245,7 @@ lemma ιTensorObj₃'_tensorHom (f₁ : X₁ ⟶ Y₁) (f₂ : X₂ ⟶ Y₂) (f ← tensorHom_id, ← tensorHom_id, ← MonoidalCategory.tensor_comp_assoc, id_comp, ι_tensorHom, ← MonoidalCategory.tensor_comp_assoc, comp_id] -@[ext] +@[ext (iff := false)] lemma tensorObj₃'_ext {j : I} {A : C} (f g : tensorObj (tensorObj X₁ X₂) X₃ j ⟶ A) [H : HasGoodTensor₁₂Tensor X₁ X₂ X₃] (h : ∀ (i₁ i₂ i₃ : I) (h : i₁ + i₂ + i₃ = j), @@ -309,7 +309,7 @@ abbrev _root_.CategoryTheory.GradedObject.HasLeftTensor₃ObjExt (j : I) := Pres variable {X₁ X₂ X₃} variable [HasTensor X₂ X₃] [HasTensor X₁ (tensorObj X₂ X₃)] -@[ext] +@[ext (iff := false)] lemma left_tensor_tensorObj₃_ext {j : I} {A : C} (Z : C) (f g : Z ⊗ tensorObj X₁ (tensorObj X₂ X₃) j ⟶ A) [H : HasGoodTensorTensor₂₃ X₁ X₂ X₃] @@ -356,7 +356,7 @@ abbrev _root_.CategoryTheory.GradedObject.HasTensor₄ObjExt := variable {X₁ X₂ X₃ X₄} -@[ext] +@[ext (iff := false)] lemma tensorObj₄_ext {j : I} {A : C} (f g : tensorObj X₁ (tensorObj X₂ (tensorObj X₃ X₄)) j ⟶ A) [HasGoodTensorTensor₂₃ X₂ X₃ X₄] [H : HasTensor₄ObjExt X₁ X₂ X₃ X₄] diff --git a/Mathlib/CategoryTheory/Grothendieck.lean b/Mathlib/CategoryTheory/Grothendieck.lean index 32aca1eaf5581..a4c2ef1ad1a44 100644 --- a/Mathlib/CategoryTheory/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Grothendieck.lean @@ -74,7 +74,7 @@ structure Hom (X Y : Grothendieck F) where /-- The morphism from the pushforward to the source fiber object to the target fiber object. -/ fiber : (F.map base).obj X.fiber ⟶ Y.fiber -@[ext] +@[ext (iff := false)] theorem ext {X Y : Grothendieck F} (f g : Hom X Y) (w_base : f.base = g.base) (w_fiber : eqToHom (by rw [w_base]) ≫ f.fiber = g.fiber) : f = g := by cases f; cases g diff --git a/Mathlib/CategoryTheory/GuitartExact/Basic.lean b/Mathlib/CategoryTheory/GuitartExact/Basic.lean index 38d0b2188fb18..c87d10938f0ce 100644 --- a/Mathlib/CategoryTheory/GuitartExact/Basic.lean +++ b/Mathlib/CategoryTheory/GuitartExact/Basic.lean @@ -70,7 +70,7 @@ variable {T L R B} @[ext] lemma ext (w w' : TwoSquare T L R B) (h : ∀ (X : C₁), w.app X = w'.app X) : w = w' := - NatTrans.ext _ _ (funext h) + NatTrans.ext (funext h) variable (w : TwoSquare T L R B) diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index 966ea0c370650..395386e40d32b 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -52,7 +52,7 @@ variable {C} attribute [reassoc (attr := simp)] idem -@[ext] +@[ext (iff := false)] theorem ext {P Q : Karoubi C} (h_X : P.X = Q.X) (h_p : P.p ≫ eqToHom h_X = eqToHom h_X ≫ Q.p) : P = Q := by cases P diff --git a/Mathlib/CategoryTheory/IsConnected.lean b/Mathlib/CategoryTheory/IsConnected.lean index 0e54bd78c74ca..88545d2a24ec0 100644 --- a/Mathlib/CategoryTheory/IsConnected.lean +++ b/Mathlib/CategoryTheory/IsConnected.lean @@ -86,7 +86,7 @@ private def liftToDiscrete {α : Type u₂} (F : J ⥤ Discrete α) : J ⥤ Disc obj j := have := Nonempty.intro j Discrete.mk (Function.invFun F.obj (F.obj j)) map {j _} f := have := Nonempty.intro j - ⟨⟨congr_arg (Function.invFun F.obj) (Discrete.ext _ _ (Discrete.eq_of_hom (F.map f)))⟩⟩ + ⟨⟨congr_arg (Function.invFun F.obj) (Discrete.ext (Discrete.eq_of_hom (F.map f)))⟩⟩ /-- Implementation detail of `isoConstant`. -/ private def factorThroughDiscrete {α : Type u₂} (F : J ⥤ Discrete α) : @@ -245,7 +245,7 @@ lemma isConnected_iff_of_equivalence {K : Type u₂} [Category.{v₂} K] (e : J instance isPreconnected_op [IsPreconnected J] : IsPreconnected Jᵒᵖ where iso_constant := fun {α} F X => ⟨NatIso.ofComponents fun Y => - eqToIso (Discrete.ext _ _ (Discrete.eq_of_hom ((Nonempty.some + eqToIso (Discrete.ext (Discrete.eq_of_hom ((Nonempty.some (IsPreconnected.iso_constant (F.rightOp ⋙ (Discrete.opposite α).functor) (unop X))).app (unop Y)).hom))⟩ diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean index b5b5357ff8075..9bab8cb39bdb4 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Ulift.lean @@ -47,7 +47,7 @@ instance : PreservesLimitsOfSize.{w', w} uliftFunctor.{v, u} where intro s hs obtain ⟨x, hx₁, hx₂⟩ := (Types.isLimit_iff c).mp ⟨hc⟩ _ ((sectionsEquiv K).symm ⟨s, hs⟩).2 exact ⟨⟨x⟩, fun i => ULift.ext _ _ (hx₁ i), - fun y hy => ULift.ext _ _ (hx₂ y.down fun i ↦ (ULift.ext_iff _ _).mp (hy i))⟩ } } + fun y hy => ULift.ext _ _ (hx₂ y.down fun i ↦ ULift.ext_iff.mp (hy i))⟩ } } /-- The functor `uliftFunctor : Type u ⥤ Type (max u v)` creates `u`-small limits. diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index f1a28c7b04482..ecb6dfd87d021 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -257,7 +257,7 @@ attribute [-simp, nolint simpNF] IsBilimit.mk.injEq attribute [local ext] Bicone.IsBilimit instance subsingleton_isBilimit {f : J → C} {c : Bicone f} : Subsingleton c.IsBilimit := - ⟨fun _ _ => Bicone.IsBilimit.ext _ _ (Subsingleton.elim _ _) (Subsingleton.elim _ _)⟩ + ⟨fun _ _ => Bicone.IsBilimit.ext (Subsingleton.elim _ _) (Subsingleton.elim _ _)⟩ section Whisker diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean index 9b93287aa4512..b403910cf1751 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean @@ -99,7 +99,7 @@ variable {f} /-- The morphism `m` in a factorisation `f = e ≫ m` through a monomorphism is uniquely determined. -/ -@[ext] +@[ext (iff := false)] theorem ext {F F' : MonoFactorisation f} (hI : F.I = F'.I) (hm : F.m = eqToHom hI ≫ F'.m) : F = F' := by cases' F with _ Fm _ _ Ffac; cases' F' with _ Fm' _ _ Ffac' @@ -383,7 +383,7 @@ theorem imageMonoIsoSource_hom_self [Mono f] : (imageMonoIsoSource f).hom ≫ f -- This is the proof that `factorThruImage f` is an epimorphism -- from https://en.wikipedia.org/wiki/Image_%28category_theory%29, which is in turn taken from: -- Mitchell, Barry (1965), Theory of categories, MR 0202787, p.12, Proposition 10.1 -@[ext] +@[ext (iff := false)] theorem image.ext [HasImage f] {W : C} {g h : image f ⟶ W} [HasLimit (parallelPair g h)] (w : factorThruImage f ≫ g = factorThruImage f ≫ h) : g = h := by let q := equalizer.ι g h @@ -703,7 +703,7 @@ theorem ImageMap.mk.injEq' {f g : Arrow C} [HasImage f.hom] [HasImage g.hom] {sq instance : Subsingleton (ImageMap sq) := Subsingleton.intro fun a b => - ImageMap.ext a b <| ImageMap.map_uniq a b + ImageMap.ext <| ImageMap.map_uniq a b end diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean index 0fabb5e85e4e5..6fe88cd02ae8f 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean @@ -11,6 +11,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.Multiequalizer import Mathlib.CategoryTheory.ConcreteCategory.Basic import Mathlib.Tactic.CategoryTheory.Elementwise import Mathlib.Data.Set.Subsingleton +import Mathlib.Logic.Relation /-! # Special shapes for limits in `Type`. diff --git a/Mathlib/CategoryTheory/Limits/Types.lean b/Mathlib/CategoryTheory/Limits/Types.lean index 76eaffd79c90a..10095a65f2e14 100644 --- a/Mathlib/CategoryTheory/Limits/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Types.lean @@ -258,9 +258,6 @@ theorem limit_ext' (F : J ⥤ Type v) (x y : limit F) (w : ∀ j, limit.π F j x x = y := limit_ext F x y w -theorem limit_ext_iff (x y : limit F) : x = y ↔ ∀ j, limit.π F j x = limit.π F j y := - ⟨fun t _ => t ▸ rfl, limit_ext _ _ _⟩ - theorem limit_ext_iff' (F : J ⥤ Type v) (x y : limit F) : x = y ↔ ∀ j, limit.π F j x = limit.π F j y := ⟨fun t _ => t ▸ rfl, limit_ext' _ _ _⟩ diff --git a/Mathlib/CategoryTheory/Limits/VanKampen.lean b/Mathlib/CategoryTheory/Limits/VanKampen.lean index 231dd4766d4da..783314187e22d 100644 --- a/Mathlib/CategoryTheory/Limits/VanKampen.lean +++ b/Mathlib/CategoryTheory/Limits/VanKampen.lean @@ -747,7 +747,7 @@ theorem isPullback_initial_to_of_cofan_isVanKampen [HasInitial C] {ι : Type*} { subst this have : ∀ i, Subsingleton (⊥_ C ⟶ (Discrete.functor f).obj i) := inferInstance convert isPullback_of_cofan_isVanKampen hc i.as j.as - exact (if_neg (mt (Discrete.ext _ _) hi.symm)).symm + exact (if_neg (mt Discrete.ext hi.symm)).symm theorem mono_of_cofan_isVanKampen [HasInitial C] {ι : Type*} {F : Discrete ι ⥤ C} {c : Cocone F} (hc : IsVanKampenColimit c) (i : Discrete ι) : Mono (c.ι.app i) := by diff --git a/Mathlib/CategoryTheory/Localization/Resolution.lean b/Mathlib/CategoryTheory/Localization/Resolution.lean index fce7904a80544..8afa0c3016b5f 100644 --- a/Mathlib/CategoryTheory/Localization/Resolution.lean +++ b/Mathlib/CategoryTheory/Localization/Resolution.lean @@ -128,7 +128,7 @@ lemma comp_f {R R' R'' : Φ.RightResolution X₂} (φ : R ⟶ R') (ψ : R' ⟶ R @[ext] lemma hom_ext {R R' : Φ.RightResolution X₂} {φ₁ φ₂ : R ⟶ R'} (h : φ₁.f = φ₂.f) : φ₁ = φ₂ := - Hom.ext _ _ h + Hom.ext h end RightResolution @@ -177,7 +177,7 @@ lemma comp_f {L L' L'' : Φ.LeftResolution X₂} (φ : L ⟶ L') (ψ : L' ⟶ L' @[ext] lemma hom_ext {L L' : Φ.LeftResolution X₂} {φ₁ φ₂ : L ⟶ L'} (h : φ₁.f = φ₂.f) : φ₁ = φ₂ := - Hom.ext _ _ h + Hom.ext h end LeftResolution diff --git a/Mathlib/CategoryTheory/Monad/Algebra.lean b/Mathlib/CategoryTheory/Monad/Algebra.lean index d7a341173df0e..7997a3d8c91e9 100644 --- a/Mathlib/CategoryTheory/Monad/Algebra.lean +++ b/Mathlib/CategoryTheory/Monad/Algebra.lean @@ -84,7 +84,7 @@ instance : CategoryStruct (Algebra T) where -- Porting note (#11041): Adding this `ext` lemma to help automation below. @[ext] -lemma Hom.ext' (X Y : Algebra T) (f g : X ⟶ Y) (h : f.f = g.f) : f = g := Hom.ext _ _ h +lemma Hom.ext' (X Y : Algebra T) (f g : X ⟶ Y) (h : f.f = g.f) : f = g := Hom.ext h @[simp] theorem comp_eq_comp {A A' A'' : Algebra T} (f : A ⟶ A') (g : A' ⟶ A'') : @@ -320,7 +320,7 @@ instance : CategoryStruct (Coalgebra G) where -- Porting note (#11041): Adding `ext` lemma to help automation below. @[ext] -lemma Hom.ext' (X Y : Coalgebra G) (f g : X ⟶ Y) (h : f.f = g.f) : f = g := Hom.ext _ _ h +lemma Hom.ext' (X Y : Coalgebra G) (f g : X ⟶ Y) (h : f.f = g.f) : f = g := Hom.ext h @[simp] theorem comp_eq_comp {A A' A'' : Coalgebra G} (f : A ⟶ A') (g : A' ⟶ A'') : diff --git a/Mathlib/CategoryTheory/Monad/Basic.lean b/Mathlib/CategoryTheory/Monad/Basic.lean index e2979953c0812..911f458280a80 100644 --- a/Mathlib/CategoryTheory/Monad/Basic.lean +++ b/Mathlib/CategoryTheory/Monad/Basic.lean @@ -115,12 +115,12 @@ instance : Quiver (Comonad C) where -- Porting note (#10688): added to ease automation @[ext] lemma MonadHom.ext' {T₁ T₂ : Monad C} (f g : T₁ ⟶ T₂) (h : f.app = g.app) : f = g := - MonadHom.ext f g h + MonadHom.ext h -- Porting note (#10688): added to ease automation @[ext] lemma ComonadHom.ext' {T₁ T₂ : Comonad C} (f g : T₁ ⟶ T₂) (h : f.app = g.app) : f = g := - ComonadHom.ext f g h + ComonadHom.ext h instance : Category (Monad C) where id M := { toNatTrans := 𝟙 (M : C ⥤ C) } @@ -129,9 +129,9 @@ instance : Category (Monad C) where { app := fun X => f.app X ≫ g.app X naturality := fun X Y h => by rw [assoc, f.1.naturality_assoc, g.1.naturality] } } -- `aesop_cat` can fill in these proofs, but is unfortunately slightly slow. - id_comp _ := MonadHom.ext _ _ (by funext; simp only [NatTrans.id_app, id_comp]) - comp_id _ := MonadHom.ext _ _ (by funext; simp only [NatTrans.id_app, comp_id]) - assoc _ _ _ := MonadHom.ext _ _ (by funext; simp only [assoc]) + id_comp _ := MonadHom.ext (by funext; simp only [NatTrans.id_app, id_comp]) + comp_id _ := MonadHom.ext (by funext; simp only [NatTrans.id_app, comp_id]) + assoc _ _ _ := MonadHom.ext (by funext; simp only [assoc]) instance : Category (Comonad C) where id M := { toNatTrans := 𝟙 (M : C ⥤ C) } @@ -140,9 +140,9 @@ instance : Category (Comonad C) where { app := fun X => f.app X ≫ g.app X naturality := fun X Y h => by rw [assoc, f.1.naturality_assoc, g.1.naturality] } } -- `aesop_cat` can fill in these proofs, but is unfortunately slightly slow. - id_comp _ := ComonadHom.ext _ _ (by funext; simp only [NatTrans.id_app, id_comp]) - comp_id _ := ComonadHom.ext _ _ (by funext; simp only [NatTrans.id_app, comp_id]) - assoc _ _ _ := ComonadHom.ext _ _ (by funext; simp only [assoc]) + id_comp _ := ComonadHom.ext (by funext; simp only [NatTrans.id_app, id_comp]) + comp_id _ := ComonadHom.ext (by funext; simp only [NatTrans.id_app, comp_id]) + assoc _ _ _ := ComonadHom.ext (by funext; simp only [assoc]) instance {T : Monad C} : Inhabited (MonadHom T T) := ⟨𝟙 T⟩ diff --git a/Mathlib/CategoryTheory/Monad/Coequalizer.lean b/Mathlib/CategoryTheory/Monad/Coequalizer.lean index e4f0c48251831..77f0a28c07d7a 100644 --- a/Mathlib/CategoryTheory/Monad/Coequalizer.lean +++ b/Mathlib/CategoryTheory/Monad/Coequalizer.lean @@ -60,7 +60,7 @@ def FreeCoequalizer.π : (Monad.free T).obj X.A ⟶ X where theorem FreeCoequalizer.condition : FreeCoequalizer.topMap X ≫ FreeCoequalizer.π X = FreeCoequalizer.bottomMap X ≫ FreeCoequalizer.π X := - Algebra.Hom.ext _ _ X.assoc.symm + Algebra.Hom.ext X.assoc.symm instance : IsReflexivePair (FreeCoequalizer.topMap X) (FreeCoequalizer.bottomMap X) := by apply IsReflexivePair.mk' _ _ _ diff --git a/Mathlib/CategoryTheory/Monad/Equalizer.lean b/Mathlib/CategoryTheory/Monad/Equalizer.lean index e7f4b9bb20b87..a898227f6d7c1 100644 --- a/Mathlib/CategoryTheory/Monad/Equalizer.lean +++ b/Mathlib/CategoryTheory/Monad/Equalizer.lean @@ -61,7 +61,7 @@ def CofreeEqualizer.ι : X ⟶ (Comonad.cofree T).obj X.A where theorem CofreeEqualizer.condition : CofreeEqualizer.ι X ≫ CofreeEqualizer.topMap X = CofreeEqualizer.ι X ≫ CofreeEqualizer.bottomMap X := - Coalgebra.Hom.ext _ _ X.coassoc.symm + Coalgebra.Hom.ext X.coassoc.symm instance : IsCoreflexivePair (CofreeEqualizer.topMap X) (CofreeEqualizer.bottomMap X) := by apply IsCoreflexivePair.mk' _ _ _ diff --git a/Mathlib/CategoryTheory/Monoidal/Bimod.lean b/Mathlib/CategoryTheory/Monoidal/Bimod.lean index bb652506f104a..a29f338d9a4d7 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimod.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimod.lean @@ -123,7 +123,7 @@ instance : Category (Bimod A B) where -- Porting note: added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {M N : Bimod A B} (f g : M ⟶ N) (h : f.hom = g.hom) : f = g := - Hom.ext _ _ h + Hom.ext h @[simp] theorem id_hom' (M : Bimod A B) : (𝟙 M : Hom M M).hom = 𝟙 M.X := diff --git a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean index 0a42bb9f51ec0..2a5a87a7c6305 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean @@ -44,7 +44,7 @@ namespace Bimon_ instance : Category (Bimon_ C) := inferInstanceAs (Category (Comon_ (Mon_ C))) @[ext] lemma ext {X Y : Bimon_ C} {f g : X ⟶ Y} (w : f.hom.hom = g.hom.hom) : f = g := - Comon_.Hom.ext _ _ (Mon_.Hom.ext _ _ w) + Comon_.Hom.ext (Mon_.Hom.ext w) @[simp] theorem id_hom' (M : Bimon_ C) : Comon_.Hom.hom (𝟙 M) = 𝟙 M.X := rfl @@ -68,6 +68,7 @@ def toComon_ : Bimon_ C ⥤ Comon_ C := (Mon_.forgetMonoidal C).toOplaxMonoidalF @[simp] theorem toComon_forget : toComon_ C ⋙ Comon_.forget C = forget C := rfl +set_option maxHeartbeats 400000 in /-- The object level part of the forward direction of `Comon_ (Mon_ C) ≌ Mon_ (Comon_ C)` -/ def toMon_Comon_obj (M : Bimon_ C) : Mon_ (Comon_ C) where X := (toComon_ C).obj M diff --git a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean index 303870ea16c5c..f0ba170eb47a1 100644 --- a/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Braided/Basic.lean @@ -402,7 +402,7 @@ instance categoryLaxBraidedFunctor : Category (LaxBraidedFunctor C D) := -- Porting note: added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : LaxBraidedFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := - MonoidalNatTrans.ext _ _ (funext w) + MonoidalNatTrans.ext (funext w) @[simp] theorem comp_toNatTrans {F G H : LaxBraidedFunctor C D} {α : F ⟶ G} {β : G ⟶ H} : @@ -468,7 +468,7 @@ instance categoryBraidedFunctor : Category (BraidedFunctor C D) := -- Porting note: added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : BraidedFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := - MonoidalNatTrans.ext _ _ (funext w) + MonoidalNatTrans.ext (funext w) @[simp] theorem comp_toNatTrans {F G H : BraidedFunctor C D} {α : F ⟶ G} {β : G ⟶ H} : diff --git a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean index 84d6e637f216e..7eaef1920b0fc 100644 --- a/Mathlib/CategoryTheory/Monoidal/CommMon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/CommMon_.lean @@ -54,7 +54,7 @@ theorem comp_hom {R S T : CommMon_ C} (f : R ⟶ S) (g : S ⟶ T) : -- for morphisms in `CommMon_ C` @[ext] lemma hom_ext {A B : CommMon_ C} (f g : A ⟶ B) (h : f.hom = g.hom) : f = g := - Mon_.Hom.ext _ _ h + Mon_.Hom.ext h -- Porting note (#10688): the following two lemmas `id'` and `comp'` -- have been added to ease automation; diff --git a/Mathlib/CategoryTheory/Monoidal/Comon_.lean b/Mathlib/CategoryTheory/Monoidal/Comon_.lean index 1c08b51af7ae2..7da6f2e2289d8 100644 --- a/Mathlib/CategoryTheory/Monoidal/Comon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Comon_.lean @@ -109,7 +109,7 @@ instance : Category (Comon_ C) where id := id comp f g := comp f g -@[ext] lemma ext {X Y : Comon_ C} {f g : X ⟶ Y} (w : f.hom = g.hom) : f = g := Hom.ext _ _ w +@[ext] lemma ext {X Y : Comon_ C} {f g : X ⟶ Y} (w : f.hom = g.hom) : f = g := Hom.ext w @[simp] theorem id_hom' (M : Comon_ C) : (𝟙 M : Hom M M).hom = 𝟙 M.X := rfl diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean index d7f2f0c27cab2..4752b8432ea1e 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean @@ -83,9 +83,9 @@ theorem inclusion_obj (X : N C) : @[simp] theorem inclusion_map {X Y : N C} (f : X ⟶ Y) : - inclusion.map f = eqToHom (congr_arg _ (Discrete.ext _ _ (Discrete.eq_of_hom f))) := by + inclusion.map f = eqToHom (congr_arg _ (Discrete.ext (Discrete.eq_of_hom f))) := by rcases f with ⟨⟨⟩⟩ - cases Discrete.ext _ _ (by assumption) + cases Discrete.ext (by assumption) apply inclusion.map_id /-- Auxiliary definition for `normalize`. -/ @@ -248,7 +248,7 @@ theorem normalizeObj_congr (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟶ Y theorem normalize_naturality (n : NormalMonoidalObject C) {X Y : F C} (f : X ⟶ Y) : inclusionObj n ◁ f ≫ (normalizeIsoApp' C Y n).hom = (normalizeIsoApp' C X n).hom ≫ - inclusion.map (eqToHom (Discrete.ext _ _ (normalizeObj_congr n f))) := by + inclusion.map (eqToHom (Discrete.ext (normalizeObj_congr n f))) := by revert n induction f using Hom.inductionOn case comp f g ihf ihg => simp [ihg, reassoc_of% (ihf _)] diff --git a/Mathlib/CategoryTheory/Monoidal/Mod_.lean b/Mathlib/CategoryTheory/Monoidal/Mod_.lean index aad48367e827b..79ddde1156f2d 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mod_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mod_.lean @@ -60,7 +60,7 @@ instance : Category (Mod_ A) where -- Porting note (#5229): added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {M N : Mod_ A} (f₁ f₂ : M ⟶ N) (h : f₁.hom = f₂.hom) : f₁ = f₂ := - Hom.ext _ _ h + Hom.ext h @[simp] theorem id_hom' (M : Mod_ A) : (𝟙 M : M ⟶ M).hom = 𝟙 M.X := by diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index 3688c4218f1ce..c19a846886021 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -107,7 +107,7 @@ instance : Category (Mon_ C) where -- Porting note: added, as `Hom.ext` does not apply to a morphism. @[ext] lemma ext {X Y : Mon_ C} {f g : X ⟶ Y} (w : f.hom = g.hom) : f = g := - Hom.ext _ _ w + Hom.ext w @[simp] theorem id_hom' (M : Mon_ C) : (𝟙 M : Hom M M).hom = 𝟙 M.X := diff --git a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean index 52c9649551f6f..c10f52bc77441 100644 --- a/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean +++ b/Mathlib/CategoryTheory/Monoidal/NaturalTransformation.lean @@ -84,7 +84,7 @@ instance categoryMonoidalFunctor : Category (MonoidalFunctor C D) := -- Porting note: added, as `MonoidalNatTrans.ext` does not apply to morphisms. @[ext] lemma ext' {F G : LaxMonoidalFunctor C D} {α β : F ⟶ G} (w : ∀ X : C, α.app X = β.app X) : α = β := - MonoidalNatTrans.ext _ _ (funext w) + MonoidalNatTrans.ext (funext w) @[simp] theorem comp_toNatTrans {F G H : MonoidalFunctor C D} {α : F ⟶ G} {β : G ⟶ H} : diff --git a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean index e0b1bcef272d7..9a2f2fa079c42 100644 --- a/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Rigid/Basic.lean @@ -148,6 +148,13 @@ attribute [instance] HasLeftDual.exact open ExactPairing HasRightDual HasLeftDual MonoidalCategory +#adaptation_note +/-- +The overlapping notation for `leftDual` and `leftAdjointMate` become more problematic in +after https://github.com/leanprover/lean4/pull/4596, and we sometimes have to disambiguate with +e.g. `(ᘁX : C)` where previously just `ᘁX` was enough. +-/ + @[inherit_doc] prefix:1024 "ᘁ" => leftDual @[inherit_doc] postfix:1024 "ᘁ" => rightDual @@ -196,7 +203,7 @@ theorem rightAdjointMate_comp {X Y Z : C} [HasRightDual X] [HasRightDual Y] {f : (ρ_ (Yᘁ)).inv ≫ _ ◁ η_ X (Xᘁ) ≫ _ ◁ (f ⊗ g) ≫ (α_ (Yᘁ) Y Z).inv ≫ ε_ Y (Yᘁ) ▷ _ ≫ (λ_ Z).hom := calc - _ = 𝟙 _ ⊗≫ Yᘁ ◁ η_ X Xᘁ ≫ Yᘁ ◁ f ▷ Xᘁ ⊗≫ (ε_ Y Yᘁ ▷ Xᘁ ≫ 𝟙_ C ◁ g) ⊗≫ 𝟙 _ := by + _ = 𝟙 _ ⊗≫ (Yᘁ : C) ◁ η_ X Xᘁ ≫ Yᘁ ◁ f ▷ Xᘁ ⊗≫ (ε_ Y Yᘁ ▷ Xᘁ ≫ 𝟙_ C ◁ g) ⊗≫ 𝟙 _ := by dsimp only [rightAdjointMate]; coherence _ = _ := by rw [← whisker_exchange, tensorHom_def]; coherence @@ -205,9 +212,9 @@ theorem leftAdjointMate_comp {X Y Z : C} [HasLeftDual X] [HasLeftDual Y] {f : X {g : (ᘁX) ⟶ Z} : (ᘁf) ≫ g = (λ_ _).inv ≫ - η_ (ᘁX) X ▷ _ ≫ (g ⊗ f) ▷ _ ≫ (α_ _ _ _).hom ≫ _ ◁ ε_ _ _ ≫ (ρ_ _).hom := + η_ (ᘁX : C) X ▷ _ ≫ (g ⊗ f) ▷ _ ≫ (α_ _ _ _).hom ≫ _ ◁ ε_ _ _ ≫ (ρ_ _).hom := calc - _ = 𝟙 _ ⊗≫ η_ (ᘁX) X ▷ (ᘁY) ⊗≫ (ᘁX) ◁ f ▷ (ᘁY) ⊗≫ ((ᘁX) ◁ ε_ (ᘁY) Y ≫ g ▷ 𝟙_ C) ⊗≫ 𝟙 _ := by + _ = 𝟙 _ ⊗≫ η_ (ᘁX : C) X ▷ (ᘁY) ⊗≫ (ᘁX) ◁ f ▷ (ᘁY) ⊗≫ ((ᘁX) ◁ ε_ (ᘁY) Y ≫ g ▷ 𝟙_ C) ⊗≫ 𝟙 _ := by dsimp only [leftAdjointMate]; coherence _ = _ := by rw [whisker_exchange, tensorHom_def']; coherence @@ -395,7 +402,7 @@ theorem tensorLeftHomEquiv_symm_coevaluation_comp_whiskerRight {X Y : C} [HasRig @[simp] theorem tensorRightHomEquiv_symm_coevaluation_comp_whiskerLeft {X Y : C} [HasLeftDual X] [HasLeftDual Y] (f : X ⟶ Y) : - (tensorRightHomEquiv _ (ᘁY) _ _).symm (η_ (ᘁX) X ≫ (ᘁX) ◁ f) = (λ_ _).hom ≫ ᘁf := by + (tensorRightHomEquiv _ (ᘁY) _ _).symm (η_ (ᘁX : C) X ≫ (ᘁX : C) ◁ f) = (λ_ _).hom ≫ ᘁf := by dsimp [tensorRightHomEquiv, leftAdjointMate] simp @@ -414,7 +421,7 @@ theorem tensorRightHomEquiv_symm_coevaluation_comp_whiskerRight {Y Y' Z : C} [Ex theorem tensorLeftHomEquiv_whiskerLeft_comp_evaluation {Y Z : C} [HasLeftDual Z] (f : Y ⟶ ᘁZ) : (tensorLeftHomEquiv _ _ _ _) (Z ◁ f ≫ ε_ _ _) = f ≫ (ρ_ _).inv := calc - _ = 𝟙 _ ⊗≫ (η_ (ᘁZ) Z ▷ Y ≫ ((ᘁZ) ⊗ Z) ◁ f) ⊗≫ (ᘁZ) ◁ ε_ (ᘁZ) Z := by + _ = 𝟙 _ ⊗≫ (η_ (ᘁZ : C) Z ▷ Y ≫ ((ᘁZ) ⊗ Z) ◁ f) ⊗≫ (ᘁZ) ◁ ε_ (ᘁZ) Z := by dsimp [tensorLeftHomEquiv]; coherence _ = f ⊗≫ (η_ (ᘁZ) Z ▷ (ᘁZ) ⊗≫ (ᘁZ) ◁ ε_ (ᘁZ) Z) := by rw [← whisker_exchange]; coherence @@ -429,7 +436,7 @@ theorem tensorLeftHomEquiv_whiskerRight_comp_evaluation {X Y : C} [HasLeftDual X @[simp] theorem tensorRightHomEquiv_whiskerLeft_comp_evaluation {X Y : C} [HasRightDual X] [HasRightDual Y] - (f : X ⟶ Y) : (tensorRightHomEquiv _ _ _ _) ((Yᘁ) ◁ f ≫ ε_ _ _) = fᘁ ≫ (λ_ _).inv := by + (f : X ⟶ Y) : (tensorRightHomEquiv _ _ _ _) ((Yᘁ : C) ◁ f ≫ ε_ _ _) = fᘁ ≫ (λ_ _).inv := by dsimp [tensorRightHomEquiv, rightAdjointMate] simp diff --git a/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean b/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean index 06576c510ac39..e4d78ec3864b8 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/IsInvertedBy.lean @@ -109,7 +109,7 @@ instance (W : MorphismProperty C) (D : Type*) [Category D] : Category (FunctorsI @[ext] lemma FunctorsInverting.hom_ext {W : MorphismProperty C} {F₁ F₂ : FunctorsInverting W D} {α β : F₁ ⟶ F₂} (h : α.app = β.app) : α = β := - NatTrans.ext _ _ h + NatTrans.ext h /-- A constructor for `W.FunctorsInverting D` -/ def FunctorsInverting.mk {W : MorphismProperty C} {D : Type*} [Category D] (F : C ⥤ D) diff --git a/Mathlib/CategoryTheory/PathCategory.lean b/Mathlib/CategoryTheory/PathCategory.lean index 280ad96bd3c84..e27f6b40b20c9 100644 --- a/Mathlib/CategoryTheory/PathCategory.lean +++ b/Mathlib/CategoryTheory/PathCategory.lean @@ -49,7 +49,7 @@ def of : V ⥤q Paths V where obj X := X map f := f.toPath -attribute [local ext] Functor.ext +attribute [local ext (iff := false)] Functor.ext /-- Any prefunctor from `V` lifts to a functor from `paths V` -/ def lift {C} [Category C] (φ : V ⥤q C) : Paths V ⥤ C where @@ -109,7 +109,7 @@ theorem lift_unique {C} [Category C] (φ : V ⥤q C) (Φ : Paths V ⥤ C) rw [this, ih] /-- Two functors out of a path category are equal when they agree on singleton paths. -/ -@[ext] +@[ext (iff := false)] theorem ext_functor {C} [Category C] {F G : Paths V ⥤ C} (h_obj : F.obj = G.obj) (h : ∀ (a b : V) (e : a ⟶ b), F.map e.toPath = eqToHom (congr_fun h_obj a) ≫ G.map e.toPath ≫ eqToHom (congr_fun h_obj.symm b)) : diff --git a/Mathlib/CategoryTheory/Quotient.lean b/Mathlib/CategoryTheory/Quotient.lean index 45179ca56d0d5..5a1f9831b2055 100644 --- a/Mathlib/CategoryTheory/Quotient.lean +++ b/Mathlib/CategoryTheory/Quotient.lean @@ -211,7 +211,7 @@ variable {r} lemma natTrans_ext {F G : Quotient r ⥤ D} (τ₁ τ₂ : F ⟶ G) (h : whiskerLeft (Quotient.functor r) τ₁ = whiskerLeft (Quotient.functor r) τ₂) : τ₁ = τ₂ := - NatTrans.ext _ _ (by ext1 ⟨X⟩; exact NatTrans.congr_app h X) + NatTrans.ext (by ext1 ⟨X⟩; exact NatTrans.congr_app h X) variable (r) diff --git a/Mathlib/CategoryTheory/Shift/SingleFunctors.lean b/Mathlib/CategoryTheory/Shift/SingleFunctors.lean index a0b3341361b49..9b845c8a62a64 100644 --- a/Mathlib/CategoryTheory/Shift/SingleFunctors.lean +++ b/Mathlib/CategoryTheory/Shift/SingleFunctors.lean @@ -143,7 +143,7 @@ variable {F G H} lemma comp_hom (f : F ⟶ G) (g : G ⟶ H) (a : A) : (f ≫ g).hom a = f.hom a ≫ g.hom a := rfl @[ext] -lemma hom_ext (f g : F ⟶ G) (h : f.hom = g.hom) : f = g := Hom.ext f g h +lemma hom_ext (f g : F ⟶ G) (h : f.hom = g.hom) : f = g := Hom.ext h /-- Construct an isomorphism in `SingleFunctors C D A` by giving level-wise isomorphisms and checking compatibility only in the forward direction. -/ diff --git a/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean b/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean index 5c19b0445d423..3f6ecdafd76ae 100644 --- a/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean +++ b/Mathlib/CategoryTheory/Sites/ConcreteSheafification.lean @@ -554,8 +554,8 @@ variable (D) noncomputable def plusPlusSheaf : (Cᵒᵖ ⥤ D) ⥤ Sheaf J D where obj P := ⟨J.sheafify P, J.sheafify_isSheaf P⟩ map η := ⟨J.sheafifyMap η⟩ - map_id _ := Sheaf.Hom.ext _ _ <| J.sheafifyMap_id _ - map_comp _ _ := Sheaf.Hom.ext _ _ <| J.sheafifyMap_comp _ _ + map_id _ := Sheaf.Hom.ext <| J.sheafifyMap_id _ + map_comp _ _ := Sheaf.Hom.ext <| J.sheafifyMap_comp _ _ instance plusPlusSheaf_preservesZeroMorphisms [Preadditive D] : (plusPlusSheaf J D).PreservesZeroMorphisms where @@ -571,7 +571,7 @@ noncomputable def plusPlusAdjunction : plusPlusSheaf J D ⊣ sheafToPresheaf J D { homEquiv := fun P Q => { toFun := fun e => J.toSheafify P ≫ e.val invFun := fun e => ⟨J.sheafifyLift e Q.2⟩ - left_inv := fun e => Sheaf.Hom.ext _ _ <| (J.sheafifyLift_unique _ _ _ rfl).symm + left_inv := fun e => Sheaf.Hom.ext <| (J.sheafifyLift_unique _ _ _ rfl).symm right_inv := fun e => J.toSheafify_sheafifyLift _ _ } homEquiv_naturality_left_symm := by intro P Q R η γ; ext1; dsimp; symm diff --git a/Mathlib/CategoryTheory/Sites/CoverLifting.lean b/Mathlib/CategoryTheory/Sites/CoverLifting.lean index 8ef2cd208aca1..3755ebea1a0f2 100644 --- a/Mathlib/CategoryTheory/Sites/CoverLifting.lean +++ b/Mathlib/CategoryTheory/Sites/CoverLifting.lean @@ -230,8 +230,8 @@ variable (A J) def Functor.sheafPushforwardCocontinuous : Sheaf J A ⥤ Sheaf K A where obj ℱ := ⟨G.op.ran.obj ℱ.val, ran_isSheaf_of_isCocontinuous _ K ℱ⟩ map f := ⟨G.op.ran.map f.val⟩ - map_id ℱ := Sheaf.Hom.ext _ _ <| (ran G.op).map_id ℱ.val - map_comp f g := Sheaf.Hom.ext _ _ <| (ran G.op).map_comp f.val g.val + map_id ℱ := Sheaf.Hom.ext <| (ran G.op).map_id ℱ.val + map_comp f g := Sheaf.Hom.ext <| (ran G.op).map_comp f.val g.val /-- `G.sheafPushforwardCocontinuous A J K : Sheaf J A ⥤ Sheaf K A` is induced by the right Kan extension functor `G.op.ran` on presheaves. -/ diff --git a/Mathlib/CategoryTheory/Sites/Coverage.lean b/Mathlib/CategoryTheory/Sites/Coverage.lean index c39afa796527f..8077b8d450f7f 100644 --- a/Mathlib/CategoryTheory/Sites/Coverage.lean +++ b/Mathlib/CategoryTheory/Sites/Coverage.lean @@ -242,7 +242,7 @@ instance : PartialOrder (Coverage C) where le A B := A.covering ≤ B.covering le_refl A X := le_refl _ le_trans A B C h1 h2 X := le_trans (h1 X) (h2 X) - le_antisymm A B h1 h2 := Coverage.ext A B <| funext <| + le_antisymm A B h1 h2 := Coverage.ext <| funext <| fun X => le_antisymm (h1 X) (h2 X) variable (C) in diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean index 8b5e5da1ab096..f5b5acba282c3 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean @@ -478,7 +478,7 @@ lemma isContinuous [G.IsLocallyFaithful K] (Hp : CoverPreserving J K G) : G.IsCo instance full_sheafPushforwardContinuous [G.IsContinuous J K] : Full (G.sheafPushforwardContinuous A J K) where - map_surjective α := ⟨⟨sheafHom α.val⟩, Sheaf.Hom.ext _ _ <| sheafHom_restrict_eq α.val⟩ + map_surjective α := ⟨⟨sheafHom α.val⟩, Sheaf.Hom.ext <| sheafHom_restrict_eq α.val⟩ instance faithful_sheafPushforwardContinuous [G.IsContinuous J K] : Faithful (G.sheafPushforwardContinuous A J K) where diff --git a/Mathlib/CategoryTheory/Sites/Limits.lean b/Mathlib/CategoryTheory/Sites/Limits.lean index 8569c1c410eb3..d85824a7edd78 100644 --- a/Mathlib/CategoryTheory/Sites/Limits.lean +++ b/Mathlib/CategoryTheory/Sites/Limits.lean @@ -138,7 +138,7 @@ theorem isSheaf_of_isLimit (F : K ⥤ Sheaf J D) (E : Cone (F ⋙ sheafToPreshea instance (F : K ⥤ Sheaf J D) : CreatesLimit F (sheafToPresheaf J D) := createsLimitOfReflectsIso fun E hE => { liftedCone := ⟨⟨E.pt, isSheaf_of_isLimit _ _ hE⟩, - ⟨fun t => ⟨E.π.app _⟩, fun u v e => Sheaf.Hom.ext _ _ <| E.π.naturality _⟩⟩ + ⟨fun t => ⟨E.π.app _⟩, fun u v e => Sheaf.Hom.ext <| E.π.naturality _⟩⟩ validLift := Cones.ext (eqToIso rfl) fun j => by dsimp simp diff --git a/Mathlib/CategoryTheory/Sites/Pretopology.lean b/Mathlib/CategoryTheory/Sites/Pretopology.lean index 4f0a883be4355..8163bb4bac21b 100644 --- a/Mathlib/CategoryTheory/Sites/Pretopology.lean +++ b/Mathlib/CategoryTheory/Sites/Pretopology.lean @@ -85,7 +85,7 @@ instance : PartialOrder (Pretopology C) := { Pretopology.LE with le_refl := fun K => le_def.mpr le_rfl le_trans := fun K₁ K₂ K₃ h₁₂ h₂₃ => le_def.mpr (le_trans h₁₂ h₂₃) - le_antisymm := fun K₁ K₂ h₁₂ h₂₁ => Pretopology.ext _ _ (le_antisymm h₁₂ h₂₁) } + le_antisymm := fun K₁ K₂ h₁₂ h₂₁ => Pretopology.ext (le_antisymm h₁₂ h₂₁) } instance : OrderTop (Pretopology C) where top := diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index ce44a0f42b685..1eb3f3d744895 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -323,9 +323,9 @@ instance instCategorySheaf : Category (Sheaf J A) 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 _ _ _ + id_comp _ := Hom.ext <| id_comp _ + comp_id _ := Hom.ext <| comp_id _ + assoc _ _ _ := Hom.ext <| assoc _ _ _ -- Let's make the inhabited linter happy.../sips instance (X : Sheaf J A) : Inhabited (Hom X X) := @@ -334,7 +334,7 @@ instance (X : Sheaf J A) : Inhabited (Hom X X) := -- Porting note: added because `Sheaf.Hom.ext` was not triggered automatically @[ext] lemma hom_ext {X Y : Sheaf J A} (x y : X ⟶ Y) (h : x.val = y.val) : x = y := - Sheaf.Hom.ext _ _ h + Sheaf.Hom.ext h end Sheaf @@ -476,7 +476,7 @@ theorem Sheaf.Hom.add_app (f g : P ⟶ Q) (U) : (f + g).1.app U = f.1.app U + g. instance Sheaf.Hom.addCommGroup : AddCommGroup (P ⟶ Q) := Function.Injective.addCommGroup (fun f : Sheaf.Hom P Q => f.1) - (fun _ _ h => Sheaf.Hom.ext _ _ h) rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) + (fun _ _ h => Sheaf.Hom.ext h) rfl (fun _ _ => rfl) (fun _ => rfl) (fun _ _ => rfl) (fun _ _ => by aesop_cat) (fun _ _ => by aesop_cat) instance : Preadditive (Sheaf J A) where diff --git a/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean b/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean index 5997ddf34abf9..db8d37ff2bf28 100644 --- a/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean +++ b/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean @@ -223,16 +223,16 @@ 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 _ _ _ + id_comp _ := Hom.ext <| id_comp _ + comp_id _ := Hom.ext <| comp_id _ + assoc _ _ _ := Hom.ext <| assoc _ _ _ -- Porting note (#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 f g w + Hom.ext w -- Let's make the inhabited linter happy... instance (X : SheafOfTypes J) : Inhabited (Hom X X) := diff --git a/Mathlib/CategoryTheory/Sites/Sieves.lean b/Mathlib/CategoryTheory/Sites/Sieves.lean index eeffcb4055728..16cee3b83af67 100644 --- a/Mathlib/CategoryTheory/Sites/Sieves.lean +++ b/Mathlib/CategoryTheory/Sites/Sieves.lean @@ -267,9 +267,6 @@ theorem arrows_ext : ∀ {R S : Sieve X}, R.arrows = S.arrows → R = S := by protected theorem ext {R S : Sieve X} (h : ∀ ⦃Y⦄ (f : Y ⟶ X), R f ↔ S f) : R = S := arrows_ext <| funext fun _ => funext fun f => propext <| h f -protected theorem ext_iff {R S : Sieve X} : R = S ↔ ∀ ⦃Y⦄ (f : Y ⟶ X), R f ↔ S f := - ⟨fun h _ _ => h ▸ Iff.rfl, Sieve.ext⟩ - open Lattice /-- The supremum of a collection of sieves: the union of them all. -/ diff --git a/Mathlib/CategoryTheory/Sites/Subsheaf.lean b/Mathlib/CategoryTheory/Sites/Subsheaf.lean index 67b98eb299d71..a58fe64c5db0c 100644 --- a/Mathlib/CategoryTheory/Sites/Subsheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Subsheaf.lean @@ -54,7 +54,7 @@ structure Subpresheaf (F : Cᵒᵖ ⥤ Type w) where variable {F F' F'' : Cᵒᵖ ⥤ Type w} (G G' : Subpresheaf F) instance : PartialOrder (Subpresheaf F) := - PartialOrder.lift Subpresheaf.obj Subpresheaf.ext + PartialOrder.lift Subpresheaf.obj (fun _ _ => Subpresheaf.ext) instance : Top (Subpresheaf F) := ⟨⟨fun U => ⊤, @fun U V _ x _ => by aesop_cat⟩⟩ @@ -84,8 +84,8 @@ instance {U} : CoeHead (G.toPresheaf.obj U) (F.obj U) where def Subpresheaf.ι : G.toPresheaf ⟶ F where app U x := x instance : Mono G.ι := - ⟨@fun _ f₁ f₂ e => - NatTrans.ext f₁ f₂ <| + ⟨@fun _ _ _ e => + NatTrans.ext <| funext fun U => funext fun x => Subtype.ext <| congr_fun (congr_app e U) x⟩ /-- The inclusion of a subpresheaf to a larger subpresheaf -/ @@ -94,8 +94,8 @@ def Subpresheaf.homOfLe {G G' : Subpresheaf F} (h : G ≤ G') : G.toPresheaf ⟶ app U x := ⟨x, h U x.prop⟩ instance {G G' : Subpresheaf F} (h : G ≤ G') : Mono (Subpresheaf.homOfLe h) := - ⟨fun f₁ f₂ e => - NatTrans.ext f₁ f₂ <| + ⟨fun _ _ e => + NatTrans.ext <| funext fun U => funext fun x => Subtype.ext <| (congr_arg Subtype.val <| (congr_fun (congr_app e U) x : _) : _)⟩ diff --git a/Mathlib/CategoryTheory/Sites/Types.lean b/Mathlib/CategoryTheory/Sites/Types.lean index a692056bb5cd3..3e4363a236d3d 100644 --- a/Mathlib/CategoryTheory/Sites/Types.lean +++ b/Mathlib/CategoryTheory/Sites/Types.lean @@ -151,9 +151,8 @@ noncomputable def typeEquiv : Type u ≌ SheafOfTypes typesGrothendieckTopology fun _ => rfl) (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) + SheafOfTypes.Hom.ext <| NatTrans.ext <| + funext fun α => funext fun s => funext fun x => eval_app S₁ S₂ f (unop α) s x) theorem subcanonical_typesGrothendieckTopology : Sheaf.Subcanonical typesGrothendieckTopology.{u} := Sheaf.Subcanonical.of_yoneda_isSheaf _ fun _ => isSheaf_yoneda' diff --git a/Mathlib/CategoryTheory/Sites/Whiskering.lean b/Mathlib/CategoryTheory/Sites/Whiskering.lean index 8bdffc11d654a..f80010696a0b9 100644 --- a/Mathlib/CategoryTheory/Sites/Whiskering.lean +++ b/Mathlib/CategoryTheory/Sites/Whiskering.lean @@ -46,8 +46,8 @@ variable [J.HasSheafCompose F] [J.HasSheafCompose G] [J.HasSheafCompose H] def sheafCompose : Sheaf J A ⥤ Sheaf J B where obj G := ⟨G.val ⋙ F, GrothendieckTopology.HasSheafCompose.isSheaf G.val G.2⟩ map η := ⟨whiskerRight η.val _⟩ - map_id _ := Sheaf.Hom.ext _ _ <| whiskerRight_id _ - map_comp _ _ := Sheaf.Hom.ext _ _ <| whiskerRight_comp _ _ _ + map_id _ := Sheaf.Hom.ext <| whiskerRight_id _ + map_comp _ _ := Sheaf.Hom.ext <| whiskerRight_comp _ _ _ instance [F.Faithful] : (sheafCompose J F ⋙ sheafToPresheaf _ _).Faithful := show (sheafToPresheaf _ _ ⋙ (whiskeringRight Cᵒᵖ A B).obj F).Faithful from inferInstance diff --git a/Mathlib/CategoryTheory/Square.lean b/Mathlib/CategoryTheory/Square.lean index 075f8994550ea..3db04feeca7ef 100644 --- a/Mathlib/CategoryTheory/Square.lean +++ b/Mathlib/CategoryTheory/Square.lean @@ -115,7 +115,7 @@ instance category : Category (Square C) where lemma hom_ext {sq₁ sq₂ : Square C} {f g : sq₁ ⟶ sq₂} (h₁ : f.τ₁ = g.τ₁) (h₂ : f.τ₂ = g.τ₂) (h₃ : f.τ₃ = g.τ₃) (h₄ : f.τ₄ = g.τ₄) : f = g := - Hom.ext _ _ h₁ h₂ h₃ h₄ + Hom.ext h₁ h₂ h₃ h₄ /-- Constructor for isomorphisms in `Square c` -/ def isoMk {sq₁ sq₂ : Square C} (e₁ : sq₁.X₁ ≅ sq₂.X₁) (e₂ : sq₁.X₂ ≅ sq₂.X₂) diff --git a/Mathlib/CategoryTheory/Subobject/Basic.lean b/Mathlib/CategoryTheory/Subobject/Basic.lean index 9d8d34786e6d6..076606a7a3d8f 100644 --- a/Mathlib/CategoryTheory/Subobject/Basic.lean +++ b/Mathlib/CategoryTheory/Subobject/Basic.lean @@ -250,7 +250,7 @@ theorem mk_le_of_comm {B A : C} {X : Subobject B} {f : A ⟶ B} [Mono f] (g : A /-- To show that two subobjects are equal, it suffices to exhibit an isomorphism commuting with the arrows. -/ -@[ext] +@[ext (iff := false)] theorem eq_of_comm {B : C} {X Y : Subobject B} (f : (X : C) ≅ (Y : C)) (w : f.hom ≫ Y.arrow = X.arrow) : X = Y := le_antisymm (le_of_comm f.hom w) <| le_of_comm f.inv <| f.inv_comp_eq.2 w.symm diff --git a/Mathlib/CategoryTheory/Thin.lean b/Mathlib/CategoryTheory/Thin.lean index 3ec821c540a68..e44b1c223c030 100644 --- a/Mathlib/CategoryTheory/Thin.lean +++ b/Mathlib/CategoryTheory/Thin.lean @@ -44,7 +44,7 @@ variable [Quiver.IsThin C] /-- If `C` is a thin category, then `D ⥤ C` is a thin category. -/ instance functor_thin : Quiver.IsThin (D ⥤ C) := fun _ _ => - ⟨fun α β => NatTrans.ext α β (by subsingleton)⟩ + ⟨fun α β => NatTrans.ext (by subsingleton)⟩ /-- To show `X ≅ Y` in a thin category, it suffices to just give any morphism in each direction. -/ def iso_of_both_ways {X Y : C} (f : X ⟶ Y) (g : Y ⟶ X) : diff --git a/Mathlib/CategoryTheory/Triangulated/Basic.lean b/Mathlib/CategoryTheory/Triangulated/Basic.lean index 2742e034453fb..308f5456a6a29 100644 --- a/Mathlib/CategoryTheory/Triangulated/Basic.lean +++ b/Mathlib/CategoryTheory/Triangulated/Basic.lean @@ -148,7 +148,7 @@ instance triangleCategory : Category (Triangle C) where @[ext] lemma Triangle.hom_ext {A B : Triangle C} (f g : A ⟶ B) (h₁ : f.hom₁ = g.hom₁) (h₂ : f.hom₂ = g.hom₂) (h₃ : f.hom₃ = g.hom₃) : f = g := - TriangleMorphism.ext _ _ h₁ h₂ h₃ + TriangleMorphism.ext h₁ h₂ h₃ @[simp] lemma id_hom₁ (A : Triangle C) : TriangleMorphism.hom₁ (𝟙 A) = 𝟙 _ := rfl diff --git a/Mathlib/CategoryTheory/Yoneda.lean b/Mathlib/CategoryTheory/Yoneda.lean index 4028d93a81389..f3cbf94ce4ae3 100644 --- a/Mathlib/CategoryTheory/Yoneda.lean +++ b/Mathlib/CategoryTheory/Yoneda.lean @@ -344,7 +344,7 @@ def yonedaPairing : Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁) ⥤ Type max u₁ v₁ := @[ext] lemma yonedaPairingExt {X : Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁)} {x y : (yonedaPairing C).obj X} (w : ∀ Y, x.app Y = y.app Y) : x = y := - NatTrans.ext _ _ (funext w) + NatTrans.ext (funext w) @[simp] theorem yonedaPairing_map (P Q : Cᵒᵖ × (Cᵒᵖ ⥤ Type v₁)) (α : P ⟶ Q) (β : (yonedaPairing C).obj P) : @@ -511,7 +511,7 @@ def coyonedaPairing : C × (C ⥤ Type v₁) ⥤ Type max u₁ v₁ := @[ext] lemma coyonedaPairingExt {X : C × (C ⥤ Type v₁)} {x y : (coyonedaPairing C).obj X} (w : ∀ Y, x.app Y = y.app Y) : x = y := - NatTrans.ext _ _ (funext w) + NatTrans.ext (funext w) @[simp] theorem coyonedaPairing_map (P Q : C × (C ⥤ Type v₁)) (α : P ⟶ Q) (β : (coyonedaPairing C).obj P) : diff --git a/Mathlib/Combinatorics/Colex.lean b/Mathlib/Combinatorics/Colex.lean index 67ffe05f45d2b..e4b1b34ce802f 100644 --- a/Mathlib/Combinatorics/Colex.lean +++ b/Mathlib/Combinatorics/Colex.lean @@ -121,7 +121,7 @@ private lemma antisymm_aux (hst : toColex s ≤ toColex t) (hts : toColex t ≤ instance instPartialOrder : PartialOrder (Colex α) where le_refl s a ha ha' := (ha' ha).elim - le_antisymm s t hst hts := Colex.ext _ _ <| (antisymm_aux hst hts).antisymm (antisymm_aux hts hst) + le_antisymm s t hst hts := Colex.ext <| (antisymm_aux hst hts).antisymm (antisymm_aux hts hst) le_trans s t u hst htu a has hau := by by_cases hat : a ∈ ofColex t · have ⟨b, hbu, hbt, hab⟩ := htu hat hau @@ -217,7 +217,7 @@ section DecidableEq variable [DecidableEq α] instance instDecidableEq : DecidableEq (Colex α) := fun s t ↦ - decidable_of_iff' (s.ofColex = t.ofColex) <| Colex.ext_iff _ _ + decidable_of_iff' (s.ofColex = t.ofColex) Colex.ext_iff instance instDecidableLE [@DecidableRel α (· ≤ ·)] : @DecidableRel (Colex α) (· ≤ ·) := fun s t ↦ decidable_of_iff' diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index 5f3f66ff4d4ec..401c788269c57 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -187,7 +187,7 @@ theorem sizeUpTo_zero : c.sizeUpTo 0 = 0 := by simp [sizeUpTo] theorem sizeUpTo_ofLength_le (i : ℕ) (h : c.length ≤ i) : c.sizeUpTo i = n := by dsimp [sizeUpTo] convert c.blocks_sum - exact take_all_of_le h + exact take_of_length_le h @[simp] theorem sizeUpTo_length : c.sizeUpTo c.length = n := @@ -868,7 +868,7 @@ theorem mem_boundaries_iff_exists_blocks_sum_take_eq {j : Fin (n + 1)} : exact this.symm theorem blocks_sum : c.blocks.sum = n := by - have : c.blocks.take c.length = c.blocks := take_all_of_le (by simp [blocks]) + have : c.blocks.take c.length = c.blocks := take_of_length_le (by simp [blocks]) rw [← this, c.blocks_partial_sum c.length_lt_card_boundaries, c.boundary_length] rfl diff --git a/Mathlib/Combinatorics/Enumerative/Partition.lean b/Mathlib/Combinatorics/Enumerative/Partition.lean index eaba453a5c731..d88c1fba8060d 100644 --- a/Mathlib/Combinatorics/Enumerative/Partition.lean +++ b/Mathlib/Combinatorics/Enumerative/Partition.lean @@ -61,7 +61,7 @@ namespace Partition -- TODO: This should be automatically derived, see lean4#2914 instance decidableEqPartition {n : ℕ} : DecidableEq (Partition n) := - fun _ _ => decidable_of_iff' _ <| Partition.ext_iff _ _ + fun _ _ => decidable_of_iff' _ Partition.ext_iff /-- A composition induces a partition (just convert the list to a multiset). -/ @[simps] @@ -73,7 +73,7 @@ def ofComposition (n : ℕ) (c : Composition n) : Partition n where theorem ofComposition_surj {n : ℕ} : Function.Surjective (ofComposition n) := by rintro ⟨b, hb₁, hb₂⟩ induction b using Quotient.inductionOn with | _ b => ?_ - exact ⟨⟨b, hb₁, by simpa using hb₂⟩, Partition.ext _ _ rfl⟩ + exact ⟨⟨b, hb₁, by simpa using hb₂⟩, Partition.ext rfl⟩ -- The argument `n` is kept explicit here since it is useful in tactic mode proofs to generate the -- proof obligation `l.sum = n`. @@ -132,7 +132,7 @@ instance {n : ℕ} : Inhabited (Partition n) := ⟨indiscrete n⟩ eq_zero_of_forall_not_mem fun _ h => (p.parts_pos h).ne' <| sum_eq_zero_iff.1 p.parts_sum _ h instance UniquePartitionZero : Unique (Partition 0) where - uniq _ := Partition.ext _ _ <| by simp + uniq _ := Partition.ext <| by simp @[simp] lemma partition_one_parts (p : Partition 1) : p.parts = {1} := by have h : p.parts = replicate (card p.parts) 1 := eq_replicate_card.2 fun x hx => @@ -141,7 +141,7 @@ instance UniquePartitionZero : Unique (Partition 0) where rw [h, h', replicate_one] instance UniquePartitionOne : Unique (Partition 1) where - uniq _ := Partition.ext _ _ <| by simp + uniq _ := Partition.ext <| by simp @[simp] lemma ofSym_one (s : Sym σ 1) : ofSym s = indiscrete 1 := by ext; simp diff --git a/Mathlib/Combinatorics/Quiver/Basic.lean b/Mathlib/Combinatorics/Quiver/Basic.lean index 3f5a48d934be3..32fe7fd00a050 100644 --- a/Mathlib/Combinatorics/Quiver/Basic.lean +++ b/Mathlib/Combinatorics/Quiver/Basic.lean @@ -67,7 +67,7 @@ lemma mk_obj {V W : Type*} [Quiver V] [Quiver W] {obj : V → W} {map} {X : V} : lemma mk_map {V W : Type*} [Quiver V] [Quiver W] {obj : V → W} {map} {X Y : V} {f : X ⟶ Y} : (Prefunctor.mk obj map).map f = map f := rfl -@[ext] +@[ext (iff := false)] theorem ext {V : Type u} [Quiver.{v₁} V] {W : Type u₂} [Quiver.{v₂} W] {F G : Prefunctor V W} (h_obj : ∀ X, F.obj X = G.obj X) (h_map : ∀ (X Y : V) (f : X ⟶ Y), diff --git a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean index 2cb9bcfec2287..6b8117d6925da 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean @@ -170,12 +170,11 @@ lemma IsTree.card_edgeFinset [Fintype V] [Fintype G.edgeSet] (hG : G.IsTree) : · exact (congrArg (·.fst) h) · have h1 : ((f a).firstDart <| not_nil_of_ne (by simpa using ha)).snd = b := congrArg (·.snd) h - have h3 := congrArg length (hf' _ ((f _).tail.copy h1 rfl) ?_) - · rw [length_copy, ← add_left_inj 1, - length_tail_add_one (not_nil_of_ne (by simpa using ha))] at h3 + have h3 := congrArg length (hf' _ (((f _).tail _).copy h1 rfl) ?_) + · rw [length_copy, ← add_left_inj 1, length_tail_add_one] at h3 omega · simp only [ne_eq, eq_mp_eq_cast, id_eq, isPath_copy] - exact (hf _).tail (not_nil_of_ne (by simpa using ha)) + exact (hf _).tail _ case surj => simp only [mem_edgeFinset, Finset.mem_compl, Finset.mem_singleton, Sym2.forall, mem_edgeSet] intros x y h @@ -189,7 +188,7 @@ lemma IsTree.card_edgeFinset [Fintype V] [Fintype G.edgeSet] (hG : G.IsTree) : length_cons, length_nil] at h' simp [Nat.le_zero, Nat.one_ne_zero] at h' rw [← hf' _ (.cons h.symm (f x)) ((cons_isPath_iff _ _).2 ⟨hf _, fun hy => ?contra⟩)] - · simp only [firstDart_toProd, getVert_cons_succ, getVert_zero, Prod.swap_prod_mk] + · rfl case contra => suffices (f x).takeUntil y hy = .cons h .nil by rw [← take_spec _ hy] at h' diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index 682cff527451b..fff296e425ccc 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -187,7 +187,7 @@ theorem ne_of_adj_of_not_adj {v w x : V} (h : G.Adj v x) (hn : ¬G.Adj w x) : v hn (h' ▸ h) theorem adj_injective : Injective (Adj : SimpleGraph V → V → V → Prop) := - SimpleGraph.ext + fun _ _ => SimpleGraph.ext @[simp] theorem adj_inj {G H : SimpleGraph V} : G.Adj = H.Adj ↔ G = H := diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity/Subgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity/Subgraph.lean index b19fdc540c1bb..e73b22eb31d5a 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity/Subgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity/Subgraph.lean @@ -189,9 +189,9 @@ theorem toSubgraph_adj_getVert {u v} (w : G.Walk u v) {i : ℕ} (hi : i < w.leng | nil => cases hi | cons hxy i' ih => cases i - · simp only [Walk.toSubgraph, Walk.getVert_zero, zero_add, getVert_cons_succ, Subgraph.sup_adj, + · simp only [Walk.toSubgraph, Walk.getVert_zero, zero_add, cons_getVert_succ, Subgraph.sup_adj, subgraphOfAdj_adj, true_or] - · simp only [Walk.toSubgraph, getVert_cons_succ, Subgraph.sup_adj, subgraphOfAdj_adj, Sym2.eq, + · simp only [Walk.toSubgraph, cons_getVert_succ, Subgraph.sup_adj, subgraphOfAdj_adj, Sym2.eq, Sym2.rel_iff', Prod.mk.injEq, Prod.swap_prod_mk] right exact ih (Nat.succ_lt_succ_iff.mp hi) @@ -211,7 +211,7 @@ theorem toSubgraph_adj_iff {u v u' v'} (w : G.Walk u v) : cases hadj with | inl hl => use 0 - simp only [Walk.getVert_zero, zero_add, getVert_cons_succ] + simp only [Walk.getVert_zero, zero_add, cons_getVert_succ] refine ⟨?_, by simp only [length_cons, Nat.zero_lt_succ]⟩ simp only [Sym2.eq, Sym2.rel_iff', Prod.mk.injEq, Prod.swap_prod_mk] cases hl with @@ -220,7 +220,7 @@ theorem toSubgraph_adj_iff {u v u' v'} (w : G.Walk u v) : | inr hr => obtain ⟨i, hi⟩ := (toSubgraph_adj_iff _).mp hr use i + 1 - simp only [getVert_cons_succ] + simp only [cons_getVert_succ] constructor · exact hi.1 · simp only [Walk.length_cons, add_lt_add_iff_right, Nat.add_lt_add_right hi.2 1] diff --git a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean index ce4ffb9224f02..1f9b4d949285c 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean @@ -67,7 +67,7 @@ end /-- A hamiltonian cycle is a cycle that visits every vertex once. -/ structure IsHamiltonianCycle (p : G.Walk a a) extends p.IsCycle : Prop := - isHamiltonian_tail : p.tail.IsHamiltonian + isHamiltonian_tail : (p.tail toIsCycle.not_nil).IsHamiltonian variable {p : G.Walk a a} @@ -79,28 +79,27 @@ lemma IsHamiltonianCycle.map {H : SimpleGraph β} (f : G →g H) (hf : Bijective toIsCycle := hp.isCycle.map hf.injective isHamiltonian_tail := by simp only [IsHamiltonian, support_tail, support_map, ne_eq, List.map_eq_nil, support_ne_nil, - not_false_eq_true, List.count_tail, hf.surjective.forall, hf.injective, - List.count_map_of_injective] + not_false_eq_true, List.count_tail, List.head_map, beq_iff_eq, hf.surjective.forall, + hf.injective, List.count_map_of_injective] intro x rcases p with (_ | ⟨y, p⟩) · cases hp.ne_nil rfl - simp only [map_cons, getVert_cons_succ, tail_cons_eq, support_copy,support_map] - rw [List.count_map_of_injective _ _ hf.injective, ← support_copy, ← tail_cons_eq] + simp only [support_cons, List.count_cons, beq_iff_eq, List.head_cons, hf.injective.eq_iff, + add_tsub_cancel_right] exact hp.isHamiltonian_tail _ lemma isHamiltonianCycle_isCycle_and_isHamiltonian_tail : - p.IsHamiltonianCycle ↔ p.IsCycle ∧ p.tail.IsHamiltonian := + p.IsHamiltonianCycle ↔ ∃ h : p.IsCycle, (p.tail h.not_nil).IsHamiltonian := ⟨fun ⟨h, h'⟩ ↦ ⟨h, h'⟩, fun ⟨h, h'⟩ ↦ ⟨h, h'⟩⟩ lemma isHamiltonianCycle_iff_isCycle_and_support_count_tail_eq_one : p.IsHamiltonianCycle ↔ p.IsCycle ∧ ∀ a, (support p).tail.count a = 1 := by - simp (config := { contextual := true }) [isHamiltonianCycle_isCycle_and_isHamiltonian_tail, - IsHamiltonian, support_tail, IsCycle.not_nil, exists_prop] + simp only [isHamiltonianCycle_isCycle_and_isHamiltonian_tail, IsHamiltonian, support_tail, + exists_prop] /-- A hamiltonian cycle visits every vertex. -/ lemma IsHamiltonianCycle.mem_support (hp : p.IsHamiltonianCycle) (b : α) : - b ∈ p.support := - List.mem_of_mem_tail <| support_tail p hp.1.not_nil ▸ hp.isHamiltonian_tail.mem_support _ + b ∈ p.support := List.mem_of_mem_tail <| support_tail p _ ▸ hp.isHamiltonian_tail.mem_support _ /-- The length of a hamiltonian cycle is the number of vertices. -/ lemma IsHamiltonianCycle.length_eq [Fintype α] (hp : p.IsHamiltonianCycle) : @@ -111,11 +110,11 @@ lemma IsHamiltonianCycle.length_eq [Fintype α] (hp : p.IsHamiltonianCycle) : lemma IsHamiltonianCycle.count_support_self (hp : p.IsHamiltonianCycle) : p.support.count a = 2 := by - rw [support_eq_cons, List.count_cons_self, ← support_tail _ hp.1.not_nil, hp.isHamiltonian_tail] + rw [support_eq_cons, List.count_cons_self, ← support_tail, hp.isHamiltonian_tail] lemma IsHamiltonianCycle.support_count_of_ne (hp : p.IsHamiltonianCycle) (h : a ≠ b) : p.support.count b = 1 := by - rw [← cons_support_tail p hp.1.not_nil, List.count_cons_of_ne h.symm, hp.isHamiltonian_tail] + rw [← cons_support_tail p, List.count_cons_of_ne h.symm, hp.isHamiltonian_tail] end Walk diff --git a/Mathlib/Combinatorics/SimpleGraph/Maps.lean b/Mathlib/Combinatorics/SimpleGraph/Maps.lean index 8ab1314618108..8ccd1909b18de 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Maps.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Maps.lean @@ -74,10 +74,10 @@ theorem map_monotone (f : V ↪ W) : Monotone (SimpleGraph.map f) := by exact ⟨_, _, h ha, rfl, rfl⟩ @[simp] lemma map_id : G.map (Function.Embedding.refl _) = G := - SimpleGraph.ext _ _ <| Relation.map_id_id _ + SimpleGraph.ext <| Relation.map_id_id _ @[simp] lemma map_map (f : V ↪ W) (g : W ↪ X) : (G.map f).map g = G.map (f.trans g) := - SimpleGraph.ext _ _ <| Relation.map_map _ _ _ _ _ + SimpleGraph.ext <| Relation.map_map _ _ _ _ _ /-- Given a function, there is a contravariant induced map on graphs by pulling back the adjacency relation. @@ -92,7 +92,7 @@ protected def comap (f : V → W) (G : SimpleGraph W) : SimpleGraph V where @[simp] lemma comap_adj {G : SimpleGraph W} {f : V → W} : (G.comap f).Adj u v ↔ G.Adj (f u) (f v) := Iff.rfl -@[simp] lemma comap_id {G : SimpleGraph V} : G.comap id = G := SimpleGraph.ext _ _ rfl +@[simp] lemma comap_id {G : SimpleGraph V} : G.comap id = G := SimpleGraph.ext rfl @[simp] lemma comap_comap {G : SimpleGraph X} (f : V → W) (g : W → X) : (G.comap g).comap f = G.comap (g ∘ f) := rfl diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index 0c7b176088328..7dd0f77cbd996 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -236,7 +236,7 @@ theorem cons_isCycle_iff {u v : V} (p : G.Walk v u) (h : G.Adj u v) : have : p.support.Nodup → p.edges.Nodup := edges_nodup_of_support_nodup tauto -lemma IsPath.tail {p : G.Walk u v} (hp : p.IsPath) (hp' : ¬ p.Nil) : p.tail.IsPath := by +lemma IsPath.tail {p : G.Walk u v} (hp : p.IsPath) (hp' : ¬ p.Nil) : (p.tail hp').IsPath := by rw [Walk.isPath_def] at hp ⊢ rw [← cons_support_tail _ hp', List.nodup_cons] at hp exact hp.2 diff --git a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean index bb0369d4dba9b..68777353d2638 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Subgraph.lean @@ -245,7 +245,7 @@ def copy (G' : Subgraph G) (V'' : Set V) (hV : V'' = G'.verts) theorem copy_eq (G' : Subgraph G) (V'' : Set V) (hV : V'' = G'.verts) (adj' : V → V → Prop) (hadj : adj' = G'.Adj) : G'.copy V'' hV adj' hadj = G' := - Subgraph.ext _ _ hV hadj + Subgraph.ext hV hadj /-- The union of two subgraphs. -/ instance : Sup G.Subgraph where @@ -381,7 +381,7 @@ theorem verts_spanningCoe_injective : (fun G' : Subgraph G => (G'.verts, G'.spanningCoe)).Injective := by intro G₁ G₂ h rw [Prod.ext_iff] at h - exact Subgraph.ext _ _ h.1 (spanningCoe_inj.1 h.2) + exact Subgraph.ext h.1 (spanningCoe_inj.1 h.2) /-- For subgraphs `G₁`, `G₂`, `G₁ ≤ G₂` iff `G₁.verts ⊆ G₂.verts` and `∀ a b, G₁.adj a b → G₂.adj a b`. -/ @@ -419,7 +419,7 @@ def completelyDistribLatticeMinimalAxioms : CompletelyDistribLattice.MinimalAxio le_sInf := fun s G' hG' => ⟨Set.subset_iInter₂ fun H hH => (hG' _ hH).1, fun a b hab => ⟨fun H hH => (hG' _ hH).2 hab, G'.adj_sub hab⟩⟩ - iInf_iSup_eq := fun f => Subgraph.ext _ _ (by simpa using iInf_iSup_eq) + iInf_iSup_eq := fun f => Subgraph.ext (by simpa using iInf_iSup_eq) (by ext; simp [Classical.skolem]) } instance : CompletelyDistribLattice G.Subgraph := @@ -847,7 +847,7 @@ theorem neighborSet_subgraphOfAdj_of_ne_of_ne {u v w : V} (hvw : G.Adj v w) (hv theorem neighborSet_subgraphOfAdj [DecidableEq V] {u v w : V} (hvw : G.Adj v w) : (G.subgraphOfAdj hvw).neighborSet u = (if u = v then {w} else ∅) ∪ if u = w then {v} else ∅ := by - split_ifs <;> subst_vars <;> simp [*, Set.singleton_def] + split_ifs <;> subst_vars <;> simp [*] theorem singletonSubgraph_fst_le_subgraphOfAdj {u v : V} {h : G.Adj u v} : G.singletonSubgraph u ≤ G.subgraphOfAdj h := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index c6972002e049b..a72ceadbd3573 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -177,21 +177,16 @@ theorem adj_getVert_succ {u v} (w : G.Walk u v) {i : ℕ} (hi : i < w.length) : · simp [getVert, hxy] · exact ih (Nat.succ_lt_succ_iff.1 hi) -lemma getVert_cons_one {u v w} (q : G.Walk v w) (hadj : G.Adj u v) : - (q.cons hadj).getVert 1 = v := by - have : (q.cons hadj).getVert 1 = q.getVert 0 := rfl - simpa [getVert_zero] using this - @[simp] -lemma getVert_cons_succ {u v w n} (p : G.Walk v w) (h : G.Adj u v) : +lemma cons_getVert_succ {u v w n} (p : G.Walk v w) (h : G.Adj u v) : (p.cons h).getVert (n + 1) = p.getVert n := rfl -lemma getVert_cons {u v w n} (p : G.Walk v w) (h : G.Adj u v) (hn : n ≠ 0) : +lemma cons_getVert {u v w n} (p : G.Walk v w) (h : G.Adj u v) (hn : n ≠ 0) : (p.cons h).getVert n = p.getVert (n - 1) := by obtain ⟨i, hi⟩ : ∃ (i : ℕ), i.succ = n := by use n - 1; exact Nat.succ_pred_eq_of_ne_zero hn rw [← hi] - simp only [Nat.succ_eq_add_one, getVert_cons_succ, Nat.add_sub_cancel] + simp only [Nat.succ_eq_add_one, cons_getVert_succ, Nat.add_sub_cancel] @[simp] theorem cons_append {u v w x : V} (h : G.Adj u v) (p : G.Walk v w) (q : G.Walk w x) : @@ -746,9 +741,6 @@ lemma nil_iff_support_eq {p : G.Walk v w} : p.Nil ↔ p.support = [v] := by lemma nil_iff_length_eq {p : G.Walk v w} : p.Nil ↔ p.length = 0 := by cases p <;> simp -lemma not_nil_iff_lt_length {p : G.Walk v w} : ¬ p.Nil ↔ 0 < p.length := by - cases p <;> simp - lemma not_nil_iff {p : G.Walk v w} : ¬ p.Nil ↔ ∃ (u : V) (h : G.Adj v u) (q : G.Walk u w), p = cons h q := by cases p <;> simp [*] @@ -773,77 +765,61 @@ lemma notNilRec_cons {motive : {u w : V} → (p : G.Walk u w) → ¬ p.Nil → S motive (q.cons h) Walk.not_nil_cons) (h' : G.Adj u v) (q' : G.Walk v w) : @Walk.notNilRec _ _ _ _ _ cons _ _ = cons h' q' := by rfl -@[simp] lemma adj_getVert_one {p : G.Walk v w} (hp : ¬ p.Nil) : - G.Adj v (p.getVert 1) := by - simpa using adj_getVert_succ p (by simpa [not_nil_iff_lt_length] using hp : 0 < p.length) +/-- The second vertex along a non-nil walk. -/ +def sndOfNotNil (p : G.Walk v w) (hp : ¬ p.Nil) : V := + p.notNilRec (@fun _ u _ _ _ => u) hp -/-- The walk obtained by removing the first `n` darts of a walk. -/ -def drop {u v : V} (p : G.Walk u v) (n : ℕ) : G.Walk (p.getVert n) v := - match p, n with - | .nil, _ => .nil - | p, 0 => p.copy (getVert_zero p).symm rfl - | .cons h q, (n + 1) => (q.drop n).copy (getVert_cons_succ _ h).symm rfl +@[simp] lemma adj_sndOfNotNil {p : G.Walk v w} (hp : ¬ p.Nil) : + G.Adj v (p.sndOfNotNil hp) := + p.notNilRec (fun h _ => h) hp /-- The walk obtained by removing the first dart of a non-nil walk. -/ -def tail (p : G.Walk u v) : G.Walk (p.getVert 1) v := p.drop 1 - -@[simp] -lemma tail_cons_nil (h : G.Adj u v) : (Walk.cons h .nil).tail = .nil := by rfl - -lemma tail_cons_eq (h : G.Adj u v) (p : G.Walk v w) : - (p.cons h).tail = p.copy (getVert_zero p).symm rfl := by - match p with - | .nil => rfl - | .cons h q => rfl +def tail (p : G.Walk u v) (hp : ¬ p.Nil) : G.Walk (p.sndOfNotNil hp) v := + p.notNilRec (fun _ q => q) hp /-- The first dart of a walk. -/ @[simps] def firstDart (p : G.Walk v w) (hp : ¬ p.Nil) : G.Dart where fst := v - snd := p.getVert 1 - adj := p.adj_getVert_one hp + snd := p.sndOfNotNil hp + adj := p.adj_sndOfNotNil hp lemma edge_firstDart (p : G.Walk v w) (hp : ¬ p.Nil) : - (p.firstDart hp).edge = s(v, p.getVert 1) := rfl + (p.firstDart hp).edge = s(v, p.sndOfNotNil hp) := rfl variable {x y : V} -- TODO: rename to u, v, w instead? -lemma cons_tail_eq (p : G.Walk x y) (hp : ¬ p.Nil) : - cons (p.adj_getVert_one hp) p.tail = p := by - cases p with - | nil => simp only [nil_nil, not_true_eq_false] at hp - | cons h q => - simp only [getVert_cons_succ, tail_cons_eq, cons_copy, copy_rfl_rfl] +@[simp] lemma cons_tail_eq (p : G.Walk x y) (hp : ¬ p.Nil) : + cons (p.adj_sndOfNotNil hp) (p.tail hp) = p := + p.notNilRec (fun _ _ => rfl) hp @[simp] lemma cons_support_tail (p : G.Walk x y) (hp : ¬p.Nil) : - x :: p.tail.support = p.support := by - rw [← support_cons, cons_tail_eq _ hp] + x :: (p.tail hp).support = p.support := by + rw [← support_cons, cons_tail_eq] @[simp] lemma length_tail_add_one {p : G.Walk x y} (hp : ¬ p.Nil) : - p.tail.length + 1 = p.length := by - rw [← length_cons, cons_tail_eq _ hp] + (p.tail hp).length + 1 = p.length := by + rw [← length_cons, cons_tail_eq] @[simp] lemma nil_copy {x' y' : V} {p : G.Walk x y} (hx : x = x') (hy : y = y') : (p.copy hx hy).Nil = p.Nil := by subst_vars; rfl -@[simp] lemma support_tail (p : G.Walk v v) (hp : ¬ p.Nil) : - p.tail.support = p.support.tail := by +@[simp] lemma support_tail (p : G.Walk v v) (hp) : + (p.tail hp).support = p.support.tail := by rw [← cons_support_tail p hp, List.tail_cons] @[simp] lemma tail_cons {t u v} (p : G.Walk u v) (h : G.Adj t u) : - (p.cons h).tail = p.copy (getVert_zero p).symm rfl := by - match p with - | .nil => rfl - | .cons h q => rfl + (p.cons h).tail not_nil_cons = p := by + unfold Walk.tail; simp only [notNilRec_cons] -lemma support_tail_of_not_nil (p : G.Walk u v) (hnp : ¬p.Nil) : - p.tail.support = p.support.tail := by - match p with - | .nil => simp only [nil_nil, not_true_eq_false] at hnp - | .cons h q => - simp only [tail_cons, getVert_cons_succ, support_copy, support_cons, List.tail_cons] +lemma tail_support_eq_support_tail (p : G.Walk u v) (hnp : ¬p.Nil) : + (p.tail hnp).support = p.support.tail := + p.notNilRec (by + intro u v w huv q + unfold Walk.tail + simp only [notNilRec_cons, Walk.support_cons, List.tail_cons]) hnp /-! ### Walk decompositions -/ @@ -925,7 +901,7 @@ theorem count_edges_takeUntil_le_one {u v w : V} (p : G.Walk v w) (h : u ∈ p.s simp · rw [edges_cons, List.count_cons] split_ifs with h'' - · rw [Sym2.eq_iff] at h'' + · simp only [beq_iff_eq, Sym2.eq, Sym2.rel_iff', Prod.mk.injEq, Prod.swap_prod_mk] at h'' obtain ⟨rfl, rfl⟩ | ⟨rfl, rfl⟩ := h'' · exact (h' rfl).elim · cases p' <;> simp! @@ -1019,23 +995,17 @@ theorem exists_boundary_dart {u v : V} (p : G.Walk u v) (S : Set V) (uS : u ∈ exact ⟨d, List.Mem.tail _ hd, hcd⟩ · exact ⟨⟨(x, y), a⟩, List.Mem.head _, uS, h⟩ -@[simp] lemma getVert_copy {u v w x : V} (p : G.Walk u v) (i : ℕ) (h : u = w) (h' : v = x) : - (p.copy h h').getVert i = p.getVert i := by - subst_vars - match p, i with - | .nil, _ => - rw [getVert_of_length_le _ (by simp only [length_nil, Nat.zero_le] : nil.length ≤ _)] - rw [getVert_of_length_le _ (by simp only [length_copy, length_nil, Nat.zero_le])] - | .cons hadj q, 0 => simp only [copy_rfl_rfl, getVert_zero] - | .cons hadj q, (n + 1) => simp only [copy_cons, getVert_cons_succ]; rfl - -@[simp] lemma getVert_tail {u v n} (p : G.Walk u v) (hnp: ¬ p.Nil) : - p.tail.getVert n = p.getVert (n + 1) := by - match p with - | .nil => rfl - | .cons h q => - simp only [getVert_cons_succ, tail_cons_eq, getVert_cons] - exact getVert_copy q n (getVert_zero q).symm rfl +lemma getVert_tail {u v n} (p : G.Walk u v) (hnp: ¬ p.Nil) : + (p.tail hnp).getVert n = p.getVert (n + 1) := + p.notNilRec (fun _ _ ↦ by simp only [tail_cons, cons_getVert_succ]) hnp + +@[simp] +lemma cons_sndOfNotNil (q : G.Walk v w) (hadj : G.Adj u v) : + (q.cons hadj).sndOfNotNil not_nil_cons = v := by + unfold sndOfNotNil; simp only [notNilRec_cons] + +lemma getVert_one (p : G.Walk u v) (hnp : ¬ p.Nil) : p.getVert 1 = p.sndOfNotNil hnp := + p.notNilRec (fun _ _ ↦ by simp only [cons_getVert_succ, getVert_zero, cons_sndOfNotNil]) hnp /-- Given a walk `w` and a node in the support, there exists a natural `n`, such that given node is the `n`-th node (zero-indexed) in the walk. In addition, `n` is at most the length of the path. @@ -1062,18 +1032,13 @@ theorem mem_support_iff_exists_getVert {u v w : V} {p : G.Walk v w} : rw [@nil_iff_length_eq] have : 1 ≤ p.length := by omega exact Nat.not_eq_zero_of_lt this - rw [← support_tail_of_not_nil _ hnp] + rw [← tail_support_eq_support_tail _ hnp] rw [mem_support_iff_exists_getVert] use n - 1 - simp only [Nat.sub_le_iff_le_add] - rw [getVert_tail _ hnp, length_tail_add_one hnp] - have : (n - 1 + 1) = n:= by omega + simp only [Nat.sub_le_iff_le_add, length_tail_add_one, getVert_tail] + have : n - 1 + 1 = n := by omega rwa [this] termination_by p.length -decreasing_by -· simp_wf - rw [@Nat.lt_iff_add_one_le] - rw [length_tail_add_one hnp] end Walk diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index 4394986d3be11..28f554c1f9089 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -763,7 +763,7 @@ private theorem list_foldl' {f : α → List β} {g : α → σ} {h : α → σ hG) suffices ∀ a n, F a n = (((f a).take n).foldl (fun s b => h a (s, b)) (g a), (f a).drop n) by refine hF.of_eq fun a => ?_ - rw [this, List.take_all_of_le (length_le_encode _)] + rw [this, List.take_of_length_le (length_le_encode _)] introv dsimp only [F] generalize f a = l diff --git a/Mathlib/Control/Traversable/Basic.lean b/Mathlib/Control/Traversable/Basic.lean index 3f7fa9369afdf..6c338a28da53e 100644 --- a/Mathlib/Control/Traversable/Basic.lean +++ b/Mathlib/Control/Traversable/Basic.lean @@ -117,10 +117,6 @@ theorem ext ⦃η η' : ApplicativeTransformation F G⦄ (h : ∀ (α : Type u) ext1 α exact funext (h α) -theorem ext_iff {η η' : ApplicativeTransformation F G} : - η = η' ↔ ∀ (α : Type u) (x : F α), η x = η' x := - ⟨fun h _ _ => h ▸ rfl, fun h => ext h⟩ - section Preserves variable (η : ApplicativeTransformation F G) diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 0a58507f00062..a24c1a1d55375 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -57,9 +57,6 @@ theorem ext : ∀ {z w : ℂ}, z.re = w.re → z.im = w.im → z = w attribute [local ext] Complex.ext -protected theorem ext_iff {z w : ℂ} : z = w ↔ z.re = w.re ∧ z.im = w.im := - ⟨fun H => by simp [H], fun h => ext h.1 h.2⟩ - theorem re_surjective : Surjective re := fun x => ⟨⟨x, 0⟩, rfl⟩ theorem im_surjective : Surjective im := fun y => ⟨⟨0, y⟩, rfl⟩ diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index 471b289ab0ed9..a2b9e02f3bc43 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -120,7 +120,7 @@ theorem antidiagonalTuple_zero_right : ∀ k, antidiagonalTuple k 0 = [0] @[simp] theorem antidiagonalTuple_one (n : ℕ) : antidiagonalTuple 1 n = [![n]] := by simp_rw [antidiagonalTuple, antidiagonal, List.range_succ, List.map_append, List.map_singleton, - tsub_self, List.append_bind, List.bind_singleton, List.bind_map] + tsub_self, List.bind_append, List.bind_singleton, List.bind_map] conv_rhs => rw [← List.nil_append [![n]]] congr 1 simp_rw [List.bind_eq_nil, List.mem_range, List.map_eq_nil] diff --git a/Mathlib/Data/Finmap.lean b/Mathlib/Data/Finmap.lean index d6c023b907e27..fb4a761948449 100644 --- a/Mathlib/Data/Finmap.lean +++ b/Mathlib/Data/Finmap.lean @@ -143,9 +143,6 @@ theorem induction_on₃ {C : Finmap β → Finmap β → Finmap β → Prop} (s theorem ext : ∀ {s t : Finmap β}, s.entries = t.entries → s = t | ⟨l₁, h₁⟩, ⟨l₂, _⟩, H => by congr -protected theorem ext_iff {s t : Finmap β} : s = t ↔ s.entries = t.entries := - ⟨congr_arg _, ext⟩ - @[simp] theorem ext_iff' {s t : Finmap β} : s.entries = t.entries ↔ s = t := Finmap.ext_iff.symm diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 0814c689f966e..b0dfbf24bd6a0 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -2806,9 +2806,7 @@ theorem mem_toList {a : α} {s : Finset α} : a ∈ s.toList ↔ a ∈ s := theorem toList_eq_nil {s : Finset α} : s.toList = [] ↔ s = ∅ := Multiset.toList_eq_nil.trans val_eq_zero -@[simp] -theorem empty_toList {s : Finset α} : s.toList.isEmpty ↔ s = ∅ := - List.isEmpty_iff_eq_nil.trans toList_eq_nil +theorem empty_toList {s : Finset α} : s.toList.isEmpty ↔ s = ∅ := by simp @[simp] theorem toList_empty : (∅ : Finset α).toList = [] := diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 279030ab43eec..c58fab65e2979 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -212,9 +212,6 @@ noncomputable def _root_.Equiv.finsuppUnique {ι : Type*} [Unique ι] : (ι → theorem unique_ext [Unique α] {f g : α →₀ M} (h : f default = g default) : f = g := ext fun a => by rwa [Unique.eq_default a] -protected theorem unique_ext_iff [Unique α] {f g : α →₀ M} : f = g ↔ f default = g default := - ⟨fun h => h ▸ rfl, unique_ext⟩ - end Basic /-! ### Declarations about `single` -/ diff --git a/Mathlib/Data/Fintype/Fin.lean b/Mathlib/Data/Fintype/Fin.lean index e0b09df15d902..bdd198e51a477 100644 --- a/Mathlib/Data/Fintype/Fin.lean +++ b/Mathlib/Data/Fintype/Fin.lean @@ -64,10 +64,10 @@ theorem card_filter_univ_succ (p : Fin (n + 1) → Prop) [DecidablePred p] : (card_filter_univ_succ' p).trans (by split_ifs <;> simp [add_comm 1]) theorem card_filter_univ_eq_vector_get_eq_count [DecidableEq α] (a : α) (v : Vector α n) : - (univ.filter fun i => a = v.get i).card = v.toList.count a := by + (univ.filter fun i => v.get i = a).card = v.toList.count a := by induction' v with n x xs hxs · simp · simp_rw [card_filter_univ_succ', Vector.get_cons_zero, Vector.toList_cons, Function.comp, - Vector.get_cons_succ, hxs, List.count_cons, add_comm (ite (a = x) 1 0)] + Vector.get_cons_succ, hxs, List.count_cons, add_comm (ite (x = a) 1 0), beq_iff_eq] end Fin diff --git a/Mathlib/Data/List/AList.lean b/Mathlib/Data/List/AList.lean index 9c586d105a495..1498c2afd3a09 100644 --- a/Mathlib/Data/List/AList.lean +++ b/Mathlib/Data/List/AList.lean @@ -60,9 +60,6 @@ namespace AList theorem ext : ∀ {s t : AList β}, s.entries = t.entries → s = t | ⟨l₁, h₁⟩, ⟨l₂, _⟩, H => by congr -protected theorem ext_iff {s t : AList β} : s = t ↔ s.entries = t.entries := - ⟨congr_arg _, ext⟩ - instance [DecidableEq α] [∀ a, DecidableEq (β a)] : DecidableEq (AList β) := fun xs ys => by rw [AList.ext_iff]; infer_instance diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index bd2aa377378ce..2bbfd9513f043 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -12,7 +12,6 @@ import Mathlib.Init.Data.List.Lemmas import Mathlib.Logic.Unique import Mathlib.Order.Basic import Mathlib.Tactic.Common -import Batteries.Data.List.EraseIdx import Batteries.Data.List.Perm /-! @@ -98,7 +97,7 @@ attribute [simp] List.mem_bind /-! ### length -/ -alias ⟨ne_nil_of_length_pos, length_pos_of_ne_nil⟩ := length_pos +alias ⟨_, length_pos_of_ne_nil⟩ := length_pos theorem length_pos_iff_ne_nil {l : List α} : 0 < length l ↔ l ≠ [] := ⟨ne_nil_of_length_pos, length_pos_of_ne_nil⟩ @@ -278,9 +277,6 @@ theorem replicate_left_injective (a : α) : Injective (replicate · a) := theorem replicate_left_inj {a : α} {n m : ℕ} : replicate n a = replicate m a ↔ n = m := (replicate_left_injective a).eq_iff -@[simp] theorem head_replicate (n : ℕ) (a : α) (h) : head (replicate n a) h = a := by - cases n <;> simp [replicate_succ] at h ⊢ - /-! ### pure -/ theorem mem_pure (x y : α) : x ∈ (pure y : List α) ↔ x = y := by simp @@ -327,9 +323,6 @@ theorem reverse_bijective : Bijective (@reverse α) := theorem reverse_inj {l₁ l₂ : List α} : reverse l₁ = reverse l₂ ↔ l₁ = l₂ := reverse_injective.eq_iff -theorem reverse_eq_iff {l l' : List α} : l.reverse = l' ↔ l = l'.reverse := - reverse_involutive.eq_iff - theorem concat_eq_reverse_cons (a : α) (l : List α) : concat l a = reverse (a :: reverse l) := by simp only [concat_eq_append, reverse_cons, reverse_reverse] @@ -348,28 +341,23 @@ theorem isEmpty_iff_eq_nil {l : List α} : l.isEmpty ↔ l = [] := by cases l <; /-! ### getLast -/ -@[simp] -theorem getLast_cons {a : α} {l : List α} : - ∀ h : l ≠ nil, getLast (a :: l) (cons_ne_nil a l) = getLast l h := by - induction l <;> intros - · contradiction - · rfl +attribute [simp] getLast_cons theorem getLast_append_singleton {a : α} (l : List α) : - getLast (l ++ [a]) (append_ne_nil_of_ne_nil_right l _ (cons_ne_nil a _)) = a := by - simp only [getLast_append] + getLast (l ++ [a]) (append_ne_nil_of_right_ne_nil l (cons_ne_nil a _)) = a := by + simp [getLast_append] -- Porting note: name should be fixed upstream theorem getLast_append' (l₁ l₂ : List α) (h : l₂ ≠ []) : - getLast (l₁ ++ l₂) (append_ne_nil_of_ne_nil_right l₁ l₂ h) = getLast l₂ h := by + getLast (l₁ ++ l₂) (append_ne_nil_of_right_ne_nil l₁ h) = getLast l₂ h := by induction' l₁ with _ _ ih · simp · simp only [cons_append] rw [List.getLast_cons] exact ih -theorem getLast_concat' {a : α} (l : List α) : getLast (concat l a) (concat_ne_nil a l) = a := - getLast_concat .. +theorem getLast_concat' {a : α} (l : List α) : getLast (concat l a) (concat_ne_nil a l) = a := by + simp @[simp] theorem getLast_singleton' (a : α) : getLast [a] (cons_ne_nil a []) = a := rfl @@ -471,7 +459,7 @@ theorem getLastI_eq_getLast? [Inhabited α] : ∀ l : List α, l.getLastI = l.ge | [a, b, c] => rfl | _ :: _ :: c :: l => by simp [getLastI, getLastI_eq_getLast? (c :: l)] -@[simp] +#adaptation_note /-- 2024-07-10: removed `@[simp]` since the LHS simplifies using the simp set. -/ theorem getLast?_append_cons : ∀ (l₁ : List α) (a : α) (l₂ : List α), getLast? (l₁ ++ a :: l₂) = getLast? (a :: l₂) | [], a, l₂ => rfl @@ -484,7 +472,7 @@ theorem getLast?_append_of_ne_nil (l₁ : List α) : | [], hl₂ => by contradiction | b :: l₂, _ => getLast?_append_cons l₁ b l₂ -theorem getLast?_append {l₁ l₂ : List α} {x : α} (h : x ∈ l₂.getLast?) : +theorem mem_getLast?_append_of_mem_getLast? {l₁ l₂ : List α} {x : α} (h : x ∈ l₂.getLast?) : x ∈ (l₁ ++ l₂).getLast? := by cases l₂ · contradiction @@ -528,7 +516,8 @@ theorem head!_append [Inhabited α] (t : List α) {s : List α} (h : s ≠ []) : · contradiction · rfl -theorem head?_append {s t : List α} {x : α} (h : x ∈ s.head?) : x ∈ (s ++ t).head? := by +theorem mem_head?_append_of_mem_head? {s t : List α} {x : α} (h : x ∈ s.head?) : + x ∈ (s ++ t).head? := by cases s · contradiction · exact h @@ -560,18 +549,11 @@ theorem head!_mem_self [Inhabited α] {l : List α} (h : l ≠ nil) : l.head! have h' := mem_cons_self l.head! l.tail rwa [cons_head!_tail h] at h' -theorem head_mem {l : List α} : ∀ (h : l ≠ nil), l.head h ∈ l := by - cases l <;> simp - theorem tail_append_of_ne_nil (l l' : List α) (h : l ≠ []) : (l ++ l').tail = l.tail ++ l' := by cases l · contradiction · simp -theorem getElem_eq_getElem? (l : List α) (i : Nat) (h : i < l.length) : - l[i] = l[i]?.get (by simp [getElem?_eq_getElem, h]) := by - simp [getElem_eq_iff] - theorem get_eq_get? (l : List α) (i : Fin l.length) : l.get i = (l.get? i).get (by simp [getElem?_eq_getElem]) := by simp [getElem_eq_iff] @@ -738,7 +720,7 @@ theorem sublist_cons_of_sublist (a : α) (h : l₁ <+ l₂) : l₁ <+ a :: l₂ theorem tail_sublist : ∀ l : List α, tail l <+ l | [] => .slnil - | a::l => sublist_cons a l + | a::l => sublist_cons_self a l @[gcongr] protected theorem Sublist.tail : ∀ {l₁ l₂ : List α}, l₁ <+ l₂ → tail l₁ <+ tail l₂ | _, _, slnil => .slnil @@ -763,13 +745,6 @@ alias sublist_nil_iff_eq_nil := sublist_nil @[simp] lemma sublist_singleton {l : List α} {a : α} : l <+ [a] ↔ l = [] ∨ l = [a] := by constructor <;> rintro (_ | _) <;> aesop -theorem sublist_replicate_iff {l : List α} {a : α} {n : ℕ} : - l <+ replicate n a ↔ ∃ k ≤ n, l = replicate k a := - ⟨fun h => - ⟨l.length, h.length_le.trans_eq (length_replicate _ _), - eq_replicate_length.mpr fun b hb => eq_of_mem_replicate (h.subset hb)⟩, - by rintro ⟨k, h, rfl⟩; exact (replicate_sublist_replicate _).mpr h⟩ - theorem Sublist.antisymm (s₁ : l₁ <+ l₂) (s₂ : l₂ <+ l₁) : l₁ = l₂ := s₁.eq_of_length_le s₂.length_le @@ -1107,8 +1082,6 @@ theorem get_set_of_ne {l : List α} {i j : ℕ} (h : i ≠ j) (a : α) @[deprecated (since := "2024-06-21")] alias map_congr := map_congr_left -@[deprecated (since := "2024-06-21")] alias map_eq_map_iff := map_inj_left - theorem bind_pure_eq_map (f : α → β) (l : List α) : l.bind (pure ∘ f) = map f l := .symm <| map_eq_bind .. @@ -1335,18 +1308,6 @@ theorem foldl_fixed {a : α} : ∀ l : List β, foldl (fun a _ => a) a l = a := theorem foldr_fixed {b : β} : ∀ l : List α, foldr (fun _ b => b) b l = b := foldr_fixed' fun _ => rfl -@[simp] -theorem foldl_join (f : α → β → α) : - ∀ (a : α) (L : List (List β)), foldl f a (join L) = foldl (foldl f) a L - | a, [] => rfl - | a, l :: L => by simp only [join, foldl_append, foldl_cons, foldl_join f (foldl f a l) L] - -@[simp] -theorem foldr_join (f : α → β → β) : - ∀ (b : β) (L : List (List α)), foldr f b (join L) = foldr (fun l b => foldr f b l) b L - | a, [] => rfl - | a, l :: L => by simp only [join, foldr_append, foldr_join f a L, foldr_cons] - -- Porting note (#10618): simp can prove this -- @[simp] theorem foldr_eta : ∀ l : List α, foldr cons [] l = l := by @@ -1672,7 +1633,7 @@ variable (p : α → Bool) (xs ys : List α) (ls : List (List α)) (f : List α theorem splitAt_eq_take_drop (n : ℕ) (l : List α) : splitAt n l = (take n l, drop n l) := by by_cases h : n < l.length <;> rw [splitAt, go_eq_take_drop] · rw [if_pos h]; rfl - · rw [if_neg h, take_all_of_le <| le_of_not_lt h, drop_eq_nil_of_le <| le_of_not_lt h] + · rw [if_neg h, take_of_length_le <| le_of_not_lt h, drop_eq_nil_of_le <| le_of_not_lt h] where go_eq_take_drop (n : ℕ) (l xs : List α) (acc : Array α) : splitAt.go l xs n acc = if n < xs.length then (acc.toList ++ take n xs, drop n xs) else (l, []) := by @@ -1747,7 +1708,7 @@ theorem splitOnP_spec (as : List α) : · rw [if_pos h, h, map, cons_append, zipWith, nil_append, join, cons_append, cons_inj_right] exact ih · rw [if_neg h, eq_false_of_ne_true h, join_zipWith (splitOnP_ne_nil _ _) - (append_ne_nil_of_ne_nil_right _ [[]] (cons_ne_nil [] [])), cons_inj_right] + (append_ne_nil_of_right_ne_nil _ (cons_ne_nil [] [])), cons_inj_right] exact ih where join_zipWith {xs ys : List (List α)} {a : α} (hxs : xs ≠ []) (hys : ys ≠ []) : @@ -1829,7 +1790,7 @@ theorem modifyLast.go_append_one (f : α → α) (a : α) (tl : List α) (r : Ar | cons hd tl => simp only [cons_append] rw [modifyLast.go, modifyLast.go] - case x_3 | x_3 => exact append_ne_nil_of_ne_nil_right tl [a] (cons_ne_nil a []) + case x_3 | x_3 => exact append_ne_nil_of_right_ne_nil tl (cons_ne_nil a []) rw [modifyLast.go_append_one _ _ tl _, modifyLast.go_append_one _ _ tl (Array.push #[] hd)] simp only [Array.toListAppend_eq, Array.push_data, Array.data_toArray, nil_append, append_assoc] @@ -1841,7 +1802,7 @@ theorem modifyLast_append_one (f : α → α) (a : α) (l : List α) : | cons _ tl => simp only [cons_append, modifyLast] rw [modifyLast.go] - case x_3 => exact append_ne_nil_of_ne_nil_right tl [a] (cons_ne_nil a []) + case x_3 => exact append_ne_nil_of_right_ne_nil tl (cons_ne_nil a []) rw [modifyLast.go_append_one, Array.toListAppend_eq, Array.push_data, Array.data_toArray, nil_append, cons_append, nil_append, cons_inj_right] exact modifyLast_append_one _ _ tl @@ -1863,8 +1824,6 @@ end ModifyLast /-! ### map for partial functions -/ -@[simp] lemma attach_nil : ([] : List α).attach = [] := rfl - theorem sizeOf_lt_sizeOf_of_mem [SizeOf α] {x : α} {l : List α} (hx : x ∈ l) : SizeOf.sizeOf x < SizeOf.sizeOf l := by induction' l with h t ih <;> cases hx <;> rw [cons.sizeOf_spec] @@ -1872,121 +1831,8 @@ theorem sizeOf_lt_sizeOf_of_mem [SizeOf α] {x : α} {l : List α} (hx : x ∈ l · specialize ih ‹_› omega -@[simp] -theorem pmap_eq_map (p : α → Prop) (f : α → β) (l : List α) (H) : - @pmap _ _ p (fun a _ => f a) l H = map f l := by - induction l <;> [rfl; simp only [*, pmap, map]] - -theorem pmap_congr {p q : α → Prop} {f : ∀ a, p a → β} {g : ∀ a, q a → β} (l : List α) {H₁ H₂} - (h : ∀ a ∈ l, ∀ (h₁ h₂), f a h₁ = g a h₂) : pmap f l H₁ = pmap g l H₂ := by - induction' l with _ _ ih - · rfl - · rw [pmap, pmap, h _ (mem_cons_self _ _), ih fun a ha => h a (mem_cons_of_mem _ ha)] - -theorem map_pmap {p : α → Prop} (g : β → γ) (f : ∀ a, p a → β) (l H) : - map g (pmap f l H) = pmap (fun a h => g (f a h)) l H := by - induction l <;> [rfl; simp only [*, pmap, map]] - -theorem pmap_map {p : β → Prop} (g : ∀ b, p b → γ) (f : α → β) (l H) : - pmap g (map f l) H = pmap (fun a h => g (f a) h) l fun a h => H _ (mem_map_of_mem _ h) := by - induction l <;> [rfl; simp only [*, pmap, map]] - -theorem pmap_eq_map_attach {p : α → Prop} (f : ∀ a, p a → β) (l H) : - pmap f l H = l.attach.map fun x => f x.1 (H _ x.2) := by - rw [attach, attachWith, map_pmap]; exact pmap_congr l fun _ _ _ _ => rfl - --- @[simp] -- Porting note (#10959): lean 4 simp can't rewrite with this -theorem attach_map_coe' (l : List α) (f : α → β) : - (l.attach.map fun (i : {i // i ∈ l}) => f i) = l.map f := by - rw [attach, attachWith, map_pmap]; exact pmap_eq_map _ _ _ _ - -theorem attach_map_val' (l : List α) (f : α → β) : (l.attach.map fun i => f i.val) = l.map f := - attach_map_coe' _ _ - -@[simp] -theorem attach_map_val (l : List α) : l.attach.map Subtype.val = l := - (attach_map_coe' _ _).trans l.map_id --- Porting note: coe is expanded eagerly, so "attach_map_coe" would have the same syntactic form. - -@[simp] -theorem mem_attach (l : List α) : ∀ x, x ∈ l.attach - | ⟨a, h⟩ => by - have := mem_map.1 (by rw [attach_map_val] <;> exact h) - rcases this with ⟨⟨_, _⟩, m, rfl⟩ - exact m - -@[simp] -theorem mem_pmap {p : α → Prop} {f : ∀ a, p a → β} {l H b} : - b ∈ pmap f l H ↔ ∃ (a : _) (h : a ∈ l), f a (H a h) = b := by - simp only [pmap_eq_map_attach, mem_map, mem_attach, true_and_iff, Subtype.exists, eq_comm] - -@[simp] -theorem length_pmap {p : α → Prop} {f : ∀ a, p a → β} {l H} : length (pmap f l H) = length l := by - induction l <;> [rfl; simp only [*, pmap, length]] - -@[simp] -theorem length_attach (L : List α) : L.attach.length = L.length := - length_pmap - -@[simp] -theorem pmap_eq_nil {p : α → Prop} {f : ∀ a, p a → β} {l H} : pmap f l H = [] ↔ l = [] := by - rw [← length_eq_zero, length_pmap, length_eq_zero] - -@[simp] -theorem attach_eq_nil (l : List α) : l.attach = [] ↔ l = [] := - pmap_eq_nil - -theorem getLast_pmap (p : α → Prop) (f : ∀ a, p a → β) (l : List α) - (hl₁ : ∀ a ∈ l, p a) (hl₂ : l ≠ []) : - (l.pmap f hl₁).getLast (mt List.pmap_eq_nil.1 hl₂) = - f (l.getLast hl₂) (hl₁ _ (List.getLast_mem hl₂)) := by - induction' l with l_hd l_tl l_ih - · apply (hl₂ rfl).elim - · by_cases hl_tl : l_tl = [] - · simp [hl_tl] - · simp only [pmap] - rw [getLast_cons, l_ih _ hl_tl] - simp only [getLast_cons hl_tl] - -theorem getElem?_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) (n : ℕ) : - (pmap f l h)[n]? = Option.pmap f l[n]? fun x H => h x (getElem?_mem H) := by - induction' l with hd tl hl generalizing n - · simp - · cases' n with n - · simp only [Option.pmap] - split <;> simp_all - · simp only [hl, pmap, Option.pmap, getElem?_cons_succ] - split <;> rename_i h₁ _ <;> split <;> rename_i h₂ _ - · simp_all - · simp at h₂ - simp_all - · simp_all - · simp_all - -theorem get?_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) (n : ℕ) : - get? (pmap f l h) n = Option.pmap f (get? l n) fun x H => h x (get?_mem H) := by - simp only [get?_eq_getElem?] - simp [getElem?_pmap, h] - -theorem getElem_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) {n : ℕ} - (hn : n < (pmap f l h).length) : - (pmap f l h)[n] = - f (l[n]'(@length_pmap _ _ p f l h ▸ hn)) - (h _ (getElem_mem l n (@length_pmap _ _ p f l h ▸ hn))) := by - induction' l with hd tl hl generalizing n - · simp only [length, pmap] at hn - exact absurd hn (not_lt_of_le n.zero_le) - · cases n - · simp - · simp [hl] - -theorem get_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) {n : ℕ} - (hn : n < (pmap f l h).length) : - get (pmap f l h) ⟨n, hn⟩ = - f (get l ⟨n, @length_pmap _ _ p f l h ▸ hn⟩) - (h _ (get_mem l n (@length_pmap _ _ p f l h ▸ hn))) := by - simp only [get_eq_getElem] - simp [getElem_pmap] +@[deprecated attach_map_coe (since := "2024-07-29")] alias attach_map_coe' := attach_map_coe +@[deprecated attach_map_val (since := "2024-07-29")] alias attach_map_val' := attach_map_val set_option linter.deprecated false in @[deprecated get_pmap (since := "2023-01-05")] @@ -1997,22 +1843,6 @@ theorem nthLe_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : (h _ (get_mem l n (@length_pmap _ _ p f l h ▸ hn))) := get_pmap .. -theorem pmap_append {p : ι → Prop} (f : ∀ a : ι, p a → α) (l₁ l₂ : List ι) - (h : ∀ a ∈ l₁ ++ l₂, p a) : - (l₁ ++ l₂).pmap f h = - (l₁.pmap f fun a ha => h a (mem_append_left l₂ ha)) ++ - l₂.pmap f fun a ha => h a (mem_append_right l₁ ha) := by - induction' l₁ with _ _ ih - · rfl - · dsimp only [pmap, cons_append] - rw [ih] - -theorem pmap_append' {p : α → Prop} (f : ∀ a : α, p a → β) (l₁ l₂ : List α) - (h₁ : ∀ a ∈ l₁, p a) (h₂ : ∀ a ∈ l₂, p a) : - ((l₁ ++ l₂).pmap f fun a ha => (List.mem_append.1 ha).elim (h₁ a) (h₂ a)) = - l₁.pmap f h₁ ++ l₂.pmap f h₂ := - pmap_append f l₁ l₂ _ - /-! ### find -/ section find? @@ -2121,9 +1951,6 @@ attribute [simp 1100] filterMap_cons_none -- removing attribute `nolint simpNF` attribute [simp 1100] filterMap_cons_some -theorem Sublist.map (f : α → β) {l₁ l₂ : List α} (s : l₁ <+ l₂) : map f l₁ <+ map f l₂ := - filterMap_eq_map f ▸ s.filterMap _ - theorem filterMap_eq_bind_toList (f : α → Option β) (l : List α) : l.filterMap f = l.bind fun a ↦ (f a).toList := by induction' l with a l ih <;> simp [filterMap_cons] @@ -2168,14 +1995,19 @@ theorem filter_eq_foldr (p : α → Bool) (l : List α) : filter p l = foldr (fun a out => bif p a then a :: out else out) [] l := by induction l <;> simp [*, filter]; rfl +#adaptation_note +/-- +This has to be temporarily renamed to avoid an unintentional collision. +The prime should be removed at nightly-2024-07-27. +-/ @[simp] -theorem filter_subset (l : List α) : filter p l ⊆ l := +theorem filter_subset' (l : List α) : filter p l ⊆ l := (filter_sublist l).subset theorem of_mem_filter {a : α} {l} (h : a ∈ filter p l) : p a := (mem_filter.1 h).2 theorem mem_of_mem_filter {a : α} {l} (h : a ∈ filter p l) : a ∈ l := - filter_subset l h + filter_subset' l h theorem mem_filter_of_mem {a : α} {l} (h₁ : a ∈ l) (h₂ : p a) : a ∈ filter p l := mem_filter.2 ⟨h₁, h₂⟩ @@ -2193,13 +2025,13 @@ theorem monotone_filter_right (l : List α) ⦃p q : α → Bool⦄ induction' l with hd tl IH · rfl · by_cases hp : p hd - · rw [filter_cons_of_pos _ hp, filter_cons_of_pos _ (h _ hp)] + · rw [filter_cons_of_pos hp, filter_cons_of_pos (h _ hp)] exact IH.cons_cons hd - · rw [filter_cons_of_neg _ hp] + · rw [filter_cons_of_neg hp] by_cases hq : q hd - · rw [filter_cons_of_pos _ hq] + · rw [filter_cons_of_pos hq] exact sublist_cons_of_sublist hd IH - · rw [filter_cons_of_neg _ hq] + · rw [filter_cons_of_neg hq] exact IH -- TODO rename to `map_filter` when the deprecated `map_filter` is removed from Lean. @@ -2221,7 +2053,7 @@ lemma filter_attach (l : List α) (p : α → Bool) : (l.filter p).attach.map (Subtype.map id fun x => mem_of_mem_filter) := map_injective_iff.2 Subtype.coe_injective <| by simp_rw [map_map, (· ∘ ·), Subtype.map, id, ← Function.comp_apply (g := Subtype.val), - ← filter_map, attach_map_val] + ← filter_map, attach_map_subtype_val] lemma filter_comm (q) (l : List α) : filter p (filter q l) = filter q (filter p l) := by simp [and_comm] @@ -2267,14 +2099,6 @@ theorem dropWhile_eq_nil_iff : dropWhile p l = [] ↔ ∀ x ∈ l, p x := by · simp [dropWhile] · by_cases hp : p x <;> simp [hp, dropWhile, IH] -theorem takeWhile_cons_of_pos {x : α} (h : p x) : - List.takeWhile p (x :: l) = x :: takeWhile p l := by - simp [takeWhile_cons, h] - -theorem takeWhile_cons_of_neg {x : α} (h : ¬ p x) : - List.takeWhile p (x :: l) = [] := by - simp [takeWhile_cons, h] - @[simp] theorem takeWhile_eq_self_iff : takeWhile p l = l ↔ ∀ x ∈ l, p x := by induction' l with x xs IH @@ -2389,7 +2213,7 @@ theorem erase_diff_erase_sublist_of_sublist {a : α} : | b :: l₁, l₂, h => if heq : b = a then by simp only [heq, erase_cons_head, diff_cons]; rfl else by - simp only [erase_cons_head b l₁, erase_cons_tail l₁ (not_beq_of_ne heq), + simp only [erase_cons_head b l₁, erase_cons_tail (not_beq_of_ne heq), diff_cons ((List.erase l₂ a)) (List.erase l₁ a) b, diff_cons l₂ l₁ b, erase_comm a b l₂] have h' := h.erase b rw [erase_cons_head] at h' @@ -2677,15 +2501,6 @@ end Forall /-! ### Miscellaneous lemmas -/ -theorem getLast_reverse {l : List α} (hl : l.reverse ≠ []) - (hl' : 0 < l.length := (by - contrapose! hl - simpa [length_eq_zero] using hl)) : - l.reverse.getLast hl = l.get ⟨0, hl'⟩ := by - rw [getLast_eq_get, get_reverse'] - · simp - · simpa using hl' - @[simp] theorem getElem_attach (L : List α) (i : Nat) (h : i < L.attach.length) : L.attach[i].1 = L[i]'(length_attach L ▸ h) := @@ -2726,6 +2541,7 @@ theorem length_dropSlice_lt (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs (List.dropSlice i j xs).length < xs.length := by simp; omega +set_option linter.deprecated false in @[deprecated (since := "2024-07-25")] theorem sizeOf_dropSlice_lt [SizeOf α] (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs.length) : SizeOf.sizeOf (List.dropSlice i j xs) < SizeOf.sizeOf xs := by diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index d4fd5e0473708..e8e436dc37099 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -284,7 +284,7 @@ theorem chain'_iff_get {R} : ∀ {l : List α}, Chain' R l ↔ | [] => iff_of_true (by simp) (fun _ h => by simp at h) | [a] => iff_of_true (by simp) (fun _ h => by simp at h) | a :: b :: t => by - rw [← and_forall_succ, chain'_cons, chain'_iff_get] + rw [← and_forall_add_one, chain'_cons, chain'_iff_get] simp /-- If `l₁ l₂` and `l₃` are lists and `l₁ ++ l₂` and `l₂ ++ l₃` both satisfy @@ -351,8 +351,7 @@ reflexive transitive closure of `r`. The converse of `exists_chain_of_relationRe -/ theorem relationReflTransGen_of_exists_chain (l : List α) (hl₁ : Chain r a l) (hl₂ : getLast (a :: l) (cons_ne_nil _ _) = b) : Relation.ReflTransGen r a b := --- Porting note: `p` behaves like an implicit argument to `Chain.induction_head` but it is explicit. - Chain.induction_head l hl₁ hl₂ (fun _ _ => Relation.ReflTransGen.head) + Chain.induction_head _ l hl₁ hl₂ (fun _ _ => Relation.ReflTransGen.head) Relation.ReflTransGen.refl theorem Chain'.cons_of_le [LinearOrder α] {a : α} {as m : List α} diff --git a/Mathlib/Data/List/Count.lean b/Mathlib/Data/List/Count.lean index 4ac84e6b0ab7c..1e128c3d42e2b 100644 --- a/Mathlib/Data/List/Count.lean +++ b/Mathlib/Data/List/Count.lean @@ -3,7 +3,7 @@ Copyright (c) 2014 Parikshit Khanna. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Mario Carneiro -/ -import Mathlib.Data.List.Basic +import Mathlib.Data.Nat.Defs /-! # Counting in lists @@ -23,21 +23,6 @@ variable {α : Type*} {l : List α} namespace List -section CountP - -variable (p q : α → Bool) - -theorem length_filter_lt_length_iff_exists (l) : - length (filter p l) < length l ↔ ∃ x ∈ l, ¬p x := by - simpa [length_eq_countP_add_countP p l, countP_eq_length_filter] using - countP_pos (fun x => ¬p x) (l := l) - --- Porting note: `Lean.Internal.coeM` forces us to type-ascript `{x // x ∈ l}` -lemma countP_attach (l : List α) : l.attach.countP (fun a : {x // x ∈ l} ↦ p a) = l.countP p := by - simp_rw [← Function.comp_apply (g := Subtype.val), ← countP_map, attach_map_val] - -end CountP - /-! ### count -/ section Count @@ -55,10 +40,6 @@ theorem count_cons' (a b : α) (l : List α) : simp only [count, beq_iff_eq, countP_cons, Nat.add_right_inj] simp only [eq_comm] -@[simp] -lemma count_attach (a : {x // x ∈ l}) : l.attach.count a = l.count ↑a := - Eq.trans (countP_congr fun _ _ => by simp [Subtype.ext_iff]) <| countP_attach _ _ - end Count end List diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 54aaf664d6058..dcf48a84137c4 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -168,7 +168,7 @@ theorem next_getLast_cons (h : x ∈ l) (y : α) (h : x ∈ y :: l) (hy : x ≠ rw [← get?_eq_get, dropLast_eq_take, get?_eq_getElem?, getElem?_take, getElem?_cons_zero, Option.some_inj] at hk' · exact hy (Eq.symm hk') - rw [length_cons, Nat.pred_succ] + rw [length_cons] exact length_pos_of_mem (by assumption) suffices k + 1 = l.length by simp [this] at hk cases' l with hd tl @@ -181,7 +181,7 @@ theorem next_getLast_cons (h : x ∈ l) (y : α) (h : x ∈ y :: l) (hy : x ≠ rw [← get?_eq_get, dropLast_eq_take, get?_eq_getElem?, getElem?_take, getElem?_cons_succ, getElem?_eq_getElem, Option.some_inj] at hk' · rw [get_eq_getElem, hk'] - simp only [getLast_eq_get, length_cons, Nat.succ_eq_add_one, Nat.succ_sub_succ_eq_sub, + simp only [getLast_eq_getElem, length_cons, Nat.succ_eq_add_one, Nat.succ_sub_succ_eq_sub, Nat.sub_zero, get_eq_getElem, getElem_cons_succ] simpa using hk @@ -256,7 +256,7 @@ theorem next_get : ∀ (l : List α) (_h : Nodup l) (i : Fin l.length), · simp [hi', get] · rw [get_cons_succ]; exact get_mem _ _ _ · exact hx' - · simp [getLast_eq_get] + · simp [getLast_eq_getElem] · exact hn.of_cons · rw [next_ne_head_ne_getLast _ _ _ _ _ hx'] · simp only [get_cons_succ] @@ -267,7 +267,7 @@ theorem next_get : ∀ (l : List α) (_h : Nodup l) (i : Fin l.length), Nat.mod_eq_of_lt (Nat.succ_lt_succ_iff.2 (Nat.succ_lt_succ_iff.2 hi'))] · simp [Nat.mod_eq_of_lt (Nat.succ_lt_succ_iff.2 hi'), hi'] · exact hn.of_cons - · rw [getLast_eq_get] + · rw [getLast_eq_getElem] intro h have := nodup_iff_injective_get.1 hn h simp at this; simp [this] at hi' diff --git a/Mathlib/Data/List/Destutter.lean b/Mathlib/Data/List/Destutter.lean index 714f67f548c69..5f80cabc72631 100644 --- a/Mathlib/Data/List/Destutter.lean +++ b/Mathlib/Data/List/Destutter.lean @@ -60,7 +60,7 @@ theorem destutter'_sublist (a) : l.destutter' R a <+ a :: l := by rw [destutter'] split_ifs · exact Sublist.cons₂ a (hl b) - · exact (hl a).trans ((l.sublist_cons b).cons_cons a) + · exact (hl a).trans ((l.sublist_cons_self b).cons_cons a) theorem mem_destutter' (a) : a ∈ l.destutter' R a := by induction' l with b l hl diff --git a/Mathlib/Data/List/DropRight.lean b/Mathlib/Data/List/DropRight.lean index 2bd4a1f3bc4d5..6a591c57257f2 100644 --- a/Mathlib/Data/List/DropRight.lean +++ b/Mathlib/Data/List/DropRight.lean @@ -106,8 +106,8 @@ theorem rdropWhile_singleton (x : α) : rdropWhile p [x] = if p x then [] else [ theorem rdropWhile_last_not (hl : l.rdropWhile p ≠ []) : ¬p ((rdropWhile p l).getLast hl) := by simp_rw [rdropWhile] - rw [getLast_reverse] - exact dropWhile_nthLe_zero_not _ _ _ + rw [getLast_reverse, head_dropWhile_not p] + simp theorem rdropWhile_prefix : l.rdropWhile p <+: l := by rw [← reverse_suffix, rdropWhile, reverse_reverse] @@ -130,7 +130,7 @@ theorem dropWhile_eq_self_iff : dropWhile p l = l ↔ ∀ hl : 0 < l.length, ¬p refine ⟨fun h => ?_, fun h => ?_⟩ · intro _ H rw [get] at H - refine (cons_ne_self hd tl) (Sublist.antisymm ?_ (sublist_cons _ _)) + refine (cons_ne_self hd tl) (Sublist.antisymm ?_ (sublist_cons_self _ _)) rw [← h] simp only [H] exact List.IsSuffix.sublist (dropWhile_suffix p) @@ -142,7 +142,7 @@ theorem dropWhile_eq_self_iff : dropWhile p l = l ↔ ∀ hl : 0 < l.length, ¬p the `l ≠ []` condition if `hl` is not `intro`'d yet -/ @[simp] theorem rdropWhile_eq_self_iff : rdropWhile p l = l ↔ ∀ hl : l ≠ [], ¬p (l.getLast hl) := by - simp only [rdropWhile, reverse_eq_iff, dropWhile_eq_self_iff, getLast_eq_get] + simp only [rdropWhile, reverse_eq_iff, dropWhile_eq_self_iff, getLast_eq_getElem] refine ⟨fun h hl => ?_, fun h hl => ?_⟩ · rw [← length_pos, ← length_reverse] at hl have := h hl @@ -197,8 +197,8 @@ theorem rtakeWhile_eq_nil_iff : rtakeWhile p l = [] ↔ ∀ hl : l ≠ [], ¬p ( induction' l using List.reverseRecOn with l a · simp only [rtakeWhile, takeWhile, reverse_nil, true_iff] intro f; contradiction - · simp only [rtakeWhile, reverse_append, takeWhile, reverse_eq_nil_iff, getLast_append, ne_eq, - append_eq_nil, and_false, not_false_eq_true, forall_true_left] + · simp only [rtakeWhile, reverse_append, takeWhile, ne_eq, not_false_eq_true, + getLast_append_of_ne_nil, getLast_singleton] refine ⟨fun h => ?_ , fun h => ?_⟩ · split at h <;> simp_all · simp [h] diff --git a/Mathlib/Data/List/EditDistance/Defs.lean b/Mathlib/Data/List/EditDistance/Defs.lean index ee121a75e2bf8..ed13538ca6e25 100644 --- a/Mathlib/Data/List/EditDistance/Defs.lean +++ b/Mathlib/Data/List/EditDistance/Defs.lean @@ -279,7 +279,7 @@ theorem levenshtein_nil_cons (y) (ys) : levenshtein C [] (y :: ys) = C.insert y + levenshtein C [] ys := by dsimp (config := { unfoldPartialApp := true }) [levenshtein, suffixLevenshtein, impl] congr - rw [List.getLast_eq_get] + rw [List.getLast_eq_getElem] congr rw [show (List.length _) = 1 from _] induction ys <;> simp diff --git a/Mathlib/Data/List/Enum.lean b/Mathlib/Data/List/Enum.lean index 08ad0a86d2413..e7bd524eb313e 100644 --- a/Mathlib/Data/List/Enum.lean +++ b/Mathlib/Data/List/Enum.lean @@ -3,7 +3,9 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Yakov Pechersky, Eric Wieser -/ -import Mathlib.Data.List.Basic +import Batteries.Tactic.Alias +import Mathlib.Tactic.TypeStar +import Mathlib.Data.Nat.Notation /-! # Properties of `List.enum` @@ -13,53 +15,21 @@ namespace List variable {α β : Type*} -@[simp] -theorem getElem?_enumFrom : - ∀ n (l : List α) m, (enumFrom n l)[m]? = l[m]?.map fun a => (n + m, a) - | n, [], m => rfl - | n, a :: l, 0 => by simp - | n, a :: l, m + 1 => by - simp only [enumFrom_cons, getElem?_cons_succ] - exact (getElem?_enumFrom (n + 1) l m).trans <| by rw [Nat.add_right_comm]; rfl - theorem get?_enumFrom (n) (l : List α) (m) : get? (enumFrom n l) m = (get? l m).map fun a => (n + m, a) := by simp @[deprecated (since := "2024-04-06")] alias enumFrom_get? := get?_enumFrom -@[simp] -theorem getElem?_enum (l : List α) (n : Nat) : (enum l)[n]? = l[n]?.map fun a => (n, a) := by - rw [enum, getElem?_enumFrom, Nat.zero_add] - theorem get?_enum (l : List α) (n) : get? (enum l) n = (get? l n).map fun a => (n, a) := by simp @[deprecated (since := "2024-04-06")] alias enum_get? := get?_enum -@[simp] -theorem enumFrom_map_snd : ∀ (n) (l : List α), map Prod.snd (enumFrom n l) = l - | _, [] => rfl - | _, _ :: _ => congr_arg (cons _) (enumFrom_map_snd _ _) - -@[simp] -theorem enum_map_snd (l : List α) : map Prod.snd (enum l) = l := - enumFrom_map_snd _ _ - -@[simp] -theorem getElem_enumFrom (l : List α) (n) (i : Nat) (h : i < (l.enumFrom n).length) : - (l.enumFrom n)[i] = (n + i, l[i]'(by simpa [enumFrom_length] using h)) := by - simp [getElem_eq_getElem?] - theorem get_enumFrom (l : List α) (n) (i : Fin (l.enumFrom n).length) : (l.enumFrom n).get i = (n + i, l.get (i.cast enumFrom_length)) := by simp -@[simp] -theorem getElem_enum (l : List α) (i : Nat) (h : i < l.enum.length) : - l.enum[i] = (i, l[i]'(by simpa [enum_length] using h)) := by - simp [enum] - theorem get_enum (l : List α) (i : Fin l.enum.length) : l.enum.get i = (i.1, l.get (i.cast enum_length)) := by simp @@ -83,77 +53,4 @@ theorem mk_mem_enum_iff_get? {i : ℕ} {x : α} {l : List α} : (i, x) ∈ enum theorem mem_enum_iff_get? {x : ℕ × α} {l : List α} : x ∈ enum l ↔ l.get? x.1 = x.2 := mk_mem_enum_iff_get? -theorem le_fst_of_mem_enumFrom {x : ℕ × α} {n : ℕ} {l : List α} (h : x ∈ enumFrom n l) : - n ≤ x.1 := - (mk_mem_enumFrom_iff_le_and_get?_sub.1 h).1 - -theorem fst_lt_add_of_mem_enumFrom {x : ℕ × α} {n : ℕ} {l : List α} (h : x ∈ enumFrom n l) : - x.1 < n + length l := by - rcases mem_iff_get.1 h with ⟨i, rfl⟩ - simpa using i.is_lt - -theorem fst_lt_of_mem_enum {x : ℕ × α} {l : List α} (h : x ∈ enum l) : x.1 < length l := by - simpa using fst_lt_add_of_mem_enumFrom h - -theorem snd_mem_of_mem_enumFrom {x : ℕ × α} {n : ℕ} {l : List α} (h : x ∈ enumFrom n l) : x.2 ∈ l := - enumFrom_map_snd n l ▸ mem_map_of_mem _ h - -theorem snd_mem_of_mem_enum {x : ℕ × α} {l : List α} (h : x ∈ enum l) : x.2 ∈ l := - snd_mem_of_mem_enumFrom h - -theorem mem_enumFrom {x : α} {i j : ℕ} (xs : List α) (h : (i, x) ∈ xs.enumFrom j) : - j ≤ i ∧ i < j + xs.length ∧ x ∈ xs := - ⟨le_fst_of_mem_enumFrom h, fst_lt_add_of_mem_enumFrom h, snd_mem_of_mem_enumFrom h⟩ - -@[simp] -theorem enumFrom_singleton (x : α) (n : ℕ) : enumFrom n [x] = [(n, x)] := - rfl - -@[simp] -theorem enum_singleton (x : α) : enum [x] = [(0, x)] := - rfl - -theorem enumFrom_append (xs ys : List α) (n : ℕ) : - enumFrom n (xs ++ ys) = enumFrom n xs ++ enumFrom (n + xs.length) ys := by - induction' xs with x xs IH generalizing ys n - · simp - · rw [cons_append, enumFrom_cons, IH, ← cons_append, ← enumFrom_cons, length, Nat.add_right_comm, - Nat.add_assoc] - -theorem enum_append (xs ys : List α) : enum (xs ++ ys) = enum xs ++ enumFrom xs.length ys := by - simp [enum, enumFrom_append] - -theorem map_fst_add_enumFrom_eq_enumFrom (l : List α) (n k : ℕ) : - map (Prod.map (· + n) id) (enumFrom k l) = enumFrom (n + k) l := - ext_get? fun i ↦ by simp [(· ∘ ·), Nat.add_comm, Nat.add_left_comm] - -theorem map_fst_add_enum_eq_enumFrom (l : List α) (n : ℕ) : - map (Prod.map (· + n) id) (enum l) = enumFrom n l := - map_fst_add_enumFrom_eq_enumFrom l _ _ - -theorem enumFrom_cons' (n : ℕ) (x : α) (xs : List α) : - enumFrom n (x :: xs) = (n, x) :: (enumFrom n xs).map (Prod.map Nat.succ id) := by - rw [enumFrom_cons, Nat.add_comm, ← map_fst_add_enumFrom_eq_enumFrom] - -theorem enum_cons' (x : α) (xs : List α) : - enum (x :: xs) = (0, x) :: (enum xs).map (Prod.map Nat.succ id) := - enumFrom_cons' _ _ _ - -theorem enumFrom_map (n : ℕ) (l : List α) (f : α → β) : - enumFrom n (l.map f) = (enumFrom n l).map (Prod.map id f) := by - induction' l with hd tl IH - · rfl - · rw [map_cons, enumFrom_cons', enumFrom_cons', map_cons, map_map, IH, map_map] - rfl - -theorem enum_map (l : List α) (f : α → β) : (l.map f).enum = l.enum.map (Prod.map id f) := - enumFrom_map _ _ _ - -@[simp] -theorem enumFrom_eq_nil {n : ℕ} {l : List α} : List.enumFrom n l = [] ↔ l = [] := by - cases l <;> simp - -@[simp] -theorem enum_eq_nil {l : List α} : List.enum l = [] ↔ l = [] := enumFrom_eq_nil - end List diff --git a/Mathlib/Data/List/Forall2.lean b/Mathlib/Data/List/Forall2.lean index b05292dc574dd..132891a19defd 100644 --- a/Mathlib/Data/List/Forall2.lean +++ b/Mathlib/Data/List/Forall2.lean @@ -258,10 +258,10 @@ theorem rel_filter {p : α → Bool} {q : β → Bool} dsimp [LiftFun] at hpq by_cases h : p a · have : q b := by rwa [← hpq h₁] - simp only [filter_cons_of_pos _ h, filter_cons_of_pos _ this, forall₂_cons, h₁, true_and_iff, + simp only [filter_cons_of_pos h, filter_cons_of_pos this, forall₂_cons, h₁, true_and_iff, rel_filter hpq h₂] · have : ¬q b := by rwa [← hpq h₁] - simp only [filter_cons_of_neg _ h, filter_cons_of_neg _ this, rel_filter hpq h₂] + simp only [filter_cons_of_neg h, filter_cons_of_neg this, rel_filter hpq h₂] theorem rel_filterMap : ((R ⇒ Option.Rel P) ⇒ Forall₂ R ⇒ Forall₂ P) filterMap filterMap | _, _, _, _, _, Forall₂.nil => Forall₂.nil diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index 398439cebfcea..07ad3fab49f97 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -6,7 +6,6 @@ Mario Carneiro -/ import Mathlib.Data.List.Defs import Mathlib.Data.Option.Basic -import Mathlib.Data.Nat.Defs import Mathlib.Init.Data.List.Basic import Mathlib.Util.AssertExists diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index c5d556fcb5944..6da70619cfb4b 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -5,6 +5,7 @@ Authors: Jannis Limperg -/ import Mathlib.Data.List.OfFn import Mathlib.Data.List.Range +import Mathlib.Data.List.Zip /-! # Lemmas about List.*Idx functions. diff --git a/Mathlib/Data/List/Intervals.lean b/Mathlib/Data/List/Intervals.lean index f809824bf8436..cefbcfd118662 100644 --- a/Mathlib/Data/List/Intervals.lean +++ b/Mathlib/Data/List/Intervals.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Mathlib.Data.List.Lattice -import Mathlib.Data.List.Range import Mathlib.Data.Bool.Basic import Mathlib.Init.Data.Nat.Lemmas +import Mathlib.Order.Lattice /-! # Intervals in ℕ diff --git a/Mathlib/Data/List/Iterate.lean b/Mathlib/Data/List/Iterate.lean index 76fffcce7b531..dc981a78f8e38 100644 --- a/Mathlib/Data/List/Iterate.lean +++ b/Mathlib/Data/List/Iterate.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Miyahara Kō -/ -import Mathlib.Data.List.Range import Mathlib.Algebra.Order.Ring.Nat +import Mathlib.Data.List.Defs +import Mathlib.Data.Set.Function /-! # iterate diff --git a/Mathlib/Data/List/Join.lean b/Mathlib/Data/List/Join.lean index a98388a8dbd2b..c8ecd55aad8c5 100644 --- a/Mathlib/Data/List/Join.lean +++ b/Mathlib/Data/List/Join.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Floris van Doorn, Mario Carneiro, Martin Dvorak -/ import Mathlib.Data.List.Basic -import Batteries.Data.Nat.Lemmas /-! # Join of a list of lists @@ -24,18 +23,9 @@ namespace List -- @[simp] theorem join_singleton (l : List α) : [l].join = l := by rw [join, join, append_nil] -@[simp] -theorem join_eq_nil : ∀ {L : List (List α)}, join L = [] ↔ ∀ l ∈ L, l = [] - | [] => iff_of_true rfl (forall_mem_nil _) - | l :: L => by simp only [join, append_eq_nil, join_eq_nil, forall_mem_cons] - -@[simp] -theorem join_append (L₁ L₂ : List (List α)) : join (L₁ ++ L₂) = join L₁ ++ join L₂ := by - induction L₁ - · rfl - · simp [*] +@[deprecated join_eq_nil_iff (since := "2024-07-10")] +theorem join_eq_nil : ∀ {L : List (List α)}, join L = [] ↔ ∀ l ∈ L, l = [] := join_eq_nil_iff -theorem join_concat (L : List (List α)) (l : List α) : join (L.concat l) = join L ++ l := by simp @[simp] theorem join_filter_not_isEmpty : @@ -51,10 +41,8 @@ theorem join_filter_not_isEmpty : @[simp] theorem join_filter_ne_nil [DecidablePred fun l : List α => l ≠ []] {L : List (List α)} : join (L.filter fun l => l ≠ []) = L.join := by - simp [join_filter_not_isEmpty, ← isEmpty_iff_eq_nil] - -theorem join_join (l : List (List (List α))) : l.join.join = (l.map join).join := by - induction l <;> simp [*] + simp only [ne_eq, ← isEmpty_iff_eq_nil, Bool.not_eq_true, Bool.decide_eq_false, + join_filter_not_isEmpty] /-- See `List.length_join` for the corresponding statement using `List.sum`. -/ lemma length_join' (L : List (List α)) : length (join L) = Nat.sum (map length L) := by @@ -84,7 +72,7 @@ lemma count_bind' [BEq β] (l : List α) (f : α → List β) (x : β) : @[simp] theorem bind_eq_nil {l : List α} {f : α → List β} : List.bind l f = [] ↔ ∀ x ∈ l, f x = [] := - join_eq_nil.trans <| by + join_eq_nil_iff.trans <| by simp only [mem_map, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂] /-- In a join, taking the first elements up to an index which is the sum of the lengths of the @@ -183,17 +171,6 @@ theorem append_join_map_append (L : List (List α)) (x : List α) : · rw [map_nil, join, append_nil, map_nil, join, nil_append] · rw [map_cons, join, map_cons, join, append_assoc, ih, append_assoc, append_assoc] -/-- Reversing a join is the same as reversing the order of parts and reversing all parts. -/ -theorem reverse_join (L : List (List α)) : - L.join.reverse = (L.map reverse).reverse.join := by - induction' L with _ _ ih - · rfl - · rw [join, reverse_append, ih, map_cons, reverse_cons', join_concat] - -/-- Joining a reverse is the same as reversing all parts and reversing the joined result. -/ -theorem join_reverse (L : List (List α)) : - L.reverse.join = (L.map reverse).join.reverse := by - simpa [reverse_reverse, map_reverse] using congr_arg List.reverse (reverse_join L.reverse) /-- Any member of `L : List (List α))` is a sublist of `L.join` -/ lemma sublist_join (L : List (List α)) {s : List α} (hs : s ∈ L) : diff --git a/Mathlib/Data/List/Lattice.lean b/Mathlib/Data/List/Lattice.lean index f0eb742372dca..181f14ce99536 100644 --- a/Mathlib/Data/List/Lattice.lean +++ b/Mathlib/Data/List/Lattice.lean @@ -122,7 +122,7 @@ theorem mem_inter_of_mem_of_mem (h₁ : a ∈ l₁) (h₂ : a ∈ l₂) : a ∈ mem_filter_of_mem h₁ <| by simpa using h₂ theorem inter_subset_left {l₁ l₂ : List α} : l₁ ∩ l₂ ⊆ l₁ := - filter_subset _ + filter_subset' _ theorem inter_subset_right {l₁ l₂ : List α} : l₁ ∩ l₂ ⊆ l₂ := fun _ => mem_of_mem_inter_right @@ -200,10 +200,12 @@ theorem count_bagInter {a : α} : by_cases hb : b ∈ l₂ · rw [cons_bagInter_of_pos _ hb, count_cons, count_cons, count_bagInter, count_erase, ← Nat.add_min_add_right] - by_cases ab : a = b - · rw [if_pos ab, Nat.sub_add_cancel] - rwa [succ_le_iff, count_pos_iff_mem, ab] - · rw [if_neg ab, Nat.sub_zero, Nat.add_zero, Nat.add_zero] + by_cases ba : b = a + · simp only [beq_iff_eq] + rw [if_pos ba, Nat.sub_add_cancel] + rwa [succ_le_iff, count_pos_iff_mem, ← ba] + · simp only [beq_iff_eq] + rw [if_neg ba, Nat.sub_zero, Nat.add_zero, Nat.add_zero] · rw [cons_bagInter_of_neg _ hb, count_bagInter] by_cases ab : a = b · rw [← ab] at hb diff --git a/Mathlib/Data/List/NatAntidiagonal.lean b/Mathlib/Data/List/NatAntidiagonal.lean index 856fdacc2be03..fbbf06a9c47bb 100644 --- a/Mathlib/Data/List/NatAntidiagonal.lean +++ b/Mathlib/Data/List/NatAntidiagonal.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ import Mathlib.Data.List.Nodup -import Mathlib.Data.List.Range /-! # Antidiagonals in ℕ × ℕ as lists diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index 792bedc48e51c..5bfb47206bb95 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -23,17 +23,6 @@ variable {α : Type u} {β : Type v} {l l₁ l₂ : List α} {r : α → α → namespace List -@[simp] -theorem forall_mem_ne {a : α} {l : List α} : (∀ a' : α, a' ∈ l → ¬a = a') ↔ a ∉ l := - ⟨fun h m => h _ m rfl, fun h _ m e => h (e.symm ▸ m)⟩ - -@[simp] -theorem nodup_nil : @Nodup α [] := - Pairwise.nil - -@[simp] -theorem nodup_cons {a : α} {l : List α} : Nodup (a :: l) ↔ a ∉ l ∧ Nodup l := by - simp only [Nodup, pairwise_cons, forall_mem_ne] protected theorem Pairwise.nodup {l : List α} {r : α → α → Prop} [IsIrrefl α r] (h : Pairwise r l) : Nodup l := @@ -60,8 +49,6 @@ theorem Nodup.not_mem (h : (a :: l).Nodup) : a ∉ l := theorem not_nodup_cons_of_mem : a ∈ l → ¬Nodup (a :: l) := imp_not_comm.1 Nodup.not_mem -protected theorem Nodup.sublist : l₁ <+ l₂ → Nodup l₂ → Nodup l₁ := - Pairwise.sublist theorem not_nodup_pair (a : α) : ¬Nodup [a, a] := not_nodup_cons_of_mem <| mem_singleton_self _ @@ -171,13 +158,6 @@ theorem nodup_iff_count_eq_one [DecidableEq α] : Nodup l ↔ ∀ a ∈ l, count ⟨fun H h => H.antisymm (count_pos_iff_mem.mpr h), fun H => if h : _ then (H h).le else (count_eq_zero.mpr h).trans_le (Nat.zero_le 1)⟩ -theorem nodup_replicate (a : α) : ∀ {n : ℕ}, Nodup (replicate n a) ↔ n ≤ 1 - | 0 => by simp [Nat.zero_le] - | 1 => by simp - | n + 2 => - iff_of_false - (fun H => nodup_iff_sublist.1 H a ((replicate_sublist_replicate _).2 (Nat.le_add_left 2 n))) - (not_le_of_lt <| Nat.le_add_left 2 n) @[simp] theorem count_eq_one_of_mem [DecidableEq α] {a : α} {l : List α} (d : Nodup l) (h : a ∈ l) : @@ -245,8 +225,8 @@ theorem nodup_map_iff {f : α → β} {l : List α} (hf : Injective f) : Nodup ( @[simp] theorem nodup_attach {l : List α} : Nodup (attach l) ↔ Nodup l := - ⟨fun h => attach_map_val l ▸ h.map fun _ _ => Subtype.eq, fun h => - Nodup.of_map Subtype.val ((attach_map_val l).symm ▸ h)⟩ + ⟨fun h => attach_map_subtype_val l ▸ h.map fun _ _ => Subtype.eq, fun h => + Nodup.of_map Subtype.val ((attach_map_subtype_val l).symm ▸ h)⟩ alias ⟨Nodup.of_attach, Nodup.attach⟩ := nodup_attach @@ -265,20 +245,6 @@ theorem Nodup.filter (p : α → Bool) {l} : Nodup l → Nodup (filter p l) := b theorem nodup_reverse {l : List α} : Nodup (reverse l) ↔ Nodup l := pairwise_reverse.trans <| by simp only [Nodup, Ne, eq_comm] -theorem Nodup.erase_eq_filter [DecidableEq α] {l} (d : Nodup l) (a : α) : - l.erase a = l.filter (· ≠ a) := by - induction' d with b l m _ IH; · rfl - by_cases h : b = a - · subst h - rw [erase_cons_head, filter_cons_of_neg _ (by simp)] - symm - rw [filter_eq_self] - simpa [@eq_comm α] using m - · rw [erase_cons_tail _ (not_beq_of_ne h), filter_cons_of_pos, IH] - simp [h] - -theorem Nodup.erase [DecidableEq α] (a : α) : Nodup l → Nodup (l.erase a) := - Nodup.sublist <| erase_sublist _ _ theorem Nodup.erase_getElem [DecidableEq α] {l : List α} (hl : l.Nodup) (i : Nat) (h : i < l.length) : l.erase l[i] = l.eraseIdx ↑i := by @@ -303,11 +269,6 @@ theorem Nodup.erase_get [DecidableEq α] {l : List α} (hl : l.Nodup) (i : Fin l theorem Nodup.diff [DecidableEq α] : l₁.Nodup → (l₁.diff l₂).Nodup := Nodup.sublist <| diff_sublist _ _ -theorem Nodup.mem_erase_iff [DecidableEq α] (d : Nodup l) : a ∈ l.erase b ↔ a ≠ b ∧ a ∈ l := by - rw [d.erase_eq_filter, mem_filter, and_comm, decide_eq_true_iff] - -theorem Nodup.not_mem_erase [DecidableEq α] (h : Nodup l) : a ∉ l.erase a := fun H => - (h.mem_erase_iff.1 H).1 rfl theorem nodup_join {L : List (List α)} : Nodup (join L) ↔ (∀ l ∈ L, Nodup l) ∧ Pairwise Disjoint L := by @@ -341,7 +302,7 @@ theorem Nodup.sigma {σ : α → Type*} {l₂ : ∀ a , List (σ a)} (d₁ : Nod protected theorem Nodup.filterMap {f : α → Option β} (h : ∀ a a' b, b ∈ f a → b ∈ f a' → a = a') : Nodup l → Nodup (filterMap f l) := - (Pairwise.filter_map f) @fun a a' n b bm b' bm' e => n <| h a a' b' (by rw [← e]; exact bm) bm' + (Pairwise.filterMap f) @fun a a' n b bm b' bm' e => n <| h a a' b' (by rw [← e]; exact bm) bm' protected theorem Nodup.concat (h : a ∉ l) (h' : l.Nodup) : (l.concat a).Nodup := by rw [concat_eq_append]; exact h'.append (nodup_singleton _) (disjoint_singleton.2 h) @@ -363,8 +324,8 @@ theorem Nodup.diff_eq_filter [DecidableEq α] : | l₁, [], _ => by simp | l₁, a :: l₂, hl₁ => by rw [diff_cons, (hl₁.erase _).diff_eq_filter, hl₁.erase_eq_filter, filter_filter] - simp only [decide_not, Bool.not_eq_true', decide_eq_false_iff_not, ne_eq, and_comm, - Bool.decide_and, find?, mem_cons, not_or] + simp only [decide_not, Bool.not_eq_true', decide_eq_false_iff_not, bne_iff_ne, ne_eq, and_comm, + Bool.decide_and, mem_cons, not_or] theorem Nodup.mem_diff_iff [DecidableEq α] (hl₁ : l₁.Nodup) : a ∈ l₁.diff l₂ ↔ a ∈ l₁ ∧ a ∉ l₂ := by rw [hl₁.diff_eq_filter, mem_filter, decide_eq_true_iff] @@ -420,7 +381,7 @@ theorem Nodup.take_eq_filter_mem [DecidableEq α] : | [], n, _ => by simp | b::l, 0, _ => by simp | b::l, n+1, hl => by - rw [take_cons, Nodup.take_eq_filter_mem (Nodup.of_cons hl), List.filter_cons_of_pos _ (by simp)] + rw [take_cons, Nodup.take_eq_filter_mem (Nodup.of_cons hl), List.filter_cons_of_pos (by simp)] congr 1 refine List.filter_congr ?_ intro x hx diff --git a/Mathlib/Data/List/OfFn.lean b/Mathlib/Data/List/OfFn.lean index 6174f2fa40e72..99f86c4e1dc1c 100644 --- a/Mathlib/Data/List/OfFn.lean +++ b/Mathlib/Data/List/OfFn.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Fin.Tuple.Basic -import Mathlib.Data.List.Join +import Mathlib.Data.List.Basic /-! # Lists from functions @@ -133,7 +133,7 @@ theorem ofFn_eq_nil_iff {n : ℕ} {f : Fin n → α} : ofFn f = [] ↔ n = 0 := theorem last_ofFn {n : ℕ} (f : Fin n → α) (h : ofFn f ≠ []) (hn : n - 1 < n := Nat.pred_lt <| ofFn_eq_nil_iff.not.mp h) : - getLast (ofFn f) h = f ⟨n - 1, hn⟩ := by simp [getLast_eq_get] + getLast (ofFn f) h = f ⟨n - 1, hn⟩ := by simp [getLast_eq_getElem] theorem last_ofFn_succ {n : ℕ} (f : Fin n.succ → α) (h : ofFn f ≠ [] := mt ofFn_eq_nil_iff.mp (Nat.succ_ne_zero _)) : @@ -164,7 +164,8 @@ theorem ofFn_mul {m n} (f : Fin (m * n) → α) : _ ≤ _ := Nat.mul_le_mul_right _ i.prop⟩) := by induction' m with m IH · simp [ofFn_zero, Nat.zero_mul, ofFn_zero, join] - · simp_rw [ofFn_succ', succ_mul, join_concat, ofFn_add, IH] + · simp_rw [ofFn_succ', succ_mul] + simp [join_concat, ofFn_add, IH] rfl /-- This breaks a list of `m*n` items into `n` groups each containing `m` elements. -/ diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index 4d7e517d9638e..74b4ead49e16c 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -135,7 +135,7 @@ attribute [simp] nil_subperm theorem subperm_nil : List.Subperm l [] ↔ l = [] := ⟨fun h ↦ length_eq_zero.1 <| Nat.le_zero.1 h.length_le, by rintro rfl; rfl⟩ -lemma subperm_cons_self : l <+~ a :: l := ⟨l, Perm.refl _, sublist_cons _ _⟩ +lemma subperm_cons_self : l <+~ a :: l := ⟨l, Perm.refl _, sublist_cons_self _ _⟩ lemma count_eq_count_filter_add [DecidableEq α] (P : α → Prop) [DecidablePred P] (l : List α) (a : α) : @@ -241,14 +241,17 @@ theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h l ~ replicate m a ++ replicate n b ↔ count a l = m ∧ count b l = n ∧ l ⊆ [a, b] := by rw [perm_iff_count, ← Decidable.and_forall_ne a, ← Decidable.and_forall_ne b] suffices l ⊆ [a, b] ↔ ∀ c, c ≠ b → c ≠ a → c ∉ l by - simp (config := { contextual := true }) [count_replicate, h, h.symm, this, count_eq_zero] + simp (config := { contextual := true }) [count_replicate, h, this, count_eq_zero, Ne.symm] trans ∀ c, c ∈ l → c = b ∨ c = a · simp [subset_def, or_comm] · exact forall_congr' fun _ => by rw [← and_imp, ← not_or, not_imp_not] theorem Perm.dedup {l₁ l₂ : List α} (p : l₁ ~ l₂) : dedup l₁ ~ dedup l₂ := perm_iff_count.2 fun a => - if h : a ∈ l₁ then by simp [nodup_dedup, h, p.subset h] else by simp [h, mt p.mem_iff.2 h] + if h : a ∈ l₁ then by + simp [h, nodup_dedup, p.subset h] + else by + simp [h, count_eq_zero_of_not_mem, mt p.mem_iff.2] theorem Perm.inter_append {l t₁ t₂ : List α} (h : Disjoint t₁ t₂) : l ∩ (t₁ ++ t₂) ~ l ∩ t₁ ++ l ∩ t₂ := by @@ -333,7 +336,7 @@ theorem Perm.drop_inter [DecidableEq α] {xs ys : List α} (n : ℕ) (h : xs ~ y have h₀ : n = xs.length - n' := by rwa [Nat.sub_sub_self] have h₁ : n' ≤ xs.length := Nat.sub_le .. have h₂ : xs.drop n = (xs.reverse.take n').reverse := by - rw [take_reverse _ h₁, h₀, reverse_reverse] + rw [take_reverse h₁, h₀, reverse_reverse] rw [h₂] apply (reverse_perm _).trans rw [inter_reverse] @@ -520,7 +523,7 @@ theorem count_permutations'Aux_self [DecidableEq α] (l : List α) (x : α) : count (x :: l) (permutations'Aux x l) = length (takeWhile (x = ·) l) + 1 := by induction' l with y l IH generalizing x · simp [takeWhile, count] - · rw [permutations'Aux, DecEq_eq, count_cons_self] + · rw [permutations'Aux, count_cons_self] by_cases hx : x = y · subst hx simpa [takeWhile, Nat.succ_inj', DecEq_eq] using IH _ diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index 833ec5889075a..64aec7b5b00f9 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -4,10 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kenny Lau, Scott Morrison -/ import Mathlib.Data.List.Chain -import Mathlib.Data.List.Enum import Mathlib.Data.List.Nodup import Mathlib.Data.List.Pairwise -import Mathlib.Data.List.Zip import Batteries.Data.Nat.Lemmas /-! @@ -29,16 +27,6 @@ namespace List variable {α : Type u} -@[simp] theorem range'_one {s step : ℕ} : range' s 1 step = [s] := rfl - -theorem pairwise_lt_range' : ∀ s n (step := 1) (_ : 0 < step := by simp), - Pairwise (· < ·) (range' s n step) - | _, 0, _, _ => Pairwise.nil - | s, n + 1, _, h => chain_iff_pairwise.1 (chain_lt_range' s n h) - -theorem nodup_range' (s n : ℕ) (step := 1) (h : 0 < step := by simp) : Nodup (range' s n step) := - (pairwise_lt_range' s n step h).imp _root_.ne_of_lt - set_option linter.deprecated false in @[simp] theorem nthLe_range' {n m step} (i) (H : i < (range' n m step).length) : @@ -48,20 +36,6 @@ set_option linter.deprecated false in theorem nthLe_range'_1 {n m} (i) (H : i < (range' n m).length) : nthLe (range' n m) i H = n + i := by simp -theorem pairwise_lt_range (n : ℕ) : Pairwise (· < ·) (range n) := by - simp (config := {decide := true}) only [range_eq_range', pairwise_lt_range'] - -theorem pairwise_le_range (n : ℕ) : Pairwise (· ≤ ·) (range n) := - Pairwise.imp (@le_of_lt ℕ _) (pairwise_lt_range _) - -theorem take_range (m n : ℕ) : take m (range n) = range (min m n) := by - apply List.ext_getElem - · simp - · simp (config := { contextual := true }) [← getElem_take, Nat.lt_min] - -theorem nodup_range (n : ℕ) : Nodup (range n) := by - simp (config := {decide := true}) only [range_eq_range', nodup_range'] - theorem chain'_range_succ (r : ℕ → ℕ → Prop) (n : ℕ) : Chain' r (range n.succ) ↔ ∀ m < n, r m m.succ := by rw [range_succ] @@ -77,12 +51,6 @@ theorem chain_range_succ (r : ℕ → ℕ → Prop) (n a : ℕ) : rw [range_succ_eq_map, chain_cons, and_congr_right_iff, ← chain'_range_succ, range_succ_eq_map] exact fun _ => Iff.rfl -theorem pairwise_gt_iota (n : ℕ) : Pairwise (· > ·) (iota n) := by - simpa only [iota_eq_reverse_range', pairwise_reverse] using pairwise_lt_range' 1 n - -theorem nodup_iota (n : ℕ) : Nodup (iota n) := - (pairwise_gt_iota n).imp _root_.ne_of_gt - /-- All elements of `Fin n`, from `0` to `n-1`. The corresponding finset is `Finset.univ`. -/ def finRange (n : ℕ) : List (Fin n) := (range n).pmap Fin.mk fun _ => List.mem_range.1 @@ -115,21 +83,6 @@ theorem pairwise_lt_finRange (n : ℕ) : Pairwise (· < ·) (finRange n) := theorem pairwise_le_finRange (n : ℕ) : Pairwise (· ≤ ·) (finRange n) := (List.pairwise_le_range n).pmap (by simp) (by simp) -theorem enum_eq_zip_range (l : List α) : l.enum = (range l.length).zip l := - zip_of_prod (enum_map_fst _) (enum_map_snd _) - -@[simp] -theorem unzip_enum_eq_prod (l : List α) : l.enum.unzip = (range l.length, l) := by - simp only [enum_eq_zip_range, unzip_zip, length_range] - -theorem enumFrom_eq_zip_range' (l : List α) {n : ℕ} : l.enumFrom n = (range' n l.length).zip l := - zip_of_prod (enumFrom_map_fst _ _) (enumFrom_map_snd _ _) - -@[simp] -theorem unzip_enumFrom_eq_prod (l : List α) {n : ℕ} : - (l.enumFrom n).unzip = (range' n l.length, l) := by - simp only [enumFrom_eq_zip_range', unzip_zip, length_range'] - set_option linter.deprecated false in @[simp] theorem nthLe_range {n} (i) (H : i < (range n).length) : nthLe (range n) i H = i := diff --git a/Mathlib/Data/List/Rotate.lean b/Mathlib/Data/List/Rotate.lean index f2baca0bc0d74..854cb1f9533fe 100644 --- a/Mathlib/Data/List/Rotate.lean +++ b/Mathlib/Data/List/Rotate.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yakov Pechersky -/ import Mathlib.Data.List.Nodup -import Mathlib.Data.List.Zip import Mathlib.Data.Nat.Defs import Mathlib.Data.List.Infix @@ -181,8 +180,8 @@ theorem zipWith_rotate_distrib {β γ : Type*} (f : α → β → γ) (l : List (h : l.length = l'.length) : (zipWith f l l').rotate n = zipWith f (l.rotate n) (l'.rotate n) := by rw [rotate_eq_drop_append_take_mod, rotate_eq_drop_append_take_mod, - rotate_eq_drop_append_take_mod, h, zipWith_append, ← zipWith_distrib_drop, ← - zipWith_distrib_take, List.length_zipWith, h, min_self] + rotate_eq_drop_append_take_mod, h, zipWith_append, ← drop_zipWith, ← + take_zipWith, List.length_zipWith, h, min_self] rw [length_drop, length_drop, h] attribute [local simp] rotate_cons_succ diff --git a/Mathlib/Data/List/Sigma.lean b/Mathlib/Data/List/Sigma.lean index 4ef118bc82680..4419034921ee1 100644 --- a/Mathlib/Data/List/Sigma.lean +++ b/Mathlib/Data/List/Sigma.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Sean Leather -/ -import Mathlib.Data.List.Range import Mathlib.Data.List.Perm +import Mathlib.Data.List.Pairwise /-! # Utilities for lists of sigmas diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index e005a4ac2ee7d..7866907505fca 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -248,8 +248,9 @@ theorem perm_orderedInsert (a) : ∀ l : List α, orderedInsert r a l ~ a :: l · simpa [orderedInsert, h] using ((perm_orderedInsert a l).cons _).trans (Perm.swap _ _ _) theorem orderedInsert_count [DecidableEq α] (L : List α) (a b : α) : - count a (L.orderedInsert r b) = count a L + if a = b then 1 else 0 := by + count a (L.orderedInsert r b) = count a L + if b = a then 1 else 0 := by rw [(L.perm_orderedInsert r b).count_eq, count_cons] + simp theorem perm_insertionSort : ∀ l : List α, insertionSort r l ~ l | [] => Perm.nil @@ -300,7 +301,7 @@ theorem orderedInsert_erase [DecidableEq α] [IsAntisymm α r] (x : α) (xs : Li rw [orderedInsert, if_pos (hxs.1 _ (.head zs))] · rw [mem_cons] at hx replace hx := hx.resolve_left hxy - rw [erase_cons_tail _ (not_beq_of_ne hxy.symm), orderedInsert, ih _ hx hxs.2, if_neg] + rw [erase_cons_tail (not_beq_of_ne hxy.symm), orderedInsert, ih _ hx hxs.2, if_neg] refine mt (fun hrxy => ?_) hxy exact antisymm hrxy (hxs.1 _ hx) diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index f933fca99a9b2..7df5ef9099e10 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -119,7 +119,7 @@ theorem sublistsAux_eq_bind : List.reverseRecOn r (by simp [sublistsAux]) (fun r l ih => by - rw [append_bind, ← ih, bind_singleton, sublistsAux, foldl_append] + rw [bind_append, ← ih, bind_singleton, sublistsAux, foldl_append] simp [sublistsAux]) @[csimp] theorem sublists_eq_sublistsFast : @sublists = @sublistsFast := by @@ -177,7 +177,7 @@ theorem length_sublists (l : List α) : length (sublists l) = 2 ^ length l := by theorem map_pure_sublist_sublists (l : List α) : map pure l <+ sublists l := by induction' l using reverseRecOn with l a ih <;> simp only [map, map_append, sublists_concat] - · simp only [sublists_nil, sublist_cons] + · simp only [sublists_nil, sublist_cons_self] exact ((append_sublist_append_left _).2 <| singleton_sublist.2 <| mem_map.2 ⟨[], mem_sublists.2 (nil_sublist _), by rfl⟩).trans ((append_sublist_append_right _).2 ih) @@ -411,7 +411,7 @@ theorem range_bind_sublistsLen_perm (l : List α) : simp_rw [← List.map_bind, ← cons_append] rw [← List.singleton_append, ← List.sublistsLen_zero tl] refine Perm.append ?_ (l_ih.map _) - rw [List.range_succ, append_bind, bind_singleton, + rw [List.range_succ, bind_append, bind_singleton, sublistsLen_of_length_lt (Nat.lt_succ_self _), append_nil, ← List.bind_map Nat.succ fun n => sublistsLen n tl, ← bind_cons 0 _ fun n => sublistsLen n tl, ← range_succ_eq_map] diff --git a/Mathlib/Data/List/Sym.lean b/Mathlib/Data/List/Sym.lean index ff9b6fc3ddfba..eaf32bb1f3904 100644 --- a/Mathlib/Data/List/Sym.lean +++ b/Mathlib/Data/List/Sym.lean @@ -269,7 +269,7 @@ protected theorem Sublist.sym (n : ℕ) {xs ys : List α} (h : xs <+ ys) : xs.sy · exact h.sym (n + 1) theorem sym_sublist_sym_cons {a : α} : xs.sym n <+ (a :: xs).sym n := - (sublist_cons a xs).sym n + (sublist_cons_self a xs).sym n theorem mem_of_mem_of_mem_sym {n : ℕ} {xs : List α} {a : α} {z : Sym α n} (ha : a ∈ z) (hz : z ∈ xs.sym n) : a ∈ xs := diff --git a/Mathlib/Data/List/TFAE.lean b/Mathlib/Data/List/TFAE.lean index dc6adea7587cc..f9bce9204f88b 100644 --- a/Mathlib/Data/List/TFAE.lean +++ b/Mathlib/Data/List/TFAE.lean @@ -79,7 +79,7 @@ example (P₁ P₂ P₃ : ℕ → Prop) (H : ∀ n, [P₁ n, P₂ n, P₃ n].TFA -/ theorem forall_tfae {α : Type*} (l : List (α → Prop)) (H : ∀ a : α, (l.map (fun p ↦ p a)).TFAE) : (l.map (fun p ↦ ∀ a, p a)).TFAE := by - simp only [TFAE, List.forall_mem_map_iff] + simp only [TFAE, List.forall_mem_map] intros p₁ hp₁ p₂ hp₂ exact forall_congr' fun a ↦ H a (p₁ a) (mem_map_of_mem (fun p ↦ p a) hp₁) (p₂ a) (mem_map_of_mem (fun p ↦ p a) hp₂) @@ -98,7 +98,7 @@ example (P₁ P₂ P₃ : ℕ → Prop) (H : ∀ n, [P₁ n, P₂ n, P₃ n].TFA -/ theorem exists_tfae {α : Type*} (l : List (α → Prop)) (H : ∀ a : α, (l.map (fun p ↦ p a)).TFAE) : (l.map (fun p ↦ ∃ a, p a)).TFAE := by - simp only [TFAE, List.forall_mem_map_iff] + simp only [TFAE, List.forall_mem_map] intros p₁ hp₁ p₂ hp₂ exact exists_congr fun a ↦ H a (p₁ a) (mem_map_of_mem (fun p ↦ p a) hp₁) (p₂ a) (mem_map_of_mem (fun p ↦ p a) hp₂) diff --git a/Mathlib/Data/List/Zip.lean b/Mathlib/Data/List/Zip.lean index 230b2eac0da02..b67d88264a517 100644 --- a/Mathlib/Data/List/Zip.lean +++ b/Mathlib/Data/List/Zip.lean @@ -44,80 +44,10 @@ theorem forall_zipWith {f : α → β → γ} {p : γ → Prop} : simp only [length_cons, succ_inj'] at h simp [forall_zipWith h] -theorem lt_length_left_of_zipWith {f : α → β → γ} {i : ℕ} {l : List α} {l' : List β} - (h : i < (zipWith f l l').length) : i < l.length := by rw [length_zipWith] at h; omega - -theorem lt_length_right_of_zipWith {f : α → β → γ} {i : ℕ} {l : List α} {l' : List β} - (h : i < (zipWith f l l').length) : i < l'.length := by rw [length_zipWith] at h; omega - -theorem lt_length_left_of_zip {i : ℕ} {l : List α} {l' : List β} (h : i < (zip l l').length) : - i < l.length := - lt_length_left_of_zipWith h - -theorem lt_length_right_of_zip {i : ℕ} {l : List α} {l' : List β} (h : i < (zip l l').length) : - i < l'.length := - lt_length_right_of_zipWith h - -theorem mem_zip {a b} : ∀ {l₁ : List α} {l₂ : List β}, (a, b) ∈ zip l₁ l₂ → a ∈ l₁ ∧ b ∈ l₂ - | _ :: l₁, _ :: l₂, h => by - cases' h with _ _ _ h - · simp - · have := mem_zip h - exact ⟨Mem.tail _ this.1, Mem.tail _ this.2⟩ - -theorem unzip_eq_map : ∀ l : List (α × β), unzip l = (l.map Prod.fst, l.map Prod.snd) - | [] => rfl - | (a, b) :: l => by simp only [unzip_cons, map_cons, unzip_eq_map l] - -theorem unzip_left (l : List (α × β)) : (unzip l).1 = l.map Prod.fst := by simp only [unzip_eq_map] - -theorem unzip_right (l : List (α × β)) : (unzip l).2 = l.map Prod.snd := by simp only [unzip_eq_map] - theorem unzip_swap (l : List (α × β)) : unzip (l.map Prod.swap) = (unzip l).swap := by simp only [unzip_eq_map, map_map] rfl -theorem zip_unzip : ∀ l : List (α × β), zip (unzip l).1 (unzip l).2 = l - | [] => rfl - | (a, b) :: l => by simp only [unzip_cons, zip_cons_cons, zip_unzip l] - -theorem unzip_zip_left : - ∀ {l₁ : List α} {l₂ : List β}, length l₁ ≤ length l₂ → (unzip (zip l₁ l₂)).1 = l₁ - | [], l₂, _ => rfl - | l₁, [], h => by rw [eq_nil_of_length_eq_zero (Nat.eq_zero_of_le_zero h)]; rfl - | a :: l₁, b :: l₂, h => by - simp only [zip_cons_cons, unzip_cons, unzip_zip_left (le_of_succ_le_succ h)] - -theorem unzip_zip_right {l₁ : List α} {l₂ : List β} (h : length l₂ ≤ length l₁) : - (unzip (zip l₁ l₂)).2 = l₂ := by rw [← zip_swap, unzip_swap]; exact unzip_zip_left h - -theorem unzip_zip {l₁ : List α} {l₂ : List β} (h : length l₁ = length l₂) : - unzip (zip l₁ l₂) = (l₁, l₂) := by - rw [← Prod.mk.eta (p := unzip (zip l₁ l₂)), - unzip_zip_left (le_of_eq h), unzip_zip_right (ge_of_eq h)] - -theorem zip_of_prod {l : List α} {l' : List β} {lp : List (α × β)} (hl : lp.map Prod.fst = l) - (hr : lp.map Prod.snd = l') : lp = l.zip l' := by - rw [← hl, ← hr, ← zip_unzip lp, ← unzip_left, ← unzip_right, zip_unzip, zip_unzip] - -theorem map_prod_left_eq_zip {l : List α} (f : α → β) : - (l.map fun x => (x, f x)) = l.zip (l.map f) := by - rw [← zip_map'] - congr - exact map_id _ - -theorem map_prod_right_eq_zip {l : List α} (f : α → β) : - (l.map fun x => (f x, x)) = (l.map f).zip l := by - rw [← zip_map'] - congr - exact map_id _ - -theorem zipWith_comm (f : α → β → γ) : - ∀ (la : List α) (lb : List β), zipWith f la lb = zipWith (fun b a => f a b) lb la - | [], _ => List.zipWith_nil_right.symm - | _ :: _, [] => rfl - | _ :: as, _ :: bs => congr_arg _ (zipWith_comm f as bs) - @[congr] theorem zipWith_congr (f g : α → β → γ) (la : List α) (lb : List β) (h : List.Forall₂ (fun a b => f a b = g a b) la lb) : zipWith f la lb = zipWith g la lb := by @@ -125,16 +55,6 @@ theorem zipWith_congr (f g : α → β → γ) (la : List α) (lb : List β) · rfl · exact congr_arg₂ _ hfg ih -theorem zipWith_comm_of_comm (f : α → α → β) (comm : ∀ x y : α, f x y = f y x) (l l' : List α) : - zipWith f l l' = zipWith f l' l := by - rw [zipWith_comm] - simp only [comm] - -@[simp] -theorem zipWith_same (f : α → α → δ) : ∀ l : List α, zipWith f l l = l.map fun a => f a a - | [] => rfl - | _ :: xs => congr_arg _ (zipWith_same f xs) - theorem zipWith_zipWith_left (f : δ → γ → ε) (g : α → β → δ) : ∀ (la : List α) (lb : List β) (lc : List γ), zipWith f (zipWith g la lb) lc = zipWith3 (fun a b c => f (g a b) c) la lb lc @@ -185,11 +105,11 @@ theorem unzip_revzip (l : List α) : (revzip l).unzip = (l, l.reverse) := @[simp] theorem revzip_map_fst (l : List α) : (revzip l).map Prod.fst = l := by - rw [← unzip_left, unzip_revzip] + rw [← unzip_fst, unzip_revzip] @[simp] theorem revzip_map_snd (l : List α) : (revzip l).map Prod.snd = l.reverse := by - rw [← unzip_right, unzip_revzip] + rw [← unzip_snd, unzip_revzip] theorem reverse_revzip (l : List α) : reverse l.revzip = revzip l.reverse := by rw [← zip_unzip (revzip l).reverse] @@ -197,14 +117,6 @@ theorem reverse_revzip (l : List α) : reverse l.revzip = revzip l.reverse := by theorem revzip_swap (l : List α) : (revzip l).map Prod.swap = revzip l.reverse := by simp [revzip] -theorem getElem?_zipWith' (f : α → β → γ) (l₁ : List α) (l₂ : List β) (i : ℕ) : - (zipWith f l₁ l₂)[i]? = (l₁[i]?.map f).bind fun g => l₂[i]?.map g := by - induction' l₁ with head tail generalizing l₂ i - · rw [zipWith] <;> simp - · cases l₂ - · simp - · cases i <;> simp_all - @[deprecated (since := "2024-07-29")] alias getElem?_zip_with := getElem?_zipWith' theorem get?_zipWith' (f : α → β → γ) (l₁ : List α) (l₂ : List β) (i : ℕ) : @@ -212,14 +124,6 @@ theorem get?_zipWith' (f : α → β → γ) (l₁ : List α) (l₂ : List β) ( simp [getElem?_zipWith'] @[deprecated (since := "2024-07-29")] alias get?_zip_with := get?_zipWith' - -theorem getElem?_zipWith_eq_some (f : α → β → γ) (l₁ : List α) (l₂ : List β) (z : γ) (i : ℕ) : - (zipWith f l₁ l₂)[i]? = some z ↔ - ∃ x y, l₁[i]? = some x ∧ l₂[i]? = some y ∧ f x y = z := by - induction l₁ generalizing l₂ i - · simp - · cases l₂ <;> cases i <;> simp_all - @[deprecated (since := "2024-07-29")] alias getElem?_zip_with_eq_some := getElem?_zipWith_eq_some theorem get?_zipWith_eq_some (f : α → β → γ) (l₁ : List α) (l₂ : List β) (z : γ) (i : ℕ) : @@ -229,30 +133,10 @@ theorem get?_zipWith_eq_some (f : α → β → γ) (l₁ : List α) (l₂ : Lis @[deprecated (since := "2024-07-29")] alias get?_zip_with_eq_some := get?_zipWith_eq_some -theorem getElem?_zip_eq_some (l₁ : List α) (l₂ : List β) (z : α × β) (i : ℕ) : - (zip l₁ l₂)[i]? = some z ↔ l₁[i]? = some z.1 ∧ l₂[i]? = some z.2 := by - cases z - rw [zip, getElem?_zipWith_eq_some]; constructor - · rintro ⟨x, y, h₀, h₁, h₂⟩ - simpa [h₀, h₁] using h₂ - · rintro ⟨h₀, h₁⟩ - exact ⟨_, _, h₀, h₁, rfl⟩ - theorem get?_zip_eq_some (l₁ : List α) (l₂ : List β) (z : α × β) (i : ℕ) : (zip l₁ l₂).get? i = some z ↔ l₁.get? i = some z.1 ∧ l₂.get? i = some z.2 := by simp [getElem?_zip_eq_some] -@[simp] -theorem getElem_zipWith {f : α → β → γ} {l : List α} {l' : List β} - {i : Nat} {h : i < (zipWith f l l').length} : - (zipWith f l l')[i] = - f (l[i]'(lt_length_left_of_zipWith h)) - (l'[i]'(lt_length_right_of_zipWith h)) := by - rw [← Option.some_inj, ← getElem?_eq_getElem, getElem?_zipWith_eq_some] - exact - ⟨l[i]'(lt_length_left_of_zipWith h), l'[i]'(lt_length_right_of_zipWith h), - by rw [getElem?_eq_getElem], by rw [getElem?_eq_getElem]; exact ⟨rfl, rfl⟩⟩ - @[deprecated getElem_zipWith (since := "2024-06-12")] theorem get_zipWith {f : α → β → γ} {l : List α} {l' : List β} {i : Fin (zipWith f l l').length} : (zipWith f l l').get i = @@ -268,12 +152,6 @@ theorem nthLe_zipWith {f : α → β → γ} {l : List α} {l' : List β} {i : f (l.nthLe i (lt_length_left_of_zipWith h)) (l'.nthLe i (lt_length_right_of_zipWith h)) := get_zipWith (i := ⟨i, h⟩) -@[simp] -theorem getElem_zip {l : List α} {l' : List β} {i : Nat} {h : i < (zip l l').length} : - (zip l l')[i] = - (l[i]'(lt_length_left_of_zip h), l'[i]'(lt_length_right_of_zip h)) := - getElem_zipWith (h := h) - @[deprecated getElem_zip (since := "2024-06-12")] theorem get_zip {l : List α} {l' : List β} {i : Fin (zip l l').length} : (zip l l').get i = @@ -312,23 +190,4 @@ theorem map_uncurry_zip_eq_zipWith (f : α → β → γ) (l : List α) (l' : Li · simp · simp [hl] -section Distrib - -/-! ### Operations that can be applied before or after a `zipWith` -/ - - -variable (f : α → β → γ) (l : List α) (l' : List β) (n : ℕ) - -theorem zipWith_distrib_reverse (h : l.length = l'.length) : - (zipWith f l l').reverse = zipWith f l.reverse l'.reverse := by - induction' l with hd tl hl generalizing l' - · simp - · cases' l' with hd' tl' - · simp - · simp only [Nat.add_left_inj, length] at h - have : tl.reverse.length = tl'.reverse.length := by simp [h] - simp [hl _ h, zipWith_append _ _ _ _ _ this] - -end Distrib - end List diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 576c387fed13e..523808b4def1c 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -416,9 +416,7 @@ theorem coe_toList (s : Multiset α) : (s.toList : Multiset α) = s := theorem toList_eq_nil {s : Multiset α} : s.toList = [] ↔ s = 0 := by rw [← coe_eq_zero, coe_toList] -@[simp] -theorem empty_toList {s : Multiset α} : s.toList.isEmpty ↔ s = 0 := - isEmpty_iff_eq_nil.trans toList_eq_nil +theorem empty_toList {s : Multiset α} : s.toList.isEmpty ↔ s = 0 := by simp @[simp] theorem toList_zero : (Multiset.toList 0 : List α) = [] := @@ -498,7 +496,8 @@ theorem le_zero : s ≤ 0 ↔ s = 0 := theorem lt_cons_self (s : Multiset α) (a : α) : s < a ::ₘ s := Quot.inductionOn s fun l => suffices l <+~ a :: l ∧ ¬l ~ a :: l by simpa [lt_iff_le_and_ne] - ⟨(sublist_cons _ _).subperm, fun p => _root_.ne_of_lt (lt_succ_self (length l)) p.length_eq⟩ + ⟨(sublist_cons_self _ _).subperm, + fun p => _root_.ne_of_lt (lt_succ_self (length l)) p.length_eq⟩ theorem le_cons_self (s : Multiset α) (a : α) : s ≤ a ::ₘ s := le_of_lt <| lt_cons_self _ _ @@ -905,7 +904,7 @@ theorem erase_cons_head (a : α) (s : Multiset α) : (a ::ₘ s).erase a = s := @[simp] theorem erase_cons_tail {a b : α} (s : Multiset α) (h : b ≠ a) : (b ::ₘ s).erase a = b ::ₘ s.erase a := - Quot.inductionOn s fun l => congr_arg _ <| List.erase_cons_tail l (not_beq_of_ne h) + Quot.inductionOn s fun _ => congr_arg _ <| List.erase_cons_tail (not_beq_of_ne h) @[simp] theorem erase_singleton (a : α) : ({a} : Multiset α).erase a = 0 := @@ -1026,7 +1025,7 @@ theorem map_hcongr {β' : Type v} {m : Multiset α} {f : α → β} {f' : α → theorem forall_mem_map_iff {f : α → β} {p : β → Prop} {s : Multiset α} : (∀ y ∈ s.map f, p y) ↔ ∀ x ∈ s, p (f x) := - Quotient.inductionOn' s fun _L => List.forall_mem_map_iff + Quotient.inductionOn' s fun _L => List.forall_mem_map @[simp, norm_cast] lemma map_coe (f : α → β) (l : List α) : map f l = l.map f := rfl @@ -1348,7 +1347,7 @@ theorem pmap_eq_map_attach {p : α → Prop} (f : ∀ a, p a → β) (s) : -- @[simp] -- Porting note: Left hand does not simplify theorem attach_map_val' (s : Multiset α) (f : α → β) : (s.attach.map fun i => f i.val) = s.map f := - Quot.inductionOn s fun l => congr_arg _ <| List.attach_map_coe' l f + Quot.inductionOn s fun l => congr_arg _ <| List.attach_map_coe l f @[simp] theorem attach_map_val (s : Multiset α) : s.attach.map Subtype.val = s := @@ -1713,15 +1712,15 @@ variable {p} @[simp] theorem filter_cons_of_pos {a : α} (s) : p a → filter p (a ::ₘ s) = a ::ₘ filter p s := - Quot.inductionOn s fun l h => congr_arg ofList <| List.filter_cons_of_pos l <| by simpa using h + Quot.inductionOn s fun l h => congr_arg ofList <| List.filter_cons_of_pos <| by simpa using h @[simp] theorem filter_cons_of_neg {a : α} (s) : ¬p a → filter p (a ::ₘ s) = filter p s := - Quot.inductionOn s fun l h => congr_arg ofList <| List.filter_cons_of_neg l <| by simpa using h + Quot.inductionOn s fun l h => congr_arg ofList <| List.filter_cons_of_neg <| by simpa using h @[simp] theorem mem_filter {a : α} {s} : a ∈ filter p s ↔ a ∈ s ∧ p a := - Quot.inductionOn s fun _l => by simpa using List.mem_filter (p := (p ·)) + Quot.inductionOn s fun _l => by simp theorem of_mem_filter {a : α} {s} (h : a ∈ filter p s) : p a := (mem_filter.1 h).2 @@ -1862,12 +1861,12 @@ theorem filterMap_zero (f : α → Option β) : filterMap f 0 = 0 := @[simp] theorem filterMap_cons_none {f : α → Option β} (a : α) (s : Multiset α) (h : f a = none) : filterMap f (a ::ₘ s) = filterMap f s := - Quot.inductionOn s fun l => congr_arg ofList <| List.filterMap_cons_none a l h + Quot.inductionOn s fun _ => congr_arg ofList <| List.filterMap_cons_none h @[simp] theorem filterMap_cons_some (f : α → Option β) (a : α) (s : Multiset α) {b : β} (h : f a = some b) : filterMap f (a ::ₘ s) = b ::ₘ filterMap f s := - Quot.inductionOn s fun l => congr_arg ofList <| List.filterMap_cons_some f a l h + Quot.inductionOn s fun _ => congr_arg ofList <| List.filterMap_cons_some h theorem filterMap_eq_map (f : α → β) : filterMap (some ∘ f) = map f := funext fun s => @@ -1906,7 +1905,7 @@ theorem filterMap_some (s : Multiset α) : filterMap some s = s := @[simp] theorem mem_filterMap (f : α → Option β) (s : Multiset α) {b : β} : b ∈ filterMap f s ↔ ∃ a, a ∈ s ∧ f a = some b := - Quot.inductionOn s fun l => List.mem_filterMap f l + Quot.inductionOn s fun _ => List.mem_filterMap theorem map_filterMap_of_inv (f : α → Option β) (g : β → α) (H : ∀ x : α, (f x).map g = some x) (s : Multiset α) : map g (filterMap f s) = s := @@ -2139,9 +2138,10 @@ theorem count_replicate_self (a : α) (n : ℕ) : count a (replicate n a) = n := convert List.count_replicate_self a n rw [← coe_count, coe_replicate] -theorem count_replicate (a b : α) (n : ℕ) : count a (replicate n b) = if a = b then n else 0 := by +theorem count_replicate (a b : α) (n : ℕ) : count a (replicate n b) = if b = a then n else 0 := by convert List.count_replicate a b n rw [← coe_count, coe_replicate] + simp @[simp] theorem count_erase_self (a : α) (s : Multiset α) : count a (erase s a) = count a s - 1 := @@ -2266,7 +2266,7 @@ theorem replicate_inter (n : ℕ) (x : α) (s : Multiset α) : replicate n x ∩ s = replicate (min n (s.count x)) x := by ext y rw [count_inter, count_replicate, count_replicate] - by_cases h : y = x + by_cases h : x = y · simp only [h, if_true] · simp only [h, if_false, Nat.zero_min] diff --git a/Mathlib/Data/Multiset/FinsetOps.lean b/Mathlib/Data/Multiset/FinsetOps.lean index 54883cdea421c..0e2b4e57ed01b 100644 --- a/Mathlib/Data/Multiset/FinsetOps.lean +++ b/Mathlib/Data/Multiset/FinsetOps.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Multiset.Dedup +import Mathlib.Data.List.Infix /-! # Preparations for defining operations on `Finset`. diff --git a/Mathlib/Data/Multiset/Nodup.lean b/Mathlib/Data/Multiset/Nodup.lean index bea9d56803934..7e24031fb1f8c 100644 --- a/Mathlib/Data/Multiset/Nodup.lean +++ b/Mathlib/Data/Multiset/Nodup.lean @@ -3,8 +3,8 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Data.List.Range import Mathlib.Data.Multiset.Range +import Mathlib.Data.List.Pairwise /-! # The `Nodup` predicate for multisets without duplicate elements. @@ -148,7 +148,7 @@ instance nodupDecidable [DecidableEq α] (s : Multiset α) : Decidable (Nodup s) theorem Nodup.erase_eq_filter [DecidableEq α] (a : α) {s} : Nodup s → s.erase a = Multiset.filter (· ≠ a) s := Quot.induction_on s fun _ d => - congr_arg ((↑) : List α → Multiset α) <| List.Nodup.erase_eq_filter d a + congr_arg ((↑) : List α → Multiset α) <| by simpa using List.Nodup.erase_eq_filter d a theorem Nodup.erase [DecidableEq α] (a : α) {l} : Nodup l → Nodup (l.erase a) := nodup_of_le (erase_le _ _) diff --git a/Mathlib/Data/Multiset/Powerset.lean b/Mathlib/Data/Multiset/Powerset.lean index 42c6e7e46e363..e344eddc1b02f 100644 --- a/Mathlib/Data/Multiset/Powerset.lean +++ b/Mathlib/Data/Multiset/Powerset.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.List.Sublists +import Mathlib.Data.List.Zip import Mathlib.Data.Multiset.Bind /-! diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index c95a6d273e5f4..db2875604c859 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -60,9 +60,6 @@ protected theorem coe_injective : Injective ((↑) : ℚ≥0 → ℚ) := theorem coe_inj : (p : ℚ) = q ↔ p = q := Subtype.coe_inj -theorem ext_iff : p = q ↔ (p : ℚ) = q := - Subtype.ext_iff - theorem ne_iff {x y : ℚ≥0} : (x : ℚ) ≠ (y : ℚ) ↔ x ≠ y := NNRat.coe_inj.not diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index a74d74281b931..6741f66bbca51 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -92,9 +92,6 @@ instance canLift : CanLift ℝ ℝ≥0 toReal fun r => 0 ≤ r := @[ext] protected theorem eq {n m : ℝ≥0} : (n : ℝ) = (m : ℝ) → n = m := Subtype.eq -protected theorem eq_iff {n m : ℝ≥0} : n = m ↔ (n : ℝ) = (m : ℝ) := - Subtype.ext_iff - theorem ne_iff {x y : ℝ≥0} : (x : ℝ) ≠ (y : ℝ) ↔ x ≠ y := NNReal.eq_iff.symm.not diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index d06b2350353f5..4670f85b6aeb7 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -9,6 +9,7 @@ import Mathlib.Tactic.Cases import Mathlib.Tactic.GCongr.Core import Mathlib.Tactic.PushNeg import Mathlib.Util.AssertExists +import Batteries.Data.Nat.Basic /-! # Basic operations on the natural numbers @@ -156,16 +157,6 @@ lemma pred_eq_of_eq_succ (H : m = n.succ) : m.pred = n := by simp [H] @[simp] lemma pred_eq_succ_iff : n - 1 = m + 1 ↔ n = m + 2 := by cases n <;> constructor <;> rintro ⟨⟩ <;> rfl --- Porting note: this doesn't work as a simp lemma in Lean 4 -lemma and_forall_succ : p 0 ∧ (∀ n, p (n + 1)) ↔ ∀ n, p n := - ⟨fun h n ↦ Nat.casesOn n h.1 h.2, fun h ↦ ⟨h _, fun _ ↦ h _⟩⟩ - --- Porting note: this doesn't work as a simp lemma in Lean 4 -lemma or_exists_succ : p 0 ∨ (∃ n, p (n + 1)) ↔ ∃ n, p n := - ⟨fun h ↦ h.elim (fun h0 ↦ ⟨0, h0⟩) fun ⟨n, hn⟩ ↦ ⟨n + 1, hn⟩, by - rintro ⟨_ | n, hn⟩ - exacts [Or.inl hn, Or.inr ⟨n, hn⟩]⟩ - lemma forall_lt_succ : (∀ m < n + 1, p m) ↔ (∀ m < n, p m) ∧ p n := by simp only [Nat.lt_succ_iff, Nat.le_iff_lt_or_eq, or_comm, forall_eq_or_imp, and_comm] diff --git a/Mathlib/Data/Nat/Nth.lean b/Mathlib/Data/Nat/Nth.lean index efcdc267a571c..0ad148511a5d1 100644 --- a/Mathlib/Data/Nat/Nth.lean +++ b/Mathlib/Data/Nat/Nth.lean @@ -94,8 +94,8 @@ theorem nth_injOn (hf : (setOf p).Finite) : (Set.Iio hf.toFinset.card).InjOn (nt (nth_strictMonoOn hf).injOn theorem range_nth_of_finite (hf : (setOf p).Finite) : Set.range (nth p) = insert 0 (setOf p) := by - simpa only [← List.getD_eq_getElem?, ← nth_eq_getD_sort hf, mem_sort, Set.Finite.mem_toFinset] - using Set.range_list_getD (hf.toFinset.sort (· ≤ ·)) 0 + simpa only [← List.getD_eq_getElem?_getD, ← nth_eq_getD_sort hf, mem_sort, + Set.Finite.mem_toFinset] using Set.range_list_getD (hf.toFinset.sort (· ≤ ·)) 0 @[simp] theorem image_nth_Iio_card (hf : (setOf p).Finite) : nth p '' Set.Iio hf.toFinset.card = setOf p := diff --git a/Mathlib/Data/Option/Basic.lean b/Mathlib/Data/Option/Basic.lean index 6b335011414fd..db3660d4695ba 100644 --- a/Mathlib/Data/Option/Basic.lean +++ b/Mathlib/Data/Option/Basic.lean @@ -253,13 +253,9 @@ theorem orElse_none' (x : Option α) : x.orElse (fun _ ↦ none) = x := by cases theorem exists_ne_none {p : Option α → Prop} : (∃ x ≠ none, p x) ↔ (∃ x : α, p x) := by simp only [← exists_prop, bex_ne_none] -@[simp] -theorem isSome_map (f : α → β) (o : Option α) : isSome (o.map f) = isSome o := by - cases o <;> rfl - @[simp] theorem get_map (f : α → β) {o : Option α} (h : isSome (o.map f)) : - (o.map f).get h = f (o.get (by rwa [← isSome_map])) := by + (o.map f).get h = f (o.get (by rwa [← isSome_map'])) := by cases o <;> [simp at h; rfl] theorem iget_mem [Inhabited α] : ∀ {o : Option α}, isSome o → o.iget ∈ o @@ -357,14 +353,6 @@ theorem elim_comp₂ (h : α → β → γ) {f : γ → α} {x : α} {g : γ → theorem elim_apply {f : γ → α → β} {x : α → β} {i : Option γ} {y : α} : i.elim x f y = i.elim (x y) fun j => f j y := by rw [elim_comp fun f : α → β => f y] -@[simp] -theorem get!_some [Inhabited α] (a : α) : (some a).get! = a := - rfl - -@[simp] -theorem get!_none [Inhabited α] : (none : Option α).get! = default := - rfl - @[simp] lemma bnot_isSome (a : Option α) : (! a.isSome) = a.isNone := by funext diff --git a/Mathlib/Data/PEquiv.lean b/Mathlib/Data/PEquiv.lean index f4c6c1a6a45b5..539fbcf90563e 100644 --- a/Mathlib/Data/PEquiv.lean +++ b/Mathlib/Data/PEquiv.lean @@ -81,9 +81,6 @@ theorem coe_mk_apply (f₁ : α → Option β) (f₂ : β → Option α) (h) (x @[ext] theorem ext {f g : α ≃. β} (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h -theorem ext_iff {f g : α ≃. β} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - /-- The identity map as a partial equivalence. -/ @[refl] protected def refl (α : Type*) : α ≃. α where diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index f4eb80b1f38dc..c384a20b1de7c 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -98,9 +98,6 @@ lemma mk_inj_left {a : α} {b₁ b₂ : β} : (a, b₁) = (a, b₂) ↔ b₁ = b lemma mk_inj_right {a₁ a₂ : α} {b : β} : (a₁, b) = (a₂, b) ↔ a₁ = a₂ := (mk.inj_right _).eq_iff -protected theorem ext_iff {p q : α × β} : p = q ↔ p.1 = q.1 ∧ p.2 = q.2 := by - rw [mk.inj_iff] - theorem map_def {f : α → γ} {g : β → δ} : Prod.map f g = fun p : α × β ↦ (f p.1, g p.2) := funext fun p ↦ Prod.ext (map_fst f g p) (map_snd f g p) diff --git a/Mathlib/Data/Seq/Seq.lean b/Mathlib/Data/Seq/Seq.lean index dc2dcb81e4797..974c780ddf181 100644 --- a/Mathlib/Data/Seq/Seq.lean +++ b/Mathlib/Data/Seq/Seq.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro -/ import Mathlib.Data.Option.NAry import Mathlib.Data.Seq.Computation -import Batteries.Data.LazyList /-! # Possibly infinite lists @@ -402,42 +401,39 @@ def ofStream (s : Stream' α) : Seq α := instance coeStream : Coe (Stream' α) (Seq α) := ⟨ofStream⟩ -section LazyList +section MLList -set_option linter.deprecated false - -/-- Embed a `LazyList α` as a sequence. Note that even though this +/-- Embed a `MLList α` as a sequence. Note that even though this is non-meta, it will produce infinite sequences if used with - cyclic `LazyList`s created by meta constructions. -/ -@[deprecated (since := "2024-07-22")] -def ofLazyList : LazyList α → Seq α := + cyclic `MLList`s created by meta constructions. -/ +def ofMLList : MLList Id α → Seq α := corec fun l => - match l with - | LazyList.nil => none - | LazyList.cons a l' => some (a, l'.get) - -@[deprecated (since := "2024-07-22")] -instance coeLazyList : Coe (LazyList α) (Seq α) := - ⟨ofLazyList⟩ - -/-- Translate a sequence into a `LazyList`. Since `LazyList` and `List` - are isomorphic as non-meta types, this function is necessarily meta. -/ -@[deprecated (since := "2024-07-22")] -unsafe def toLazyList : Seq α → LazyList α + match l.uncons with + | .none => none + | .some (a, l') => some (a, l') + +@[deprecated (since := "2024-07-26")] alias ofLazyList := ofMLList + +instance coeMLList : Coe (MLList Id α) (Seq α) := + ⟨ofMLList⟩ + +@[deprecated (since := "2024-07-26")] alias coeLazyList := coeMLList + +/-- Translate a sequence into a `MLList`. -/ +unsafe def toMLList : Seq α → MLList Id α | s => match destruct s with - | none => LazyList.nil - | some (a, s') => LazyList.cons a (toLazyList s') + | none => .nil + | some (a, s') => .cons a (toMLList s') -end LazyList +@[deprecated (since := "2024-07-26")] alias toLazyList := toMLList + +end MLList /-- Translate a sequence to a list. This function will run forever if run on an infinite sequence. -/ -unsafe def forceToList : Seq α → List α - | s => - match destruct s with - | none => [] - | some (a, s') => a :: forceToList s' +unsafe def forceToList (s : Seq α) : List α := + (toMLList s).force /-- The sequence of natural numbers some 0, some 1, ... -/ def nats : Seq ℕ := diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index ba81c9b14efe6..740784235067c 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -187,9 +187,6 @@ variable {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {a b : α} {s s instance : Inhabited (Set α) := ⟨∅⟩ -protected theorem ext_iff {s t : Set α} : s = t ↔ ∀ x, x ∈ s ↔ x ∈ t := - ⟨fun h x => by rw [h], ext⟩ - @[trans] theorem mem_of_mem_of_subset {x : α} {s t : Set α} (hx : x ∈ s) (h : s ⊆ t) : x ∈ t := h hx diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 66d623ea92bd4..c0e717e96f023 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -5,6 +5,7 @@ Authors: Jeremy Avigad, Leonardo de Moura -/ import Mathlib.Data.Set.Subsingleton import Mathlib.Order.WithBot +import Mathlib.Tactic.Use import Batteries.Tactic.Congr /-! diff --git a/Mathlib/Data/Set/List.lean b/Mathlib/Data/Set/List.lean index 430a15171d4c3..bbd932e0d0918 100644 --- a/Mathlib/Data/Set/List.lean +++ b/Mathlib/Data/Set/List.lean @@ -20,10 +20,10 @@ variable {α β : Type*} (l : List α) namespace Set theorem range_list_map (f : α → β) : range (map f) = { l | ∀ x ∈ l, x ∈ range f } := by - refine antisymm (range_subset_iff.2 fun l => forall_mem_map_iff.2 fun y _ => mem_range_self _) + refine antisymm (range_subset_iff.2 fun l => forall_mem_map.2 fun y _ => mem_range_self _) fun l hl => ?_ induction' l with a l ihl; · exact ⟨[], rfl⟩ - rcases ihl fun x hx => hl x <| subset_cons _ _ hx with ⟨l, rfl⟩ + rcases ihl fun x hx => hl x <| subset_cons_self _ _ hx with ⟨l, rfl⟩ rcases hl a (mem_cons_self _ _) with ⟨a, rfl⟩ exact ⟨a :: l, map_cons _ _ _⟩ diff --git a/Mathlib/Data/Set/Pairwise/Basic.lean b/Mathlib/Data/Set/Pairwise/Basic.lean index cd21fb647a7f0..64e482ceb8858 100644 --- a/Mathlib/Data/Set/Pairwise/Basic.lean +++ b/Mathlib/Data/Set/Pairwise/Basic.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl -/ import Mathlib.Data.Set.Function import Mathlib.Logic.Pairwise +import Mathlib.Logic.Relation /-! # Relations holding pairwise diff --git a/Mathlib/Data/Sigma/Basic.lean b/Mathlib/Data/Sigma/Basic.lean index 826c42ce454b7..238a6f84ac8a5 100644 --- a/Mathlib/Data/Sigma/Basic.lean +++ b/Mathlib/Data/Sigma/Basic.lean @@ -61,9 +61,6 @@ theorem mk.inj_iff {a₁ a₂ : α} {b₁ : β a₁} {b₂ : β a₂} : theorem eta : ∀ x : Σa, β a, Sigma.mk x.1 x.2 = x | ⟨_, _⟩ => rfl -protected theorem ext_iff {x₀ x₁ : Sigma β} : x₀ = x₁ ↔ x₀.1 = x₁.1 ∧ HEq x₀.2 x₁.2 := by - cases x₀; cases x₁; exact Sigma.mk.inj_iff - /-- A version of `Iff.mp Sigma.ext_iff` for functions from a nonempty type to a sigma type. -/ theorem _root_.Function.eq_of_sigmaMk_comp {γ : Type*} [Nonempty γ] {a b : α} {f : γ → β a} {g : γ → β b} (h : Sigma.mk a ∘ f = Sigma.mk b ∘ g) : @@ -78,10 +75,6 @@ theorem subtype_ext {β : Type*} {p : α → β → Prop} : ∀ {x₀ x₁ : Σa, Subtype (p a)}, x₀.fst = x₁.fst → (x₀.snd : β) = x₁.snd → x₀ = x₁ | ⟨_, _, _⟩, ⟨_, _, _⟩, rfl, rfl => rfl -theorem subtype_ext_iff {β : Type*} {p : α → β → Prop} {x₀ x₁ : Σa, Subtype (p a)} : - x₀ = x₁ ↔ x₀.fst = x₁.fst ∧ (x₀.snd : β) = x₁.snd := - ⟨fun h ↦ h ▸ ⟨rfl, rfl⟩, fun ⟨h₁, h₂⟩ ↦ subtype_ext h₁ h₂⟩ - @[simp] theorem «forall» {p : (Σa, β a) → Prop} : (∀ x, p x) ↔ ∀ a b, p ⟨a, b⟩ := ⟨fun h a b ↦ h ⟨a, b⟩, fun h ⟨a, b⟩ ↦ h a b⟩ @@ -239,15 +232,16 @@ theorem mk.inj_iff {a₁ a₂ : α} {b₁ : β a₁} {b₂ : β a₂} : match a₁, a₂, b₁, b₂, h₁, h₂ with | _, _, _, _, Eq.refl _, HEq.refl _ => rfl -theorem ext_iff {x₀ x₁ : PSigma β} : x₀ = x₁ ↔ x₀.1 = x₁.1 ∧ HEq x₀.2 x₁.2 := by - cases x₀; cases x₁; exact PSigma.mk.inj_iff - @[simp] theorem «forall» {p : (Σ'a, β a) → Prop} : (∀ x, p x) ↔ ∀ a b, p ⟨a, b⟩ := ⟨fun h a b ↦ h ⟨a, b⟩, fun h ⟨a, b⟩ ↦ h a b⟩ +#adaptation_note +/-- +This should be renamed back to `exists` after `nightly-2024-07-31`. +-/ @[simp] -theorem «exists» {p : (Σ'a, β a) → Prop} : (∃ x, p x) ↔ ∃ a b, p ⟨a, b⟩ := +theorem exists' {p : (Σ'a, β a) → Prop} : (∃ x, p x) ↔ ∃ a b, p ⟨a, b⟩ := ⟨fun ⟨⟨a, b⟩, h⟩ ↦ ⟨a, b, h⟩, fun ⟨a, b, h⟩ ↦ ⟨⟨a, b⟩, h⟩⟩ /-- A specialized ext lemma for equality of `PSigma` types over an indexed subtype. -/ @@ -256,10 +250,6 @@ theorem subtype_ext {β : Sort*} {p : α → β → Prop} : ∀ {x₀ x₁ : Σ'a, Subtype (p a)}, x₀.fst = x₁.fst → (x₀.snd : β) = x₁.snd → x₀ = x₁ | ⟨_, _, _⟩, ⟨_, _, _⟩, rfl, rfl => rfl -theorem subtype_ext_iff {β : Sort*} {p : α → β → Prop} {x₀ x₁ : Σ'a, Subtype (p a)} : - x₀ = x₁ ↔ x₀.fst = x₁.fst ∧ (x₀.snd : β) = x₁.snd := - ⟨fun h ↦ h ▸ ⟨rfl, rfl⟩, fun ⟨h₁, h₂⟩ ↦ subtype_ext h₁ h₂⟩ - variable {α₁ : Sort*} {α₂ : Sort*} {β₁ : α₁ → Sort*} {β₂ : α₂ → Sort*} /-- Map the left and right components of a sigma -/ diff --git a/Mathlib/Data/Sigma/Order.lean b/Mathlib/Data/Sigma/Order.lean index a9e61853bbb7c..f5acd3a0b8adc 100644 --- a/Mathlib/Data/Sigma/Order.lean +++ b/Mathlib/Data/Sigma/Order.lean @@ -106,7 +106,7 @@ instance [∀ i, PartialOrder (α i)] : PartialOrder (Σi, α i) := { Sigma.preorder with le_antisymm := by rintro _ _ ⟨i, a, b, hab⟩ ⟨_, _, _, hba⟩ - exact ext rfl (heq_of_eq <| hab.antisymm hba) } + exact congr_arg (Sigma.mk _ ·) <| hab.antisymm hba } instance [∀ i, Preorder (α i)] [∀ i, DenselyOrdered (α i)] : DenselyOrdered (Σi, α i) where dense := by diff --git a/Mathlib/Data/String/Defs.lean b/Mathlib/Data/String/Defs.lean index 9afdc68d3848f..3e121b543fc6f 100644 --- a/Mathlib/Data/String/Defs.lean +++ b/Mathlib/Data/String/Defs.lean @@ -3,8 +3,8 @@ Copyright (c) 2019 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Keeley Hoek, Floris van Doorn, Chris Bailey -/ -import Batteries.Data.List.Basic import Batteries.Data.String.Basic +import Batteries.Tactic.Alias /-! # Definitions for `String` diff --git a/Mathlib/Data/String/Lemmas.lean b/Mathlib/Data/String/Lemmas.lean index 6e421a20e6f3b..8fb57239faa51 100644 --- a/Mathlib/Data/String/Lemmas.lean +++ b/Mathlib/Data/String/Lemmas.lean @@ -6,7 +6,6 @@ Authors: Chris Bailey import Mathlib.Data.Nat.Notation import Mathlib.Data.String.Defs import Mathlib.Tactic.Basic -import Batteries.Data.List.Lemmas /-! # Miscellaneous lemmas about strings diff --git a/Mathlib/Data/Subtype.lean b/Mathlib/Data/Subtype.lean index e4559ef97ec4e..48aaa0d08f165 100644 --- a/Mathlib/Data/Subtype.lean +++ b/Mathlib/Data/Subtype.lean @@ -37,35 +37,21 @@ initialize_simps_projections Subtype (val → coe) theorem prop (x : Subtype p) : p x := x.2 -@[simp] -protected theorem «forall» {q : { a // p a } → Prop} : (∀ x, q x) ↔ ∀ a b, q ⟨a, b⟩ := - ⟨fun h a b ↦ h ⟨a, b⟩, fun h ⟨a, b⟩ ↦ h a b⟩ - /-- An alternative version of `Subtype.forall`. This one is useful if Lean cannot figure out `q` when using `Subtype.forall` from right to left. -/ protected theorem forall' {q : ∀ x, p x → Prop} : (∀ x h, q x h) ↔ ∀ x : { a // p a }, q x x.2 := (@Subtype.forall _ _ fun x ↦ q x.1 x.2).symm -@[simp] -protected theorem «exists» {q : { a // p a } → Prop} : (∃ x, q x) ↔ ∃ a b, q ⟨a, b⟩ := - ⟨fun ⟨⟨a, b⟩, h⟩ ↦ ⟨a, b, h⟩, fun ⟨a, b, h⟩ ↦ ⟨⟨a, b⟩, h⟩⟩ - /-- An alternative version of `Subtype.exists`. This one is useful if Lean cannot figure out `q` when using `Subtype.exists` from right to left. -/ protected theorem exists' {q : ∀ x, p x → Prop} : (∃ x h, q x h) ↔ ∃ x : { a // p a }, q x x.2 := (@Subtype.exists _ _ fun x ↦ q x.1 x.2).symm -@[ext] -protected theorem ext : ∀ {a1 a2 : { x // p x }}, (a1 : α) = (a2 : α) → a1 = a2 - | ⟨_, _⟩, ⟨_, _⟩, rfl => rfl - -protected theorem ext_iff {a1 a2 : { x // p x }} : a1 = a2 ↔ (a1 : α) = (a2 : α) := - ⟨congr_arg _, Subtype.ext⟩ - theorem heq_iff_coe_eq (h : ∀ x, p x ↔ q x) {a1 : { x // p x }} {a2 : { x // q x }} : HEq a1 a2 ↔ (a1 : α) = (a2 : α) := - Eq.rec (motive := fun (pp : (α → Prop)) _ ↦ ∀ a2' : {x // pp x}, HEq a1 a2' ↔ (a1 : α) = (a2' : α)) - (fun _ ↦ heq_iff_eq.trans Subtype.ext_iff) (funext <| fun x ↦ propext (h x)) a2 + Eq.rec + (motive := fun (pp : (α → Prop)) _ ↦ ∀ a2' : {x // pp x}, HEq a1 a2' ↔ (a1 : α) = (a2' : α)) + (fun _ ↦ heq_iff_eq.trans Subtype.ext_iff) (funext <| fun x ↦ propext (h x)) a2 lemma heq_iff_coe_heq {α β : Sort _} {p : α → Prop} {q : β → Prop} {a : {x // p x}} {b : {y // q y}} (h : α = β) (h' : HEq p q) : HEq a b ↔ HEq (a : α) (b : β) := by diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index 4612fa17121c9..aa85d9183ee69 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -522,7 +522,7 @@ theorem fill_filterNe [DecidableEq α] (a : α) (m : Sym α n) : rw [count_add, count_filter, Sym.coe_replicate, count_replicate] obtain rfl | h := eq_or_ne a b · rw [if_pos rfl, if_neg (not_not.2 rfl), zero_add] - · rw [if_pos h, if_neg h.symm, add_zero]) + · rw [if_pos h, if_neg h, add_zero]) theorem filter_ne_fill [DecidableEq α] (a : α) (m : Σi : Fin (n + 1), Sym α (n - i)) (h : a ∉ m.2) : (m.2.fill a m.1).filterNe a = m := diff --git a/Mathlib/Data/ULift.lean b/Mathlib/Data/ULift.lean index 2534752afb819..45fa5cad19423 100644 --- a/Mathlib/Data/ULift.lean +++ b/Mathlib/Data/ULift.lean @@ -135,7 +135,4 @@ theorem «exists» {p : ULift α → Prop} : (∃ x, p x) ↔ ∃ x : α, p (ULi theorem ext (x y : ULift α) (h : x.down = y.down) : x = y := congrArg up h -theorem ext_iff {α : Type*} (x y : ULift α) : x = y ↔ x.down = y.down := - ⟨congrArg _, ULift.ext _ _⟩ - end ULift diff --git a/Mathlib/Data/Vector/Zip.lean b/Mathlib/Data/Vector/Zip.lean index 5074615cb13be..eee02c2029439 100644 --- a/Mathlib/Data/Vector/Zip.lean +++ b/Mathlib/Data/Vector/Zip.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Mathlib.Data.Vector.Basic -import Mathlib.Data.List.Zip /-! # The `zipWith` operation on vectors. diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index 7464246e2fb52..7e51244947644 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -160,9 +160,6 @@ theorem map_one_add (x : ℝ) : f (1 + x) = 1 + f x := by rw [add_comm, map_add_ theorem ext ⦃f g : CircleDeg1Lift⦄ (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h -protected theorem ext_iff {f g : CircleDeg1Lift} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - instance : Monoid CircleDeg1Lift where mul f g := { toOrderHom := f.1.comp g.1 diff --git a/Mathlib/Geometry/Euclidean/Angle/Sphere.lean b/Mathlib/Geometry/Euclidean/Angle/Sphere.lean index 59b4d58ff7acb..dd2a20922dbdb 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Sphere.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Sphere.lean @@ -281,7 +281,7 @@ theorem circumsphere_eq_of_dist_of_oangle (t : Triangle ℝ P) {i₁ i₂ i₃ : o.rotation (π / 2 : ℝ) (t.points i₃ -ᵥ t.points i₁) +ᵥ midpoint ℝ (t.points i₁) (t.points i₃), dist (t.points i₁) (t.points i₃) / |Real.Angle.sin (∡ (t.points i₁) (t.points i₂) (t.points i₃))| / 2⟩ := - t.circumsphere.ext _ + t.circumsphere.ext (t.inv_tan_div_two_smul_rotation_pi_div_two_vadd_midpoint_eq_circumcenter h₁₂ h₁₃ h₂₃).symm (t.dist_div_sin_oangle_div_two_eq_circumradius h₁₂ h₁₃ h₂₃).symm diff --git a/Mathlib/Geometry/Euclidean/Circumcenter.lean b/Mathlib/Geometry/Euclidean/Circumcenter.lean index adf0c6844407d..991e11164415f 100644 --- a/Mathlib/Geometry/Euclidean/Circumcenter.lean +++ b/Mathlib/Geometry/Euclidean/Circumcenter.lean @@ -793,7 +793,7 @@ theorem exists_circumsphere_eq_of_cospherical_subset {s : AffineSubspace ℝ P} ∃ c : Sphere P, ∀ sx : Simplex ℝ P n, Set.range sx.points ⊆ ps → sx.circumsphere = c := by obtain ⟨r, hr⟩ := exists_circumradius_eq_of_cospherical_subset h hd hc obtain ⟨c, hc⟩ := exists_circumcenter_eq_of_cospherical_subset h hd hc - exact ⟨⟨c, r⟩, fun sx hsx => Sphere.ext _ _ (hc sx hsx) (hr sx hsx)⟩ + exact ⟨⟨c, r⟩, fun sx hsx => Sphere.ext (hc sx hsx) (hr sx hsx)⟩ /-- Two n-simplices among cospherical points in an n-dimensional subspace have the same circumsphere. -/ diff --git a/Mathlib/Geometry/Euclidean/Sphere/Basic.lean b/Mathlib/Geometry/Euclidean/Sphere/Basic.lean index 480ea95c717cd..8b1bfe714bc4f 100644 --- a/Mathlib/Geometry/Euclidean/Sphere/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Sphere/Basic.lean @@ -108,7 +108,7 @@ theorem Sphere.ne_iff {s₁ s₂ : Sphere P} : theorem Sphere.center_eq_iff_eq_of_mem {s₁ s₂ : Sphere P} {p : P} (hs₁ : p ∈ s₁) (hs₂ : p ∈ s₂) : s₁.center = s₂.center ↔ s₁ = s₂ := by - refine ⟨fun h => Sphere.ext _ _ h ?_, fun h => h ▸ rfl⟩ + refine ⟨fun h => Sphere.ext h ?_, fun h => h ▸ rfl⟩ rw [mem_sphere] at hs₁ hs₂ rw [← hs₁, ← hs₂, h] diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index b6bfd88f9c7bc..2cfef4d3ff400 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -539,7 +539,7 @@ theorem mfderivWithin_congr (hs : UniqueMDiffWithinAt I s x) (hL : ∀ x ∈ s, theorem tangentMapWithin_congr (h : ∀ x ∈ s, f x = f₁ x) (p : TangentBundle I M) (hp : p.1 ∈ s) (hs : UniqueMDiffWithinAt I s p.1) : tangentMapWithin I I' f s p = tangentMapWithin I I' f₁ s p := by - refine TotalSpace.ext _ _ (h p.1 hp) ?_ + refine TotalSpace.ext (h p.1 hp) ?_ -- This used to be `simp only`, but we need `erw` after leanprover/lean4#2644 rw [tangentMapWithin, h p.1 hp, tangentMapWithin, mfderivWithin_congr hs h (h _ hp)] diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean index 25e4dfacdf7a2..340c62513dedd 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace.lean @@ -80,7 +80,7 @@ instance : Quiver LocallyRingedSpace := ⟨Hom⟩ @[ext] lemma Hom.ext' (X Y : LocallyRingedSpace.{u}) {f g : X ⟶ Y} (h : f.val = g.val) : f = g := - Hom.ext _ _ h + Hom.ext h /-- A morphism of locally ringed spaces `f : X ⟶ Y` induces a local ring homomorphism from `Y.stalk (f x)` to `X.stalk x` for any `x : X`. @@ -115,9 +115,9 @@ instance : Category LocallyRingedSpace.{u} where Hom := Hom id := id comp {X Y Z} f g := comp f g - comp_id {X Y} f := Hom.ext _ _ <| by simp [comp] - id_comp {X Y} f := Hom.ext _ _ <| by simp [comp] - assoc {_ _ _ _} f g h := Hom.ext _ _ <| by simp [comp] + comp_id {X Y} f := Hom.ext <| by simp [comp] + id_comp {X Y} f := Hom.ext <| by simp [comp] + assoc {_ _ _ _} f g h := Hom.ext <| by simp [comp] /-- The forgetful functor from `LocallyRingedSpace` to `SheafedSpace CommRing`. -/ @[simps] @@ -126,7 +126,7 @@ def forgetToSheafedSpace : LocallyRingedSpace.{u} ⥤ SheafedSpace CommRingCat.{ map {X Y} f := f.1 instance : forgetToSheafedSpace.Faithful where - map_injective {_ _} _ _ h := Hom.ext _ _ h + map_injective {_ _} _ _ h := Hom.ext h /-- The forgetful functor from `LocallyRingedSpace` to `Top`. -/ @[simps!] @@ -178,13 +178,13 @@ def isoOfSheafedSpaceIso {X Y : LocallyRingedSpace.{u}} (f : X.toSheafedSpace X ≅ Y where hom := homOfSheafedSpaceHomOfIsIso f.hom inv := homOfSheafedSpaceHomOfIsIso f.inv - hom_inv_id := Hom.ext _ _ f.hom_inv_id - inv_hom_id := Hom.ext _ _ f.inv_hom_id + hom_inv_id := Hom.ext f.hom_inv_id + inv_hom_id := Hom.ext f.inv_hom_id instance : forgetToSheafedSpace.ReflectsIsomorphisms where reflects {_ _} f i := { out := ⟨homOfSheafedSpaceHomOfIsIso (CategoryTheory.inv (forgetToSheafedSpace.map f)), - Hom.ext _ _ (IsIso.hom_inv_id (I := i)), Hom.ext _ _ (IsIso.inv_hom_id (I := i))⟩ } + Hom.ext (IsIso.hom_inv_id (I := i)), Hom.ext (IsIso.inv_hom_id (I := i))⟩ } instance is_sheafedSpace_iso {X Y : LocallyRingedSpace.{u}} (f : X ⟶ Y) [IsIso f] : IsIso f.1 := LocallyRingedSpace.forgetToSheafedSpace.map_isIso f diff --git a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean index a90c40ac98dc7..b94c92ea5a9a2 100644 --- a/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/LocallyRingedSpace/HasColimits.lean @@ -107,10 +107,10 @@ noncomputable def coproductCofanIsColimit : IsColimit (coproductCofan F) where (((forgetToSheafedSpace.mapCocone s).ι.app i).stalkMap y) := (s.ι.app i).2 y infer_instance⟩ - fac s j := LocallyRingedSpace.Hom.ext _ _ + fac s j := LocallyRingedSpace.Hom.ext (colimit.ι_desc (C := SheafedSpace.{u+1, u, u} CommRingCatMax.{u, u}) _ _) uniq s f h := - LocallyRingedSpace.Hom.ext _ _ + LocallyRingedSpace.Hom.ext (IsColimit.uniq _ (forgetToSheafedSpace.mapCocone s) f.1 fun j => congr_arg LocallyRingedSpace.Hom.val (h j)) @@ -254,7 +254,7 @@ noncomputable def coequalizerCofork : Cofork f g := @Cofork.ofπ _ _ _ _ f g (coequalizer f g) ⟨coequalizer.π f.1 g.1, -- Porting note: this used to be automatic HasCoequalizer.coequalizer_π_stalk_isLocalRingHom _ _⟩ - (LocallyRingedSpace.Hom.ext _ _ (coequalizer.condition f.1 g.1)) + (LocallyRingedSpace.Hom.ext (coequalizer.condition f.1 g.1)) theorem isLocalRingHom_stalkMap_congr {X Y : RingedSpace} (f g : X ⟶ Y) (H : f = g) (x) (h : IsLocalRingHom (f.stalkMap x)) : @@ -282,7 +282,7 @@ noncomputable def coequalizerCoforkIsColimit : IsColimit (coequalizerCofork f g) apply isLocalRingHom_stalkMap_congr _ _ (coequalizer.π_desc s.π.1 e).symm y infer_instance constructor - · exact LocallyRingedSpace.Hom.ext _ _ (coequalizer.π_desc _ _) + · exact LocallyRingedSpace.Hom.ext (coequalizer.π_desc _ _) intro m h replace h : (coequalizerCofork f g).π.1 ≫ m.1 = s.π.1 := by rw [← h]; rfl apply LocallyRingedSpace.Hom.ext diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 7d86d2963f23e..edd9564421caf 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -124,7 +124,7 @@ noncomputable def isoRestrict : X ≅ Y.restrict H.base_open := @[reassoc (attr := simp)] theorem isoRestrict_hom_ofRestrict : (isoRestrict f).hom ≫ Y.ofRestrict _ = f := by -- Porting note: `ext` did not pick up `NatTrans.ext` - refine PresheafedSpace.Hom.ext _ _ rfl <| NatTrans.ext _ _ <| funext fun x => ?_ + refine PresheafedSpace.Hom.ext _ _ rfl <| NatTrans.ext <| funext fun x => ?_ simp only [isoRestrict_hom_c_app, NatTrans.comp_app, eqToHom_refl, ofRestrict_c_app, Category.assoc, whiskerRight_id'] erw [Category.comp_id, comp_c_app, f.c.naturality_assoc, ← X.presheaf.map_comp] @@ -334,7 +334,7 @@ def pullbackConeOfLeftFst : theorem pullback_cone_of_left_condition : pullbackConeOfLeftFst f g ≫ f = Y.ofRestrict _ ≫ g := by -- Porting note: `ext` did not pick up `NatTrans.ext` - refine PresheafedSpace.Hom.ext _ _ ?_ <| NatTrans.ext _ _ <| funext fun U => ?_ + refine PresheafedSpace.Hom.ext _ _ ?_ <| NatTrans.ext <| funext fun U => ?_ · simpa using pullback.condition · induction U using Opposite.rec' -- Porting note: `NatTrans.comp_app` is not picked up by `dsimp` @@ -392,7 +392,7 @@ def pullbackConeOfLeftLift : s.pt ⟶ (pullbackConeOfLeft f g).pt where theorem pullbackConeOfLeftLift_fst : pullbackConeOfLeftLift f g s ≫ (pullbackConeOfLeft f g).fst = s.fst := by -- Porting note: `ext` did not pick up `NatTrans.ext` - refine PresheafedSpace.Hom.ext _ _ ?_ <| NatTrans.ext _ _ <| funext fun x => ?_ + refine PresheafedSpace.Hom.ext _ _ ?_ <| NatTrans.ext <| funext fun x => ?_ · change pullback.lift _ _ _ ≫ pullback.fst _ _ = _ simp · induction x using Opposite.rec' with | h x => ?_ @@ -411,7 +411,7 @@ theorem pullbackConeOfLeftLift_fst : theorem pullbackConeOfLeftLift_snd : pullbackConeOfLeftLift f g s ≫ (pullbackConeOfLeft f g).snd = s.snd := by -- Porting note: `ext` did not pick up `NatTrans.ext` - refine PresheafedSpace.Hom.ext _ _ ?_ <| NatTrans.ext _ _ <| funext fun x => ?_ + refine PresheafedSpace.Hom.ext _ _ ?_ <| NatTrans.ext <| funext fun x => ?_ · change pullback.lift _ _ _ ≫ pullback.snd _ _ = _ simp · change (_ ≫ _ ≫ _) ≫ _ = _ @@ -960,7 +960,7 @@ def pullbackConeOfLeft : PullbackCone f g := by rw [← IsIso.eq_inv_comp] at this rw [this] infer_instance - · exact LocallyRingedSpace.Hom.ext _ _ + · exact LocallyRingedSpace.Hom.ext (PresheafedSpace.IsOpenImmersion.pullback_cone_of_left_condition _ _) instance : LocallyRingedSpace.IsOpenImmersion (pullbackConeOfLeft f g).snd := @@ -971,9 +971,9 @@ def pullbackConeOfLeftIsLimit : IsLimit (pullbackConeOfLeft f g) := PullbackCone.isLimitAux' _ fun s => by refine ⟨LocallyRingedSpace.Hom.mk (PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftLift f.1 g.1 (PullbackCone.mk _ _ (congr_arg LocallyRingedSpace.Hom.val s.condition))) ?_, - LocallyRingedSpace.Hom.ext _ _ + LocallyRingedSpace.Hom.ext (PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftLift_fst f.1 g.1 _), - LocallyRingedSpace.Hom.ext _ _ + LocallyRingedSpace.Hom.ext (PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftLift_snd f.1 g.1 _), ?_⟩ · intro x have := @@ -987,7 +987,7 @@ def pullbackConeOfLeftIsLimit : IsLimit (pullbackConeOfLeft f g) := infer_instance · intro m _ h₂ rw [← cancel_mono (pullbackConeOfLeft f g).snd] - exact h₂.trans <| LocallyRingedSpace.Hom.ext _ _ + exact h₂.trans <| LocallyRingedSpace.Hom.ext (PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftLift_snd f.1 g.1 <| PullbackCone.mk s.fst.1 s.snd.1 <| congr_arg LocallyRingedSpace.Hom.val s.condition).symm diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean index 6127625f49c64..b94f7d8b249c7 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace.lean @@ -94,7 +94,7 @@ structure Hom (X Y : PresheafedSpace C) where -- rather than `Hom X Y`, this one was renamed `Hom.ext` instead of `ext`, -- and the more practical lemma `ext` is defined just after the definition -- of the `Category` instance -@[ext] +@[ext (iff := false)] theorem Hom.ext {X Y : PresheafedSpace C} (α β : Hom X Y) (w : α.base = β.base) (h : α.c ≫ whiskerRight (eqToHom (by rw [w])) _ = β.c) : α = β := by rcases α with ⟨base, c⟩ @@ -158,7 +158,7 @@ instance categoryOfPresheafedSpaces : Category (PresheafedSpace C) where variable {C} -- Porting note (#5229): adding an `ext` lemma. -@[ext] +@[ext (iff := false)] theorem ext {X Y : PresheafedSpace C} (α β : X ⟶ Y) (w : α.base = β.base) (h : α.c ≫ whiskerRight (eqToHom (by rw [w])) _ = β.c) : α = β := Hom.ext α β w h diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean index 2b08ae57ebd2f..1a3dd8fdd7585 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/HasColimits.lean @@ -102,14 +102,14 @@ def pushforwardDiagramToColimit (F : J ⥤ PresheafedSpace.{_, _, v} C) : (pushforwardEq (colimit.w (F ⋙ PresheafedSpace.forget C) f) (F.obj j).presheaf).hom).op map_id j := by apply (opEquiv _ _).injective - refine NatTrans.ext _ _ (funext fun U => ?_) + refine NatTrans.ext (funext fun U => ?_) induction U with | h U => simp [opEquiv] rfl map_comp {j₁ j₂ j₃} f g := by apply (opEquiv _ _).injective - refine NatTrans.ext _ _ (funext fun U => ?_) + refine NatTrans.ext (funext fun U => ?_) dsimp [opEquiv] have : op ((Opens.map (F.map g).base).obj @@ -253,7 +253,7 @@ def colimitCoconeIsColimit (F : J ⥤ PresheafedSpace.{_, _, v} C) : simp ext : 1 · exact t - · refine NatTrans.ext _ _ (funext fun U => limit_obj_ext fun j => ?_) + · refine NatTrans.ext (funext fun U => limit_obj_ext fun j => ?_) simp [desc, descCApp, PresheafedSpace.congr_app (w (unop j)).symm U, NatTrans.congr (limit.π (pushforwardDiagramToColimit F).leftOp j) diff --git a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean index 2b6137deed9c7..5ec030cd1628c 100644 --- a/Mathlib/Geometry/RingedSpace/SheafedSpace.lean +++ b/Mathlib/Geometry/RingedSpace/SheafedSpace.lean @@ -81,8 +81,8 @@ instance : Category (SheafedSpace C) := infer_instance -- Porting note (#5229): adding an `ext` lemma. -@[ext] -theorem ext {X Y : SheafedSpace C} {α β : X ⟶ Y} (w : α.base = β.base) +@[ext (iff := false)] +theorem ext {X Y : SheafedSpace C} (α β : X ⟶ Y) (w : α.base = β.base) (h : α.c ≫ whiskerRight (eqToHom (by rw [w])) _ = β.c) : α = β := PresheafedSpace.ext α β w h diff --git a/Mathlib/GroupTheory/Congruence/Basic.lean b/Mathlib/GroupTheory/Congruence/Basic.lean index 7f1772576e807..2c1b0c3b3405a 100644 --- a/Mathlib/GroupTheory/Congruence/Basic.lean +++ b/Mathlib/GroupTheory/Congruence/Basic.lean @@ -171,11 +171,6 @@ relation is injective."] theorem toSetoid_inj {c d : Con M} (H : c.toSetoid = d.toSetoid) : c = d := ext <| ext_iff.1 H -/-- Iff version of extensionality rule for congruence relations. -/ -@[to_additive "Iff version of extensionality rule for additive congruence relations."] -protected theorem ext_iff {c d : Con M} : c = d ↔ (∀ x y, c x y ↔ d x y) := - ⟨fun h _ _ => h ▸ Iff.rfl, ext⟩ - /-- Two congruence relations are equal iff their underlying binary relations are equal. -/ @[to_additive "Two additive congruence relations are equal iff their underlying binary relations are equal."] diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index 7bc3fa28299da..719e0b7b18fae 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -357,7 +357,7 @@ theorem rcons_inj {i} : Function.Injective (rcons : Pair M i → Word M) := by heq_iff_eq, ← Subtype.ext_iff_val] using he rcases this with ⟨rfl, h⟩ congr - exact Word.ext _ _ h + exact Word.ext h theorem mem_rcons_iff {i j : ι} (p : Pair M i) (m : M j) : ⟨_, m⟩ ∈ (rcons p).toList ↔ ⟨_, m⟩ ∈ p.tail.toList ∨ @@ -526,7 +526,7 @@ theorem mem_smul_iff {i j : ι} {m₁ : M i} {m₂ : M j} {w : Word M} : intro hm1 split_ifs with h · rcases h with ⟨hnil, rfl⟩ - simp only [List.head?_eq_head _ hnil, Option.some.injEq, ne_eq] + simp only [List.head?_eq_head hnil, Option.some.injEq, ne_eq] constructor · rintro rfl exact Or.inl ⟨_, rfl, rfl⟩ @@ -605,7 +605,7 @@ def equiv : CoprodI M ≃ Word M where rw [prod_smul, mul_smul, ih] instance : DecidableEq (Word M) := - Function.Injective.decidableEq Word.ext + Function.Injective.decidableEq fun _ _ => Word.ext instance : DecidableEq (CoprodI M) := Equiv.decidableEq Word.equiv @@ -637,7 +637,7 @@ def toList : ∀ {i j} (_w : NeWord M i j), List (Σi, M i) theorem toList_ne_nil {i j} (w : NeWord M i j) : w.toList ≠ List.nil := by induction w · rintro ⟨rfl⟩ - · apply List.append_ne_nil_of_ne_nil_left + · apply List.append_ne_nil_of_left_ne_nil assumption /-- The first letter of a `NeWord` -/ @@ -658,7 +658,7 @@ theorem toList_head? {i j} (w : NeWord M i j) : w.toList.head? = Option.some ⟨ induction w · rw [Option.mem_def] rfl - · exact List.head?_append (by assumption) + · exact List.mem_head?_append_of_mem_head? (by assumption) @[simp] theorem toList_getLast? {i j} (w : NeWord M i j) : w.toList.getLast? = Option.some ⟨j, w.last⟩ := by @@ -666,7 +666,7 @@ theorem toList_getLast? {i j} (w : NeWord M i j) : w.toList.getLast? = Option.so induction w · rw [Option.mem_def] rfl - · exact List.getLast?_append (by assumption) + · exact List.mem_getLast?_append_of_mem_getLast? (by assumption) /-- The `Word M` represented by a `NeWord M i j` -/ def toWord {i j} (w : NeWord M i j) : Word M where diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index 1c942c9080420..90a3048bbb6de 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -250,7 +250,7 @@ theorem getD_rightInvSeq (ω : List B) (j : ℕ) : · dsimp only [rightInvSeq] rcases j with _ | j' · simp [getD_cons_zero] - · simp only [getD_eq_getElem?, get?_eq_getElem?] at ih + · simp only [getD_eq_getElem?_getD, get?_eq_getElem?] at ih simp [getD_cons_succ, ih j'] theorem getD_leftInvSeq (ω : List B) (j : ℕ) : @@ -299,11 +299,11 @@ theorem leftInvSeq_take (ω : List B) (j : ℕ) : lis (ω.take j) = (lis ω).take j := by obtain le | ge := Nat.le_or_ge j ω.length · simp only [leftInvSeq_eq_reverse_rightInvSeq_reverse] - rw [List.take_reverse j (by simpa)] + rw [List.take_reverse (by simpa)] nth_rw 1 [← List.reverse_reverse ω] - rw [List.take_reverse j (by simpa)] + rw [List.take_reverse (by simpa)] simp [rightInvSeq_drop] - · rw [take_length_le ge, take_length_le (by simpa)] + · rw [take_of_length_le ge, take_of_length_le (by simpa)] theorem isReflection_of_mem_rightInvSeq (ω : List B) {t : W} (ht : t ∈ ris ω) : cs.IsReflection t := by @@ -400,7 +400,7 @@ theorem IsReduced.nodup_rightInvSeq {ω : List B} (rω : cs.IsReduced ω) : List have h₃ : t' = (ris ω).getD j' 1 := by rw [h₂, cs.getD_rightInvSeq, cs.getD_rightInvSeq, (Nat.sub_add_cancel (by omega) : j' - 1 + 1 = j'), eraseIdx_eq_take_drop_succ, - drop_append_eq_append_drop, drop_length_le (by simp [j_lt_j'.le]), length_take, + drop_append_eq_append_drop, drop_of_length_le (by simp [j_lt_j'.le]), length_take, drop_drop, nil_append, min_eq_left_of_lt (j_lt_j'.trans j'_lt_length), ← add_assoc, Nat.sub_add_cancel (by omega), mul_left_inj, mul_right_inj] congr 2 diff --git a/Mathlib/GroupTheory/GroupAction/Hom.lean b/Mathlib/GroupTheory/GroupAction/Hom.lean index 8ba622aa3480f..d843581eba0b4 100644 --- a/Mathlib/GroupTheory/GroupAction/Hom.lean +++ b/Mathlib/GroupTheory/GroupAction/Hom.lean @@ -156,10 +156,6 @@ theorem ext {f g : X →ₑ[φ] Y} : (∀ x, f x = g x) → f = g := DFunLike.ext f g -protected theorem ext_iff {f g : X →ₑ[φ] Y} : - f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - protected theorem congr_fun {f g : X →ₑ[φ] Y} (h : f = g) (x : X) : f x = g x := DFunLike.congr_fun h _ @@ -428,9 +424,6 @@ theorem coe_fn_coe' (f : A →ₑ+[φ] B) : ⇑(f : A →ₑ[φ] B) = f := theorem ext {f g : A →ₑ+[φ] B} : (∀ x, f x = g x) → f = g := DFunLike.ext f g -theorem ext_iff {f g : A →ₑ+[φ] B} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - protected theorem congr_fun {f g : A →ₑ+[φ] B} (h : f = g) (x : A) : f x = g x := DFunLike.congr_fun h _ @@ -547,8 +540,6 @@ theorem ext_ring {f g : R →ₑ+[σ] N'} (h : f 1 = g 1) : f = g := by ext x rw [← mul_one x, ← smul_eq_mul R, f.map_smulₑ, g.map_smulₑ, h] -theorem ext_ring_iff {f g : R →ₑ+[σ] N'} : f = g ↔ f 1 = g 1 := - ⟨fun h => h ▸ rfl, ext_ring⟩ end Semiring @@ -676,9 +667,6 @@ theorem coe_fn_coe' (f : R →ₑ+*[φ] S) : ⇑(f : R →ₑ+[φ] S) = f := theorem ext {f g : R →ₑ+*[φ] S} : (∀ x, f x = g x) → f = g := DFunLike.ext f g -theorem ext_iff {f g : R →ₑ+*[φ] S} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - protected theorem map_zero (f : R →ₑ+*[φ] S) : f 0 = 0 := map_zero f diff --git a/Mathlib/GroupTheory/HNNExtension.lean b/Mathlib/GroupTheory/HNNExtension.lean index c300d2371ccc6..eaf0e222c3698 100644 --- a/Mathlib/GroupTheory/HNNExtension.lean +++ b/Mathlib/GroupTheory/HNNExtension.lean @@ -645,8 +645,8 @@ theorem exists_normalWord_prod_eq have : a.fst = -a.fst := by have hl : l ≠ [] := by rintro rfl; simp_all have : a.fst = (l.head hl).fst := (List.chain'_cons'.1 chain).1 (l.head hl) - (List.head?_eq_head _ _) hS - rwa [List.head?_eq_head _ hl, Option.map_some', ← this, Option.some_inj] at hx' + (List.head?_eq_head _) hS + rwa [List.head?_eq_head hl, Option.map_some', ← this, Option.some_inj] at hx' simp at this erw [List.map_cons, mul_smul, of_smul_eq_smul, NormalWord.group_smul_def, t_pow_smul_eq_unitsSMul, unitsSMul, dif_neg this, ← hw'2] diff --git a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean index 487dffa39c60e..8089e5d60ce04 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean @@ -386,10 +386,6 @@ theorem ext {f g : LocalizationMap S N} (h : ∀ x, f.toMap x = g.toMap x) : f = simp only [mk.injEq, MonoidHom.mk.injEq] exact OneHom.ext h -@[to_additive] -protected theorem ext_iff {f g : LocalizationMap S N} : f = g ↔ ∀ x, f.toMap x = g.toMap x := - ⟨fun h _ ↦ h ▸ rfl, ext⟩ - @[to_additive] theorem toMap_injective : Function.Injective (@LocalizationMap.toMap _ _ S N _) := fun _ _ h ↦ ext <| DFunLike.ext_iff.1 h diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean index 31593924565f6..851351d1966d0 100644 --- a/Mathlib/GroupTheory/PushoutI.lean +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -279,9 +279,6 @@ theorem ext {w₁ w₂ : NormalWord d} (hhead : w₁.head = w₂.head) rcases w₂ with ⟨⟨_, _, _⟩, _, _⟩ simp_all -theorem ext_iff {w₁ w₂ : NormalWord d} : w₁ = w₂ ↔ w₁.head = w₂.head ∧ w₁.toList = w₂.toList := - ⟨fun h => by simp [h], fun ⟨h₁, h₂⟩ => ext h₁ h₂⟩ - open Subgroup.IsComplement /-- Given a word in `CoprodI`, if every letter is in the transversal and when @@ -320,12 +317,12 @@ theorem eq_one_of_smul_normalized (w : CoprodI.Word G) {i : ι} (h : H) dsimp split_ifs with hep · rcases hep with ⟨hnil, rfl⟩ - rw [head?_eq_head _ hnil] + rw [head?_eq_head hnil] simp_all · push_neg at hep by_cases hw : w.toList = [] · simp [hw, Word.fstIdx] - · simp [head?_eq_head _ hw, Word.fstIdx, hep hw] + · simp [head?_eq_head hw, Word.fstIdx, hep hw] theorem ext_smul {w₁ w₂ : NormalWord d} (i : ι) (h : CoprodI.of (φ i w₁.head) • w₁.toWord = diff --git a/Mathlib/GroupTheory/SemidirectProduct.lean b/Mathlib/GroupTheory/SemidirectProduct.lean index 66e43e8051a1c..9d702707e2f17 100644 --- a/Mathlib/GroupTheory/SemidirectProduct.lean +++ b/Mathlib/GroupTheory/SemidirectProduct.lean @@ -87,10 +87,10 @@ theorem inv_left (a : N ⋊[φ] G) : a⁻¹.left = φ a.right⁻¹ a.left⁻¹ : theorem inv_right (a : N ⋊[φ] G) : a⁻¹.right = a.right⁻¹ := rfl instance : Group (N ⋊[φ] G) where - mul_assoc a b c := SemidirectProduct.ext _ _ (by simp [mul_assoc]) (by simp [mul_assoc]) - one_mul a := SemidirectProduct.ext _ _ (by simp) (one_mul a.2) - mul_one a := SemidirectProduct.ext _ _ (by simp) (mul_one _) - mul_left_inv a := SemidirectProduct.ext _ _ (by simp) (by simp) + mul_assoc a b c := SemidirectProduct.ext (by simp [mul_assoc]) (by simp [mul_assoc]) + one_mul a := SemidirectProduct.ext (by simp) (one_mul a.2) + mul_one a := SemidirectProduct.ext (by simp) (mul_one _) + mul_left_inv a := SemidirectProduct.ext (by simp) (by simp) instance : Inhabited (N ⋊[φ] G) := ⟨1⟩ diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index 251ce43267056..578db97ec19ca 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -67,9 +67,6 @@ instance : CoeOut (Sylow p G) (Subgroup G) := @[ext] theorem ext {P Q : Sylow p G} (h : (P : Subgroup G) = Q) : P = Q := by cases P; cases Q; congr -theorem ext_iff {P Q : Sylow p G} : P = Q ↔ (P : Subgroup G) = Q := - ⟨congr_arg _, ext⟩ - instance : SetLike (Sylow p G) G where coe := (↑) coe_injective' _ _ h := ext (SetLike.coe_injective h) diff --git a/Mathlib/Init/Data/Nat/GCD.lean b/Mathlib/Init/Data/Nat/GCD.lean deleted file mode 100644 index 74154704c6ff9..0000000000000 --- a/Mathlib/Init/Data/Nat/GCD.lean +++ /dev/null @@ -1,32 +0,0 @@ -/- -Copyright (c) 2014 Jeremy Avigad. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro - --/ - -import Batteries.Data.Nat.Gcd -import Mathlib.Data.Nat.Notation - -/-! -# Note about `Mathlib/Init/` -The files in `Mathlib/Init` are leftovers from the port from Mathlib3. -(They contain content moved from lean3 itself that Mathlib needed but was not moved to lean4.) - -We intend to move all the content of these files out into the main `Mathlib` directory structure. -Contributions assisting with this are appreciated. - -# Definitions and properties of gcd, lcm, and coprime --/ - - -open WellFounded - -namespace Nat - -/-! gcd -/ - -theorem gcd_def (x y : ℕ) : gcd x y = if x = 0 then y else gcd (y % x) x := by - cases x <;> simp [Nat.gcd_succ] - -end Nat diff --git a/Mathlib/Lean/Name.lean b/Mathlib/Lean/Name.lean index 41efc2a395626..84f0943244ae3 100644 --- a/Mathlib/Lean/Name.lean +++ b/Mathlib/Lean/Name.lean @@ -3,10 +3,9 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Batteries.Data.HashMap.Basic -import Batteries.Lean.SMap import Lean.Meta.Match.MatcherInfo import Lean.Meta.Tactic.Delta +import Std.Data.HashMap.Basic /-! # Additional functions on `Lean.Name`. @@ -38,12 +37,12 @@ def allNames (p : Name → Bool) : CoreM (Array Name) := do Retrieve all names in the environment satisfying a predicate, gathered together into a `HashMap` according to the module they are defined in. -/ -def allNamesByModule (p : Name → Bool) : CoreM (Batteries.HashMap Name (Array Name)) := do - (← getEnv).constants.foldM (init := Batteries.HashMap.empty) fun names n _ => do +def allNamesByModule (p : Name → Bool) : CoreM (Std.HashMap Name (Array Name)) := do + (← getEnv).constants.foldM (init := Std.HashMap.empty) fun names n _ => do if p n && !(← isBlackListed n) then let some m ← findModuleOf? n | return names - -- TODO use `Batteries.HashMap.modify` when we bump Batteries (or `alter` if that is written). - match names.find? m with + -- TODO use `modify` and/or `alter` when available + match names[m]? with | some others => return names.insert m (others.push n) | none => return names.insert m #[n] else diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean index bfc1d0dd2af44..2bc4d04b8a70c 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean @@ -130,9 +130,6 @@ theorem linearMap_vsub (f : P1 →ᵃ[k] P2) (p1 p2 : P1) : f.linear (p1 -ᵥ p2 theorem ext {f g : P1 →ᵃ[k] P2} (h : ∀ p, f p = g p) : f = g := DFunLike.ext _ _ h -theorem ext_iff {f g : P1 →ᵃ[k] P2} : f = g ↔ ∀ p, f p = g p := - ⟨fun h _ => h ▸ rfl, ext⟩ - theorem coeFn_injective : @Function.Injective (P1 →ᵃ[k] P2) (P1 → P2) (⇑) := DFunLike.coe_injective @@ -740,7 +737,7 @@ theorem pi_ext_nonempty [Nonempty ι] (h : ∀ i x, f (Pi.single i x) = g (Pi.si /-- This is used as the ext lemma instead of `AffineMap.pi_ext_nonempty` for reasons explained in note [partially-applied ext lemmas]. Analogous to `LinearMap.pi_ext'`-/ -@[ext] +@[ext (iff := false)] theorem pi_ext_nonempty' [Nonempty ι] (h : ∀ i, f.comp (LinearMap.single i).toAffineMap = g.comp (LinearMap.single i).toAffineMap) : f = g := by refine pi_ext_nonempty fun i x => ?_ diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean index 613c9d30be672..fc963c50a4cdc 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean @@ -305,11 +305,10 @@ theorem vsub_left_mem_direction_iff_mem {s : AffineSubspace k P} {p : P} (hp : p theorem coe_injective : Function.Injective ((↑) : AffineSubspace k P → Set P) := SetLike.coe_injective -@[ext] +@[ext (iff := false)] theorem ext {p q : AffineSubspace k P} (h : ∀ x, x ∈ p ↔ x ∈ q) : p = q := SetLike.ext h --- Porting note: removed `simp`, proof is `simp only [SetLike.ext'_iff]` protected theorem ext_iff (s₁ s₂ : AffineSubspace k P) : s₁ = s₂ ↔ (s₁ : Set P) = s₂ := SetLike.ext'_iff diff --git a/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean index 66086cf98e0e9..ab0bbf619ae27 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean @@ -114,9 +114,6 @@ initialize_simps_projections ContinuousLinearMap (toAffineEquiv_toFun → apply, theorem ext {e e' : P₁ ≃ᵃL[k] P₂} (h : ∀ x, e x = e' x) : e = e' := DFunLike.ext _ _ h -theorem ext_iff {e e' : P₁ ≃ᵃL[k] P₂} : e = e' ↔ ∀ x, e x = e' x := - DFunLike.ext_iff - @[continuity] protected theorem continuous (e : P₁ ≃ᵃL[k] P₂) : Continuous e := e.2 diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index 9f50f7aa5e011..f5ffb8af6f5a6 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -798,8 +798,7 @@ theorem ext {n : ℕ} {s1 s2 : Simplex k P n} (h : ∀ i, s1.points i = s2.point exact h i /-- Two simplices are equal if and only if they have the same points. -/ -theorem ext_iff {n : ℕ} (s1 s2 : Simplex k P n) : s1 = s2 ↔ ∀ i, s1.points i = s2.points i := - ⟨fun h _ => h ▸ rfl, ext⟩ +add_decl_doc Affine.Simplex.ext_iff /-- A face of a simplex is a simplex with the given subset of points. -/ diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index e0a190fc0a344..d838fd0ac45a3 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -125,9 +125,6 @@ theorem coe_inj {f g : M [⋀^ι]→ₗ[R] N} : (f : (ι → M) → N) = g ↔ f theorem ext {f f' : M [⋀^ι]→ₗ[R] N} (H : ∀ x, f x = f' x) : f = f' := DFunLike.ext _ _ H -protected theorem ext_iff {f g : M [⋀^ι]→ₗ[R] N} : f = g ↔ ∀ x, f x = g x := - ⟨fun h _ => h ▸ rfl, fun h => ext h⟩ - attribute [coe] AlternatingMap.toMultilinearMap instance coe : Coe (M [⋀^ι]→ₗ[R] N) (MultilinearMap R (fun _ : ι => M) N) := diff --git a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean index 230bc4bee75e8..0768933dedc46 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean @@ -101,8 +101,6 @@ 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 _ _ -theorem ext_iff : B = D ↔ ∀ x y, B x y = D x y := ext_iff₂ - @[deprecated (since := "2024-04-14")] theorem coe_zero : ⇑(0 : BilinForm R M) = 0 := rfl diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Even.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Even.lean index 6e88c7312e40e..ab5a28a5dcedc 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Even.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Even.lean @@ -245,8 +245,8 @@ multiplication. -/ def even.lift : EvenHom Q A ≃ (CliffordAlgebra.even Q →ₐ[R] A) where toFun f := AlgHom.ofLinearMap (aux f) (aux_one f) (aux_mul f) invFun F := (even.ι Q).compr₂ F - left_inv f := EvenHom.ext _ _ <| LinearMap.ext₂ <| even.lift.aux_ι f - right_inv _ := even.algHom_ext Q <| EvenHom.ext _ _ <| LinearMap.ext₂ <| even.lift.aux_ι _ + left_inv f := EvenHom.ext <| LinearMap.ext₂ <| even.lift.aux_ι f + right_inv _ := even.algHom_ext Q <| EvenHom.ext <| LinearMap.ext₂ <| even.lift.aux_ι _ -- @[simp] -- Porting note: simpNF linter times out on this one theorem even.lift_ι (f : EvenHom Q A) (m₁ m₂ : M) : diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean index 5dd04e4ade682..ff16afc3d7f70 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/EvenEquiv.lean @@ -162,7 +162,7 @@ theorem ofEven_ι (x y : M × R) : theorem toEven_comp_ofEven : (toEven Q).comp (ofEven Q) = AlgHom.id R _ := even.algHom_ext (Q' Q) <| - EvenHom.ext _ _ <| + EvenHom.ext <| LinearMap.ext fun m₁ => LinearMap.ext fun m₂ => Subtype.ext <| diff --git a/Mathlib/LinearAlgebra/LinearPMap.lean b/Mathlib/LinearAlgebra/LinearPMap.lean index 8c4e8333b4463..1438fc123e93f 100644 --- a/Mathlib/LinearAlgebra/LinearPMap.lean +++ b/Mathlib/LinearAlgebra/LinearPMap.lean @@ -57,7 +57,7 @@ instance : CoeFun (E →ₗ.[R] F) fun f : E →ₗ.[R] F => f.domain → F := theorem toFun_eq_coe (f : E →ₗ.[R] F) (x : f.domain) : f.toFun x = f x := rfl -@[ext] +@[ext (iff := false)] theorem ext {f g : E →ₗ.[R] F} (h : f.domain = g.domain) (h' : ∀ ⦃x : f.domain⦄ ⦃y : g.domain⦄ (_h : (x : E) = y), f x = g y) : f = g := by rcases f with ⟨f_dom, f⟩ diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 0dffd74338f4a..0037b0403d11e 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -138,9 +138,6 @@ theorem coe_inj {f g : MultilinearMap R M₁ M₂} : (f : (∀ i, M₁ i) → M theorem ext {f f' : MultilinearMap R M₁ M₂} (H : ∀ x, f x = f' x) : f = f' := DFunLike.ext _ _ H -protected theorem ext_iff {f g : MultilinearMap R M₁ M₂} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - @[simp] theorem mk_coe (f : MultilinearMap R M₁ M₂) (h₁ h₂) : (⟨f, h₁, h₂⟩ : MultilinearMap R M₁ M₂) = f := rfl diff --git a/Mathlib/LinearAlgebra/Pi.lean b/Mathlib/LinearAlgebra/Pi.lean index f01f5eeed6c48..59907ee20f3e1 100644 --- a/Mathlib/LinearAlgebra/Pi.lean +++ b/Mathlib/LinearAlgebra/Pi.lean @@ -172,9 +172,6 @@ theorem pi_ext' (h : ∀ i, f.comp (single i) = g.comp (single i)) : f = g := by refine pi_ext fun i x => ?_ convert LinearMap.congr_fun (h i) x -theorem pi_ext'_iff : f = g ↔ ∀ i, f.comp (single i) = g.comp (single i) := - ⟨fun h _ => h ▸ rfl, pi_ext'⟩ - end Ext section diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 45d6fc2fec4a7..8686fa1fe4bb1 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -188,9 +188,6 @@ theorem ext (H : ∀ x : M, Q x = Q' x) : Q = Q' := theorem congr_fun (h : Q = Q') (x : M) : Q x = Q' x := DFunLike.congr_fun h _ -theorem ext_iff : Q = Q' ↔ ∀ x, Q x = Q' x := - DFunLike.ext_iff - /-- Copy of a `QuadraticMap` with a new `toFun` equal to the old one. Useful to fix definitional equalities. -/ protected def copy (Q : QuadraticMap R M N) (Q' : M → N) (h : Q' = ⇑Q) : QuadraticMap R M N where diff --git a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean index 03f3abd4b9267..a86dc75313b8c 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean @@ -56,15 +56,15 @@ instance category : Category (QuadraticModuleCat.{v} R) where Hom M N := Hom M N id M := ⟨Isometry.id M.form⟩ comp f g := ⟨Isometry.comp g.toIsometry f.toIsometry⟩ - id_comp g := Hom.ext _ _ <| Isometry.id_comp g.toIsometry - comp_id f := Hom.ext _ _ <| Isometry.comp_id f.toIsometry - assoc f g h := Hom.ext _ _ <| Isometry.comp_assoc h.toIsometry g.toIsometry f.toIsometry + id_comp g := Hom.ext <| Isometry.id_comp g.toIsometry + comp_id f := Hom.ext <| Isometry.comp_id f.toIsometry + assoc f g h := Hom.ext <| Isometry.comp_assoc h.toIsometry g.toIsometry f.toIsometry -- TODO: if `Quiver.Hom` and the instance above were `reducible`, this wouldn't be needed. @[ext] lemma hom_ext {M N : QuadraticModuleCat.{v} R} (f g : M ⟶ N) (h : f.toIsometry = g.toIsometry) : f = g := - Hom.ext _ _ h + Hom.ext h /-- Typecheck a `QuadraticForm.Isometry` as a morphism in `Module R`. -/ abbrev ofHom {X : Type v} [AddCommGroup X] [Module R X] @@ -112,8 +112,8 @@ variable {Q₁ : QuadraticForm R X} {Q₂ : QuadraticForm R Y} {Q₃ : Quadratic def ofIso (e : Q₁.IsometryEquiv Q₂) : QuadraticModuleCat.of Q₁ ≅ QuadraticModuleCat.of Q₂ where hom := ⟨e.toIsometry⟩ inv := ⟨e.symm.toIsometry⟩ - hom_inv_id := Hom.ext _ _ <| DFunLike.ext _ _ e.left_inv - inv_hom_id := Hom.ext _ _ <| DFunLike.ext _ _ e.right_inv + hom_inv_id := Hom.ext <| DFunLike.ext _ _ e.left_inv + inv_hom_id := Hom.ext <| DFunLike.ext _ _ e.right_inv @[simp] theorem ofIso_refl : ofIso (IsometryEquiv.refl Q₁) = .refl _ := rfl diff --git a/Mathlib/LinearAlgebra/TensorPower.lean b/Mathlib/LinearAlgebra/TensorPower.lean index 9cab25a722e01..0e0eed359fc0c 100644 --- a/Mathlib/LinearAlgebra/TensorPower.lean +++ b/Mathlib/LinearAlgebra/TensorPower.lean @@ -45,7 +45,7 @@ namespace PiTensorProduct /-- Two dependent pairs of tensor products are equal if their index is equal and the contents are equal after a canonical reindexing. -/ -@[ext] +@[ext (iff := false)] theorem gradedMonoid_eq_of_reindex_cast {ιι : Type*} {ι : ιι → Type*} : ∀ {a b : GradedMonoid fun ii => ⨂[R] _ : ι ii, M} (h : a.fst = b.fst), reindex R (fun _ ↦ M) (Equiv.cast <| congr_arg ι h) a.snd = b.snd → a = b @@ -118,7 +118,7 @@ theorem cast_cast {i j k} (h : i = j) (h' : j = k) (a : ⨂[R]^i M) : cast R M h' (cast R M h a) = cast R M (h.trans h') a := reindex_reindex _ _ _ -@[ext] +@[ext (iff := false)] theorem gradedMonoid_eq_of_cast {a b : GradedMonoid fun n => ⨂[R] _ : Fin n, M} (h : a.fst = b.fst) (h2 : cast R M h a.snd = b.snd) : a = b := by refine gradedMonoid_eq_of_reindex_cast h ?_ diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 9b8f4c65a0df5..1d1460adf46e1 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -577,22 +577,12 @@ theorem exists_apply_eq (a : α) (b : β) : ∃ f : α → β, f a = b := ⟨fun ⟨fun ⟨_, ⟨a, b, hab⟩, hc⟩ ↦ ⟨a, b, hab.symm ▸ hc⟩, fun ⟨a, b, hab⟩ ↦ ⟨f a b, ⟨a, b, rfl⟩, hab⟩⟩ -@[simp] theorem exists_or_eq_left (y : α) (p : α → Prop) : ∃ x : α, x = y ∨ p x := ⟨y, .inl rfl⟩ - -@[simp] theorem exists_or_eq_right (y : α) (p : α → Prop) : ∃ x : α, p x ∨ x = y := ⟨y, .inr rfl⟩ - -@[simp] theorem exists_or_eq_left' (y : α) (p : α → Prop) : ∃ x : α, y = x ∨ p x := ⟨y, .inl rfl⟩ - -@[simp] theorem exists_or_eq_right' (y : α) (p : α → Prop) : ∃ x : α, p x ∨ y = x := ⟨y, .inr rfl⟩ - theorem forall_apply_eq_imp_iff' {f : α → β} {p : β → Prop} : (∀ a b, f a = b → p b) ↔ ∀ a, p (f a) := by simp theorem forall_eq_apply_imp_iff' {f : α → β} {p : β → Prop} : (∀ a b, b = f a → p b) ↔ ∀ a, p (f a) := by simp -@[simp] theorem exists_eq_right' {a' : α} : (∃ a, p a ∧ a' = a) ↔ p a' := by simp [@eq_comm _ a'] - theorem exists₂_comm {ι₁ ι₂ : Sort*} {κ₁ : ι₁ → Sort*} {κ₂ : ι₂ → Sort*} {p : ∀ i₁, κ₁ i₁ → ∀ i₂, κ₂ i₂ → Prop} : (∃ i₁ j₁ i₂ j₂, p i₁ j₁ i₂ j₂) ↔ ∃ i₂ j₂ i₁ j₁, p i₁ j₁ i₂ j₂ := by @@ -638,9 +628,6 @@ theorem Prop.exists_iff {p : Prop → Prop} : (∃ h, p h) ↔ p False ∨ p Tru theorem Prop.forall_iff {p : Prop → Prop} : (∀ h, p h) ↔ p False ∧ p True := ⟨fun H ↦ ⟨H _, H _⟩, fun ⟨h₁, h₂⟩ h ↦ by by_cases H : h <;> simpa only [H]⟩ -theorem exists_prop_of_true {p : Prop} {q : p → Prop} (h : p) : (∃ h' : p, q h') ↔ q h := - @exists_const (q h) p ⟨h⟩ - theorem exists_iff_of_forall {p : Prop} {q : p → Prop} (h : ∀ h, q h) : (∃ h, q h) ↔ p := ⟨Exists.fst, fun H ↦ ⟨H, h H⟩⟩ @@ -650,14 +637,7 @@ theorem exists_unique_prop_of_true {p : Prop} {q : p → Prop} (h : p) : (∃! h theorem exists_prop_of_false {p : Prop} {q : p → Prop} : ¬p → ¬∃ h' : p, q h' := mt Exists.fst -@[congr] -theorem exists_prop_congr {p p' : Prop} {q q' : p → Prop} (hq : ∀ h, q h ↔ q' h) (hp : p ↔ p') : - Exists q ↔ ∃ h : p', q' (hp.2 h) := - ⟨fun ⟨_, _⟩ ↦ ⟨hp.1 ‹_›, (hq _).1 ‹_›⟩, fun ⟨_, _⟩ ↦ ⟨_, (hq _).2 ‹_›⟩⟩ - -/-- See `IsEmpty.exists_iff` for the `False` version. -/ -@[simp] theorem exists_true_left (p : True → Prop) : (∃ x, p x) ↔ p True.intro := - exists_prop_of_true _ +/- See `IsEmpty.exists_iff` for the `False` version of `exists_true_left`. -/ -- Porting note: `@[congr]` commented out for now. -- @[congr] diff --git a/Mathlib/Logic/Embedding/Basic.lean b/Mathlib/Logic/Embedding/Basic.lean index 34e5615211fd9..0f91bfa54fc08 100644 --- a/Mathlib/Logic/Embedding/Basic.lean +++ b/Mathlib/Logic/Embedding/Basic.lean @@ -105,10 +105,6 @@ instance {α β : Sort*} [IsEmpty α] : Unique (α ↪ β) where default := ⟨isEmptyElim, Function.injective_of_subsingleton _⟩ uniq := by intro; ext v; exact isEmptyElim v --- Porting note : in Lean 3 `DFunLike.ext_iff.symm` works -theorem ext_iff {α β} {f g : Embedding α β} : (∀ x, f x = g x) ↔ f = g := - Iff.symm (DFunLike.ext_iff) - @[simp] theorem toFun_eq_coe {α β} (f : α ↪ β) : toFun f = f := rfl diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 7609a278df16b..d87713b0db049 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1341,7 +1341,7 @@ def subtypeQuotientEquivQuotientSubtype (p₁ : α → Prop) {s₁ : Setoid α} Quotient.liftOn a (fun a => (⟨⟦a.1⟧, (hp₂ _).1 a.2⟩ : { x // p₂ x })) fun a b hab => Subtype.ext_val (Quotient.sound ((h _ _).1 hab)) left_inv := by exact fun ⟨a, ha⟩ => Quotient.inductionOn a (fun b hb => rfl) ha - right_inv a := Quotient.inductionOn a fun ⟨a, ha⟩ => rfl + right_inv a := by exact Quotient.inductionOn a fun ⟨a, ha⟩ => rfl @[simp] theorem subtypeQuotientEquivQuotientSubtype_mk (p₁ : α → Prop) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 8ae3f8c50283b..b42323a8b7ae0 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -124,8 +124,6 @@ protected theorem congr_arg {f : Equiv α β} {x x' : α} : x = x' → f x = f x protected theorem congr_fun {f g : Equiv α β} (h : f = g) (x : α) : f x = g x := DFunLike.congr_fun h x -theorem ext_iff {f g : Equiv α β} : f = g ↔ ∀ x, f x = g x := DFunLike.ext_iff - @[ext] theorem Perm.ext {σ τ : Equiv.Perm α} (H : ∀ x, σ x = τ x) : σ = τ := Equiv.ext H protected theorem Perm.congr_arg {f : Equiv.Perm α} {x x' : α} : x = x' → f x = f x' := @@ -134,8 +132,6 @@ protected theorem Perm.congr_arg {f : Equiv.Perm α} {x x' : α} : x = x' → f protected theorem Perm.congr_fun {f g : Equiv.Perm α} (h : f = g) (x : α) : f x = g x := Equiv.congr_fun h x -protected theorem Perm.ext_iff {σ τ : Equiv.Perm α} : σ = τ ↔ ∀ x, σ x = τ x := Equiv.ext_iff - /-- Any type is equivalent to itself. -/ @[refl] protected def refl (α : Sort*) : α ≃ α := ⟨id, id, fun _ => rfl, fun _ => rfl⟩ diff --git a/Mathlib/Logic/Nontrivial/Basic.lean b/Mathlib/Logic/Nontrivial/Basic.lean index 0c1f612c668cf..fb49821633797 100644 --- a/Mathlib/Logic/Nontrivial/Basic.lean +++ b/Mathlib/Logic/Nontrivial/Basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.Data.Prod.Basic -import Mathlib.Data.Subtype import Mathlib.Logic.Function.Basic import Mathlib.Logic.Nontrivial.Defs import Mathlib.Logic.Unique diff --git a/Mathlib/Logic/Relation.lean b/Mathlib/Logic/Relation.lean index f9a7e8a977eb5..9f30f820360f2 100644 --- a/Mathlib/Logic/Relation.lean +++ b/Mathlib/Logic/Relation.lean @@ -229,11 +229,8 @@ inductive ReflGen (r : α → α → Prop) (a : α) : α → Prop | refl : ReflGen r a a | single {b} : r a b → ReflGen r a b -/-- `TransGen r`: transitive closure of `r` -/ -@[mk_iff] -inductive TransGen (r : α → α → Prop) (a : α) : α → Prop - | single {b} : r a b → TransGen r a b - | tail {b c} : TransGen r a b → r b c → TransGen r a c +attribute [mk_iff] TransGen + attribute [refl] ReflGen.refl @@ -401,17 +398,6 @@ theorem head'_iff : TransGen r a c ↔ ∃ b, r a b ∧ ReflTransGen r b c := by end TransGen -theorem _root_.Acc.TransGen (h : Acc r a) : Acc (TransGen r) a := by - induction' h with x _ H - refine Acc.intro x fun y hy ↦ ?_ - cases' hy with _ hyx z _ hyz hzx - exacts [H y hyx, (H z hzx).inv hyz] - -theorem _root_.acc_transGen_iff : Acc (TransGen r) a ↔ Acc r a := - ⟨Subrelation.accessible TransGen.single, Acc.TransGen⟩ - -theorem _root_.WellFounded.transGen (h : WellFounded r) : WellFounded (TransGen r) := - ⟨fun a ↦ (h.apply a).TransGen⟩ section reflGen diff --git a/Mathlib/Mathport/Syntax.lean b/Mathlib/Mathport/Syntax.lean index 9acddedd901a7..247469d916e96 100644 --- a/Mathlib/Mathport/Syntax.lean +++ b/Mathlib/Mathport/Syntax.lean @@ -114,7 +114,6 @@ estimated difficulty of writing the tactic. The key is as follows: namespace Mathlib.Tactic open Lean Parser.Tactic -/- N -/ elab (name := include) "include" (ppSpace ident)+ : command => pure () /- N -/ elab (name := omit) "omit" (ppSpace ident)+ : command => pure () /- N -/ syntax (name := parameter) "parameter" (ppSpace bracketedBinder)+ : command diff --git a/Mathlib/MeasureTheory/Decomposition/Jordan.lean b/Mathlib/MeasureTheory/Decomposition/Jordan.lean index 64dc6877bf0da..d7d28a7b8bd02 100644 --- a/Mathlib/MeasureTheory/Decomposition/Jordan.lean +++ b/Mathlib/MeasureTheory/Decomposition/Jordan.lean @@ -73,7 +73,7 @@ instance instInhabited : Inhabited (JordanDecomposition α) where default := 0 instance instInvolutiveNeg : InvolutiveNeg (JordanDecomposition α) where neg j := ⟨j.negPart, j.posPart, j.mutuallySingular.symm⟩ - neg_neg _ := JordanDecomposition.ext _ _ rfl rfl + neg_neg _ := JordanDecomposition.ext rfl rfl instance instSMul : SMul ℝ≥0 (JordanDecomposition α) where smul r j := diff --git a/Mathlib/MeasureTheory/Function/AEEqFun.lean b/Mathlib/MeasureTheory/Function/AEEqFun.lean index a38ca7eef51a0..78474106e00ff 100644 --- a/Mathlib/MeasureTheory/Function/AEEqFun.lean +++ b/Mathlib/MeasureTheory/Function/AEEqFun.lean @@ -154,9 +154,6 @@ theorem mk_coeFn (f : α →ₘ[μ] β) : mk f f.aestronglyMeasurable = f := by theorem ext {f g : α →ₘ[μ] β} (h : f =ᵐ[μ] g) : f = g := by rwa [← f.mk_coeFn, ← g.mk_coeFn, mk_eq_mk] -theorem ext_iff {f g : α →ₘ[μ] β} : f = g ↔ f =ᵐ[μ] g := - ⟨fun h => by rw [h], fun h => ext h⟩ - theorem coeFn_mk (f : α → β) (hf) : (mk f hf : α →ₘ[μ] β) =ᵐ[μ] f := by apply (AEStronglyMeasurable.ae_eq_mk _).symm.trans exact @Quotient.mk_out' _ (μ.aeEqSetoid β) (⟨f, hf⟩ : { f // AEStronglyMeasurable f μ }) diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean index e8ae43ee6b393..2f9b8035e0291 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean @@ -580,9 +580,7 @@ theorem Lp.induction_stronglyMeasurable (hm : m ≤ m0) (hp_ne_top : p ≠ ∞) ∀ ⦃f g⦄, ∀ hf : Memℒp f p μ, ∀ hg : Memℒp g p μ, AEStronglyMeasurable' m f μ → AEStronglyMeasurable' m g μ → Disjoint (Function.support f) (Function.support g) → P (hf.toLp f) → P (hg.toLp g) → P (hf.toLp f + hg.toLp g) from - -- Porting note: `P` should be an explicit argument to `Lp.induction_stronglyMeasurable_aux`, but - -- it isn't? - Lp.induction_stronglyMeasurable_aux hm hp_ne_top h_ind h_add_ae h_closed f hf + Lp.induction_stronglyMeasurable_aux hm hp_ne_top _ h_ind h_add_ae h_closed f hf intro f g hf hg hfm hgm h_disj hPf hPg let s_f : Set α := Function.support (hfm.mk f) have hs_f : MeasurableSet[m] s_f := hfm.stronglyMeasurable_mk.measurableSet_support @@ -645,9 +643,7 @@ theorem Memℒp.induction_stronglyMeasurable (hm : m ≤ m0) (hp_ne_top : p ≠ have hfm_Lp : AEStronglyMeasurable' m f_Lp μ := hfm.congr hf.coeFn_toLp.symm refine h_ae hf.coeFn_toLp (Lp.memℒp _) ?_ change P f_Lp - -- Porting note: `P` should be an explicit argument to `Lp.induction_stronglyMeasurable`, but - -- it isn't? - refine Lp.induction_stronglyMeasurable hm hp_ne_top (P := fun f => P f) ?_ ?_ h_closed f_Lp hfm_Lp + refine Lp.induction_stronglyMeasurable hm hp_ne_top (fun f => P f) ?_ ?_ h_closed f_Lp hfm_Lp · intro c s hs hμs rw [Lp.simpleFunc.coe_indicatorConst] refine h_ae indicatorConstLp_coeFn.symm ?_ (h_ind c hs hμs) diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index 459316708b3a6..71aec0c9f8971 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -335,8 +335,9 @@ theorem addHaar_image_le_mul_of_det_lt (A : E →L[ℝ] E) {m : ℝ≥0} A '' closedBall 0 r + closedBall (f x) (ε * r) = {f x} + r • (A '' closedBall 0 1 + closedBall 0 ε) := by rw [smul_add, ← add_assoc, add_comm {f x}, add_assoc, smul_closedBall _ _ εpos.le, smul_zero, - singleton_add_closedBall_zero, ← image_smul_set ℝ E E A, smul_closedBall _ _ zero_le_one, - smul_zero, Real.norm_eq_abs, abs_of_nonneg r0, mul_one, mul_comm] + singleton_add_closedBall_zero, ← image_smul_set ℝ E E A, + _root_.smul_closedBall _ _ zero_le_one, smul_zero, Real.norm_eq_abs, abs_of_nonneg r0, + mul_one, mul_comm] rw [this] at K calc μ (f '' (s ∩ closedBall x r)) ≤ μ ({f x} + r • (A '' closedBall 0 1 + closedBall 0 ε)) := @@ -796,7 +797,7 @@ theorem addHaar_image_le_lintegral_abs_det_fderiv_aux1 (hs : MeasurableSet s) simp only [m, ENNReal.ofReal, lt_add_iff_pos_right, εpos, ENNReal.coe_lt_coe] rcases ((addHaar_image_le_mul_of_det_lt μ A I).and self_mem_nhdsWithin).exists with ⟨δ, h, δpos⟩ obtain ⟨δ', δ'pos, hδ'⟩ : ∃ (δ' : ℝ), 0 < δ' ∧ ∀ B, dist B A < δ' → dist B.det A.det < ↑ε := - continuousAt_iff.1 ContinuousLinearMap.continuous_det.continuousAt ε εpos + continuousAt_iff.1 (ContinuousLinearMap.continuous_det (E := E)).continuousAt ε εpos let δ'' : ℝ≥0 := ⟨δ' / 2, (half_pos δ'pos).le⟩ refine ⟨min δ δ'', lt_min δpos (half_pos δ'pos), ?_, ?_⟩ · intro B hB @@ -918,7 +919,7 @@ theorem lintegral_abs_det_fderiv_le_addHaar_image_aux1 (hs : MeasurableSet s) ENNReal.ofReal |A.det| * μ t ≤ μ (g '' t) + ε * μ t := by intro A obtain ⟨δ', δ'pos, hδ'⟩ : ∃ (δ' : ℝ), 0 < δ' ∧ ∀ B, dist B A < δ' → dist B.det A.det < ↑ε := - continuousAt_iff.1 ContinuousLinearMap.continuous_det.continuousAt ε εpos + continuousAt_iff.1 (ContinuousLinearMap.continuous_det (E := E)).continuousAt ε εpos let δ'' : ℝ≥0 := ⟨δ' / 2, (half_pos δ'pos).le⟩ have I'' : ∀ B : E →L[ℝ] E, ‖B - A‖ ≤ ↑δ'' → |B.det - A.det| ≤ ↑ε := by intro B hB diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index de8bc05c92f74..8ff7405363c51 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -161,9 +161,6 @@ theorem ext {f g : Lp E p μ} (h : f =ᵐ[μ] g) : f = g := by simp only [Subtype.mk_eq_mk] exact AEEqFun.ext h -theorem ext_iff {f g : Lp E p μ} : f = g ↔ f =ᵐ[μ] g := - ⟨fun h => by rw [h], fun h => ext h⟩ - theorem mem_Lp_iff_eLpNorm_lt_top {f : α →ₘ[μ] E} : f ∈ Lp E p μ ↔ eLpNorm f p μ < ∞ := Iff.rfl @[deprecated (since := "2024-07-27")] diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean b/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean index f60cb8c5a745e..24be3674a76da 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Defs.lean @@ -227,10 +227,6 @@ theorem MeasurableSpace.ext {m₁ m₂ : MeasurableSpace α} (h : ∀ s : Set α, MeasurableSet[m₁] s ↔ MeasurableSet[m₂] s) : m₁ = m₂ := measurableSet_injective <| funext fun s => propext (h s) -theorem MeasurableSpace.ext_iff {m₁ m₂ : MeasurableSpace α} : - m₁ = m₂ ↔ ∀ s : Set α, MeasurableSet[m₁] s ↔ MeasurableSet[m₂] s := - ⟨fun h _ => h ▸ Iff.rfl, MeasurableSpace.ext⟩ - /-- A typeclass mixin for `MeasurableSpace`s such that each singleton is measurable. -/ class MeasurableSingletonClass (α : Type*) [MeasurableSpace α] : Prop where /-- A singleton is a measurable set. -/ diff --git a/Mathlib/MeasureTheory/Measure/AddContent.lean b/Mathlib/MeasureTheory/Measure/AddContent.lean index dc7e2d6d48694..362cd5b634339 100644 --- a/Mathlib/MeasureTheory/Measure/AddContent.lean +++ b/Mathlib/MeasureTheory/Measure/AddContent.lean @@ -75,9 +75,6 @@ variable {m m' : AddContent C} @[ext] protected lemma AddContent.ext (h : ∀ s, m s = m' s) : m = m' := DFunLike.ext _ _ h -protected lemma AddContent.ext_iff (m m' : AddContent C) : m = m' ↔ ∀ s, m s = m' s := - DFunLike.ext_iff - @[simp] lemma addContent_empty : m ∅ = 0 := m.empty' lemma addContent_sUnion (h_ss : ↑I ⊆ C) diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean index bdfd9e9925b99..61c419b87e67f 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean @@ -126,9 +126,6 @@ theorem ext (h : ∀ s, MeasurableSet s → μ₁ s = μ₂ s) : μ₁ = μ₂ : toOuterMeasure_injective <| by rw [← trimmed, OuterMeasure.trim_congr (h _), trimmed] -theorem ext_iff : μ₁ = μ₂ ↔ ∀ s, MeasurableSet s → μ₁ s = μ₂ s := - ⟨by rintro rfl s _hs; rfl, Measure.ext⟩ - theorem ext_iff' : μ₁ = μ₂ ↔ ∀ s, μ₁ s = μ₂ s := ⟨by rintro rfl s; rfl, fun h ↦ Measure.ext (fun s _ ↦ h s)⟩ diff --git a/Mathlib/ModelTheory/Basic.lean b/Mathlib/ModelTheory/Basic.lean index 20bd4499e732e..58d599986ed9e 100644 --- a/Mathlib/ModelTheory/Basic.lean +++ b/Mathlib/ModelTheory/Basic.lean @@ -449,9 +449,6 @@ theorem toFun_eq_coe {f : M →[L] N} : f.toFun = (f : M → N) := theorem ext ⦃f g : M →[L] N⦄ (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h -theorem ext_iff {f g : M →[L] N} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - @[simp] theorem map_fun (φ : M →[L] N) {n : ℕ} (f : L.Functions n) (x : Fin n → M) : φ (funMap f x) = funMap f (φ ∘ x) := @@ -567,9 +564,6 @@ theorem coe_injective : @Function.Injective (M ↪[L] N) (M → N) (↑) theorem ext ⦃f g : M ↪[L] N⦄ (h : ∀ x, f x = g x) : f = g := coe_injective (funext h) -theorem ext_iff {f g : M ↪[L] N} : f = g ↔ ∀ x, f x = g x := - ⟨fun h _ => h ▸ rfl, fun h => ext h⟩ - theorem toHom_injective : @Function.Injective (M ↪[L] N) (M →[L] N) (·.toHom) := by intro f f' h ext @@ -765,9 +759,6 @@ theorem coe_injective : @Function.Injective (M ≃[L] N) (M → N) (↑) := theorem ext ⦃f g : M ≃[L] N⦄ (h : ∀ x, f x = g x) : f = g := coe_injective (funext h) -theorem ext_iff {f g : M ≃[L] N} : f = g ↔ ∀ x, f x = g x := - ⟨fun h _ => h ▸ rfl, fun h => ext h⟩ - theorem bijective (f : M ≃[L] N) : Function.Bijective f := EquivLike.bijective f diff --git a/Mathlib/ModelTheory/ElementaryMaps.lean b/Mathlib/ModelTheory/ElementaryMaps.lean index f9e0c790bdce4..9c25fd43ecb9e 100644 --- a/Mathlib/ModelTheory/ElementaryMaps.lean +++ b/Mathlib/ModelTheory/ElementaryMaps.lean @@ -169,9 +169,6 @@ theorem coe_injective : @Function.Injective (M ↪ₑ[L] N) (M → N) (↑) := theorem ext ⦃f g : M ↪ₑ[L] N⦄ (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h -theorem ext_iff {f g : M ↪ₑ[L] N} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - variable (L) (M) /-- The identity elementary embedding from a structure to itself -/ diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index 21a23619eb838..4c7fe629d2236 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -182,6 +182,14 @@ 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} : sigmaImp ⟨n, φ⟩ ⟨n, ψ⟩ = ⟨n, φ.imp ψ⟩ := by diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 2880121b3899a..bb4265f91c63c 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -118,9 +118,6 @@ theorem zero_apply {x : ℕ} : (0 : ArithmeticFunction R) x = 0 := theorem ext ⦃f g : ArithmeticFunction R⦄ (h : ∀ x, f x = g x) : f = g := ZeroHom.ext h -theorem ext_iff {f g : ArithmeticFunction R} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - section One variable [One R] diff --git a/Mathlib/NumberTheory/ModularForms/Basic.lean b/Mathlib/NumberTheory/ModularForms/Basic.lean index 0e9b09a1bb433..77d18bd72f999 100644 --- a/Mathlib/NumberTheory/ModularForms/Basic.lean +++ b/Mathlib/NumberTheory/ModularForms/Basic.lean @@ -392,7 +392,7 @@ def mcast {a b : ℤ} {Γ : Subgroup SL(2, ℤ)} (h : a = b) (f : ModularForm Γ holo' := f.holo' bdd_at_infty' A := h ▸ f.bdd_at_infty' A -@[ext] +@[ext (iff := false)] theorem gradedMonoid_eq_of_cast {Γ : Subgroup SL(2, ℤ)} {a b : GradedMonoid (ModularForm Γ)} (h : a.fst = b.fst) (h2 : mcast h a.snd = b.snd) : a = b := by obtain ⟨i, a⟩ := a diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index 9113b004f434c..0aad41a44a431 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -134,11 +134,6 @@ theorem ext {χ χ' : MulChar R R'} (h : ∀ a : Rˣ, χ a = χ' a) : χ = χ' : · exact h ha.unit · rw [map_nonunit χ ha, map_nonunit χ' ha] -theorem ext_iff {χ χ' : MulChar R R'} : χ = χ' ↔ ∀ a : Rˣ, χ a = χ' a := - ⟨by - rintro rfl a - rfl, ext⟩ - /-! ### Equivalence of multiplicative characters with homomorphisms on units diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index 8679043eb2c47..c917742ef44af 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -106,9 +106,6 @@ lemma coe_eq_algebraMap (x : 𝓞 K) : (x : K) = algebraMap _ _ x := rfl @[ext] theorem ext {x y : 𝓞 K} (h : (x : K) = (y : K)) : x = y := Subtype.ext h -theorem ext_iff {x y : 𝓞 K} : x = y ↔ (x : K) = (y : K) := - Subtype.ext_iff - @[norm_cast] theorem eq_iff {x y : 𝓞 K} : (x : K) = (y : K) ↔ x = y := NumberField.RingOfIntegers.ext_iff.symm diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index 3ecf88b979d6f..80b91e674b707 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -128,7 +128,7 @@ theorem prop_y (a : Solution₁ d) : d * a.y ^ 2 = a.x ^ 2 - 1 := by rw [← a.p /-- Two solutions are equal if their `x` and `y` components are equal. -/ @[ext] theorem ext {a b : Solution₁ d} (hx : a.x = b.x) (hy : a.y = b.y) : a = b := - Subtype.ext <| Zsqrtd.ext _ _ hx hy + Subtype.ext <| Zsqrtd.ext hx hy /-- Construct a solution from `x`, `y` and a proof that the equation is satisfied. -/ def mk (x y : ℤ) (prop : x ^ 2 - d * y ^ 2 = 1) : Solution₁ d where @@ -145,7 +145,7 @@ theorem y_mk (x y : ℤ) (prop : x ^ 2 - d * y ^ 2 = 1) : (mk x y prop).y = y := @[simp] theorem coe_mk (x y : ℤ) (prop : x ^ 2 - d * y ^ 2 = 1) : (↑(mk x y prop) : ℤ√d) = ⟨x, y⟩ := - Zsqrtd.ext _ _ (x_mk x y prop) (y_mk x y prop) + Zsqrtd.ext (x_mk x y prop) (y_mk x y prop) @[simp] theorem x_one : (1 : Solution₁ d).x = 1 := diff --git a/Mathlib/NumberTheory/Zsqrtd/Basic.lean b/Mathlib/NumberTheory/Zsqrtd/Basic.lean index 22c2543553c68..8afaf8159fc1c 100644 --- a/Mathlib/NumberTheory/Zsqrtd/Basic.lean +++ b/Mathlib/NumberTheory/Zsqrtd/Basic.lean @@ -217,13 +217,13 @@ theorem star_im (z : ℤ√d) : (star z).im = -z.im := rfl instance : StarRing (ℤ√d) where - star_involutive x := Zsqrtd.ext _ _ rfl (neg_neg _) + star_involutive x := Zsqrtd.ext rfl (neg_neg _) star_mul a b := by ext <;> simp <;> ring - star_add a b := Zsqrtd.ext _ _ rfl (neg_add _ _) + star_add a b := Zsqrtd.ext rfl (neg_add _ _) -- Porting note: proof was `by decide` instance nontrivial : Nontrivial (ℤ√d) := - ⟨⟨0, 1, (Zsqrtd.ext_iff 0 1).not.mpr (by simp)⟩⟩ + ⟨⟨0, 1, Zsqrtd.ext_iff.not.mpr (by simp)⟩⟩ @[simp] theorem natCast_re (n : ℕ) : (n : ℤ√d).re = n := @@ -330,7 +330,7 @@ theorem coprime_of_dvd_coprime {a b : ℤ√d} (hcoprime : IsCoprime a.re a.im) IsCoprime b.re b.im := by apply isCoprime_of_dvd · rintro ⟨hre, him⟩ - obtain rfl : b = 0 := Zsqrtd.ext b 0 hre him + obtain rfl : b = 0 := Zsqrtd.ext hre him rw [zero_dvd_iff] at hdvd simp [hdvd, zero_im, zero_re, not_isCoprime_zero_zero] at hcoprime · rintro z hz - hzdvdu hzdvdv @@ -883,7 +883,7 @@ instance : OrderedRing (ℤ√d) := by infer_instance end theorem norm_eq_zero {d : ℤ} (h_nonsquare : ∀ n : ℤ, d ≠ n * n) (a : ℤ√d) : norm a = 0 ↔ a = 0 := by - refine ⟨fun ha => (Zsqrtd.ext_iff _ _).mpr ?_, fun h => by rw [h, norm_zero]⟩ + refine ⟨fun ha => Zsqrtd.ext_iff.mpr ?_, fun h => by rw [h, norm_zero]⟩ dsimp only [norm] at ha rw [sub_eq_zero] at ha by_cases h : 0 ≤ d diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 1074443406cc3..a4e53468a767d 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -124,9 +124,6 @@ theorem filter_eq_iff : f = g ↔ f.sets = g.sets := protected theorem ext (h : ∀ s, s ∈ f ↔ s ∈ g) : f = g := by simpa [filter_eq_iff, Set.ext_iff, Filter.mem_sets] -protected theorem ext_iff : f = g ↔ ∀ s, s ∈ f ↔ s ∈ g := - ⟨by rintro rfl s; rfl, Filter.ext⟩ - /-- An extensionality lemma that is useful for filters with good lemmas about `sᶜ ∈ f` (e.g., `Filter.comap`, `Filter.coprod`, `Filter.Coprod`, `Filter.cofinite`). -/ protected theorem coext (h : ∀ s, sᶜ ∈ f ↔ sᶜ ∈ g) : f = g := diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index 0cdca7509d3e2..b2fd3766945b1 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -5,8 +5,8 @@ Authors: Jeremy Avigad, Mario Carneiro, Yury G. Kudryashov -/ import Mathlib.Data.Nat.Defs import Mathlib.Logic.IsEmpty -import Mathlib.Logic.Relation import Mathlib.Order.Basic +import Mathlib.Tactic.MkIffOfInductiveProp import Batteries.WF /-! diff --git a/Mathlib/Order/RelIso/Basic.lean b/Mathlib/Order/RelIso/Basic.lean index 55c86fb983e5c..fee223a097f83 100644 --- a/Mathlib/Order/RelIso/Basic.lean +++ b/Mathlib/Order/RelIso/Basic.lean @@ -119,9 +119,6 @@ theorem coe_fn_injective : Injective fun (f : r →r s) => (f : α → β) := theorem ext ⦃f g : r →r s⦄ (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h -theorem ext_iff {f g : r →r s} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - /-- Identity map is a relation homomorphism. -/ @[refl, simps] protected def id (r : α → α → Prop) : r →r r := @@ -251,9 +248,6 @@ theorem coe_fn_injective : Injective fun f : r ↪r s => (f : α → β) := theorem ext ⦃f g : r ↪r s⦄ (h : ∀ x, f x = g x) : f = g := DFunLike.ext _ _ h -theorem ext_iff {f g : r ↪r s} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - /-- Identity map is a relation embedding. -/ @[refl, simps!] protected def refl (r : α → α → Prop) : r ↪r r := @@ -582,9 +576,6 @@ theorem coe_fn_injective : Injective fun f : r ≃r s => (f : α → β) := theorem ext ⦃f g : r ≃r s⦄ (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h -theorem ext_iff {f g : r ≃r s} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - /-- Inverse map of a relation isomorphism is a relation isomorphism. -/ protected def symm (f : r ≃r s) : s ≃r r := ⟨f.toEquiv.symm, @fun a b => by erw [← f.map_rel_iff, f.1.apply_symm_apply, f.1.apply_symm_apply]⟩ diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index f8387b60b92a8..45f5e2f12ec4c 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -58,7 +58,7 @@ instance [Nonempty α] : Nonempty (RelSeries r) := variable {r} -@[ext] +@[ext (iff := false)] lemma ext {x y : RelSeries r} (length_eq : x.length = y.length) (toFun_eq : x.toFun = y.toFun ∘ Fin.cast (by rw [length_eq])) : x = y := by rcases x with ⟨nx, fx⟩ diff --git a/Mathlib/Probability/ConditionalExpectation.lean b/Mathlib/Probability/ConditionalExpectation.lean index 921ac10b7cd47..e9c83057f3fe8 100644 --- a/Mathlib/Probability/ConditionalExpectation.lean +++ b/Mathlib/Probability/ConditionalExpectation.lean @@ -44,7 +44,7 @@ theorem condexp_indep_eq (hle₁ : m₁ ≤ m) (hle₂ : m₂ ≤ m) [SigmaFinit stronglyMeasurable_const.aeStronglyMeasurable').symm rw [setIntegral_const] rw [← memℒp_one_iff_integrable] at hfint - refine Memℒp.induction_stronglyMeasurable hle₁ ENNReal.one_ne_top ?_ ?_ ?_ ?_ hfint ?_ + refine Memℒp.induction_stronglyMeasurable hle₁ ENNReal.one_ne_top _ ?_ ?_ ?_ ?_ hfint ?_ · exact ⟨f, hf, EventuallyEq.rfl⟩ · intro c t hmt _ rw [Indep_iff] at hindp diff --git a/Mathlib/Probability/Kernel/Basic.lean b/Mathlib/Probability/Kernel/Basic.lean index 1b00c9334020b..bb952329d7bd0 100644 --- a/Mathlib/Probability/Kernel/Basic.lean +++ b/Mathlib/Probability/Kernel/Basic.lean @@ -195,8 +195,6 @@ namespace Kernel @[ext] theorem ext {η : Kernel α β} (h : ∀ a, κ a = η a) : κ = η := DFunLike.ext _ _ h -protected theorem ext_iff {η : Kernel α β} : κ = η ↔ ∀ a, κ a = η a := DFunLike.ext_iff - theorem ext_iff' {η : Kernel α β} : κ = η ↔ ∀ a s, MeasurableSet s → κ a s = η a s := by simp_rw [Kernel.ext_iff, Measure.ext_iff] diff --git a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean index 178964e36caa6..ec56f812e2f47 100644 --- a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean +++ b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean @@ -52,9 +52,6 @@ instance instFunLike : FunLike (PMF α) α ℝ≥0∞ where protected theorem ext {p q : PMF α} (h : ∀ x, p x = q x) : p = q := DFunLike.ext p q h -theorem ext_iff {p q : PMF α} : p = q ↔ ∀ x, p x = q x := - DFunLike.ext_iff - theorem hasSum_coe_one (p : PMF α) : HasSum p 1 := p.2 diff --git a/Mathlib/RepresentationTheory/Action/Basic.lean b/Mathlib/RepresentationTheory/Action/Basic.lean index f92661d066b17..1d925a3a897c2 100644 --- a/Mathlib/RepresentationTheory/Action/Basic.lean +++ b/Mathlib/RepresentationTheory/Action/Basic.lean @@ -114,7 +114,7 @@ instance : Category (Action V G) where -- Porting note: added because `Hom.ext` is not triggered automatically @[ext] lemma hom_ext {M N : Action V G} (φ₁ φ₂ : M ⟶ N) (h : φ₁.hom = φ₂.hom) : φ₁ = φ₂ := - Hom.ext _ _ h + Hom.ext h @[simp] theorem id_hom (M : Action V G) : (𝟙 M : Hom M M).hom = 𝟙 M.V := @@ -229,7 +229,7 @@ def forget : Action V G ⥤ V where obj M := M.V map f := f.hom -instance : (forget V G).Faithful where map_injective w := Hom.ext _ _ w +instance : (forget V G).Faithful where map_injective w := Hom.ext w instance [ConcreteCategory V] : ConcreteCategory (Action V G) where forget := forget V G ⋙ ConcreteCategory.forget diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 52b7ada8a304f..3bc762a3a2580 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -297,7 +297,7 @@ noncomputable def leftRegularHomEquiv (A : Rep k G) : (Rep.ofMulAction k G G ⟶ map_smul' r x := rfl invFun x := leftRegularHom A x left_inv f := by - refine Action.Hom.ext _ _ (Finsupp.lhom_ext' fun x : G => LinearMap.ext_ring ?_) + refine Action.Hom.ext (Finsupp.lhom_ext' fun x : G => LinearMap.ext_ring ?_) have : f.hom ((ofMulAction k G G).ρ x (Finsupp.single (1 : G) (1 : k))) = A.ρ x (f.hom (Finsupp.single (1 : G) (1 : k))) := @@ -373,7 +373,7 @@ def homEquiv (A B C : Rep k G) : (A ⊗ B ⟶ C) ≃ (B ⟶ (Rep.ihom A).obj C) dsimp erw [ρ_inv_self_apply] rfl} - left_inv f := Action.Hom.ext _ _ (TensorProduct.ext' fun _ _ => rfl) + left_inv f := Action.Hom.ext (TensorProduct.ext' fun _ _ => rfl) right_inv f := by ext; rfl variable {A B C} @@ -393,9 +393,9 @@ instance : MonoidalClosed (Rep k G) where { rightAdj := Rep.ihom A adj := Adjunction.mkOfHomEquiv ( { homEquiv := Rep.homEquiv A - homEquiv_naturality_left_symm := fun _ _ => Action.Hom.ext _ _ + homEquiv_naturality_left_symm := fun _ _ => Action.Hom.ext (TensorProduct.ext' fun _ _ => rfl) - homEquiv_naturality_right := fun _ _ => Action.Hom.ext _ _ (LinearMap.ext + homEquiv_naturality_right := fun _ _ => Action.Hom.ext (LinearMap.ext fun _ => LinearMap.ext fun _ => rfl) })} @[simp] diff --git a/Mathlib/RingTheory/AdicCompletion/Basic.lean b/Mathlib/RingTheory/AdicCompletion/Basic.lean index cbf3ccfa57406..97af754373c43 100644 --- a/Mathlib/RingTheory/AdicCompletion/Basic.lean +++ b/Mathlib/RingTheory/AdicCompletion/Basic.lean @@ -294,9 +294,6 @@ theorem val_smul (n : ℕ) (r : R) (f : AdicCompletion I M) : (r • f).val n = theorem ext {x y : AdicCompletion I M} (h : ∀ n, x.val n = y.val n) : x = y := Subtype.eq <| funext h -theorem ext_iff {x y : AdicCompletion I M} : x = y ↔ ∀ n, x.val n = y.val n := - ⟨fun h n ↦ congrArg (eval I M n) h, ext⟩ - variable (I M) instance : IsHausdorff I (AdicCompletion I M) where @@ -422,9 +419,6 @@ theorem smul_apply (n : ℕ) (r : R) (f : AdicCauchySequence I M) : (r • f) n theorem ext {x y : AdicCauchySequence I M} (h : ∀ n, x n = y n) : x = y := Subtype.eq <| funext h -theorem ext_iff {x y : AdicCauchySequence I M} : x = y ↔ ∀ n, x n = y n := - ⟨fun h ↦ congrFun (congrArg Subtype.val h), ext⟩ - /-- The defining property of an adic cauchy sequence unwrapped. -/ theorem mk_eq_mk {m n : ℕ} (hmn : m ≤ n) (f : AdicCauchySequence I M) : Submodule.Quotient.mk (p := (I ^ m • ⊤ : Submodule R M)) (f n) = diff --git a/Mathlib/RingTheory/Bialgebra/Equiv.lean b/Mathlib/RingTheory/Bialgebra/Equiv.lean index 11a730028cf8a..a94379991aff9 100644 --- a/Mathlib/RingTheory/Bialgebra/Equiv.lean +++ b/Mathlib/RingTheory/Bialgebra/Equiv.lean @@ -194,9 +194,6 @@ variable {e e'} theorem ext (h : ∀ x, e x = e' x) : e = e' := DFunLike.ext _ _ h -theorem ext_iff : e = e' ↔ ∀ x, e x = e' x := - DFunLike.ext_iff - protected theorem congr_arg {x x'} : x = x' → e x = e x' := DFunLike.congr_arg e diff --git a/Mathlib/RingTheory/Bialgebra/Hom.lean b/Mathlib/RingTheory/Bialgebra/Hom.lean index 1f877e7c8e1c8..78feadf839f29 100644 --- a/Mathlib/RingTheory/Bialgebra/Hom.lean +++ b/Mathlib/RingTheory/Bialgebra/Hom.lean @@ -191,9 +191,6 @@ protected theorem congr_arg (φ : A →ₐc[R] B) {x y : A} (h : x = y) : φ x = theorem ext {φ₁ φ₂ : A →ₐc[R] B} (H : ∀ x, φ₁ x = φ₂ x) : φ₁ = φ₂ := DFunLike.ext _ _ H -theorem ext_iff {φ₁ φ₂ : A →ₐc[R] B} : φ₁ = φ₂ ↔ ∀ x, φ₁ x = φ₂ x := - DFunLike.ext_iff - @[ext high] theorem ext_of_ring {f g : R →ₐc[R] A} (h : f 1 = g 1) : f = g := coe_linearMap_injective (by ext; assumption) diff --git a/Mathlib/RingTheory/Coalgebra/Equiv.lean b/Mathlib/RingTheory/Coalgebra/Equiv.lean index 890e6cc40514b..5b15a25b51a3b 100644 --- a/Mathlib/RingTheory/Coalgebra/Equiv.lean +++ b/Mathlib/RingTheory/Coalgebra/Equiv.lean @@ -157,9 +157,6 @@ variable {e e'} theorem ext (h : ∀ x, e x = e' x) : e = e' := DFunLike.ext _ _ h -theorem ext_iff : e = e' ↔ ∀ x, e x = e' x := - DFunLike.ext_iff - protected theorem congr_arg {x x'} : x = x' → e x = e x' := DFunLike.congr_arg e diff --git a/Mathlib/RingTheory/Coalgebra/Hom.lean b/Mathlib/RingTheory/Coalgebra/Hom.lean index 877419e01f005..d138b068c895a 100644 --- a/Mathlib/RingTheory/Coalgebra/Hom.lean +++ b/Mathlib/RingTheory/Coalgebra/Hom.lean @@ -167,9 +167,6 @@ protected theorem congr_arg (φ : A →ₗc[R] B) {x y : A} (h : x = y) : φ x = theorem ext {φ₁ φ₂ : A →ₗc[R] B} (H : ∀ x, φ₁ x = φ₂ x) : φ₁ = φ₂ := DFunLike.ext _ _ H -theorem ext_iff {φ₁ φ₂ : A →ₗc[R] B} : φ₁ = φ₂ ↔ ∀ x, φ₁ x = φ₂ x := - DFunLike.ext_iff - @[ext high] theorem ext_of_ring {f g : R →ₗc[R] A} (h : f 1 = g 1) : f = g := coe_linearMap_injective (by ext; assumption) diff --git a/Mathlib/RingTheory/DedekindDomain/Factorization.lean b/Mathlib/RingTheory/DedekindDomain/Factorization.lean index 6fa3042f4d106..c07bfc797cc24 100644 --- a/Mathlib/RingTheory/DedekindDomain/Factorization.lean +++ b/Mathlib/RingTheory/DedekindDomain/Factorization.lean @@ -70,7 +70,7 @@ theorem Ideal.finite_factors {I : Ideal R} (hI : I ≠ 0) : Finite.of_injective (fun v => (⟨(v : HeightOneSpectrum R).asIdeal, v.2⟩ : { x // x ∣ I })) ?_ intro v w hvw simp? at hvw says simp only [Subtype.mk.injEq] at hvw - exact Subtype.coe_injective (HeightOneSpectrum.ext _ _ hvw) + exact Subtype.coe_injective (HeightOneSpectrum.ext hvw) /-- For every nonzero ideal `I` of `v`, there are finitely many maximal ideals `v` such that the multiplicity of `v` in the factorization of `I`, denoted `val_v(I)`, is nonzero. -/ @@ -133,7 +133,7 @@ theorem finprod_not_dvd (I : Ideal R) (hI : I ≠ 0) : have hw_prime : Prime w.asIdeal := Ideal.prime_of_isPrime w.ne_bot w.isPrime have hvw := Prime.dvd_of_dvd_pow hv_prime hvw' rw [Prime.dvd_prime_iff_associated hv_prime hw_prime, associated_iff_eq] at hvw - exact (Finset.mem_erase.mp hw).1 (HeightOneSpectrum.ext w v (Eq.symm hvw)) + exact (Finset.mem_erase.mp hw).1 (HeightOneSpectrum.ext hvw.symm) end Ideal diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 0a3b51efa73ac..d71fe641b714b 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -366,7 +366,7 @@ theorem exists_multiset_prod_cons_le_and_prod_not_le [IsDedekindDomain A] (hNF : -- Then in fact there is a `P ∈ Z` with `P ≤ M`. obtain ⟨P, hPZ, rfl⟩ := Multiset.mem_map.mp hPZ' classical - have := Multiset.map_erase PrimeSpectrum.asIdeal PrimeSpectrum.ext P Z + have := Multiset.map_erase PrimeSpectrum.asIdeal (fun _ _ => PrimeSpectrum.ext) P Z obtain ⟨hP0, hZP0⟩ : P.asIdeal ≠ ⊥ ∧ ((Z.erase P).map PrimeSpectrum.asIdeal).prod ≠ ⊥ := by rwa [Ne, ← Multiset.cons_erase hPZ', Multiset.prod_cons, Ideal.mul_eq_bot, not_or, ← this] at hprodZ diff --git a/Mathlib/RingTheory/Filtration.lean b/Mathlib/RingTheory/Filtration.lean index 237a8357cabf9..d28ab3f545103 100644 --- a/Mathlib/RingTheory/Filtration.lean +++ b/Mathlib/RingTheory/Filtration.lean @@ -161,7 +161,8 @@ theorem iInf_N {ι : Sort*} (f : ι → I.Filtration M) : (iInf f).N = ⨅ i, (f congr_arg sInf (Set.range_comp _ _).symm instance : CompleteLattice (I.Filtration M) := - Function.Injective.completeLattice Ideal.Filtration.N Ideal.Filtration.ext sup_N inf_N + Function.Injective.completeLattice Ideal.Filtration.N + (fun _ _ => Ideal.Filtration.ext) sup_N inf_N (fun _ => sSup_image) (fun _ => sInf_image) top_N bot_N instance : Inhabited (I.Filtration M) := diff --git a/Mathlib/RingTheory/HahnSeries/Basic.lean b/Mathlib/RingTheory/HahnSeries/Basic.lean index 22a4a72efcf7a..b189d32690b7d 100644 --- a/Mathlib/RingTheory/HahnSeries/Basic.lean +++ b/Mathlib/RingTheory/HahnSeries/Basic.lean @@ -52,7 +52,7 @@ section Zero variable [PartialOrder Γ] [Zero R] theorem coeff_injective : Injective (coeff : HahnSeries Γ R → Γ → R) := - HahnSeries.ext + fun _ _ => HahnSeries.ext @[simp] theorem coeff_inj {x y : HahnSeries Γ R} : x.coeff = y.coeff ↔ x = y := @@ -83,7 +83,7 @@ instance : Inhabited (HahnSeries Γ R) := ⟨0⟩ instance [Subsingleton R] : Subsingleton (HahnSeries Γ R) := - ⟨fun a b => a.ext b (by subsingleton)⟩ + ⟨fun _ _ => HahnSeries.ext (by subsingleton)⟩ @[simp] theorem zero_coeff {a : Γ} : (0 : HahnSeries Γ R).coeff a = 0 := @@ -154,7 +154,7 @@ def single (a : Γ) : ZeroHom R (HahnSeries Γ R) where toFun r := { coeff := Pi.single a r isPWO_support' := (Set.isPWO_singleton a).mono Pi.support_single_subset } - map_zero' := HahnSeries.ext _ _ (Pi.single_zero _) + map_zero' := HahnSeries.ext (Pi.single_zero _) variable {a b : Γ} {r : R} diff --git a/Mathlib/RingTheory/HahnSeries/Multiplication.lean b/Mathlib/RingTheory/HahnSeries/Multiplication.lean index 5436740e4da6f..0ddaf67084bec 100644 --- a/Mathlib/RingTheory/HahnSeries/Multiplication.lean +++ b/Mathlib/RingTheory/HahnSeries/Multiplication.lean @@ -104,10 +104,6 @@ def rec {motive : HahnModule Γ R V → Sort*} (h : ∀ x : HahnSeries Γ V, mot theorem ext (x y : HahnModule Γ R V) (h : ((of R).symm x).coeff = ((of R).symm y).coeff) : x = y := (of R).symm.injective <| HahnSeries.coeff_inj.1 h -protected theorem ext_iff (x y : HahnModule Γ R V) : - x = y ↔ ((of R).symm x).coeff = ((of R).symm y).coeff := by - simp_all only [HahnSeries.coeff_inj, EmbeddingLike.apply_eq_iff_eq] - end section SMul diff --git a/Mathlib/RingTheory/IsTensorProduct.lean b/Mathlib/RingTheory/IsTensorProduct.lean index 35a5b3b55d663..f3699b8d59848 100644 --- a/Mathlib/RingTheory/IsTensorProduct.lean +++ b/Mathlib/RingTheory/IsTensorProduct.lean @@ -174,7 +174,7 @@ nonrec theorem IsBaseChange.inductionOn (x : N) (P : N → Prop) (h₁ : P 0) (h theorem IsBaseChange.algHom_ext (g₁ g₂ : N →ₗ[S] Q) (e : ∀ x, g₁ (f x) = g₂ (f x)) : g₁ = g₂ := by ext x - refine h.inductionOn x ?_ ?_ ?_ ?_ + refine h.inductionOn x _ ?_ ?_ ?_ ?_ · rw [map_zero, map_zero] · assumption · intro s n e' @@ -396,7 +396,7 @@ noncomputable def Algebra.pushoutDesc [H : Algebra.IsPushout R S R' S'] {A : Typ · dsimp only [LinearMap.restrictScalars_apply] rw [← (algebraMap R' S').map_one, this, _root_.map_one] · intro x y - refine H.out.inductionOn x ?_ ?_ ?_ ?_ + refine H.out.inductionOn x _ ?_ ?_ ?_ ?_ · rw [zero_mul, map_zero, zero_mul] rotate_left · intro s s' e @@ -408,7 +408,7 @@ noncomputable def Algebra.pushoutDesc [H : Algebra.IsPushout R S R' S'] {A : Typ intro x dsimp rw [this] - refine H.out.inductionOn y ?_ ?_ ?_ ?_ + refine H.out.inductionOn y _ ?_ ?_ ?_ ?_ · rw [mul_zero, map_zero, mul_zero] · intro y dsimp @@ -453,12 +453,12 @@ theorem Algebra.lift_algHom_comp_right [Algebra.IsPushout R S R' S'] {A : Type*} (Algebra.pushoutDesc S' f g H).comp (toAlgHom R R' S') = g := AlgHom.ext fun x => (Algebra.pushoutDesc_right S' f g H x : _) -@[ext] +@[ext (iff := false)] theorem Algebra.IsPushout.algHom_ext [H : Algebra.IsPushout R S R' S'] {A : Type*} [Semiring A] [Algebra R A] {f g : S' →ₐ[R] A} (h₁ : f.comp (toAlgHom R R' S') = g.comp (toAlgHom R R' S')) (h₂ : f.comp (toAlgHom R S S') = g.comp (toAlgHom R S S')) : f = g := by ext x - refine H.1.inductionOn x ?_ ?_ ?_ ?_ + refine H.1.inductionOn x _ ?_ ?_ ?_ ?_ · simp only [map_zero] · exact AlgHom.congr_fun h₁ · intro s s' e diff --git a/Mathlib/RingTheory/MaximalSpectrum.lean b/Mathlib/RingTheory/MaximalSpectrum.lean index 4a881affacc76..b4996af33c24a 100644 --- a/Mathlib/RingTheory/MaximalSpectrum.lean +++ b/Mathlib/RingTheory/MaximalSpectrum.lean @@ -53,7 +53,7 @@ def toPrimeSpectrum (x : MaximalSpectrum R) : PrimeSpectrum R := ⟨x.asIdeal, x.IsMaximal.isPrime⟩ theorem toPrimeSpectrum_injective : (@toPrimeSpectrum R _).Injective := fun ⟨_, _⟩ ⟨_, _⟩ h => by - simpa only [MaximalSpectrum.mk.injEq] using (PrimeSpectrum.ext_iff _ _).mp h + simpa only [MaximalSpectrum.mk.injEq] using PrimeSpectrum.ext_iff.mp h open PrimeSpectrum Set diff --git a/Mathlib/RingTheory/MvPowerSeries/Basic.lean b/Mathlib/RingTheory/MvPowerSeries/Basic.lean index f9a9f48c3daee..626b2270b0471 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Basic.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Basic.lean @@ -143,10 +143,8 @@ theorem ext {φ ψ} (h : ∀ n : σ →₀ ℕ, coeff R n φ = coeff R n ψ) : funext h /-- Two multivariate formal power series are equal - if and only if all their coefficients are equal. -/ -protected theorem ext_iff {φ ψ : MvPowerSeries σ R} : - φ = ψ ↔ ∀ n : σ →₀ ℕ, coeff R n φ = coeff R n ψ := - Function.funext_iff +if and only if all their coefficients are equal. -/ +add_decl_doc MvPowerSeries.ext_iff theorem monomial_def [DecidableEq σ] (n : σ →₀ ℕ) : (monomial R n) = LinearMap.stdBasis R (fun _ ↦ R) n := by diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index b7826e615a803..b16ab068c4441 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -156,8 +156,7 @@ theorem ext {φ ψ : R⟦X⟧} (h : ∀ n, coeff R n φ = coeff R n ψ) : φ = rfl /-- Two formal power series are equal if all their coefficients are equal. -/ -protected theorem ext_iff {φ ψ : R⟦X⟧} : φ = ψ ↔ ∀ n, coeff R n φ = coeff R n ψ := - ⟨fun h n => congr_arg (coeff R n) h, ext⟩ +add_decl_doc PowerSeries.ext_iff instance [Subsingleton R] : Subsingleton R⟦X⟧ := by simp only [subsingleton_iff, PowerSeries.ext_iff] diff --git a/Mathlib/RingTheory/PrimeSpectrum.lean b/Mathlib/RingTheory/PrimeSpectrum.lean index a8e72685b4246..e672235438cc2 100644 --- a/Mathlib/RingTheory/PrimeSpectrum.lean +++ b/Mathlib/RingTheory/PrimeSpectrum.lean @@ -375,7 +375,7 @@ See the corresponding section at `AlgebraicGeometry/PrimeSpectrum/Basic`. -/ instance : PartialOrder (PrimeSpectrum R) := - PartialOrder.lift asIdeal (PrimeSpectrum.ext) + PartialOrder.lift asIdeal (@PrimeSpectrum.ext _ _) @[simp] theorem asIdeal_le_asIdeal (x y : PrimeSpectrum R) : x.asIdeal ≤ y.asIdeal ↔ x ≤ y := @@ -391,7 +391,7 @@ instance [IsDomain R] : OrderBot (PrimeSpectrum R) where instance {R : Type*} [Field R] : Unique (PrimeSpectrum R) where default := ⊥ - uniq x := PrimeSpectrum.ext _ _ ((IsSimpleOrder.eq_bot_or_eq_top _).resolve_right x.2.ne_top) + uniq x := PrimeSpectrum.ext ((IsSimpleOrder.eq_bot_or_eq_top _).resolve_right x.2.ne_top) end Order diff --git a/Mathlib/RingTheory/Valuation/Basic.lean b/Mathlib/RingTheory/Valuation/Basic.lean index a793d69a95a4a..1008370dd53ef 100644 --- a/Mathlib/RingTheory/Valuation/Basic.lean +++ b/Mathlib/RingTheory/Valuation/Basic.lean @@ -195,11 +195,6 @@ theorem map_sum_lt' {ι : Type*} {s : Finset ι} {f : ι → R} {g : Γ₀} (hg theorem map_pow : ∀ (x) (n : ℕ), v (x ^ n) = v x ^ n := v.toMonoidWithZeroHom.toMonoidHom.map_pow -/-- Deprecated. Use `DFunLike.ext_iff`. -/ --- @[deprecated] Porting note: using `DFunLike.ext_iff` is not viable below for now -theorem ext_iff {v₁ v₂ : Valuation R Γ₀} : v₁ = v₂ ↔ ∀ r, v₁ r = v₂ r := - DFunLike.ext_iff - -- The following definition is not an instance, because we have more than one `v` on a given `R`. -- In addition, type class inference would not be able to infer `v`. /-- A valuation gives a preorder on the underlying ring. -/ @@ -681,9 +676,6 @@ theorem map_pow : ∀ (x : R) (n : ℕ), v (x ^ n) = n • (v x) := theorem ext {v₁ v₂ : AddValuation R Γ₀} (h : ∀ r, v₁ r = v₂ r) : v₁ = v₂ := Valuation.ext h -theorem ext_iff {v₁ v₂ : AddValuation R Γ₀} : v₁ = v₂ ↔ ∀ (r : R), v₁ r = v₂ r := - Valuation.ext_iff - -- The following definition is not an instance, because we have more than one `v` on a given `R`. -- In addition, type class inference would not be able to infer `v`. /-- A valuation gives a preorder on the underlying ring. -/ diff --git a/Mathlib/RingTheory/WittVector/Defs.lean b/Mathlib/RingTheory/WittVector/Defs.lean index 35913b6865610..c50793f3e1de2 100644 --- a/Mathlib/RingTheory/WittVector/Defs.lean +++ b/Mathlib/RingTheory/WittVector/Defs.lean @@ -74,9 +74,6 @@ theorem ext {x y : 𝕎 R} (h : ∀ n, x.coeff n = y.coeff n) : x = y := by simp only at h simp [Function.funext_iff, h] -theorem ext_iff {x y : 𝕎 R} : x = y ↔ ∀ n, x.coeff n = y.coeff n := - ⟨fun h n => by rw [h], ext⟩ - variable (p) theorem coeff_mk (x : ℕ → R) : (mk p x).coeff = x := diff --git a/Mathlib/RingTheory/WittVector/Truncated.lean b/Mathlib/RingTheory/WittVector/Truncated.lean index 6ca23160ab060..ce513738df05e 100644 --- a/Mathlib/RingTheory/WittVector/Truncated.lean +++ b/Mathlib/RingTheory/WittVector/Truncated.lean @@ -81,9 +81,6 @@ def coeff (i : Fin n) (x : TruncatedWittVector p n R) : R := theorem ext {x y : TruncatedWittVector p n R} (h : ∀ i, x.coeff i = y.coeff i) : x = y := funext h -theorem ext_iff {x y : TruncatedWittVector p n R} : x = y ↔ ∀ i, x.coeff i = y.coeff i := - ⟨fun h i => by rw [h], ext⟩ - @[simp] theorem coeff_mk (x : Fin n → R) (i : Fin n) : (mk p x).coeff i = x i := rfl diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index 0c83ed384947a..3b4f62f933cd2 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -651,9 +651,6 @@ theorem toSet_subset_iff {x y : ZFSet} : x.toSet ⊆ y.toSet ↔ x ⊆ y := by theorem ext {x y : ZFSet.{u}} : (∀ z : ZFSet.{u}, z ∈ x ↔ z ∈ y) → x = y := Quotient.inductionOn₂ x y fun _ _ h => Quotient.sound (Mem.ext fun w => h ⟦w⟧) -theorem ext_iff {x y : ZFSet.{u}} : x = y ↔ ∀ z : ZFSet.{u}, z ∈ x ↔ z ∈ y := - ⟨fun h => by simp [h], ext⟩ - theorem toSet_injective : Function.Injective toSet := fun _ _ h => ext <| Set.ext_iff.1 h @[simp] @@ -1227,9 +1224,6 @@ protected def sep (p : ZFSet → Prop) (A : Class) : Class := theorem ext {x y : Class.{u}} : (∀ z : ZFSet.{u}, x z ↔ y z) → x = y := Set.ext -theorem ext_iff {x y : Class.{u}} : x = y ↔ ∀ z, x z ↔ y z := - Set.ext_iff - /-- Coerce a ZFC set into a class -/ @[coe] def ofSet (x : ZFSet.{u}) : Class.{u} := diff --git a/Mathlib/Tactic/Abel.lean b/Mathlib/Tactic/Abel.lean index c493b03da8c72..decb801e22837 100644 --- a/Mathlib/Tactic/Abel.lean +++ b/Mathlib/Tactic/Abel.lean @@ -386,7 +386,7 @@ elab_rules : tactic | `(tactic| abel1 $[!%$tk]?) => withMainContext do | throwError "abel1 requires an equality goal" trace[abel] "running on an equality `{e₁} = {e₂}`." let c ← mkContext e₁ - closeMainGoal <| ← AtomM.run tm <| ReaderT.run (r := c) do + closeMainGoal `abel1 <| ← AtomM.run tm <| ReaderT.run (r := c) do let (e₁', p₁) ← eval e₁ trace[abel] "found `{p₁}`, a proof that `{e₁} = {e₁'.e}`" let (e₂', p₂) ← eval e₂ diff --git a/Mathlib/Tactic/Attr/Core.lean b/Mathlib/Tactic/Attr/Core.lean index 55ef5e39e7b18..e08139380e134 100644 --- a/Mathlib/Tactic/Attr/Core.lean +++ b/Mathlib/Tactic/Attr/Core.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Tactic.Attr.Register -import Batteries.Logic /-! # Simp tags for core lemmas diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index c56882d30a96c..b6ea9fc557177 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ import Lean.Linter.Util -import Batteries.Data.Array.Basic import Batteries.Data.String.Matcher import Batteries.Tactic.Lint diff --git a/Mathlib/Tactic/Linter/UnusedTactic.lean b/Mathlib/Tactic/Linter/UnusedTactic.lean index fb57a6ed99fd2..425cdc0023e0e 100644 --- a/Mathlib/Tactic/Linter/UnusedTactic.lean +++ b/Mathlib/Tactic/Linter/UnusedTactic.lean @@ -5,7 +5,6 @@ Authors: Damiano Testa -/ import Lean.Elab.Command import Lean.Linter.Util -import Batteries.Data.List.Basic import Batteries.Tactic.Unreachable /-! diff --git a/Mathlib/Tactic/TFAE.lean b/Mathlib/Tactic/TFAE.lean index 52fac7f800bfc..d5ef4b815aad8 100644 --- a/Mathlib/Tactic/TFAE.lean +++ b/Mathlib/Tactic/TFAE.lean @@ -221,7 +221,7 @@ elab_rules : tactic let goal ← getMainGoal goal.withContext do let (tfaeListQ, tfaeList) ← getTFAEList (← goal.getType) - closeMainGoal <|← AtomM.run .reducible do + closeMainGoal `tfae_finish <|← AtomM.run .reducible do let is ← tfaeList.mapM AtomM.addAtom let mut hyps := #[] for hyp in ← getLocalHyps do diff --git a/Mathlib/Tactic/Widget/Calc.lean b/Mathlib/Tactic/Widget/Calc.lean index e287eb58106d4..34ea5f388d745 100644 --- a/Mathlib/Tactic/Widget/Calc.lean +++ b/Mathlib/Tactic/Widget/Calc.lean @@ -7,6 +7,8 @@ import Lean.Elab.Tactic.Calc import Mathlib.Data.String.Defs import Mathlib.Tactic.Widget.SelectPanelUtils +import Batteries.CodeAction.Attr +import Batteries.Lean.Position /-! # Calc widget diff --git a/Mathlib/Tactic/Widget/Conv.lean b/Mathlib/Tactic/Widget/Conv.lean index d3deed1858847..50cd77a425f8d 100644 --- a/Mathlib/Tactic/Widget/Conv.lean +++ b/Mathlib/Tactic/Widget/Conv.lean @@ -6,6 +6,7 @@ Authors: Robin Böhne, Wojciech Nawrocki, Patrick Massot import Mathlib.Tactic.Widget.SelectPanelUtils import Mathlib.Data.String.Defs import Batteries.Tactic.Lint +import Batteries.Lean.Position /-! # Conv widget diff --git a/Mathlib/Testing/SlimCheck/Functions.lean b/Mathlib/Testing/SlimCheck/Functions.lean index 1f5e707a062a8..7f0f7d67f8cb1 100644 --- a/Mathlib/Testing/SlimCheck/Functions.lean +++ b/Mathlib/Testing/SlimCheck/Functions.lean @@ -334,7 +334,7 @@ theorem applyId_mem_iff [DecidableEq α] {xs ys : List α} (h₀ : List.Nodup xs erw [← Option.mem_def, List.mem_dlookup_iff] at h₃ · simp only [Prod.toSigma, List.mem_map, heq_iff_eq, Prod.exists] at h₃ rcases h₃ with ⟨a, b, h₃, h₄, h₅⟩ - apply (List.mem_zip h₃).2 + apply (List.of_mem_zip h₃).2 simp only [List.NodupKeys, List.keys, comp, Prod.fst_toSigma, List.map_map] rwa [List.map_fst_zip _ _ (le_of_eq h₆)] @@ -347,7 +347,7 @@ theorem List.applyId_eq_self [DecidableEq α] {xs ys : List α} (x : α) : simp only [List.keys, not_exists, Prod.toSigma, exists_and_right, exists_eq_right, List.mem_map, Function.comp_apply, List.map_map, Prod.exists] intro y hy - exact h (List.mem_zip hy).1 + exact h (List.of_mem_zip hy).1 theorem applyId_injective [DecidableEq α] {xs ys : List α} (h₀ : List.Nodup xs) (h₁ : xs ~ ys) : Injective.{u + 1, u + 1} (List.applyId (xs.zip ys)) := by @@ -392,12 +392,12 @@ def Perm.slice [DecidableEq α] (n m : ℕ) : /-- A list, in decreasing order, of sizes that should be sliced off a list of length `n` -/ -def sliceSizes : ℕ → List ℕ+ +def sliceSizes : ℕ → MLList Id ℕ+ | n => if h : 0 < n then have : n / 2 < n := Nat.div_lt_self h (by decide : 1 < 2) - ⟨_, h⟩ :: (sliceSizes <| n / 2) - else [] + .cons ⟨_, h⟩ (sliceSizes <| n / 2) + else .nil /-- Shrink a permutation of a list, slicing a segment in the middle. @@ -409,7 +409,7 @@ protected def shrinkPerm {α : Type} [DecidableEq α] : (Σ' xs ys : List α, xs ~ ys ∧ ys.Nodup) → List (Σ' xs ys : List α, xs ~ ys ∧ ys.Nodup) | xs => do let k := xs.1.length - let n ← sliceSizes k + let n ← (sliceSizes k).force let i ← List.finRange <| k / n pure <| Perm.slice (i * n) n xs diff --git a/Mathlib/Testing/SlimCheck/Testable.lean b/Mathlib/Testing/SlimCheck/Testable.lean index 764ae3acc748c..e4ece33e569d8 100644 --- a/Mathlib/Testing/SlimCheck/Testable.lean +++ b/Mathlib/Testing/SlimCheck/Testable.lean @@ -558,7 +558,7 @@ scoped elab "mk_decorations" : tactic => do let goal ← getMainGoal let goalType ← goal.getType if let .app (.const ``Decorations.DecorationsOf _) body := goalType then - closeMainGoal (← addDecorations body) + closeMainGoal `mk_decorations (← addDecorations body) end Decorations diff --git a/Mathlib/Topology/Algebra/Algebra.lean b/Mathlib/Topology/Algebra/Algebra.lean index 250871fe500e9..5870483dbea15 100644 --- a/Mathlib/Topology/Algebra/Algebra.lean +++ b/Mathlib/Topology/Algebra/Algebra.lean @@ -155,8 +155,6 @@ initialize_simps_projections ContinuousAlgHom (toAlgHom_toFun → apply, toAlgHo @[ext] theorem ext {f g : A →A[R] B} (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h -theorem ext_iff {f g : A →A[R] B} : f = g ↔ ∀ x, f x = g x := DFunLike.ext_iff - /-- Copy of a `ContinuousAlgHom` with a new `toFun` equal to the old one. Useful to fix definitional equalities. -/ def copy (f : A →A[R] B) (f' : A → B) (h : f' = ⇑f) : A →A[R] B where @@ -190,7 +188,7 @@ protected theorem map_sum {ι : Type*} (f : A →A[R] B) (s : Finset ι) (g : ι map_sum .. /-- Any two continuous `R`-algebra morphisms from `R` are equal -/ -@[ext] +@[ext (iff := false)] theorem ext_ring [TopologicalSpace R] {f g : R →A[R] A} : f = g := coe_inj.mp (ext_id _ _ _) diff --git a/Mathlib/Topology/Algebra/ContinuousAffineMap.lean b/Mathlib/Topology/Algebra/ContinuousAffineMap.lean index e6093be661e36..39cb34e5d4264 100644 --- a/Mathlib/Topology/Algebra/ContinuousAffineMap.lean +++ b/Mathlib/Topology/Algebra/ContinuousAffineMap.lean @@ -69,9 +69,6 @@ theorem coe_injective : @Function.Injective (P →ᴬ[R] Q) (P → Q) (⇑) := theorem ext {f g : P →ᴬ[R] Q} (h : ∀ x, f x = g x) : f = g := DFunLike.ext _ _ h -theorem ext_iff {f g : P →ᴬ[R] Q} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - theorem congr_fun {f g : P →ᴬ[R] Q} (h : f = g) (x : P) : f x = g x := DFunLike.congr_fun h _ diff --git a/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean b/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean index 731fe283d9bfd..c6b9bc5f30347 100644 --- a/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Alternating/Basic.lean @@ -94,9 +94,6 @@ theorem coe_toAlternatingMap : ⇑f.toAlternatingMap = f := rfl theorem ext {f g : M [⋀^ι]→L[R] N} (H : ∀ x, f x = g x) : f = g := DFunLike.ext _ _ H -theorem ext_iff {f g : M [⋀^ι]→L[R] N} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - theorem toAlternatingMap_injective : Injective (toAlternatingMap : (M [⋀^ι]→L[R] N) → (M [⋀^ι]→ₗ[R] N)) := fun f g h => DFunLike.ext' <| by convert DFunLike.ext'_iff.1 h diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 827a8543772a5..e50b1de578992 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -425,9 +425,6 @@ initialize_simps_projections ContinuousLinearMap (toLinearMap_toFun → apply, t theorem ext {f g : M₁ →SL[σ₁₂] M₂} (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h -protected theorem ext_iff {f g : M₁ →SL[σ₁₂] M₂} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - /-- Copy of a `ContinuousLinearMap` with a new `toFun` equal to the old one. Useful to fix definitional equalities. -/ protected def copy (f : M₁ →SL[σ₁₂] M₂) (f' : M₁ → M₂) (h : f' = ⇑f) : M₁ →SL[σ₁₂] M₂ where @@ -475,9 +472,6 @@ theorem coe_coe (f : M₁ →SL[σ₁₂] M₂) : ⇑(f : M₁ →ₛₗ[σ₁ theorem ext_ring [TopologicalSpace R₁] {f g : R₁ →L[R₁] M₁} (h : f 1 = g 1) : f = g := coe_inj.1 <| LinearMap.ext_ring h -protected theorem ext_ring_iff [TopologicalSpace R₁] {f g : R₁ →L[R₁] M₁} : f = g ↔ f 1 = g 1 := - ⟨fun h => h ▸ rfl, ext_ring⟩ - /-- If two continuous linear maps are equal on a set `s`, then they are equal on the closure of the `Submodule.span` of this set. -/ theorem eqOn_closure_span [T2Space M₂] {s : Set M₁} {f g : M₁ →SL[σ₁₂] M₂} (h : Set.EqOn f g s) : diff --git a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean index d81fce82a7cb8..c4849ddb17994 100644 --- a/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Multilinear/Basic.lean @@ -101,9 +101,6 @@ theorem coe_coe : (f.toMultilinearMap : (∀ i, M₁ i) → M₂) = f := theorem ext {f f' : ContinuousMultilinearMap R M₁ M₂} (H : ∀ x, f x = f' x) : f = f' := DFunLike.ext _ _ H -theorem ext_iff {f f' : ContinuousMultilinearMap R M₁ M₂} : f = f' ↔ ∀ x, f x = f' x := by - rw [← toMultilinearMap_injective.eq_iff, MultilinearMap.ext_iff]; rfl - @[simp] theorem map_add [DecidableEq ι] (m : ∀ i, M₁ i) (i : ι) (x y : M₁ i) : f (update m i (x + y)) = f (update m i x) + f (update m i y) := diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index f4fcee7d2ef3b..3558c44c57f83 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -76,7 +76,7 @@ open Topology lemma isOpen_mk {p h₁ h₂ h₃} : IsOpen[⟨p, h₁, h₂, h₃⟩] s ↔ p s := Iff.rfl -@[ext] +@[ext (iff := false)] protected theorem TopologicalSpace.ext : ∀ {f g : TopologicalSpace X}, IsOpen[f] = IsOpen[g] → f = g | ⟨_, _, _, _⟩, ⟨_, _, _, _⟩, rfl => rfl diff --git a/Mathlib/Topology/Bornology/Basic.lean b/Mathlib/Topology/Bornology/Basic.lean index 64e6f76fbbbc0..6b8a44d2bdedf 100644 --- a/Mathlib/Topology/Bornology/Basic.lean +++ b/Mathlib/Topology/Bornology/Basic.lean @@ -75,10 +75,6 @@ lemma Bornology.ext (t t' : Bornology α) cases t' congr -lemma Bornology.ext_iff (t t' : Bornology α) : - t = t' ↔ @Bornology.cobounded α t = @Bornology.cobounded α t' := -⟨congrArg _, Bornology.ext _ _⟩ - /-- A constructor for bornologies by specifying the bounded sets, and showing that they satisfy the appropriate conditions. -/ @[simps] @@ -204,7 +200,7 @@ end theorem ext_iff' {t t' : Bornology α} : t = t' ↔ ∀ s, s ∈ @cobounded α t ↔ s ∈ @cobounded α t' := - (Bornology.ext_iff _ _).trans Filter.ext_iff + Bornology.ext_iff.trans Filter.ext_iff theorem ext_iff_isBounded {t t' : Bornology α} : t = t' ↔ ∀ s, @IsBounded α t s ↔ @IsBounded α t' s := diff --git a/Mathlib/Topology/Category/TopCat/OpenNhds.lean b/Mathlib/Topology/Category/TopCat/OpenNhds.lean index b821e2fe7784a..c5465486c0688 100644 --- a/Mathlib/Topology/Category/TopCat/OpenNhds.lean +++ b/Mathlib/Topology/Category/TopCat/OpenNhds.lean @@ -44,7 +44,7 @@ instance partialOrder (x : X) : PartialOrder (OpenNhds x) where le U V := U.1 ≤ V.1 le_refl _ := by dsimp [LE.le]; exact le_rfl le_trans _ _ _ := by dsimp [LE.le]; exact le_trans - le_antisymm _ _ i j := FullSubcategory.ext _ _ <| le_antisymm i j + le_antisymm _ _ i j := FullSubcategory.ext <| le_antisymm i j instance (x : X) : Lattice (OpenNhds x) := { OpenNhds.partialOrder x with diff --git a/Mathlib/Topology/ContinuousFunction/StoneWeierstrass.lean b/Mathlib/Topology/ContinuousFunction/StoneWeierstrass.lean index ef8e8920998bd..57b51d1724771 100644 --- a/Mathlib/Topology/ContinuousFunction/StoneWeierstrass.lean +++ b/Mathlib/Topology/ContinuousFunction/StoneWeierstrass.lean @@ -425,7 +425,7 @@ theorem polynomialFunctions.starClosure_topologicalClosure {𝕜 : Type*} [RCLik /-- Continuous algebra homomorphisms from `C(s, ℝ)` into an `ℝ`-algebra `A` which agree at `X : 𝕜[X]` (interpreted as a continuous map) are, in fact, equal. -/ -@[ext] +@[ext (iff := false)] theorem ContinuousMap.algHom_ext_map_X {A : Type*} [Ring A] [Algebra ℝ A] [TopologicalSpace A] [T2Space A] {s : Set ℝ} [CompactSpace s] {φ ψ : C(s, ℝ) →ₐ[ℝ] A} (hφ : Continuous φ) (hψ : Continuous ψ) @@ -438,7 +438,7 @@ theorem ContinuousMap.algHom_ext_map_X {A : Type*} [Ring A] /-- Continuous star algebra homomorphisms from `C(s, 𝕜)` into a star `𝕜`-algebra `A` which agree at `X : 𝕜[X]` (interpreted as a continuous map) are, in fact, equal. -/ -@[ext] +@[ext (iff := false)] theorem ContinuousMap.starAlgHom_ext_map_X {𝕜 A : Type*} [RCLike 𝕜] [Ring A] [StarRing A] [Algebra 𝕜 A] [TopologicalSpace A] [T2Space A] {s : Set 𝕜} [CompactSpace s] {φ ψ : C(s, 𝕜) →⋆ₐ[𝕜] A} (hφ : Continuous φ) (hψ : Continuous ψ) diff --git a/Mathlib/Topology/LocallyConstant/Basic.lean b/Mathlib/Topology/LocallyConstant/Basic.lean index bc79242f984a4..cd4f5070a4e86 100644 --- a/Mathlib/Topology/LocallyConstant/Basic.lean +++ b/Mathlib/Topology/LocallyConstant/Basic.lean @@ -247,8 +247,6 @@ theorem coe_inj {f g : LocallyConstant X Y} : (f : X → Y) = g ↔ f = g := theorem ext ⦃f g : LocallyConstant X Y⦄ (h : ∀ x, f x = g x) : f = g := DFunLike.ext _ _ h -theorem ext_iff {f g : LocallyConstant X Y} : f = g ↔ ∀ x, f x = g x := DFunLike.ext_iff - section CodomainTopologicalSpace variable [TopologicalSpace Y] (f : LocallyConstant X Y) diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index b981f1e70b330..ffa01e46f437f 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -110,9 +110,6 @@ protected theorem congr_arg (f : α →ᵈ β) {x y : α} (h : x = y) : f x = f theorem ext {f g : α →ᵈ β} (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g h -theorem ext_iff {f g : α →ᵈ β} : f = g ↔ ∀ x, f x = g x := - DFunLike.ext_iff - @[simp] theorem mk_coe (f : α →ᵈ β) (h) : Dilation.mk f h = f := ext fun _ => rfl diff --git a/Mathlib/Topology/Metrizable/Uniformity.lean b/Mathlib/Topology/Metrizable/Uniformity.lean index 78c9aade74fc3..84aa5a4ff45ae 100644 --- a/Mathlib/Topology/Metrizable/Uniformity.lean +++ b/Mathlib/Topology/Metrizable/Uniformity.lean @@ -61,7 +61,7 @@ noncomputable def ofPreNNDist (d : X → X → ℝ≥0) (dist_self : ∀ x, d x dist_comm x y := NNReal.coe_inj.2 <| by refine reverse_surjective.iInf_congr _ fun l ↦ ?_ - rw [← sum_reverse, zipWith_distrib_reverse, reverse_append, reverse_reverse, + rw [← sum_reverse, reverse_zipWith, reverse_append, reverse_reverse, reverse_singleton, singleton_append, reverse_cons, reverse_reverse, zipWith_comm_of_comm _ dist_comm] simp only [length, length_append] @@ -134,7 +134,7 @@ theorem le_two_mul_dist_ofPreNNDist (d : X → X → ℝ≥0) (dist_self : ∀ x intro m hm rw [← not_lt, Nat.lt_iff_add_one_le, ← hL_len] intro hLm - rw [mem_setOf_eq, take_all_of_le hLm, two_mul, add_le_iff_nonpos_left, nonpos_iff_eq_zero, + rw [mem_setOf_eq, take_of_length_le hLm, two_mul, add_le_iff_nonpos_left, nonpos_iff_eq_zero, sum_eq_zero_iff, ← forall_iff_forall_mem, forall_zipWith, ← chain_append_singleton_iff_forall₂] at hm <;> @@ -154,7 +154,7 @@ theorem le_two_mul_dist_ofPreNNDist (d : X → X → ℝ≥0) (dist_self : ∀ x have hMl' : length (take M l) = M := (length_take _ _).trans (min_eq_left hMl.le) refine (ihn _ hMl _ _ _ hMl').trans ?_ convert hMs.1.out - rw [zipWith_distrib_take, take, take_succ, getElem?_append hMl, getElem?_eq_getElem hMl, + rw [take_zipWith, take, take_succ, getElem?_append hMl, getElem?_eq_getElem hMl, ← Option.coe_def, Option.toList_some, take_append_of_le_length hMl.le, getElem_cons_succ] · exact single_le_sum (fun x _ => zero_le x) _ (mem_iff_get.2 ⟨⟨M, hM_lt⟩, getElem_zipWith⟩) · rcases hMl.eq_or_lt with (rfl | hMl) @@ -169,7 +169,7 @@ theorem le_two_mul_dist_ofPreNNDist (d : X → X → ℝ≥0) (dist_self : ∀ x rw [← sum_take_add_sum_drop L (M + 1), two_mul, add_le_add_iff_left, ← add_le_add_iff_right, sum_take_add_sum_drop, ← two_mul] at hMs' convert hMs' - rwa [zipWith_distrib_drop, drop, drop_append_of_le_length] + rwa [drop_zipWith, drop, drop_append_of_le_length] end PseudoMetricSpace diff --git a/Mathlib/Topology/PartialHomeomorph.lean b/Mathlib/Topology/PartialHomeomorph.lean index 9487aea43ddeb..ee0f1334ee46f 100644 --- a/Mathlib/Topology/PartialHomeomorph.lean +++ b/Mathlib/Topology/PartialHomeomorph.lean @@ -289,12 +289,6 @@ protected theorem ext (e' : PartialHomeomorph X Y) (h : ∀ x, e x = e' x) (hinv : ∀ x, e.symm x = e'.symm x) (hs : e.source = e'.source) : e = e' := toPartialEquiv_injective (PartialEquiv.ext h hinv hs) -protected theorem ext_iff {e e' : PartialHomeomorph X Y} : - e = e' ↔ (∀ x, e x = e' x) ∧ (∀ x, e.symm x = e'.symm x) ∧ e.source = e'.source := - ⟨by - rintro rfl - exact ⟨fun x => rfl, fun x => rfl, rfl⟩, fun h => e.ext e' h.1 h.2.1 h.2.2⟩ - @[simp, mfld_simps] theorem symm_toPartialEquiv : e.symm.toPartialEquiv = e.toPartialEquiv.symm := rfl diff --git a/Mathlib/Topology/Sheaves/Sheaf.lean b/Mathlib/Topology/Sheaves/Sheaf.lean index 35de899a80af7..c13307d1eeab1 100644 --- a/Mathlib/Topology/Sheaves/Sheaf.lean +++ b/Mathlib/Topology/Sheaves/Sheaf.lean @@ -133,7 +133,7 @@ instance forget_full : (forget C X).Full where -- Porting note: `deriving Faithful` failed instance forgetFaithful : (forget C X).Faithful where - map_injective := Sheaf.Hom.ext _ _ + map_injective := Sheaf.Hom.ext -- Note: These can be proved by simp. theorem id_app (F : Sheaf C X) (t) : (𝟙 F : F ⟶ F).1.app t = 𝟙 _ := diff --git a/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean b/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean index 6963f6e1655c0..650373c48901f 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/OpensLeCover.lean @@ -156,7 +156,7 @@ def generateEquivalenceOpensLe : (by rintro ⟨⟨_, _⟩, _⟩; dsimp; congr) (by intros; refine Over.OverMorphism.ext ?_; aesop_cat) counitIso := eqToIso <| CategoryTheory.Functor.hext - (by intro; refine FullSubcategory.ext _ _ ?_; rfl) (by intros; rfl) + (by intro; refine FullSubcategory.ext ?_; rfl) (by intros; rfl) /-- Given a family of opens `opensLeCoverCocone U` is essentially the natural cocone associated to the sieve generated by the presieve associated to `U` with indexing diff --git a/Mathlib/Topology/Sheaves/Skyscraper.lean b/Mathlib/Topology/Sheaves/Skyscraper.lean index 1112729eaf30e..542289c1577d7 100644 --- a/Mathlib/Topology/Sheaves/Skyscraper.lean +++ b/Mathlib/Topology/Sheaves/Skyscraper.lean @@ -231,8 +231,8 @@ sending every `f : a ⟶ b` to the natural transformation `α` defined as: `α(U def skyscraperSheafFunctor : C ⥤ Sheaf C X where obj c := skyscraperSheaf p₀ c map f := Sheaf.Hom.mk <| (skyscraperPresheafFunctor p₀).map f - map_id _ := Sheaf.Hom.ext _ _ <| (skyscraperPresheafFunctor p₀).map_id _ - map_comp _ _ := Sheaf.Hom.ext _ _ <| (skyscraperPresheafFunctor p₀).map_comp _ _ + map_id _ := Sheaf.Hom.ext <| (skyscraperPresheafFunctor p₀).map_id _ + map_comp _ _ := Sheaf.Hom.ext <| (skyscraperPresheafFunctor p₀).map_comp _ _ namespace StalkSkyscraperPresheafAdjunctionAuxs @@ -374,13 +374,13 @@ def stalkSkyscraperSheafAdjunction [HasColimits C] : -- Porting note (#11041): `ext1` is changed to `Sheaf.Hom.ext`, homEquiv 𝓕 c := ⟨fun f => ⟨toSkyscraperPresheaf p₀ f⟩, fun g => fromStalk p₀ g.1, fromStalk_to_skyscraper p₀, - fun g => Sheaf.Hom.ext _ _ <| to_skyscraper_fromStalk _ _⟩ + fun g => Sheaf.Hom.ext <| to_skyscraper_fromStalk _ _⟩ unit := { app := fun 𝓕 => ⟨(StalkSkyscraperPresheafAdjunctionAuxs.unit p₀).app 𝓕.1⟩ - naturality := fun 𝓐 𝓑 f => Sheaf.Hom.ext _ _ <| by + naturality := fun 𝓐 𝓑 f => Sheaf.Hom.ext <| by apply (StalkSkyscraperPresheafAdjunctionAuxs.unit p₀).naturality } counit := StalkSkyscraperPresheafAdjunctionAuxs.counit p₀ - homEquiv_unit {𝓐} c f := Sheaf.Hom.ext _ _ (skyscraperPresheafStalkAdjunction p₀).homEquiv_unit + homEquiv_unit {𝓐} c f := Sheaf.Hom.ext (skyscraperPresheafStalkAdjunction p₀).homEquiv_unit homEquiv_counit {𝓐} c f := (skyscraperPresheafStalkAdjunction p₀).homEquiv_counit instance [HasColimits C] : (skyscraperSheafFunctor p₀ : C ⥤ Sheaf C X).IsRightAdjoint := diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index 2d988c29214b8..144ab42061543 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -79,7 +79,7 @@ instance : PartialOrder (PartialRefinement u s) where ⟨Subset.trans h₁₂.1 h₂₃.1, fun i hi => (h₁₂.2 i hi).trans (h₂₃.2 i <| h₁₂.1 hi)⟩ le_antisymm v₁ v₂ h₁₂ h₂₁ := have hc : v₁.carrier = v₂.carrier := Subset.antisymm h₁₂.1 h₂₁.1 - PartialRefinement.ext _ _ + PartialRefinement.ext (funext fun x => if hx : x ∈ v₁.carrier then h₁₂.2 _ hx else (v₁.apply_eq hx).trans (Eq.symm <| v₂.apply_eq <| hc ▸ hx)) diff --git a/Mathlib/Topology/UniformSpace/Basic.lean b/Mathlib/Topology/UniformSpace/Basic.lean index 9351d243a910f..37b930b9377e6 100644 --- a/Mathlib/Topology/UniformSpace/Basic.lean +++ b/Mathlib/Topology/UniformSpace/Basic.lean @@ -340,7 +340,7 @@ def UniformSpace.ofNhdsEqComap (u : UniformSpace.Core α) (_t : TopologicalSpace __ := u nhds_eq_comap_uniformity := h -@[ext] +@[ext (iff := false)] protected theorem UniformSpace.ext {u₁ u₂ : UniformSpace α} (h : 𝓤[u₁] = 𝓤[u₂]) : u₁ = u₂ := by have : u₁.toTopologicalSpace = u₂.toTopologicalSpace := TopologicalSpace.ext_nhds fun x ↦ by rw [u₁.nhds_eq_comap_uniformity, u₂.nhds_eq_comap_uniformity] diff --git a/Mathlib/Util/AssertNoSorry.lean b/Mathlib/Util/AssertNoSorry.lean index 49987ad8dade9..e29f256c1db8b 100644 --- a/Mathlib/Util/AssertNoSorry.lean +++ b/Mathlib/Util/AssertNoSorry.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Renshaw -/ -import Lean.Elab.Print -import Lean.Environment +import Lean.Util.CollectAxioms +import Lean.Elab.Command /-! # Defines the `assert_no_sorry` command. @@ -13,12 +13,11 @@ import Lean.Environment Throws an error if the given identifier uses sorryAx. -/ -open Lean.Elab.Command +open Lean Meta Elab Command /-- Throws an error if the given identifier uses sorryAx. -/ elab "assert_no_sorry " n:ident : command => do - let env ← Lean.getEnv - let (_, s) := ((Lean.Elab.Command.CollectAxioms.collect - (← liftCoreM <| Lean.Elab.realizeGlobalConstNoOverloadWithInfo n)).run env).run {} - if s.axioms.contains ``sorryAx + let name ← liftCoreM <| Lean.Elab.realizeGlobalConstNoOverloadWithInfo n + let axioms ← Lean.collectAxioms name + if axioms.contains ``sorryAx then throwError "{n} contains sorry" diff --git a/Mathlib/Util/LongNames.lean b/Mathlib/Util/LongNames.lean index 7c1846f052e50..1551a5e5ae21e 100644 --- a/Mathlib/Util/LongNames.lean +++ b/Mathlib/Util/LongNames.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Batteries.Data.HashMap.WF import Mathlib.Lean.Name import Mathlib.Lean.Expr.Basic @@ -16,7 +15,7 @@ For finding declarations with excessively long names. open Lean Meta Elab /-- Helper function for `#long_names` and `#long_instances`. -/ -def printNameHashMap (h : Batteries.HashMap Name (Array Name)) : IO Unit := +def printNameHashMap (h : Std.HashMap Name (Array Name)) : IO Unit := for (m, names) in h.toList do IO.println "----" IO.println <| m.toString ++ ":" diff --git a/lake-manifest.json b/lake-manifest.json index d7ba94b86d2ff..a00fdb8a546f2 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "41bc768e2224d6c75128a877f1d6e198859b3178", + "rev": "dc167d260ff7ee9849b436037add06bed15104be", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "6058ab8d938c5104eace7d0fb5ac17b21cb067b1", + "rev": "ae6ea60e9d8bc2d4b37ff02115854da2e1b710d0", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -35,10 +35,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "c87908619cccadda23f71262e6898b9893bffa36", + "rev": "a96aee5245720f588876021b6a0aa73efee49c76", "name": "proofwidgets", "manifestFile": "lake-manifest.json", - "inputRev": "v0.0.40", + "inputRev": "v0.0.41", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover/lean4-cli", @@ -55,7 +55,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "e7e90d90a62e6d12cbb27cbbfc31c094ee4ecc58", + "rev": "68cd8ae0f5b996176d1243d94c56e17de570e3bf", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/lakefile.lean b/lakefile.lean index d6b5b580297a3..a4540d64f35f6 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -18,11 +18,11 @@ package mathlib where ## Mathlib dependencies on upstream projects. -/ -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.40" -require "leanprover-community" / "importGraph" @ "git#main" +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.41" +require "leanprover-community" / "importGraph" @ git "main" /-! ## Mathlib libraries diff --git a/lean-toolchain b/lean-toolchain index 7f0ea50aaa500..64981ae5f58db 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.10.0 +leanprover/lean4:v4.11.0-rc1 diff --git a/scripts/check-yaml.lean b/scripts/check-yaml.lean index 374010edebbe0..23730a3fabb2e 100644 --- a/scripts/check-yaml.lean +++ b/scripts/check-yaml.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Batteries.Lean.Util.Path import Mathlib.Lean.CoreM import Mathlib.Tactic.ToExpr diff --git a/scripts/noshake.json b/scripts/noshake.json index f5e00b5c06202..a4ef16aeb038b 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -313,7 +313,7 @@ "Mathlib.Lean.Expr.ExtraRecognizers": ["Mathlib.Data.Set.Defs"], "Mathlib.Lean.Expr.Basic": ["Batteries.Logic"], "Mathlib.Init.Data.Nat.Lemmas": ["Batteries.Data.Nat.Lemmas", "Batteries.WF"], - "Mathlib.Init.Data.Nat.GCD": ["Batteries.Data.Nat.Gcd"], + "Mathlib.Init.Data.List.Lemmas": ["Batteries.Data.List.Lemmas"], "Mathlib.Init.Data.List.Basic": ["Batteries.Data.List.Basic"], "Mathlib.Init.Data.Int.Order": ["Mathlib.Data.Int.Notation"], "Mathlib.Init.Data.Int.Basic": ["Batteries.Data.Int.Order"], diff --git a/test/DeriveToExpr.lean b/test/DeriveToExpr.lean index 4dca8978142a4..7780eef44e28c 100644 --- a/test/DeriveToExpr.lean +++ b/test/DeriveToExpr.lean @@ -40,7 +40,7 @@ run_cmd Elab.Command.liftTermElabM <| /-- error: failed to synthesize ToExpr (Bool → Nat) -use `set_option diagnostics true` to get diagnostic information +Additional diagnostic information may be available using the `set_option diagnostics true` command. -/ #guard_msgs in inductive Bar diff --git a/test/TCSynth.lean b/test/TCSynth.lean index 8e32f2f8c1cf7..dd2933a7cfc96 100644 --- a/test/TCSynth.lean +++ b/test/TCSynth.lean @@ -22,7 +22,7 @@ open Complex Filter Bornology /-- error: failed to synthesize AddMonoidHomClass (AddGroupSeminorm ℂ) ℂ ℝ -use `set_option diagnostics true` to get diagnostic information +Additional diagnostic information may be available using the `set_option diagnostics true` command. -/ #guard_msgs in set_option synthInstance.maxHeartbeats 3000 in diff --git a/test/Use.lean b/test/Use.lean index 7c9e466b84ac1..9dd4b75d1c145 100644 --- a/test/Use.lean +++ b/test/Use.lean @@ -38,7 +38,7 @@ error: failed to synthesize numerals are polymorphic in Lean, but the numeral `42` cannot be used in a context where the expected type is Nat × Nat due to the absence of the instance above -use `set_option diagnostics true` to get diagnostic information +Additional diagnostic information may be available using the `set_option diagnostics true` command. -/ #guard_msgs in example : ∃ p : Nat × Nat, p.1 = p.2 := by use 42; sorry @@ -94,7 +94,7 @@ error: failed to synthesize numerals are polymorphic in Lean, but the numeral `1` cannot be used in a context where the expected type is Option Nat due to the absence of the instance above -use `set_option diagnostics true` to get diagnostic information +Additional diagnostic information may be available using the `set_option diagnostics true` command. -/ #guard_msgs in example : Option Nat := by use 1 @@ -105,7 +105,7 @@ error: failed to synthesize numerals are polymorphic in Lean, but the numeral `1` cannot be used in a context where the expected type is Nat → Nat due to the absence of the instance above -use `set_option diagnostics true` to get diagnostic information +Additional diagnostic information may be available using the `set_option diagnostics true` command. -/ #guard_msgs in example : Nat → Nat := by use 1 diff --git a/test/Variable.lean b/test/Variable.lean index ee7dc07aa0dd3..327d36efd228e 100644 --- a/test/Variable.lean +++ b/test/Variable.lean @@ -54,7 +54,7 @@ section /-- error: failed to synthesize Semiring R -use `set_option diagnostics true` to get diagnostic information +Additional diagnostic information may be available using the `set_option diagnostics true` command. -/ #guard_msgs in variable? [Module R M] => [Module R M] diff --git a/test/eval_elab.lean b/test/eval_elab.lean index fc7c63455c8c4..7e929bdda3698 100644 --- a/test/eval_elab.lean +++ b/test/eval_elab.lean @@ -12,7 +12,7 @@ section from_zulip /-- error: failed to synthesize Lean.ToExpr (Finset (Finset ℕ)) -use `set_option diagnostics true` to get diagnostic information +Additional diagnostic information may be available using the `set_option diagnostics true` command. -/ #guard_msgs in #check eval% Finset.powerset ({1, 2, 3} : Finset ℕ) diff --git a/test/jacobiSym.lean b/test/jacobiSym.lean index bfc8b7c9deff4..916deb2aa2d2d 100644 --- a/test/jacobiSym.lean +++ b/test/jacobiSym.lean @@ -43,6 +43,6 @@ warning: declaration uses 'sorry' info: 1 -/ #guard_msgs in -#eval @legendreSym (2 ^ 11213 - 1) sorry 7 +#eval! @legendreSym (2 ^ 11213 - 1) sorry 7 end Csimp From 4d93d88d127510415f85af6aa74cc8ad7a10f33d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 5 Aug 2024 10:13:01 +0000 Subject: [PATCH 034/359] refactor: use consistent names for forgetful smul structures (#15500) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Moves: - `AddCommMonoid.natModule` -> `AddCommMonoid.toNatModule` - `AddCommGroup.intModule` -> `AddCommGroup.toIntModule` - `algebraNat` -> `Semiring.toNatAlgebra` - `algebraInt` -> `Ring.toIntAlgebra` - `algebraRat` -> `DivisionRing.toRatAlgebra` - `AddCommGroup.natModule.unique` -> `AddCommMonoid.uniqueNatModule` - `AddCommGroup.intModule.unique` -> `AddCommGroup.uniqueIntModule` The premise here is that the conversion from `AddCommGroup` to `AddCommMonoid` is not really any different from converting from `AddCommGroup` to `Module ℤ`; and so both should be named with the `to*` naming convention. Tihs also brings the naming of the `Module` and `Algebra` instances into alignment. I chose `toNatModule` instead of `toModuleNat` to match the prose, ℕ-module. This fits a common naming convention if you view `Module` as infix, as we do with various other lean functions without actual infix notation. --- Counterexamples/Pseudoelement.lean | 2 +- Mathlib/Algebra/Algebra/Basic.lean | 12 ++++++------ Mathlib/Algebra/Algebra/Rat.lean | 5 +++-- Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean | 2 +- Mathlib/Algebra/CharP/Algebra.lean | 8 ++++---- Mathlib/Algebra/FreeAlgebra.lean | 4 ++-- Mathlib/Algebra/Module/Defs.lean | 8 ++++---- Mathlib/Algebra/Module/LinearMap/Defs.lean | 2 +- Mathlib/Data/ZMod/Module.lean | 2 +- Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean | 4 ++-- Mathlib/LinearAlgebra/TensorAlgebra/Basic.lean | 4 ++-- Mathlib/LinearAlgebra/TensorProduct/Basic.lean | 4 ++-- Mathlib/NumberTheory/Cyclotomic/Basic.lean | 4 ++-- Mathlib/NumberTheory/NumberField/Basic.lean | 2 +- Mathlib/RingTheory/GradedAlgebra/Basic.lean | 3 ++- Mathlib/RingTheory/OreLocalization/Ring.lean | 4 ++-- Mathlib/RingTheory/TensorProduct/Basic.lean | 4 ++-- Mathlib/Topology/Algebra/Algebra/Rat.lean | 2 +- test/instance_diamonds.lean | 8 ++++---- .../FieldTheory/IsAlgClosed/AlgebraicClosure.lean | 8 ++++---- .../FieldTheory/SplittingField/Construction.lean | 8 ++++---- 21 files changed, 51 insertions(+), 49 deletions(-) diff --git a/Counterexamples/Pseudoelement.lean b/Counterexamples/Pseudoelement.lean index 00b5270c53f4b..89b82c0d77a28 100644 --- a/Counterexamples/Pseudoelement.lean +++ b/Counterexamples/Pseudoelement.lean @@ -67,7 +67,7 @@ theorem snd_x_pseudo_eq_snd_y : PseudoEqual _ (app biprod.snd x) (app biprod.snd simp_rw [biprod.lift_snd]; rfl -- Porting note: locally disable instance to avoid inferred/synthesized clash -attribute [-instance] AddCommGroup.intModule in +attribute [-instance] AddCommGroup.toIntModule in /-- `x` is not pseudoequal to `y`. -/ theorem x_not_pseudo_eq : ¬PseudoEqual _ x y := by intro h diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 7cba7ba8d66c3..8e377b67365e8 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -223,10 +223,10 @@ section Nat variable {R : Type*} [Semiring R] -- Lower the priority so that `Algebra.id` is picked most of the time when working with --- `ℕ`-algebras. This is only an issue since `Algebra.id` and `algebraNat` are not yet defeq. --- TODO: fix this by adding an `ofNat` field to semirings. +-- `ℕ`-algebras. +-- TODO: is this still needed? /-- Semiring ⥤ ℕ-Alg -/ -instance (priority := 99) algebraNat : Algebra ℕ R where +instance (priority := 99) Semiring.toNatAlgebra : Algebra ℕ R where commutes' := Nat.cast_commute smul_def' _ _ := nsmul_eq_mul _ _ toRingHom := Nat.castRingHom R @@ -241,10 +241,10 @@ section Int variable (R : Type*) [Ring R] -- Lower the priority so that `Algebra.id` is picked most of the time when working with --- `ℤ`-algebras. This is only an issue since `Algebra.id ℤ` and `algebraInt ℤ` are not yet defeq. --- TODO: fix this by adding an `ofInt` field to rings. +-- `ℤ`-algebras. +-- TODO: is this still needed? /-- Ring ⥤ ℤ-Alg -/ -instance (priority := 99) algebraInt : Algebra ℤ R where +instance (priority := 99) Ring.toIntAlgebra : Algebra ℤ R where commutes' := Int.cast_commute smul_def' _ _ := zsmul_eq_mul _ _ toRingHom := Int.castRingHom R diff --git a/Mathlib/Algebra/Algebra/Rat.lean b/Mathlib/Algebra/Algebra/Rat.lean index 6cb0d03ee0227..6ad708e5d3439 100644 --- a/Mathlib/Algebra/Algebra/Rat.lean +++ b/Mathlib/Algebra/Algebra/Rat.lean @@ -27,7 +27,8 @@ end RingHom section Rat -instance algebraRat {α} [DivisionRing α] [CharZero α] : Algebra ℚ α where +/-- Every division ring of characteristic zero is an algebra over the rationals. -/ +instance DivisionRing.toRatAlgebra {α} [DivisionRing α] [CharZero α] : Algebra ℚ α where smul := (· • ·) smul_def' := Rat.smul_def toRingHom := Rat.castHom α @@ -38,7 +39,7 @@ instance : Algebra NNRat ℚ := NNRat.coeHom.toAlgebra /-- The two `Algebra ℚ ℚ` instances should coincide. -/ -example : algebraRat = Algebra.id ℚ := +example : DivisionRing.toRatAlgebra = Algebra.id ℚ := rfl @[simp] theorem algebraMap_rat_rat : algebraMap ℚ ℚ = RingHom.id ℚ := rfl diff --git a/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean b/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean index 82204010ed78a..fd9a529d37985 100644 --- a/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean +++ b/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean @@ -27,7 +27,7 @@ namespace ModuleCat instance forget₂_addCommGroup_full : (forget₂ (ModuleCat ℤ) AddCommGrp.{u}).Full where map_surjective {A B} -- `AddMonoidHom.toIntLinearMap` doesn't work here because `A` and `B` are not - -- definitionally equal to the canonical `AddCommGroup.intModule` module + -- definitionally equal to the canonical `AddCommGroup.toIntModule` module -- instances it expects. f := ⟨@LinearMap.mk _ _ _ _ _ _ _ _ _ A.isModule B.isModule { toFun := f, diff --git a/Mathlib/Algebra/CharP/Algebra.lean b/Mathlib/Algebra/CharP/Algebra.lean index 626c9e01db0d4..05641583b0b03 100644 --- a/Mathlib/Algebra/CharP/Algebra.lean +++ b/Mathlib/Algebra/CharP/Algebra.lean @@ -81,8 +81,8 @@ variable (R : Type*) [Nontrivial R] /-- A nontrivial `ℚ`-algebra has `CharP` equal to zero. This cannot be a (local) instance because it would immediately form a loop with the -instance `algebraRat`. It's probably easier to go the other way: prove `CharZero R` and -automatically receive an `Algebra ℚ R` instance. +instance `DivisionRing.toRatAlgebra`. It's probably easier to go the other way: prove `CharZero R` +and automatically receive an `Algebra ℚ R` instance. -/ theorem algebraRat.charP_zero [Semiring R] [Algebra ℚ R] : CharP R 0 := charP_of_injective_algebraMap (algebraMap ℚ R).injective 0 @@ -90,8 +90,8 @@ theorem algebraRat.charP_zero [Semiring R] [Algebra ℚ R] : CharP R 0 := /-- A nontrivial `ℚ`-algebra has characteristic zero. This cannot be a (local) instance because it would immediately form a loop with the -instance `algebraRat`. It's probably easier to go the other way: prove `CharZero R` and -automatically receive an `Algebra ℚ R` instance. +instance `DivisionRing.toRatAlgebra`. It's probably easier to go the other way: prove `CharZero R` +and automatically receive an `Algebra ℚ R` instance. -/ theorem algebraRat.charZero [Ring R] [Algebra ℚ R] : CharZero R := @CharP.charP_to_charZero R _ (algebraRat.charP_zero R) diff --git a/Mathlib/Algebra/FreeAlgebra.lean b/Mathlib/Algebra/FreeAlgebra.lean index 1a8ca31f56a0f..9c12c522bad65 100644 --- a/Mathlib/Algebra/FreeAlgebra.lean +++ b/Mathlib/Algebra/FreeAlgebra.lean @@ -255,7 +255,7 @@ instance instAlgebra {A} [CommSemiring A] [Algebra R A] : Algebra R (FreeAlgebra -- verify there is no diamond at `default` transparency but we will need -- `reducible_and_instances` which currently fails #10906 variable (S : Type) [CommSemiring S] in -example : (algebraNat : Algebra ℕ (FreeAlgebra S X)) = instAlgebra _ _ := rfl +example : (Semiring.toNatAlgebra : Algebra ℕ (FreeAlgebra S X)) = instAlgebra _ _ := rfl instance {R S A} [CommSemiring R] [CommSemiring S] [CommSemiring A] [SMul R S] [Algebra R A] [Algebra S A] [IsScalarTower R S A] : @@ -277,7 +277,7 @@ instance {S : Type*} [CommRing S] : Ring (FreeAlgebra S X) := -- verify there is no diamond but we will need -- `reducible_and_instances` which currently fails #10906 variable (S : Type) [CommRing S] in -example : (algebraInt _ : Algebra ℤ (FreeAlgebra S X)) = instAlgebra _ _ := rfl +example : (Ring.toIntAlgebra _ : Algebra ℤ (FreeAlgebra S X)) = instAlgebra _ _ := rfl variable {X} diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index 22eea21dc9935..db4b297d90ae7 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -72,7 +72,7 @@ instance (priority := 100) Module.toMulActionWithZero : MulActionWithZero R M := smul_zero := smul_zero zero_smul := Module.zero_smul } -instance AddCommMonoid.natModule : Module ℕ M where +instance AddCommGroup.toNatModule : Module ℕ M where one_smul := one_nsmul mul_smul m n a := mul_nsmul' a m n smul_add n a b := nsmul_add a b n @@ -190,7 +190,7 @@ section AddCommGroup variable (R M) [Semiring R] [AddCommGroup M] -instance AddCommGroup.intModule : Module ℤ M where +instance AddCommGroup.toIntModule : Module ℤ M where one_smul := one_zsmul mul_smul m n a := mul_zsmul a m n smul_add n a b := zsmul_add a b n @@ -350,7 +350,7 @@ theorem nat_smul_eq_nsmul (h : Module ℕ M) (n : ℕ) (x : M) : @SMul.smul ℕ /-- All `ℕ`-module structures are equal. Not an instance since in mathlib all `AddCommMonoid` should normally have exactly one `ℕ`-module structure by design. -/ -def AddCommMonoid.natModule.unique : Unique (Module ℕ M) where +def AddCommMonoid.uniqueNatModule : Unique (Module ℕ M) where default := by infer_instance uniq P := (Module.ext' P _) fun n => by convert nat_smul_eq_nsmul P n @@ -391,7 +391,7 @@ theorem int_smul_eq_zsmul (h : Module ℤ M) (n : ℤ) (x : M) : @SMul.smul ℤ /-- All `ℤ`-module structures are equal. Not an instance since in mathlib all `AddCommGroup` should normally have exactly one `ℤ`-module structure by design. -/ -def AddCommGroup.intModule.unique : Unique (Module ℤ M) where +def AddCommGroup.uniqueIntModule : Unique (Module ℤ M) where default := by infer_instance uniq P := (Module.ext' P _) fun n => by convert int_smul_eq_zsmul P n diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index 1de087baccedf..4549f739077ee 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -356,7 +356,7 @@ variable (M M₂) /-- A typeclass for `SMul` structures which can be moved through a `LinearMap`. This typeclass is generated automatically from an `IsScalarTower` instance, but exists so that -we can also add an instance for `AddCommGroup.intModule`, allowing `z •` to be moved even if +we can also add an instance for `AddCommGroup.toIntModule`, allowing `z •` to be moved even if `S` does not support negation. -/ class CompatibleSMul (R S : Type*) [Semiring S] [SMul R M] [Module S M] [SMul R M₂] diff --git a/Mathlib/Data/ZMod/Module.lean b/Mathlib/Data/ZMod/Module.lean index 1697e15c69a12..aebc50352620e 100644 --- a/Mathlib/Data/ZMod/Module.lean +++ b/Mathlib/Data/ZMod/Module.lean @@ -38,7 +38,7 @@ See note [reducible non-instances]. -/ abbrev AddCommGroup.zmodModule {G : Type*} [AddCommGroup G] (h : ∀ (x : G), n • x = 0) : Module (ZMod n) G := match n with - | 0 => AddCommGroup.intModule G + | 0 => AddCommGroup.toIntModule G | _ + 1 => AddCommMonoid.zmodModule h variable {F S : Type*} [AddCommGroup M] [AddCommGroup M₁] [FunLike F M M₁] diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean index 2501f13896783..ac24ee710f6f0 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Basic.lean @@ -80,9 +80,9 @@ instance (priority := 900) instAlgebra' {R A M} [CommSemiring R] [AddCommGroup M -- verify there are no diamonds -- but doesn't work at `reducible_and_instances` #10906 -example : (algebraNat : Algebra ℕ (CliffordAlgebra Q)) = instAlgebra' _ := rfl +example : (Semiring.toNatAlgebra : Algebra ℕ (CliffordAlgebra Q)) = instAlgebra' _ := rfl -- but doesn't work at `reducible_and_instances` #10906 -example : (algebraInt _ : Algebra ℤ (CliffordAlgebra Q)) = instAlgebra' _ := rfl +example : (Ring.toIntAlgebra _ : Algebra ℤ (CliffordAlgebra Q)) = instAlgebra' _ := rfl -- shortcut instance, as the other instance is slow instance instAlgebra : Algebra R (CliffordAlgebra Q) := instAlgebra' _ diff --git a/Mathlib/LinearAlgebra/TensorAlgebra/Basic.lean b/Mathlib/LinearAlgebra/TensorAlgebra/Basic.lean index 9333633f06058..bfa9c8e59147a 100644 --- a/Mathlib/LinearAlgebra/TensorAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorAlgebra/Basic.lean @@ -73,7 +73,7 @@ instance instAlgebra {R A M} [CommSemiring R] [AddCommMonoid M] [CommSemiring A] -- verify there is no diamond -- but doesn't work at `reducible_and_instances` #10906 -example : (algebraNat : Algebra ℕ (TensorAlgebra R M)) = instAlgebra := rfl +example : (Semiring.toNatAlgebra : Algebra ℕ (TensorAlgebra R M)) = instAlgebra := rfl instance {R S A M} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [CommSemiring A] [Algebra R A] [Algebra S A] [Module R M] [Module S M] [Module A M] @@ -95,7 +95,7 @@ instance {S : Type*} [CommRing S] [Module S M] : Ring (TensorAlgebra S M) := -- verify there is no diamond -- but doesn't work at `reducible_and_instances` #10906 variable (S M : Type) [CommRing S] [AddCommGroup M] [Module S M] in -example : (algebraInt _ : Algebra ℤ (TensorAlgebra S M)) = instAlgebra := rfl +example : (Ring.toIntAlgebra _ : Algebra ℤ (TensorAlgebra S M)) = instAlgebra := rfl variable {M} diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index 393c10b371f81..0dbf00d45ee83 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -216,7 +216,7 @@ variable (R R' M N) /-- A typeclass for `SMul` structures which can be moved across a tensor product. This typeclass is generated automatically from an `IsScalarTower` instance, but exists so that -we can also add an instance for `AddCommGroup.intModule`, allowing `z •` to be moved even if +we can also add an instance for `AddCommGroup.toIntModule`, allowing `z •` to be moved even if `R` does not support negation. Note that `Module R' (M ⊗[R] N)` is available even without this typeclass on `R'`; it's only @@ -1441,7 +1441,7 @@ theorem sub_tmul (m₁ m₂ : M) (n : N) : (m₁ - m₂) ⊗ₜ n = m₁ ⊗ₜ[ (mk R M N).map_sub₂ _ _ _ /-- While the tensor product will automatically inherit a ℤ-module structure from -`AddCommGroup.intModule`, that structure won't be compatible with lemmas like `tmul_smul` unless +`AddCommGroup.toIntModule`, that structure won't be compatible with lemmas like `tmul_smul` unless we use a `ℤ-Module` instance provided by `TensorProduct.left_module`. When `R` is a `Ring` we get the required `TensorProduct.compatible_smul` instance through diff --git a/Mathlib/NumberTheory/Cyclotomic/Basic.lean b/Mathlib/NumberTheory/Cyclotomic/Basic.lean index 08a0a83d2f9f7..4027bbcc9a884 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Basic.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Basic.lean @@ -532,7 +532,7 @@ instance CyclotomicField.algebraBase : Algebra A (CyclotomicField n K) := SplittingField.algebra' (cyclotomic n K) /-- Ensure there are no diamonds when `A = ℤ` but there are `reducible_and_instances` #10906 -/ -example : algebraInt (CyclotomicField n ℚ) = CyclotomicField.algebraBase _ _ _ := rfl +example : Ring.toIntAlgebra (CyclotomicField n ℚ) = CyclotomicField.algebraBase _ _ _ := rfl instance CyclotomicField.algebra' {R : Type*} [CommRing R] [Algebra R K] : Algebra R (CyclotomicField n K) := @@ -576,7 +576,7 @@ instance algebraBase : Algebra A (CyclotomicRing n A K) := -- Ensure that there is no diamonds with ℤ. -- but there is at `reducible_and_instances` #10906 -example {n : ℕ+} : CyclotomicRing.algebraBase n ℤ ℚ = algebraInt _ := rfl +example {n : ℕ+} : CyclotomicRing.algebraBase n ℤ ℚ = Ring.toIntAlgebra _ := rfl instance : NoZeroSMulDivisors A (CyclotomicRing n A K) := (adjoin A _).noZeroSMulDivisors_bot diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index c917742ef44af..e1edb1734ec41 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -300,7 +300,7 @@ section open scoped Polynomial -attribute [-instance] algebraRat +attribute [-instance] DivisionRing.toRatAlgebra /-- The quotient of `ℚ[X]` by the ideal generated by an irreducible polynomial of `ℚ[X]` is a number field. -/ diff --git a/Mathlib/RingTheory/GradedAlgebra/Basic.lean b/Mathlib/RingTheory/GradedAlgebra/Basic.lean index 11e8d628a73b7..8f91b495bb1d4 100644 --- a/Mathlib/RingTheory/GradedAlgebra/Basic.lean +++ b/Mathlib/RingTheory/GradedAlgebra/Basic.lean @@ -31,7 +31,8 @@ See the docstring of that typeclass for more information. For now, we do not have internally-graded semirings and internally-graded rings; these can be represented with `𝒜 : ι → Submodule ℕ A` and `𝒜 : ι → Submodule ℤ A` respectively, since all -`Semiring`s are ℕ-algebras via `algebraNat`, and all `Ring`s are `ℤ`-algebras via `algebraInt`. +`Semiring`s are ℕ-algebras via `Semiring.toNatAlgebra`, and all `Ring`s are `ℤ`-algebras via +`Ring.toIntAlgebra`. ## Tags diff --git a/Mathlib/RingTheory/OreLocalization/Ring.lean b/Mathlib/RingTheory/OreLocalization/Ring.lean index 40e7844ce5caa..3d54c272d5b5d 100644 --- a/Mathlib/RingTheory/OreLocalization/Ring.lean +++ b/Mathlib/RingTheory/OreLocalization/Ring.lean @@ -102,7 +102,7 @@ lemma nsmul_eq_nsmul (n : ℕ) (x : X[S⁻¹]) : letI inst := OreLocalization.instModuleOfIsScalarTower (R₀ := ℕ) (R := R) (X := X) (S := S) HSMul.hSMul (self := @instHSMul _ _ inst.toSMul) n x = n • x := by letI inst := OreLocalization.instModuleOfIsScalarTower (R₀ := ℕ) (R := R) (X := X) (S := S) - exact congr($(AddCommMonoid.natModule.unique.2 inst).smul n x) + exact congr($(AddCommMonoid.uniqueNatModule.2 inst).smul n x) /-- The ring homomorphism from `R` to `R[S⁻¹]`, mapping `r : R` to the fraction `r /ₒ 1`. -/ @[simps!] @@ -188,7 +188,7 @@ lemma zsmul_eq_zsmul (n : ℤ) (x : X[S⁻¹]) : letI inst := OreLocalization.instModuleOfIsScalarTower (R₀ := ℤ) (R := R) (X := X) (S := S) HSMul.hSMul (self := @instHSMul _ _ inst.toSMul) n x = n • x := by letI inst := OreLocalization.instModuleOfIsScalarTower (R₀ := ℤ) (R := R) (X := X) (S := S) - exact congr($(AddCommGroup.intModule.unique.2 inst).smul n x) + exact congr($(AddCommGroup.uniqueIntModule.2 inst).smul n x) open nonZeroDivisors diff --git a/Mathlib/RingTheory/TensorProduct/Basic.lean b/Mathlib/RingTheory/TensorProduct/Basic.lean index d6dde2c19d15d..c061e1411a588 100644 --- a/Mathlib/RingTheory/TensorProduct/Basic.lean +++ b/Mathlib/RingTheory/TensorProduct/Basic.lean @@ -399,7 +399,7 @@ instance leftAlgebra [SMulCommClass R S A] : Algebra S (A ⊗[R] B) := rw [algebraMap_eq_smul_one, ← smul_tmul', smul_mul_assoc, ← one_def, one_mul] toRingHom := TensorProduct.includeLeftRingHom.comp (algebraMap S A) } -example : (algebraNat : Algebra ℕ (ℕ ⊗[ℕ] B)) = leftAlgebra := rfl +example : (Semiring.toNatAlgebra : Algebra ℕ (ℕ ⊗[ℕ] B)) = leftAlgebra := rfl -- This is for the `undergrad.yaml` list. /-- The tensor product of two `R`-algebras is an `R`-algebra. -/ @@ -560,7 +560,7 @@ theorem intCast_def' (z : ℤ) : (z : A ⊗[R] B) = (1 : A) ⊗ₜ (z : B) := by example : (instRing : Ring (A ⊗[R] B)).toAddCommGroup = addCommGroup := by with_reducible_and_instances rfl -- fails at `with_reducible_and_instances rfl` #10906 -example : (algebraInt _ : Algebra ℤ (ℤ ⊗[ℤ] B)) = leftAlgebra := rfl +example : (Ring.toIntAlgebra _ : Algebra ℤ (ℤ ⊗[ℤ] B)) = leftAlgebra := rfl end Ring diff --git a/Mathlib/Topology/Algebra/Algebra/Rat.lean b/Mathlib/Topology/Algebra/Algebra/Rat.lean index ed1fd7b6dd321..d63250277c31e 100644 --- a/Mathlib/Topology/Algebra/Algebra/Rat.lean +++ b/Mathlib/Topology/Algebra/Algebra/Rat.lean @@ -17,7 +17,7 @@ This is just a minimal stub for now! section DivisionRing -/-- The action induced by `algebraRat` is continuous. -/ +/-- The action induced by `DivisionRing.toRatAlgebra` is continuous. -/ instance DivisionRing.continuousConstSMul_rat {A} [DivisionRing A] [TopologicalSpace A] [ContinuousMul A] [CharZero A] : ContinuousConstSMul ℚ A := ⟨fun r => by simpa only [Algebra.smul_def] using continuous_const.mul continuous_id⟩ diff --git a/test/instance_diamonds.lean b/test/instance_diamonds.lean index bd1fad85f4e2b..7b746593fdb05 100644 --- a/test/instance_diamonds.lean +++ b/test/instance_diamonds.lean @@ -208,13 +208,13 @@ example [CommSemiring R] [Nontrivial R] : rfl -- fails `with_reducible_and_instances` #10906 -/-- `Polynomial.algebraOfAlgebra` is consistent with `algebraNat`. -/ -example [Semiring R] : (Polynomial.algebraOfAlgebra : Algebra ℕ R[X]) = algebraNat := +/-- `Polynomial.algebraOfAlgebra` is consistent with `Semiring.toNatAlgebra`. -/ +example [Semiring R] : (Polynomial.algebraOfAlgebra : Algebra ℕ R[X]) = Semiring.toNatAlgebra := rfl -- fails `with_reducible_and_instances` #10906 -/-- `Polynomial.algebraOfAlgebra` is consistent with `algebraInt`. -/ -example [Ring R] : (Polynomial.algebraOfAlgebra : Algebra ℤ R[X]) = algebraInt _ := +/-- `Polynomial.algebraOfAlgebra` is consistent with `Ring.toIntAlgebra`. -/ +example [Ring R] : (Polynomial.algebraOfAlgebra : Algebra ℤ R[X]) = Ring.toIntAlgebra _ := rfl end Polynomial diff --git a/test/instance_diamonds/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/test/instance_diamonds/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index 57978c08fdc3c..dffc922831a81 100644 --- a/test/instance_diamonds/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/test/instance_diamonds/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -3,18 +3,18 @@ import Mathlib.FieldTheory.IsAlgClosed.AlgebraicClosure variable {k : Type*} [Field k] -example : (AddCommMonoid.natModule : Module ℕ (AlgebraicClosure k)) = +example : (AddCommGroup.toNatModule : Module ℕ (AlgebraicClosure k)) = @Algebra.toModule _ _ _ _ (AlgebraicClosure.instAlgebra k) := by with_reducible_and_instances rfl -example : (AddCommGroup.intModule _ : Module ℤ (AlgebraicClosure k)) = +example : (AddCommGroup.toIntModule _ : Module ℤ (AlgebraicClosure k)) = @Algebra.toModule _ _ _ _ (AlgebraicClosure.instAlgebra k) := by with_reducible_and_instances rfl -example [CharZero k] : AlgebraicClosure.instAlgebra k = algebraRat := +example [CharZero k] : AlgebraicClosure.instAlgebra k = DivisionRing.toRatAlgebra := rfl -- TODO: by with_reducible_and_instances rfl fails -example : algebraInt (AlgebraicClosure ℚ) = AlgebraicClosure.instAlgebra ℚ := +example : Ring.toIntAlgebra (AlgebraicClosure ℚ) = AlgebraicClosure.instAlgebra ℚ := rfl -- TODO: by with_reducible_and_instances rfl fails diff --git a/test/instance_diamonds/FieldTheory/SplittingField/Construction.lean b/test/instance_diamonds/FieldTheory/SplittingField/Construction.lean index fc26ca3b8fbe4..c3ec562788ada 100644 --- a/test/instance_diamonds/FieldTheory/SplittingField/Construction.lean +++ b/test/instance_diamonds/FieldTheory/SplittingField/Construction.lean @@ -14,19 +14,19 @@ variable (f : K[X]) -- The algebra instance deriving from `K` should be definitionally equal to that -- deriving from the field structure on `SplittingField f`. example : - (AddCommMonoid.natModule : Module ℕ (SplittingField f)) = + (AddCommGroup.toNatModule : Module ℕ (SplittingField f)) = @Algebra.toModule _ _ _ _ (SplittingField.algebra' f) := by with_reducible_and_instances rfl example : - (AddCommGroup.intModule _ : Module ℤ (SplittingField f)) = + (AddCommGroup.toIntModule _ : Module ℤ (SplittingField f)) = @Algebra.toModule _ _ _ _ (SplittingField.algebra' f) := by with_reducible_and_instances rfl -example [CharZero K] : SplittingField.algebra' f = algebraRat := +example [CharZero K] : SplittingField.algebra' f = DivisionRing.toRatAlgebra := rfl -- TODO: by with_reducible_and_instances rfl fails -example {q : ℚ[X]} : algebraInt (SplittingField q) = SplittingField.algebra' q := +example {q : ℚ[X]} : Ring.toIntAlgebra (SplittingField q) = SplittingField.algebra' q := rfl -- TODO: by with_reducible_and_instances rfl fails From 5ff8d6041626a80e1385e3c49856513e8709d0aa Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 5 Aug 2024 11:02:59 +0000 Subject: [PATCH 035/359] chore: update lean4checker version (#15517) --- .github/workflows/lean4checker.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/lean4checker.yml b/.github/workflows/lean4checker.yml index 00588af364ac3..3152d364350df 100644 --- a/.github/workflows/lean4checker.yml +++ b/.github/workflows/lean4checker.yml @@ -69,7 +69,7 @@ jobs: run: | git clone https://github.com/leanprover/lean4checker cd lean4checker - git checkout v4.10.0 + git checkout v4.11.0-rc1 # Now that the git hash is embedded in each olean, # we need to compile lean4checker on the same toolchain cp ../lean-toolchain . From 38886f70edb714ea73774e12e111095042166c75 Mon Sep 17 00:00:00 2001 From: FR Date: Mon, 5 Aug 2024 12:37:02 +0000 Subject: [PATCH 036/359] chore(GroupTheory/Coset): deprecate some lemma (#15482) They were different in mathlib3 because of coercions. --- Mathlib/Algebra/Order/ToIntervalMod.lean | 34 +++++++++---------- Mathlib/Algebra/Pointwise/Stabilizer.lean | 2 +- Mathlib/Analysis/Fourier/AddCircle.lean | 4 +-- .../SpecialFunctions/Complex/Circle.lean | 10 +++--- Mathlib/Data/ZMod/Quotient.lean | 4 +-- Mathlib/GroupTheory/Commensurable.lean | 2 +- Mathlib/GroupTheory/Complement.lean | 2 +- Mathlib/GroupTheory/Coset.lean | 25 ++++++-------- Mathlib/GroupTheory/GroupAction/Quotient.lean | 2 +- Mathlib/GroupTheory/QuotientGroup.lean | 6 ++-- Mathlib/GroupTheory/Torsion.lean | 2 +- .../LinearAlgebra/Alternating/DomCoprod.lean | 2 +- .../NumberTheory/LSeries/HurwitzZetaEven.lean | 14 ++++---- .../NumberTheory/LSeries/HurwitzZetaOdd.lean | 14 ++++---- Mathlib/Topology/Instances/AddCircle.lean | 6 ++-- 15 files changed, 63 insertions(+), 66 deletions(-) diff --git a/Mathlib/Algebra/Order/ToIntervalMod.lean b/Mathlib/Algebra/Order/ToIntervalMod.lean index d62f8e8707e99..c77210693d946 100644 --- a/Mathlib/Algebra/Order/ToIntervalMod.lean +++ b/Mathlib/Algebra/Order/ToIntervalMod.lean @@ -665,11 +665,11 @@ end Zero @[simps symm_apply] def QuotientAddGroup.equivIcoMod (a : α) : α ⧸ AddSubgroup.zmultiples p ≃ Set.Ico a (a + p) where toFun b := - ⟨(toIcoMod_periodic hp a).lift b, QuotientAddGroup.induction_on' b <| toIcoMod_mem_Ico hp a⟩ + ⟨(toIcoMod_periodic hp a).lift b, QuotientAddGroup.induction_on b <| toIcoMod_mem_Ico hp a⟩ invFun := (↑) right_inv b := Subtype.ext <| (toIcoMod_eq_self hp).mpr b.prop left_inv b := by - induction b using QuotientAddGroup.induction_on' + induction b using QuotientAddGroup.induction_on dsimp rw [QuotientAddGroup.eq_iff_sub_mem, toIcoMod_sub_self] apply AddSubgroup.zsmul_mem_zmultiples @@ -688,11 +688,11 @@ theorem QuotientAddGroup.equivIcoMod_zero (a : α) : @[simps symm_apply] def QuotientAddGroup.equivIocMod (a : α) : α ⧸ AddSubgroup.zmultiples p ≃ Set.Ioc a (a + p) where toFun b := - ⟨(toIocMod_periodic hp a).lift b, QuotientAddGroup.induction_on' b <| toIocMod_mem_Ioc hp a⟩ + ⟨(toIocMod_periodic hp a).lift b, QuotientAddGroup.induction_on b <| toIocMod_mem_Ioc hp a⟩ invFun := (↑) right_inv b := Subtype.ext <| (toIocMod_eq_self hp).mpr b.prop left_inv b := by - induction b using QuotientAddGroup.induction_on' + induction b using QuotientAddGroup.induction_on dsimp rw [QuotientAddGroup.eq_iff_sub_mem, toIocMod_sub_self] apply AddSubgroup.zsmul_mem_zmultiples @@ -805,36 +805,36 @@ theorem btw_coe_iff {x₁ x₂ x₃ : α} : instance circularPreorder : CircularPreorder (α ⧸ AddSubgroup.zmultiples p) where btw_refl x := show _ ≤ _ by simp [sub_self, hp'.out.le] btw_cyclic_left {x₁ x₂ x₃} h := by - induction x₁ using QuotientAddGroup.induction_on' - induction x₂ using QuotientAddGroup.induction_on' - induction x₃ using QuotientAddGroup.induction_on' + induction x₁ using QuotientAddGroup.induction_on + induction x₂ using QuotientAddGroup.induction_on + induction x₃ using QuotientAddGroup.induction_on simp_rw [btw_coe_iff] at h ⊢ apply toIxxMod_cyclic_left _ h sbtw := _ sbtw_iff_btw_not_btw := Iff.rfl sbtw_trans_left {x₁ x₂ x₃ x₄} (h₁₂₃ : _ ∧ _) (h₂₃₄ : _ ∧ _) := show _ ∧ _ by - induction x₁ using QuotientAddGroup.induction_on' - induction x₂ using QuotientAddGroup.induction_on' - induction x₃ using QuotientAddGroup.induction_on' - induction x₄ using QuotientAddGroup.induction_on' + induction x₁ using QuotientAddGroup.induction_on + induction x₂ using QuotientAddGroup.induction_on + induction x₃ using QuotientAddGroup.induction_on + induction x₄ using QuotientAddGroup.induction_on simp_rw [btw_coe_iff] at h₁₂₃ h₂₃₄ ⊢ apply toIxxMod_trans _ h₁₂₃ h₂₃₄ instance circularOrder : CircularOrder (α ⧸ AddSubgroup.zmultiples p) := { QuotientAddGroup.circularPreorder with btw_antisymm := fun {x₁ x₂ x₃} h₁₂₃ h₃₂₁ => by - induction x₁ using QuotientAddGroup.induction_on' - induction x₂ using QuotientAddGroup.induction_on' - induction x₃ using QuotientAddGroup.induction_on' + induction x₁ using QuotientAddGroup.induction_on + induction x₂ using QuotientAddGroup.induction_on + induction x₃ using QuotientAddGroup.induction_on rw [btw_cyclic] at h₃₂₁ simp_rw [btw_coe_iff] at h₁₂₃ h₃₂₁ simp_rw [← modEq_iff_eq_mod_zmultiples] exact toIxxMod_antisymm _ h₁₂₃ h₃₂₁ btw_total := fun x₁ x₂ x₃ => by - induction x₁ using QuotientAddGroup.induction_on' - induction x₂ using QuotientAddGroup.induction_on' - induction x₃ using QuotientAddGroup.induction_on' + induction x₁ using QuotientAddGroup.induction_on + induction x₂ using QuotientAddGroup.induction_on + induction x₃ using QuotientAddGroup.induction_on simp_rw [btw_coe_iff] apply toIxxMod_total } diff --git a/Mathlib/Algebra/Pointwise/Stabilizer.lean b/Mathlib/Algebra/Pointwise/Stabilizer.lean index 4bdde91f4d060..14d5840e635a6 100644 --- a/Mathlib/Algebra/Pointwise/Stabilizer.lean +++ b/Mathlib/Algebra/Pointwise/Stabilizer.lean @@ -170,7 +170,7 @@ local notation " q " => ((↑) : G → Q) @[to_additive] lemma stabilizer_image_coe_quotient : stabilizer Q (q '' s) = ⊥ := by ext a - induction' a using QuotientGroup.induction_on' with a + induction' a using QuotientGroup.induction_on with a simp only [mem_stabilizer_iff, Subgroup.mem_bot, QuotientGroup.eq_one_iff] have : q a • q '' s = q '' (a • s) := (image_smul_distrib (QuotientGroup.mk' <| stabilizer G s) _ _).symm diff --git a/Mathlib/Analysis/Fourier/AddCircle.lean b/Mathlib/Analysis/Fourier/AddCircle.lean index e556c89fdbc89..732c088472bb7 100644 --- a/Mathlib/Analysis/Fourier/AddCircle.lean +++ b/Mathlib/Analysis/Fourier/AddCircle.lean @@ -123,7 +123,7 @@ theorem fourier_coe_apply' {n : ℤ} {x : ℝ} : -- @[simp] -- Porting note: simp normal form is `fourier_zero'` theorem fourier_zero {x : AddCircle T} : fourier 0 x = 1 := by - induction x using QuotientAddGroup.induction_on' + induction x using QuotientAddGroup.induction_on simp only [fourier_coe_apply] norm_num @@ -142,7 +142,7 @@ theorem fourier_one {x : AddCircle T} : fourier 1 x = toCircle x := by rw [fouri -- @[simp] -- Porting note: simp normal form is `fourier_neg'` theorem fourier_neg {n : ℤ} {x : AddCircle T} : fourier (-n) x = conj (fourier n x) := by - induction x using QuotientAddGroup.induction_on' + induction x using QuotientAddGroup.induction_on simp_rw [fourier_apply, toCircle] rw [← QuotientAddGroup.mk_zsmul, ← QuotientAddGroup.mk_zsmul] simp_rw [Function.Periodic.lift_coe, ← coe_inv_circle_eq_conj, ← expMapCircle_neg, diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean index 2254faba3e1af..72cfc152f87d1 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean @@ -150,8 +150,8 @@ theorem toCircle_apply_mk (x : ℝ) : @toCircle T x = expMapCircle (2 * π / T * theorem toCircle_add (x : AddCircle T) (y : AddCircle T) : @toCircle T (x + y) = toCircle x * toCircle y := by - induction x using QuotientAddGroup.induction_on' - induction y using QuotientAddGroup.induction_on' + induction x using QuotientAddGroup.induction_on + induction y using QuotientAddGroup.induction_on simp_rw [← coe_add, toCircle_apply_mk, mul_add, expMapCircle_add] lemma toCircle_zero : toCircle (0 : AddCircle T) = 1 := by @@ -162,8 +162,8 @@ theorem continuous_toCircle : Continuous (@toCircle T) := theorem injective_toCircle (hT : T ≠ 0) : Function.Injective (@toCircle T) := by intro a b h - induction a using QuotientAddGroup.induction_on' - induction b using QuotientAddGroup.induction_on' + induction a using QuotientAddGroup.induction_on + induction b using QuotientAddGroup.induction_on simp_rw [toCircle_apply_mk] at h obtain ⟨m, hm⟩ := expMapCircle_eq_expMapCircle.mp h.symm rw [QuotientAddGroup.eq]; simp_rw [AddSubgroup.mem_zmultiples_iff, zsmul_eq_mul] @@ -194,7 +194,7 @@ noncomputable def homeomorphCircle (hT : T ≠ 0) : AddCircle T ≃ₜ circle := theorem homeomorphCircle_apply (hT : T ≠ 0) (x : AddCircle T) : homeomorphCircle hT x = toCircle x := by - induction' x using QuotientAddGroup.induction_on' with x + induction' x using QuotientAddGroup.induction_on with x rw [homeomorphCircle, Homeomorph.trans_apply, homeomorphAddCircle_apply_mk, homeomorphCircle'_apply_mk, toCircle_apply_mk] ring_nf diff --git a/Mathlib/Data/ZMod/Quotient.lean b/Mathlib/Data/ZMod/Quotient.lean index 50242d9fc93ea..39ff9a8cdf6d4 100644 --- a/Mathlib/Data/ZMod/Quotient.lean +++ b/Mathlib/Data/ZMod/Quotient.lean @@ -92,11 +92,11 @@ noncomputable def zmultiplesQuotientStabilizerEquiv : exact isPeriodicPt_minimalPeriod (a +ᵥ ·) b)) ⟨by rw [← ker_eq_bot_iff, eq_bot_iff] - refine fun q => induction_on' q fun n hn => ?_ + refine fun q => induction_on q fun n hn => ?_ rw [mem_bot, eq_zero_iff, Int.mem_zmultiples_iff, ← zsmul_vadd_eq_iff_minimalPeriod_dvd] exact (eq_zero_iff _).mp hn, fun q => - induction_on' q fun ⟨_, n, rfl⟩ => ⟨n, rfl⟩⟩).symm.trans + induction_on q fun ⟨_, n, rfl⟩ => ⟨n, rfl⟩⟩).symm.trans (Int.quotientZMultiplesNatEquivZMod (minimalPeriod (a +ᵥ ·) b)) theorem zmultiplesQuotientStabilizerEquiv_symm_apply (n : ZMod (minimalPeriod (a +ᵥ ·) b)) : diff --git a/Mathlib/GroupTheory/Commensurable.lean b/Mathlib/GroupTheory/Commensurable.lean index c7e7b52ef5bb2..bc2c93f90a19f 100644 --- a/Mathlib/GroupTheory/Commensurable.lean +++ b/Mathlib/GroupTheory/Commensurable.lean @@ -56,7 +56,7 @@ def quotConjEquiv (H K : Subgroup G) (g : ConjAct G) : K ⧸ H.subgroupOf K ≃ (g • K).1 ⧸ (g • H).subgroupOf (g • K) := Quotient.congr (K.equivSMul g).toEquiv fun a b => by dsimp - rw [← Quotient.eq'', ← Quotient.eq'', QuotientGroup.eq', QuotientGroup.eq', + rw [← Quotient.eq'', ← Quotient.eq'', QuotientGroup.eq, QuotientGroup.eq, Subgroup.mem_subgroupOf, Subgroup.mem_subgroupOf, ← MulEquiv.map_inv, ← MulEquiv.map_mul, Subgroup.equivSMul_apply_coe] exact Subgroup.smul_mem_pointwise_smul_iff.symm diff --git a/Mathlib/GroupTheory/Complement.lean b/Mathlib/GroupTheory/Complement.lean index 1c743e1acc6b1..d7820cbb29123 100644 --- a/Mathlib/GroupTheory/Complement.lean +++ b/Mathlib/GroupTheory/Complement.lean @@ -201,7 +201,7 @@ theorem mem_leftTransversals_iff_existsUnique_quotient_mk''_eq : S ∈ leftTransversals (H : Set G) ↔ ∀ q : Quotient (QuotientGroup.leftRel H), ∃! s : S, Quotient.mk'' s.1 = q := by simp_rw [mem_leftTransversals_iff_existsUnique_inv_mul_mem, SetLike.mem_coe, ← - QuotientGroup.eq'] + QuotientGroup.eq] exact ⟨fun h q => Quotient.inductionOn' q h, fun h g => h (Quotient.mk'' g)⟩ @[to_additive] diff --git a/Mathlib/GroupTheory/Coset.lean b/Mathlib/GroupTheory/Coset.lean index fc3913d35b708..c21b21711d37b 100644 --- a/Mathlib/GroupTheory/Coset.lean +++ b/Mathlib/GroupTheory/Coset.lean @@ -371,9 +371,7 @@ theorem induction_on {C : α ⧸ s → Prop} (x : α ⧸ s) (H : ∀ z, C (Quoti instance : Coe α (α ⧸ s) := ⟨mk⟩ -@[to_additive (attr := elab_as_elim)] -theorem induction_on' {C : α ⧸ s → Prop} (x : α ⧸ s) (H : ∀ z : α, C z) : C x := - Quotient.inductionOn' x H +@[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 := @@ -397,9 +395,7 @@ protected theorem eq {a b : α} : (a : α ⧸ s) = b ↔ a⁻¹ * b ∈ s := _ ↔ @Setoid.r _ (leftRel s) a b := Quotient.eq'' _ ↔ _ := by rw [leftRel_apply] -@[to_additive] -theorem eq' {a b : α} : (mk a : α ⧸ s) = mk b ↔ a⁻¹ * b ∈ s := - QuotientGroup.eq +@[to_additive (attr := deprecated (since := "2024-08-04"))] alias eq' := QuotientGroup.eq @[to_additive] -- Porting note (#10618): `simp` can prove this. theorem out_eq' (a : α ⧸ s) : mk a.out' = a := @@ -412,13 +408,13 @@ variable (s) stated in terms of an arbitrary `h : s`, rather than the specific `h = g⁻¹ * (mk g).out'`. -/ @[to_additive QuotientAddGroup.mk_out'_eq_mul] theorem mk_out'_eq_mul (g : α) : ∃ h : s, (mk g : α ⧸ s).out' = g * h := - ⟨⟨g⁻¹ * (mk g).out', eq'.mp (mk g).out_eq'.symm⟩, by rw [mul_inv_cancel_left]⟩ + ⟨⟨g⁻¹ * (mk g).out', QuotientGroup.eq.mp (mk g).out_eq'.symm⟩, by rw [mul_inv_cancel_left]⟩ variable {s} {a b : α} @[to_additive (attr := simp)] theorem mk_mul_of_mem (a : α) (hb : b ∈ s) : (mk (a * b) : α ⧸ s) = mk a := by - rwa [eq', mul_inv_rev, inv_mul_cancel_right, s.inv_mem_iff] + rwa [QuotientGroup.eq, mul_inv_rev, inv_mul_cancel_right, s.inv_mem_iff] @[to_additive] theorem eq_class_eq_leftCoset (s : Subgroup α) (g : α) : @@ -486,8 +482,8 @@ variable {t : Subgroup α} def quotientEquivOfEq (h : s = t) : α ⧸ s ≃ α ⧸ t where toFun := Quotient.map' id fun _a _b h' => h ▸ h' invFun := Quotient.map' id fun _a _b h' => h.symm ▸ h' - left_inv q := induction_on' q fun _g => rfl - right_inv q := induction_on' q fun _g => rfl + left_inv q := induction_on q fun _g => rfl + right_inv q := induction_on q fun _g => rfl theorem quotientEquivOfEq_mk (h : s = t) (a : α) : quotientEquivOfEq h (QuotientGroup.mk a) = QuotientGroup.mk a := @@ -544,7 +540,7 @@ def quotientSubgroupOfEmbeddingOfLE (H : Subgroup α) (h : s ≤ t) : inj' := Quotient.ind₂' <| by intro a b h - simpa only [Quotient.map'_mk'', eq'] using h + simpa only [Quotient.map'_mk'', QuotientGroup.eq] using h -- Porting note: I had to add the type ascription to the right-hand side or else Lean times out. @[to_additive (attr := simp)] @@ -588,8 +584,8 @@ def quotientiInfSubgroupOfEmbedding {ι : Type*} (f : ι → Subgroup α) (H : S toFun q i := quotientSubgroupOfMapOfLE H (iInf_le f i) q inj' := Quotient.ind₂' <| by - simp_rw [funext_iff, quotientSubgroupOfMapOfLE_apply_mk, eq', mem_subgroupOf, mem_iInf, - imp_self, forall_const] + simp_rw [funext_iff, quotientSubgroupOfMapOfLE_apply_mk, QuotientGroup.eq, mem_subgroupOf, + mem_iInf, imp_self, forall_const] -- Porting note: I had to add the type ascription to the right-hand side or else Lean times out. @[to_additive (attr := simp)] @@ -605,7 +601,8 @@ def quotientiInfEmbedding {ι : Type*} (f : ι → Subgroup α) : (α ⧸ ⨅ i, toFun q i := quotientMapOfLE (iInf_le f i) q inj' := Quotient.ind₂' <| by - simp_rw [funext_iff, quotientMapOfLE_apply_mk, eq', mem_iInf, imp_self, forall_const] + simp_rw [funext_iff, quotientMapOfLE_apply_mk, QuotientGroup.eq, mem_iInf, imp_self, + forall_const] @[to_additive (attr := simp)] theorem quotientiInfEmbedding_apply_mk {ι : Type*} (f : ι → Subgroup α) (g : α) (i : ι) : diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 1d055b96f86dc..68a6f02b68d09 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -99,7 +99,7 @@ theorem Quotient.coe_smul_out' [QuotientAction β H] (b : β) (q : α ⧸ H) : theorem _root_.QuotientGroup.out'_conj_pow_minimalPeriod_mem (a : α) (q : α ⧸ H) : q.out'⁻¹ * a ^ Function.minimalPeriod (a • ·) q * q.out' ∈ H := by - rw [mul_assoc, ← QuotientGroup.eq', QuotientGroup.out_eq', ← smul_eq_mul, Quotient.mk_smul_out', + rw [mul_assoc, ← QuotientGroup.eq, QuotientGroup.out_eq', ← smul_eq_mul, Quotient.mk_smul_out', eq_comm, pow_smul_eq_iff_minimalPeriod_dvd] end QuotientAction diff --git a/Mathlib/GroupTheory/QuotientGroup.lean b/Mathlib/GroupTheory/QuotientGroup.lean index c5c111db89a9f..06dad0c62e756 100644 --- a/Mathlib/GroupTheory/QuotientGroup.lean +++ b/Mathlib/GroupTheory/QuotientGroup.lean @@ -83,7 +83,7 @@ theorem mk'_surjective : Surjective <| mk' N := @[to_additive] theorem mk'_eq_mk' {x y : G} : mk' N x = mk' N y ↔ ∃ z ∈ N, x * z = y := - QuotientGroup.eq'.trans <| by + QuotientGroup.eq.trans <| by simp only [← _root_.eq_inv_mul_iff_mul_eq, exists_prop, exists_eq_right] open scoped Pointwise in @@ -217,7 +217,7 @@ theorem map_mk' (M : Subgroup H) [M.Normal] (f : G →* H) (h : N ≤ M.comap f) @[to_additive] theorem map_id_apply (h : N ≤ Subgroup.comap (MonoidHom.id _) N := (Subgroup.comap_id N).le) (x) : map N N (MonoidHom.id _) h x = x := - induction_on' x fun _x => rfl + induction_on x fun _x => rfl @[to_additive (attr := simp)] theorem map_id (h : N ≤ Subgroup.comap (MonoidHom.id _) N := (Subgroup.comap_id N).le) : @@ -230,7 +230,7 @@ theorem map_map {I : Type*} [Group I] (M : Subgroup H) (O : Subgroup I) [M.Norma (hgf : N ≤ Subgroup.comap (g.comp f) O := hf.trans ((Subgroup.comap_mono hg).trans_eq (Subgroup.comap_comap _ _ _))) (x : G ⧸ N) : map M O g hg (map N M f hf x) = map N O (g.comp f) hgf x := by - refine induction_on' x fun x => ?_ + refine induction_on x fun x => ?_ simp only [map_mk, MonoidHom.comp_apply] @[to_additive (attr := simp)] diff --git a/Mathlib/GroupTheory/Torsion.lean b/Mathlib/GroupTheory/Torsion.lean index 6eca87ed01fd9..d3eef9e840874 100644 --- a/Mathlib/GroupTheory/Torsion.lean +++ b/Mathlib/GroupTheory/Torsion.lean @@ -372,7 +372,7 @@ variable (G) [CommGroup G] "Quotienting a group by its additive torsion subgroup yields an additive torsion free group."] theorem IsTorsionFree.quotient_torsion : IsTorsionFree <| G ⧸ torsion G := fun g hne hfin => hne <| by - induction' g using QuotientGroup.induction_on' with g + induction' g using QuotientGroup.induction_on with g obtain ⟨m, mpos, hm⟩ := hfin.exists_pow_eq_one obtain ⟨n, npos, hn⟩ := ((QuotientGroup.eq_one_iff _).mp hm).exists_pow_eq_one exact diff --git a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean index 4930aa7e1587b..6da58edbd959e 100644 --- a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean +++ b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean @@ -229,7 +229,7 @@ theorem MultilinearMap.domCoprod_alternization [DecidableEq ιa] [DecidableEq ι (QuotientGroup.leftRelDecidable (MonoidHom.range (Perm.sumCongrHom ιa ιb))) (Quotient.mk (QuotientGroup.leftRel (MonoidHom.range (Perm.sumCongrHom ιa ιb))) a) (Quotient.mk'' σ)) _ (s := Finset.univ) - fun x _ => QuotientGroup.eq' (s := MonoidHom.range (Perm.sumCongrHom ιa ιb)) (a := x) (b := σ)] + fun x _ => QuotientGroup.eq (s := MonoidHom.range (Perm.sumCongrHom ιa ιb)) (a := x) (b := σ)] -- eliminate a multiplication rw [← Finset.map_univ_equiv (Equiv.mulLeft σ), Finset.filter_map, Finset.sum_map] simp_rw [Equiv.coe_toEmbedding, Equiv.coe_mulLeft, (· ∘ ·), mul_inv_rev, inv_mul_cancel_right, diff --git a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean index 9cac82e7bfa82..7a65bfbbc95a0 100644 --- a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean +++ b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean @@ -81,7 +81,7 @@ lemma evenKernel_def (a x : ℝ) : /-- For `x ≤ 0` the defining sum diverges, so the kernel is 0. -/ lemma evenKernel_undef (a : UnitAddCircle) {x : ℝ} (hx : x ≤ 0) : evenKernel a x = 0 := by have : (I * ↑x).im ≤ 0 := by rwa [I_mul_im, ofReal_re] - induction' a using QuotientAddGroup.induction_on' with a' + induction' a using QuotientAddGroup.induction_on with a' rw [← ofReal_inj, evenKernel_def, jacobiTheta₂_undef _ this, mul_zero, ofReal_zero] /-- Cosine Hurwitz zeta kernel. See `cosKernel_def` for the defining formula, and @@ -96,7 +96,7 @@ lemma cosKernel_def (a x : ℝ) : ↑(cosKernel ↑a x) = jacobiTheta₂ a (I * conj_I, neg_mul, neg_neg, ← mul_two, mul_div_cancel_right₀ _ (two_ne_zero' ℂ)] lemma cosKernel_undef (a : UnitAddCircle) {x : ℝ} (hx : x ≤ 0) : cosKernel a x = 0 := by - induction' a using QuotientAddGroup.induction_on' with a' + induction' a using QuotientAddGroup.induction_on with a' rw [← ofReal_inj, cosKernel_def, jacobiTheta₂_undef _ (by rwa [I_mul_im, ofReal_re]), ofReal_zero] /-- For `a = 0`, both kernels agree. -/ @@ -106,17 +106,17 @@ lemma evenKernel_eq_cosKernel_of_zero : evenKernel 0 = cosKernel 0 := by zero_mul, Complex.exp_zero, one_mul, cosKernel_def] lemma evenKernel_neg (a : UnitAddCircle) (x : ℝ) : evenKernel (-a) x = evenKernel a x := by - induction' a using QuotientAddGroup.induction_on' with a' + induction' a using QuotientAddGroup.induction_on with a' simp only [← QuotientAddGroup.mk_neg, ← ofReal_inj, evenKernel_def, ofReal_neg, neg_sq, neg_mul, jacobiTheta₂_neg_left] lemma cosKernel_neg (a : UnitAddCircle) (x : ℝ) : cosKernel (-a) x = cosKernel a x := by - induction' a using QuotientAddGroup.induction_on' with a' + induction' a using QuotientAddGroup.induction_on with a' simp only [← QuotientAddGroup.mk_neg, ← ofReal_inj, cosKernel_def, ofReal_neg, jacobiTheta₂_neg_left] lemma continuousOn_evenKernel (a : UnitAddCircle) : ContinuousOn (evenKernel a) (Ioi 0) := by - induction' a using QuotientAddGroup.induction_on' with a' + induction' a using QuotientAddGroup.induction_on with a' apply continuous_re.comp_continuousOn (f := fun x ↦ (evenKernel a' x : ℂ)) simp only [evenKernel_def a'] refine ContinuousAt.continuousOn (fun x hx ↦ ((Continuous.continuousAt ?_).mul ?_)) @@ -126,7 +126,7 @@ lemma continuousOn_evenKernel (a : UnitAddCircle) : ContinuousOn (evenKernel a) · rwa [mul_im, I_re, I_im, zero_mul, one_mul, zero_add, ofReal_re] lemma continuousOn_cosKernel (a : UnitAddCircle) : ContinuousOn (cosKernel a) (Ioi 0) := by - induction' a using QuotientAddGroup.induction_on' with a' + induction' a using QuotientAddGroup.induction_on with a' apply continuous_re.comp_continuousOn (f := fun x ↦ (cosKernel a' x : ℂ)) simp only [cosKernel_def] refine ContinuousAt.continuousOn (fun x hx ↦ ?_) @@ -138,7 +138,7 @@ lemma evenKernel_functional_equation (a : UnitAddCircle) (x : ℝ) : rcases le_or_lt x 0 with hx | hx · rw [evenKernel_undef _ hx, cosKernel_undef, mul_zero] exact div_nonpos_of_nonneg_of_nonpos zero_le_one hx - induction' a using QuotientAddGroup.induction_on' with a + induction' a using QuotientAddGroup.induction_on with a rw [← ofReal_inj, ofReal_mul, evenKernel_def, cosKernel_def, jacobiTheta₂_functional_equation] have h1 : I * ↑(1 / x) = -1 / (I * x) := by push_cast diff --git a/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean b/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean index 05208a36a5677..a77aff55c0fd1 100644 --- a/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean +++ b/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean @@ -119,7 +119,7 @@ lemma oddKernel_def' (a x : ℝ) : ↑(oddKernel ↑a x) = cexp (-π * a ^ 2 * x (by ring : ↑π * I * ↑a ^ 2 * (I * ↑x) = I ^ 2 * ↑π * ↑a ^ 2 * x), I_sq, neg_one_mul] lemma oddKernel_undef (a : UnitAddCircle) {x : ℝ} (hx : x ≤ 0) : oddKernel a x = 0 := by - induction' a using QuotientAddGroup.induction_on' with a' + induction' a using QuotientAddGroup.induction_on with a' rw [← ofReal_eq_zero, oddKernel_def', jacobiTheta₂_undef, jacobiTheta₂'_undef, zero_div, zero_add, mul_zero, mul_zero] <;> rwa [I_mul_im, ofReal_re] @@ -136,12 +136,12 @@ lemma sinKernel_def (a x : ℝ) : ↑(sinKernel ↑a x) = jacobiTheta₂' a (I * simp_rw [map_mul, conj_I, conj_ofReal, map_neg, map_ofNat, neg_mul, neg_neg, add_self_div_two] lemma sinKernel_undef (a : UnitAddCircle) {x : ℝ} (hx : x ≤ 0) : sinKernel a x = 0 := by - induction' a using QuotientAddGroup.induction_on' with a' + induction' a using QuotientAddGroup.induction_on with a' rw [← ofReal_eq_zero, sinKernel_def, jacobiTheta₂'_undef _ (by rwa [I_mul_im, ofReal_re]), zero_div] lemma oddKernel_neg (a : UnitAddCircle) (x : ℝ) : oddKernel (-a) x = -oddKernel a x := by - induction' a using QuotientAddGroup.induction_on' with a' + induction' a using QuotientAddGroup.induction_on with a' rw [← ofReal_inj, ← QuotientAddGroup.mk_neg, oddKernel_def, ofReal_neg, ofReal_neg, oddKernel_def, jacobiTheta₂''_neg_left] @@ -150,7 +150,7 @@ lemma oddKernel_zero (x : ℝ) : oddKernel 0 x = 0 := by lemma sinKernel_neg (a : UnitAddCircle) (x : ℝ) : sinKernel (-a) x = -sinKernel a x := by - induction' a using QuotientAddGroup.induction_on' with a' + induction' a using QuotientAddGroup.induction_on with a' rw [← ofReal_inj, ← QuotientAddGroup.mk_neg, ofReal_neg, sinKernel_def, sinKernel_def, ofReal_neg, jacobiTheta₂'_neg_left, neg_div] @@ -159,7 +159,7 @@ lemma sinKernel_zero (x : ℝ) : sinKernel 0 x = 0 := by /-- The odd kernel is continuous on `Ioi 0`. -/ lemma continuousOn_oddKernel (a : UnitAddCircle) : ContinuousOn (oddKernel a) (Ioi 0) := by - induction' a using QuotientAddGroup.induction_on' with a + induction' a using QuotientAddGroup.induction_on with a suffices ContinuousOn (fun x ↦ (oddKernel a x : ℂ)) (Ioi 0) from (continuous_re.comp_continuousOn this).congr fun a _ ↦ (ofReal_re _).symm simp_rw [oddKernel_def' a] @@ -173,7 +173,7 @@ lemma continuousOn_oddKernel (a : UnitAddCircle) : ContinuousOn (oddKernel a) (I (by rwa [I_mul_im, ofReal_re])).comp (f := fun u : ℝ ↦ (a * I * u, I * u)) hf.continuousAt lemma continuousOn_sinKernel (a : UnitAddCircle) : ContinuousOn (sinKernel a) (Ioi 0) := by - induction' a using QuotientAddGroup.induction_on' with a + induction' a using QuotientAddGroup.induction_on with a suffices ContinuousOn (fun x ↦ (sinKernel a x : ℂ)) (Ioi 0) from (continuous_re.comp_continuousOn this).congr fun a _ ↦ (ofReal_re _).symm simp_rw [sinKernel_def] @@ -186,7 +186,7 @@ lemma oddKernel_functional_equation (a : UnitAddCircle) (x : ℝ) : -- first reduce to `0 < x` rcases le_or_lt x 0 with hx | hx · rw [oddKernel_undef _ hx, sinKernel_undef _ (one_div_nonpos.mpr hx), mul_zero] - induction' a using QuotientAddGroup.induction_on' with a + induction' a using QuotientAddGroup.induction_on with a have h1 : -1 / (I * ↑(1 / x)) = I * x := by rw [one_div, ofReal_inv, mul_comm, ← div_div, div_inv_eq_mul, div_eq_mul_inv, inv_I, mul_neg, neg_one_mul, neg_mul, neg_neg, mul_comm] have h2 : (-I * (I * ↑(1 / x))) = 1 / x := by diff --git a/Mathlib/Topology/Instances/AddCircle.lean b/Mathlib/Topology/Instances/AddCircle.lean index 63abd941d7228..9f33cbe774d18 100644 --- a/Mathlib/Topology/Instances/AddCircle.lean +++ b/Mathlib/Topology/Instances/AddCircle.lean @@ -238,12 +238,12 @@ theorem continuous_equivIoc_symm : Continuous (equivIoc p a).symm := variable {x : AddCircle p} (hx : x ≠ a) theorem continuousAt_equivIco : ContinuousAt (equivIco p a) x := by - induction x using QuotientAddGroup.induction_on' + induction x using QuotientAddGroup.induction_on rw [ContinuousAt, Filter.Tendsto, QuotientAddGroup.nhds_eq, Filter.map_map] exact (continuousAt_toIcoMod hp.out a hx).codRestrict _ theorem continuousAt_equivIoc : ContinuousAt (equivIoc p a) x := by - induction x using QuotientAddGroup.induction_on' + induction x using QuotientAddGroup.induction_on rw [ContinuousAt, Filter.Tendsto, QuotientAddGroup.nhds_eq, Filter.map_map] exact (continuousAt_toIocMod hp.out a hx).codRestrict _ @@ -409,7 +409,7 @@ theorem addOrderOf_coe_rat {q : ℚ} : addOrderOf (↑(↑q * p) : AddCircle p) theorem addOrderOf_eq_pos_iff {u : AddCircle p} {n : ℕ} (h : 0 < n) : addOrderOf u = n ↔ ∃ m < n, m.gcd n = 1 ∧ ↑(↑m / ↑n * p) = u := by - refine ⟨QuotientAddGroup.induction_on' u fun k hk => ?_, ?_⟩ + refine ⟨QuotientAddGroup.induction_on u fun k hk => ?_, ?_⟩ · rintro ⟨m, _, h₁, rfl⟩ exact addOrderOf_div_of_gcd_eq_one h h₁ have h0 := addOrderOf_nsmul_eq_zero (k : AddCircle p) From fcfce35759363fc0cbf30be2ab529283e49e740c Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 5 Aug 2024 13:16:25 +0000 Subject: [PATCH 037/359] chore: remove some open Classical, part 3 (#15417) See #15371. --- Mathlib/Algebra/Order/CompleteField.lean | 1 - Mathlib/Analysis/Analytic/Composition.lean | 2 +- Mathlib/Analysis/Analytic/Inverse.lean | 8 ++-- Mathlib/Analysis/Analytic/Linear.lean | 3 +- Mathlib/Analysis/Analytic/RadiusLiminf.lean | 2 +- Mathlib/Analysis/BoxIntegral/Basic.lean | 13 +++-- .../BoxIntegral/Box/SubboxInduction.lean | 4 +- .../BoxIntegral/DivergenceTheorem.lean | 3 +- .../Analysis/BoxIntegral/Integrability.lean | 4 +- .../BoxIntegral/Partition/Measure.lean | 3 +- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 3 +- .../Analysis/Calculus/ContDiff/Bounds.lean | 2 +- Mathlib/Analysis/Calculus/Deriv/Basic.lean | 3 +- Mathlib/Analysis/Calculus/Deriv/Mul.lean | 14 +++--- .../Analysis/Calculus/Deriv/Polynomial.lean | 5 +- Mathlib/Analysis/Calculus/Dslope.lean | 15 +++--- .../ApproximatesLinearOn.lean | 2 +- .../InverseFunctionTheorem/FDeriv.lean | 2 +- .../Analysis/Calculus/IteratedDeriv/Defs.lean | 4 +- .../Analysis/Calculus/LocalExtr/Basic.lean | 48 +++++++++++-------- .../Calculus/LocalExtr/LineDeriv.lean | 12 +++-- Mathlib/Analysis/Calculus/MeanValue.lean | 3 +- .../Complex/UpperHalfPlane/Basic.lean | 4 +- Mathlib/Analysis/Convex/Deriv.lean | 3 +- .../Analysis/InnerProductSpace/Calculus.lean | 2 - .../Analysis/InnerProductSpace/l2Space.lean | 47 ++++++++++-------- .../Normed/Module/FiniteDimension.lean | 6 +-- .../NormedSpace/OperatorNorm/Basic.lean | 2 +- .../NormedSpace/OperatorNorm/Bilinear.lean | 2 +- .../OperatorNorm/Completeness.lean | 2 +- .../NormedSpace/OperatorNorm/Mul.lean | 3 +- .../NormedSpace/OperatorNorm/NormedSpace.lean | 2 +- Mathlib/Analysis/ODE/Gronwall.lean | 6 +-- .../Analysis/SpecialFunctions/Bernstein.lean | 3 +- Mathlib/Analysis/SpecialFunctions/Exp.lean | 3 +- .../Analysis/SpecialFunctions/ExpDeriv.lean | 4 +- .../Analysis/SpecialFunctions/Pow/Deriv.lean | 3 +- .../SpecialFunctions/SmoothTransition.lean | 5 +- .../SpecialFunctions/Trigonometric/Deriv.lean | 6 +-- .../Trigonometric/InverseDeriv.lean | 8 +--- .../CategoryTheory/Groupoid/FreeGroupoid.lean | 5 +- Mathlib/Data/Complex/Exponential.lean | 2 +- Mathlib/FieldTheory/AbelRuffini.lean | 1 - Mathlib/FieldTheory/Adjoin.lean | 4 +- Mathlib/FieldTheory/Finite/Polynomial.lean | 14 ++++-- .../IsAlgClosed/AlgebraicClosure.lean | 4 +- Mathlib/FieldTheory/IsAlgClosed/Basic.lean | 4 +- .../Minpoly/IsIntegrallyClosed.lean | 3 -- Mathlib/FieldTheory/PrimitiveElement.lean | 3 +- Mathlib/FieldTheory/SeparableClosure.lean | 2 - Mathlib/FieldTheory/SeparableDegree.lean | 22 +++++++-- .../SplittingField/Construction.lean | 5 +- .../SplittingField/IsSplittingField.lean | 12 +++-- Mathlib/Geometry/Manifold/BumpFunction.lean | 2 +- .../TensorProduct/Subalgebra.lean | 2 +- .../TensorProduct/Vanishing.lean | 3 +- .../Constructions/BorelSpace/Real.lean | 3 +- Mathlib/MeasureTheory/Function/UnifTight.lean | 6 +-- .../DedekindDomain/FiniteAdeleRing.lean | 9 ++-- .../RingTheory/MvPolynomial/Localization.lean | 3 +- 60 files changed, 189 insertions(+), 187 deletions(-) diff --git a/Mathlib/Algebra/Order/CompleteField.lean b/Mathlib/Algebra/Order/CompleteField.lean index 2effd10d52e6f..6b540009fb273 100644 --- a/Mathlib/Algebra/Order/CompleteField.lean +++ b/Mathlib/Algebra/Order/CompleteField.lean @@ -42,7 +42,6 @@ archimedean. We also construct the natural map from a `LinearOrderedField` to su reals, conditionally complete, ordered field, uniqueness -/ - variable {F α β γ : Type*} noncomputable section diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 2ece61963849f..d71cf4e190305 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -70,7 +70,7 @@ variable {𝕜 : Type*} {E F G H : Type*} open Filter List -open scoped Topology Classical NNReal ENNReal +open scoped Topology NNReal ENNReal section Topological diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index 5e27e4c169c31..76ec01632ad96 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -26,8 +26,7 @@ we prove that they coincide and study their properties (notably convergence). -/ - -open scoped Classical Topology +open scoped Topology open Finset Filter @@ -91,6 +90,7 @@ term is invertible. -/ theorem leftInv_comp (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[𝕜] F) (h : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i) : (leftInv p i).comp p = id 𝕜 E := by ext (n v) + classical match n with | 0 => simp only [leftInv_coeff_zero, ContinuousMultilinearMap.zero_apply, id_apply_ne_one, Ne, @@ -192,8 +192,8 @@ theorem comp_rightInv_aux1 {n : ℕ} (hn : 0 < n) (p : FormalMultilinearSeries (q : FormalMultilinearSeries 𝕜 F E) (v : Fin n → F) : p.comp q n v = ∑ c ∈ {c : Composition n | 1 < c.length}.toFinset, - p c.length (q.applyComposition c v) + - p 1 fun _ => q n v := by + p c.length (q.applyComposition c v) + p 1 fun _ => q n v := by + classical have A : (Finset.univ : Finset (Composition n)) = {c | 1 < Composition.length c}.toFinset ∪ {Composition.single n hn} := by diff --git a/Mathlib/Analysis/Analytic/Linear.lean b/Mathlib/Analysis/Analytic/Linear.lean index c4c00f7a49b5f..6addbfac86dae 100644 --- a/Mathlib/Analysis/Analytic/Linear.lean +++ b/Mathlib/Analysis/Analytic/Linear.lean @@ -17,8 +17,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCom [NormedSpace 𝕜 E] {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} [NormedAddCommGroup G] [NormedSpace 𝕜 G] -open scoped Topology Classical NNReal ENNReal - +open scoped Topology NNReal ENNReal open Set Filter Asymptotics noncomputable section diff --git a/Mathlib/Analysis/Analytic/RadiusLiminf.lean b/Mathlib/Analysis/Analytic/RadiusLiminf.lean index 98a0a6f8d823e..0bee5b674904b 100644 --- a/Mathlib/Analysis/Analytic/RadiusLiminf.lean +++ b/Mathlib/Analysis/Analytic/RadiusLiminf.lean @@ -19,7 +19,7 @@ because this would create a circular dependency once we redefine `exp` using variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] -open scoped Topology Classical NNReal ENNReal +open scoped Topology NNReal ENNReal open Filter Asymptotics diff --git a/Mathlib/Analysis/BoxIntegral/Basic.lean b/Mathlib/Analysis/BoxIntegral/Basic.lean index b6e5359107f9c..1cb018280092e 100644 --- a/Mathlib/Analysis/BoxIntegral/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Basic.lean @@ -50,8 +50,7 @@ non-Riemann filter (e.g., Henstock-Kurzweil and McShane). integral -/ - -open scoped Classical Topology NNReal Filter Uniformity BoxIntegral +open scoped Topology NNReal Filter Uniformity BoxIntegral open Set Finset Function Filter Metric BoxIntegral.IntegrationParams @@ -101,6 +100,7 @@ theorem integralSum_inf_partition (f : ℝⁿ → E) (vol : ι →ᵇᵃ E →L[ integralSum f vol (π.infPrepartition π') = integralSum f vol π := integralSum_biUnion_partition f vol π _ fun _J hJ => h.restrict (Prepartition.le_of_mem _ hJ) +open Classical in theorem integralSum_fiberwise {α} (g : Box ι → α) (f : ℝⁿ → E) (vol : ι →ᵇᵃ E →L[ℝ] F) (π : TaggedPrepartition I) : (∑ y ∈ π.boxes.image g, integralSum f vol (π.filter (g · = y))) = integralSum f vol π := @@ -159,6 +159,7 @@ predicate. -/ def Integrable (I : Box ι) (l : IntegrationParams) (f : ℝⁿ → E) (vol : ι →ᵇᵃ E →L[ℝ] F) := ∃ y, HasIntegral I l f vol y +open Classical in /-- The integral of a function `f` over a box `I` along a filter `l` w.r.t. a volume `vol`. Returns zero on non-integrable functions. -/ def integral (I : Box ι) (l : IntegrationParams) (f : ℝⁿ → E) (vol : ι →ᵇᵃ E →L[ℝ] F) := @@ -260,8 +261,9 @@ theorem integrable_neg : Integrable I l (-f) vol ↔ Integrable I l f vol := ⟨fun h => h.of_neg, fun h => h.neg⟩ @[simp] -theorem integral_neg : integral I l (-f) vol = -integral I l f vol := - if h : Integrable I l f vol then h.hasIntegral.neg.integral_eq +theorem integral_neg : integral I l (-f) vol = -integral I l f vol := by + classical + exact if h : Integrable I l f vol then h.hasIntegral.neg.integral_eq else by rw [integral, integral, dif_neg h, dif_neg (mt Integrable.of_neg h), neg_zero] theorem HasIntegral.sub (h : HasIntegral I l f vol y) (h' : HasIntegral I l g vol y') : @@ -298,6 +300,7 @@ theorem integral_zero : integral I l (fun _ => (0 : E)) vol = 0 := theorem HasIntegral.sum {α : Type*} {s : Finset α} {f : α → ℝⁿ → E} {g : α → F} (h : ∀ i ∈ s, HasIntegral I l (f i) vol (g i)) : HasIntegral I l (fun x => ∑ i ∈ s, f i x) vol (∑ i ∈ s, g i) := by + classical induction' s using Finset.induction_on with a s ha ihs; · simp [hasIntegral_zero] simp only [Finset.sum_insert ha]; rw [Finset.forall_mem_insert] at h exact h.1.add (ihs h.2) @@ -661,6 +664,7 @@ theorem integrable_of_bounded_and_ae_continuousWithinAt [CompleteSpace E] {I : B let t₁ (J : Box ι) : ℝⁿ := (π₁.infPrepartition π₂.toPrepartition).tag J let t₂ (J : Box ι) : ℝⁿ := (π₂.infPrepartition π₁.toPrepartition).tag J let B := (π₁.toPrepartition ⊓ π₂.toPrepartition).boxes + classical let B' := B.filter (fun J ↦ J.toSet ⊆ U) have hB' : B' ⊆ B := B.filter_subset (fun J ↦ J.toSet ⊆ U) have μJ_ne_top : ∀ J ∈ B, μ J ≠ ⊤ := @@ -777,6 +781,7 @@ theorem HasIntegral.of_bRiemann_eq_false_of_forall_isLittleO (hl : l.bRiemann = rcases hs.exists_pos_forall_sum_le (div_pos ε0' H0) with ⟨εs, hεs0, hεs⟩ simp only [le_div_iff' H0, mul_sum] at hεs rcases exists_pos_mul_lt ε0' (B I) with ⟨ε', ε'0, hεI⟩ + classical set δ : ℝ≥0 → ℝⁿ → Ioi (0 : ℝ) := fun c x => if x ∈ s then δ₁ c x (εs x) else (δ₂ c) x ε' refine ⟨δ, fun c => l.rCond_of_bRiemann_eq_false hl, ?_⟩ simp only [Set.mem_iUnion, mem_inter_iff, mem_setOf_eq] diff --git a/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean b/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean index 081b2bd0d1f37..d5bc8119aefce 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean @@ -26,8 +26,7 @@ Then `p I` is true. rectangular box, induction -/ - -open Set Finset Function Filter Metric Classical Topology Filter ENNReal +open Set Function Filter Topology noncomputable section @@ -37,6 +36,7 @@ namespace Box variable {ι : Type*} {I J : Box ι} +open Classical in /-- For a box `I`, the hyperplanes passing through its center split `I` into `2 ^ card ι` boxes. `BoxIntegral.Box.splitCenterBox I s` is one of these boxes. See also `BoxIntegral.Partition.splitCenter` for the corresponding `BoxIntegral.Partition`. -/ diff --git a/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean b/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean index 88b2e6003f974..79429fdb79725 100644 --- a/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean +++ b/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean @@ -38,8 +38,7 @@ Henstock-Kurzweil integral. Henstock-Kurzweil integral, integral, Stokes theorem, divergence theorem -/ - -open scoped Classical NNReal ENNReal Topology BoxIntegral +open scoped NNReal ENNReal Topology BoxIntegral open ContinuousLinearMap (lsmul) diff --git a/Mathlib/Analysis/BoxIntegral/Integrability.lean b/Mathlib/Analysis/BoxIntegral/Integrability.lean index a4dd54bc4ffd7..8de8af41da9e9 100644 --- a/Mathlib/Analysis/BoxIntegral/Integrability.lean +++ b/Mathlib/Analysis/BoxIntegral/Integrability.lean @@ -21,8 +21,7 @@ We deduce that the same is true for the Riemann integral for continuous function integral, McShane integral, Bochner integral -/ - -open scoped Classical NNReal ENNReal Topology +open scoped NNReal ENNReal Topology universe u v @@ -63,6 +62,7 @@ theorem hasIntegralIndicatorConst (l : IntegrationParams) (hl : l.bRiemann = fal nhds_basis_closedBall.mem_iff.1 (hFc.isOpen_compl.mem_nhds fun hx' => hx.2 (hFs hx').1) exact ⟨⟨r, hr₀⟩, hr⟩ choose! rs' hrs'F using this + classical set r : (ι → ℝ) → Ioi (0 : ℝ) := s.piecewise rs rs' refine ⟨fun _ => r, fun c => l.rCond_of_bRiemann_eq_false hl, fun c π hπ hπp => ?_⟩; rw [mul_comm] /- Then the union of boxes `J ∈ π` such that `π.tag ∈ s` includes `F` and is included by `U`, diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Measure.lean b/Mathlib/Analysis/BoxIntegral/Partition/Measure.lean index 3be9288909dcd..8e7cf04ecf4b8 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Measure.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Measure.lean @@ -23,12 +23,11 @@ For the last statement, we both prove it as a proposition and define a bundled rectangular box, measure -/ - open Set noncomputable section -open scoped ENNReal Classical BoxIntegral +open scoped ENNReal BoxIntegral variable {ι : Type*} diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 3d4f7a6ddaa30..d45ba684d6e3d 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -32,10 +32,9 @@ In this file, we denote `⊤ : ℕ∞` with `∞`. derivative, differentiability, higher derivative, `C^n`, multilinear, Taylor series, formal series -/ - noncomputable section -open scoped Classical NNReal Nat +open scoped NNReal Nat local notation "∞" => (⊤ : ℕ∞) diff --git a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean index a2079964ef670..ba117c50a4db6 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean @@ -18,7 +18,7 @@ import Mathlib.Data.Nat.Choose.Multinomial noncomputable section -open scoped Classical NNReal Nat +open scoped NNReal Nat universe u uD uE uF uG diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index f323a5724a986..edda887e12373 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -90,8 +90,7 @@ universe u v w noncomputable section -open scoped Classical Topology Filter ENNReal NNReal - +open scoped Topology ENNReal NNReal open Filter Asymptotics Set open ContinuousLinearMap (smulRight smulRight_one_eq_iff) diff --git a/Mathlib/Analysis/Calculus/Deriv/Mul.lean b/Mathlib/Analysis/Calculus/Deriv/Mul.lean index 4598777c30c98..64628ceb448d0 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Mul.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Mul.lean @@ -20,12 +20,11 @@ For a more detailed overview of one-dimensional derivatives in mathlib, see the derivative, multiplication -/ - universe u v w noncomputable section -open scoped Classical Topology Filter ENNReal +open scoped Topology Filter ENNReal open Filter Asymptotics Set @@ -326,12 +325,15 @@ variable {ι : Type*} {𝔸' : Type*} [NormedCommRing 𝔸'] [NormedAlgebra 𝕜 {u : Finset ι} {f : ι → 𝕜 → 𝔸'} {f' : ι → 𝔸'} theorem DifferentiableAt.finset_prod (hd : ∀ i ∈ u, DifferentiableAt 𝕜 (f i) x) : - DifferentiableAt 𝕜 (∏ i ∈ u, f i ·) x := - (HasDerivAt.finset_prod (fun i hi ↦ DifferentiableAt.hasDerivAt (hd i hi))).differentiableAt + DifferentiableAt 𝕜 (∏ i ∈ u, f i ·) x := by + classical + exact + (HasDerivAt.finset_prod (fun i hi ↦ DifferentiableAt.hasDerivAt (hd i hi))).differentiableAt theorem DifferentiableWithinAt.finset_prod (hd : ∀ i ∈ u, DifferentiableWithinAt 𝕜 (f i) s x) : - DifferentiableWithinAt 𝕜 (∏ i ∈ u, f i ·) s x := - (HasDerivWithinAt.finset_prod (fun i hi ↦ + DifferentiableWithinAt 𝕜 (∏ i ∈ u, f i ·) s x := by + classical + exact (HasDerivWithinAt.finset_prod (fun i hi ↦ DifferentiableWithinAt.hasDerivWithinAt (hd i hi))).differentiableWithinAt theorem DifferentiableOn.finset_prod (hd : ∀ i ∈ u, DifferentiableOn 𝕜 (f i) s) : diff --git a/Mathlib/Analysis/Calculus/Deriv/Polynomial.lean b/Mathlib/Analysis/Calculus/Deriv/Polynomial.lean index da976e0255227..3174575fb194d 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Polynomial.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Polynomial.lean @@ -30,9 +30,8 @@ derivative, polynomial universe u v w -open scoped Classical Topology Filter ENNReal Polynomial - -open Filter Asymptotics Set +open scoped Topology Filter ENNReal Polynomial +open Set open ContinuousLinearMap (smulRight smulRight_one_eq_iff) diff --git a/Mathlib/Analysis/Calculus/Dslope.lean b/Mathlib/Analysis/Calculus/Dslope.lean index d2991b5fe7092..a35fb76d7fa45 100644 --- a/Mathlib/Analysis/Calculus/Dslope.lean +++ b/Mathlib/Analysis/Calculus/Dslope.lean @@ -17,26 +17,28 @@ In this file we define `dslope` and prove some basic lemmas about its continuity differentiability. -/ - -open scoped Classical Topology Filter +open scoped Topology Filter open Function Set Filter variable {𝕜 E : Type*} [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] +open Classical in /-- `dslope f a b` is defined as `slope f a b = (b - a)⁻¹ • (f b - f a)` for `a ≠ b` and `deriv f a` for `a = b`. -/ noncomputable def dslope (f : 𝕜 → E) (a : 𝕜) : 𝕜 → E := update (slope f a) a (deriv f a) @[simp] -theorem dslope_same (f : 𝕜 → E) (a : 𝕜) : dslope f a a = deriv f a := - update_same _ _ _ +theorem dslope_same (f : 𝕜 → E) (a : 𝕜) : dslope f a a = deriv f a := by + classical + exact update_same _ _ _ variable {f : 𝕜 → E} {a b : 𝕜} {s : Set 𝕜} -theorem dslope_of_ne (f : 𝕜 → E) (h : b ≠ a) : dslope f a b = slope f a b := - update_noteq h _ _ +theorem dslope_of_ne (f : 𝕜 → E) (h : b ≠ a) : dslope f a b = slope f a b := by + classical + exact update_noteq h _ _ theorem ContinuousLinearMap.dslope_comp {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] (f : E →L[𝕜] F) (g : 𝕜 → E) (a b : 𝕜) (H : a = b → DifferentiableAt 𝕜 g a) : @@ -89,6 +91,7 @@ theorem ContinuousOn.of_dslope (h : ContinuousOn (dslope f a) s) : ContinuousOn theorem continuousWithinAt_dslope_of_ne (h : b ≠ a) : ContinuousWithinAt (dslope f a) s b ↔ ContinuousWithinAt f s b := by refine ⟨ContinuousWithinAt.of_dslope, fun hc => ?_⟩ + classical simp only [dslope, continuousWithinAt_update_of_ne h] exact ((continuousWithinAt_id.sub continuousWithinAt_const).inv₀ (sub_ne_zero.2 h)).smul (hc.sub continuousWithinAt_const) diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean index 8517d2665d2bd..378ee5517382e 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean @@ -45,7 +45,7 @@ We introduce some `local notation` to make formulas shorter: open Function Set Filter Metric -open scoped Topology Classical NNReal +open scoped Topology NNReal noncomputable section diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean index 6ccbe6ae8b90a..a76f0b66ac46e 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean @@ -40,7 +40,7 @@ derivative, strictly differentiable, continuously differentiable, smooth, invers open Function Set Filter Metric -open scoped Topology Classical NNReal +open scoped Topology NNReal noncomputable section diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean index 1cbee9761a0a7..32cf207826227 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean @@ -40,11 +40,9 @@ by translating the corresponding result `iteratedFDerivWithin_succ_apply_left` f iterated Fréchet derivative. -/ - noncomputable section -open scoped Classical Topology - +open scoped Topology open Filter Asymptotics Set variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] diff --git a/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean b/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean index fec7a2aa7d864..bfa8f865b600c 100644 --- a/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean +++ b/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean @@ -58,7 +58,7 @@ universe u v open Filter Set -open scoped Topology Classical Convex +open scoped Topology Convex section Module @@ -134,9 +134,11 @@ theorem IsLocalMaxOn.hasFDerivWithinAt_nonpos (h : IsLocalMaxOn f s a) /-- If `f` has a local max on `s` at `a` and `y` belongs to the positive tangent cone of `s` at `a`, then `f' y ≤ 0`. -/ theorem IsLocalMaxOn.fderivWithin_nonpos (h : IsLocalMaxOn f s a) - (hy : y ∈ posTangentConeAt s a) : (fderivWithin ℝ f s a : E → ℝ) y ≤ 0 := - if hf : DifferentiableWithinAt ℝ f s a then h.hasFDerivWithinAt_nonpos hf.hasFDerivWithinAt hy - else by rw [fderivWithin_zero_of_not_differentiableWithinAt hf]; rfl + (hy : y ∈ posTangentConeAt s a) : (fderivWithin ℝ f s a : E → ℝ) y ≤ 0 := by + classical + exact + if hf : DifferentiableWithinAt ℝ f s a then h.hasFDerivWithinAt_nonpos hf.hasFDerivWithinAt hy + else by rw [fderivWithin_zero_of_not_differentiableWithinAt hf]; rfl /-- If `f` has a local max on `s` at `a`, `f'` is a derivative of `f` at `a` within `s`, and both `y` and `-y` belong to the positive tangent cone of `s` at `a`, then `f' y ≤ 0`. -/ @@ -149,8 +151,9 @@ theorem IsLocalMaxOn.hasFDerivWithinAt_eq_zero (h : IsLocalMaxOn f s a) of `s` at `a`, then `f' y = 0`. -/ theorem IsLocalMaxOn.fderivWithin_eq_zero (h : IsLocalMaxOn f s a) (hy : y ∈ posTangentConeAt s a) (hy' : -y ∈ posTangentConeAt s a) : - (fderivWithin ℝ f s a : E → ℝ) y = 0 := - if hf : DifferentiableWithinAt ℝ f s a then + (fderivWithin ℝ f s a : E → ℝ) y = 0 := by + classical + exact if hf : DifferentiableWithinAt ℝ f s a then h.hasFDerivWithinAt_eq_zero hf.hasFDerivWithinAt hy hy' else by rw [fderivWithin_zero_of_not_differentiableWithinAt hf]; rfl @@ -163,9 +166,11 @@ theorem IsLocalMinOn.hasFDerivWithinAt_nonneg (h : IsLocalMinOn f s a) /-- If `f` has a local min on `s` at `a` and `y` belongs to the positive tangent cone of `s` at `a`, then `0 ≤ f' y`. -/ theorem IsLocalMinOn.fderivWithin_nonneg (h : IsLocalMinOn f s a) - (hy : y ∈ posTangentConeAt s a) : (0 : ℝ) ≤ (fderivWithin ℝ f s a : E → ℝ) y := - if hf : DifferentiableWithinAt ℝ f s a then h.hasFDerivWithinAt_nonneg hf.hasFDerivWithinAt hy - else by rw [fderivWithin_zero_of_not_differentiableWithinAt hf]; rfl + (hy : y ∈ posTangentConeAt s a) : (0 : ℝ) ≤ (fderivWithin ℝ f s a : E → ℝ) y := by + classical + exact + if hf : DifferentiableWithinAt ℝ f s a then h.hasFDerivWithinAt_nonneg hf.hasFDerivWithinAt hy + else by rw [fderivWithin_zero_of_not_differentiableWithinAt hf]; rfl /-- If `f` has a local max on `s` at `a`, `f'` is a derivative of `f` at `a` within `s`, and both `y` and `-y` belong to the positive tangent cone of `s` at `a`, then `f' y ≤ 0`. -/ @@ -178,8 +183,9 @@ theorem IsLocalMinOn.hasFDerivWithinAt_eq_zero (h : IsLocalMinOn f s a) of `s` at `a`, then `f' y = 0`. -/ theorem IsLocalMinOn.fderivWithin_eq_zero (h : IsLocalMinOn f s a) (hy : y ∈ posTangentConeAt s a) (hy' : -y ∈ posTangentConeAt s a) : - (fderivWithin ℝ f s a : E → ℝ) y = 0 := - if hf : DifferentiableWithinAt ℝ f s a then + (fderivWithin ℝ f s a : E → ℝ) y = 0 := by + classical + exact if hf : DifferentiableWithinAt ℝ f s a then h.hasFDerivWithinAt_eq_zero hf.hasFDerivWithinAt hy hy' else by rw [fderivWithin_zero_of_not_differentiableWithinAt hf]; rfl @@ -191,8 +197,9 @@ theorem IsLocalMin.hasFDerivAt_eq_zero (h : IsLocalMin f a) (hf : HasFDerivAt f apply mem_univ /-- **Fermat's Theorem**: the derivative of a function at a local minimum equals zero. -/ -theorem IsLocalMin.fderiv_eq_zero (h : IsLocalMin f a) : fderiv ℝ f a = 0 := - if hf : DifferentiableAt ℝ f a then h.hasFDerivAt_eq_zero hf.hasFDerivAt +theorem IsLocalMin.fderiv_eq_zero (h : IsLocalMin f a) : fderiv ℝ f a = 0 := by + classical + exact if hf : DifferentiableAt ℝ f a then h.hasFDerivAt_eq_zero hf.hasFDerivAt else fderiv_zero_of_not_differentiableAt hf /-- **Fermat's Theorem**: the derivative of a function at a local maximum equals zero. -/ @@ -200,8 +207,9 @@ theorem IsLocalMax.hasFDerivAt_eq_zero (h : IsLocalMax f a) (hf : HasFDerivAt f neg_eq_zero.1 <| h.neg.hasFDerivAt_eq_zero hf.neg /-- **Fermat's Theorem**: the derivative of a function at a local maximum equals zero. -/ -theorem IsLocalMax.fderiv_eq_zero (h : IsLocalMax f a) : fderiv ℝ f a = 0 := - if hf : DifferentiableAt ℝ f a then h.hasFDerivAt_eq_zero hf.hasFDerivAt +theorem IsLocalMax.fderiv_eq_zero (h : IsLocalMax f a) : fderiv ℝ f a = 0 := by + classical + exact if hf : DifferentiableAt ℝ f a then h.hasFDerivAt_eq_zero hf.hasFDerivAt else fderiv_zero_of_not_differentiableAt hf /-- **Fermat's Theorem**: the derivative of a function at a local extremum equals zero. -/ @@ -249,8 +257,9 @@ theorem IsLocalMin.hasDerivAt_eq_zero (h : IsLocalMin f a) (hf : HasDerivAt f f' simpa using DFunLike.congr_fun (h.hasFDerivAt_eq_zero (hasDerivAt_iff_hasFDerivAt.1 hf)) 1 /-- **Fermat's Theorem**: the derivative of a function at a local minimum equals zero. -/ -theorem IsLocalMin.deriv_eq_zero (h : IsLocalMin f a) : deriv f a = 0 := - if hf : DifferentiableAt ℝ f a then h.hasDerivAt_eq_zero hf.hasDerivAt +theorem IsLocalMin.deriv_eq_zero (h : IsLocalMin f a) : deriv f a = 0 := by + classical + exact if hf : DifferentiableAt ℝ f a then h.hasDerivAt_eq_zero hf.hasDerivAt else deriv_zero_of_not_differentiableAt hf /-- **Fermat's Theorem**: the derivative of a function at a local maximum equals zero. -/ @@ -258,8 +267,9 @@ theorem IsLocalMax.hasDerivAt_eq_zero (h : IsLocalMax f a) (hf : HasDerivAt f f' neg_eq_zero.1 <| h.neg.hasDerivAt_eq_zero hf.neg /-- **Fermat's Theorem**: the derivative of a function at a local maximum equals zero. -/ -theorem IsLocalMax.deriv_eq_zero (h : IsLocalMax f a) : deriv f a = 0 := - if hf : DifferentiableAt ℝ f a then h.hasDerivAt_eq_zero hf.hasDerivAt +theorem IsLocalMax.deriv_eq_zero (h : IsLocalMax f a) : deriv f a = 0 := by + classical + exact if hf : DifferentiableAt ℝ f a then h.hasDerivAt_eq_zero hf.hasDerivAt else deriv_zero_of_not_differentiableAt hf /-- **Fermat's Theorem**: the derivative of a function at a local extremum equals zero. -/ diff --git a/Mathlib/Analysis/Calculus/LocalExtr/LineDeriv.lean b/Mathlib/Analysis/Calculus/LocalExtr/LineDeriv.lean index 82152d6b411b8..0a903aeacbc7e 100644 --- a/Mathlib/Analysis/Calculus/LocalExtr/LineDeriv.lean +++ b/Mathlib/Analysis/Calculus/LocalExtr/LineDeriv.lean @@ -14,7 +14,7 @@ In this file we prove several versions of this fact for line derivatives. -/ open Function Set Filter -open scoped Classical Topology +open scoped Topology section Module @@ -25,8 +25,9 @@ theorem IsExtrFilter.hasLineDerivAt_eq_zero {l : Filter E} (h : IsExtrFilter f l IsLocalExtr.hasDerivAt_eq_zero (IsExtrFilter.comp_tendsto (by simpa using h) h') hd theorem IsExtrFilter.lineDeriv_eq_zero {l : Filter E} (h : IsExtrFilter f l a) - (h' : Tendsto (fun t : ℝ ↦ a + t • b) (𝓝 0) l) : lineDeriv ℝ f a b = 0 := - if hd : LineDifferentiableAt ℝ f a b then + (h' : Tendsto (fun t : ℝ ↦ a + t • b) (𝓝 0) l) : lineDeriv ℝ f a b = 0 := by + classical + exact if hd : LineDifferentiableAt ℝ f a b then h.hasLineDerivAt_eq_zero hd.hasLineDerivAt h' else lineDeriv_zero_of_not_lineDifferentiableAt hd @@ -60,8 +61,9 @@ theorem IsExtrOn.hasLineDerivWithinAt_eq_zero (h : IsExtrOn f s a) h.hasLineDerivAt_eq_zero (hd.hasLineDerivAt' h') h' theorem IsExtrOn.lineDerivWithin_eq_zero (h : IsExtrOn f s a) - (h' : ∀ᶠ t : ℝ in 𝓝 0, a + t • b ∈ s) : lineDerivWithin ℝ f s a b = 0 := - if hd : LineDifferentiableWithinAt ℝ f s a b then + (h' : ∀ᶠ t : ℝ in 𝓝 0, a + t • b ∈ s) : lineDerivWithin ℝ f s a b = 0 := by + classical + exact if hd : LineDifferentiableWithinAt ℝ f s a b then h.hasLineDerivWithinAt_eq_zero hd.hasLineDerivWithinAt h' else lineDerivWithin_zero_of_not_lineDifferentiableWithinAt hd diff --git a/Mathlib/Analysis/Calculus/MeanValue.lean b/Mathlib/Analysis/Calculus/MeanValue.lean index 4584f80bb2cc1..a52881c81d539 100644 --- a/Mathlib/Analysis/Calculus/MeanValue.lean +++ b/Mathlib/Analysis/Calculus/MeanValue.lean @@ -66,13 +66,12 @@ In this file we prove the following facts: strictly differentiable. (This is a corollary of the mean value inequality.) -/ - variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] open Metric Set Asymptotics ContinuousLinearMap Filter -open scoped Classical Topology NNReal +open scoped Topology NNReal /-! ### One-dimensional fencing inequalities -/ diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 146016e97eba6..6afa53eec6ef1 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -23,12 +23,10 @@ We define the notation `ℍ` for the upper half plane available in the locale `UpperHalfPlane` so as not to conflict with the quaternions. -/ - noncomputable section open Matrix Matrix.SpecialLinearGroup - -open scoped Classical MatrixGroups +open scoped MatrixGroups /- Disable these instances as they are not the simp-normal form, and having them disabled ensures we state lemmas in this file without spurious `coe_fn` terms. -/ diff --git a/Mathlib/Analysis/Convex/Deriv.lean b/Mathlib/Analysis/Convex/Deriv.lean index d04d0f4953ff5..20066fe9abaa8 100644 --- a/Mathlib/Analysis/Convex/Deriv.lean +++ b/Mathlib/Analysis/Convex/Deriv.lean @@ -20,8 +20,7 @@ Here we relate convexity of functions `ℝ → ℝ` to properties of their deriv -/ open Metric Set Asymptotics ContinuousLinearMap Filter - -open scoped Classical Topology NNReal +open scoped Topology NNReal /-! ## Monotonicity of `f'` implies convexity of `f` diff --git a/Mathlib/Analysis/InnerProductSpace/Calculus.lean b/Mathlib/Analysis/InnerProductSpace/Calculus.lean index f7794212be464..eae31940da226 100644 --- a/Mathlib/Analysis/InnerProductSpace/Calculus.lean +++ b/Mathlib/Analysis/InnerProductSpace/Calculus.lean @@ -28,8 +28,6 @@ noncomputable section open RCLike Real Filter -open scoped Classical Topology - section DerivInner variable {𝕜 E F : Type*} [RCLike 𝕜] diff --git a/Mathlib/Analysis/InnerProductSpace/l2Space.lean b/Mathlib/Analysis/InnerProductSpace/l2Space.lean index 5980076f937fc..a93c696be81f4 100644 --- a/Mathlib/Analysis/InnerProductSpace/l2Space.lean +++ b/Mathlib/Analysis/InnerProductSpace/l2Space.lean @@ -79,10 +79,8 @@ We also define a *predicate* `IsHilbertSum 𝕜 G V`, where `V : Π i, G i → Hilbert space, Hilbert sum, l2, Hilbert basis, unitary equivalence, isometric isomorphism -/ - open RCLike Submodule Filter - -open scoped NNReal ENNReal Classical ComplexConjugate Topology +open scoped NNReal ENNReal ComplexConjugate Topology noncomputable section @@ -156,7 +154,8 @@ theorem inner_eq_tsum (f g : lp G 2) : ⟪f, g⟫ = ∑' i, ⟪f i, g i⟫ := theorem hasSum_inner (f g : lp G 2) : HasSum (fun i => ⟪f i, g i⟫) ⟪f, g⟫ := (summable_inner f g).hasSum -theorem inner_single_left (i : ι) (a : G i) (f : lp G 2) : ⟪lp.single 2 i a, f⟫ = ⟪a, f i⟫ := by +theorem inner_single_left [DecidableEq ι] (i : ι) (a : G i) (f : lp G 2) : + ⟪lp.single 2 i a, f⟫ = ⟪a, f i⟫ := by refine (hasSum_inner (lp.single 2 i a) f).unique ?_ convert hasSum_ite_eq i ⟪a, f i⟫ using 1 ext j @@ -165,8 +164,9 @@ theorem inner_single_left (i : ι) (a : G i) (f : lp G 2) : ⟪lp.single 2 i a, · subst h; rfl · simp -theorem inner_single_right (i : ι) (a : G i) (f : lp G 2) : ⟪f, lp.single 2 i a⟫ = ⟪f i, a⟫ := by - simpa [inner_conj_symm] using congr_arg conj (@inner_single_left _ 𝕜 _ _ _ _ i a f) +theorem inner_single_right [DecidableEq ι] (i : ι) (a : G i) (f : lp G 2) : + ⟪f, lp.single 2 i a⟫ = ⟪f i, a⟫ := by + simpa [inner_conj_symm] using congr_arg conj (inner_single_left (𝕜 := 𝕜) i a f) end lp @@ -212,7 +212,7 @@ protected theorem hasSum_linearIsometry (f : lp G 2) : (hV.summable_of_lp f).hasSum @[simp] -protected theorem linearIsometry_apply_single {i : ι} (x : G i) : +protected theorem linearIsometry_apply_single [DecidableEq ι] {i : ι} (x : G i) : hV.linearIsometry (lp.single 2 i x) = V i x := by rw [hV.linearIsometry_apply, ← tsum_ite_eq i (V i x)] congr @@ -222,8 +222,8 @@ protected theorem linearIsometry_apply_single {i : ι} (x : G i) : · subst h; simp · simp [h] -protected theorem linearIsometry_apply_dfinsupp_sum_single (W₀ : Π₀ i : ι, G i) : - hV.linearIsometry (W₀.sum (lp.single 2)) = W₀.sum fun i => V i := by +protected theorem linearIsometry_apply_dfinsupp_sum_single [DecidableEq ι] [∀ i, DecidableEq (G i)] + (W₀ : Π₀ i : ι, G i) : hV.linearIsometry (W₀.sum (lp.single 2)) = W₀.sum fun i => V i := by simp /-- The canonical linear isometry from the `lp 2` of a mutually orthogonal family of subspaces of @@ -232,6 +232,7 @@ protected theorem range_linearIsometry [∀ i, CompleteSpace (G i)] : LinearMap.range hV.linearIsometry.toLinearMap = (⨆ i, LinearMap.range (V i).toLinearMap).topologicalClosure := by -- Porting note: dot notation broken + classical refine le_antisymm ?_ ?_ · rintro x ⟨f, rfl⟩ refine mem_closure_of_tendsto (hV.hasSum_linearIsometry f) (eventually_of_forall ?_) @@ -313,15 +314,16 @@ protected theorem IsHilbertSum.hasSum_linearIsometryEquiv_symm (hV : IsHilbertSu `lp G 2`, an "elementary basis vector" in `lp G 2` supported at `i : ι` is the image of the associated element in `E`. -/ @[simp] -protected theorem IsHilbertSum.linearIsometryEquiv_symm_apply_single (hV : IsHilbertSum 𝕜 G V) - {i : ι} (x : G i) : hV.linearIsometryEquiv.symm (lp.single 2 i x) = V i x := by +protected theorem IsHilbertSum.linearIsometryEquiv_symm_apply_single + [DecidableEq ι] (hV : IsHilbertSum 𝕜 G V) {i : ι} (x : G i) : + hV.linearIsometryEquiv.symm (lp.single 2 i x) = V i x := by simp [IsHilbertSum.linearIsometryEquiv, OrthogonalFamily.linearIsometry_apply_single] /-- In the canonical isometric isomorphism between a Hilbert sum `E` of `G : ι → Type*` and `lp G 2`, a finitely-supported vector in `lp G 2` is the image of the associated finite sum of elements of `E`. -/ protected theorem IsHilbertSum.linearIsometryEquiv_symm_apply_dfinsupp_sum_single - (hV : IsHilbertSum 𝕜 G V) (W₀ : Π₀ i : ι, G i) : + [DecidableEq ι] [∀ i, DecidableEq (G i)] (hV : IsHilbertSum 𝕜 G V) (W₀ : Π₀ i : ι, G i) : hV.linearIsometryEquiv.symm (W₀.sum (lp.single 2)) = W₀.sum fun i => V i := by simp only [map_dfinsupp_sum, IsHilbertSum.linearIsometryEquiv_symm_apply_single] @@ -330,7 +332,7 @@ protected theorem IsHilbertSum.linearIsometryEquiv_symm_apply_dfinsupp_sum_singl elements of `E`. -/ @[simp] protected theorem IsHilbertSum.linearIsometryEquiv_apply_dfinsupp_sum_single - (hV : IsHilbertSum 𝕜 G V) (W₀ : Π₀ i : ι, G i) : + [DecidableEq ι] [∀ i, DecidableEq (G i)] (hV : IsHilbertSum 𝕜 G V) (W₀ : Π₀ i : ι, G i) : ((W₀.sum (γ := lp G 2) fun a b ↦ hV.linearIsometryEquiv (V a b)) : ∀ i, G i) = W₀ := by rw [← map_dfinsupp_sum] rw [← hV.linearIsometryEquiv_symm_apply_dfinsupp_sum_single] @@ -380,26 +382,30 @@ namespace HilbertBasis instance {ι : Type*} : Inhabited (HilbertBasis ι 𝕜 ℓ²(ι, 𝕜)) := ⟨ofRepr (LinearIsometryEquiv.refl 𝕜 _)⟩ +open Classical in /-- `b i` is the `i`th basis vector. -/ instance instCoeFun : CoeFun (HilbertBasis ι 𝕜 E) fun _ => ι → E where coe b i := b.repr.symm (lp.single 2 i (1 : 𝕜)) -- This is a bad `@[simp]` lemma: the RHS is a coercion containing the LHS. -protected theorem repr_symm_single (b : HilbertBasis ι 𝕜 E) (i : ι) : - b.repr.symm (lp.single 2 i (1 : 𝕜)) = b i := - rfl +protected theorem repr_symm_single [DecidableEq ι] (b : HilbertBasis ι 𝕜 E) (i : ι) : + b.repr.symm (lp.single 2 i (1 : 𝕜)) = b i := by + convert rfl -protected theorem repr_self (b : HilbertBasis ι 𝕜 E) (i : ι) : +protected theorem repr_self [DecidableEq ι] (b : HilbertBasis ι 𝕜 E) (i : ι) : b.repr (b i) = lp.single 2 i (1 : 𝕜) := by simp only [LinearIsometryEquiv.apply_symm_apply] + convert rfl protected theorem repr_apply_apply (b : HilbertBasis ι 𝕜 E) (v : E) (i : ι) : b.repr v i = ⟪b i, v⟫ := by + classical rw [← b.repr.inner_map_map (b i) v, b.repr_self, lp.inner_single_left] simp @[simp] protected theorem orthonormal (b : HilbertBasis ι 𝕜 E) : Orthonormal 𝕜 b := by + classical rw [orthonormal_iff_ite] intro i j rw [← b.repr.inner_map_map (b i) (b j), b.repr_self, b.repr_self, lp.inner_single_left, @@ -408,6 +414,7 @@ protected theorem orthonormal (b : HilbertBasis ι 𝕜 E) : Orthonormal 𝕜 b protected theorem hasSum_repr_symm (b : HilbertBasis ι 𝕜 E) (f : ℓ²(ι, 𝕜)) : HasSum (fun i => f i • b i) (b.repr.symm f) := by + classical suffices H : (fun i : ι => f i • b i) = fun b_1 : ι => b.repr.symm.toContinuousLinearEquiv <| (fun i : ι => lp.single 2 i (f i) (E := (fun _ : ι => 𝕜))) b_1 by rw [H] @@ -461,6 +468,7 @@ protected def toOrthonormalBasis [Fintype ι] (b : HilbertBasis ι 𝕜 E) : Ort OrthonormalBasis.mk b.orthonormal (by refine Eq.ge ?_ + classical have := (span 𝕜 (Finset.univ.image b : Set E)).closed_of_finiteDimensional simpa only [Finset.coe_image, Finset.coe_univ, Set.image_univ, HilbertBasis.dense_span] using this.submodule_topologicalClosure_eq.symm) @@ -476,7 +484,7 @@ protected theorem hasSum_orthogonalProjection {U : Submodule 𝕜 E} [CompleteSp simpa only [b.repr_apply_apply, inner_orthogonalProjection_eq_of_mem_left] using b.hasSum_repr (orthogonalProjection U x) -theorem finite_spans_dense (b : HilbertBasis ι 𝕜 E) : +theorem finite_spans_dense [DecidableEq E] (b : HilbertBasis ι 𝕜 E) : (⨆ J : Finset ι, span 𝕜 (J.image b : Set E)).topologicalClosure = ⊤ := eq_top_iff.mpr <| b.dense_span.ge.trans (by simp_rw [← Submodule.span_iUnion] @@ -490,7 +498,7 @@ variable {v : ι → E} (hv : Orthonormal 𝕜 v) protected def mk (hsp : ⊤ ≤ (span 𝕜 (Set.range v)).topologicalClosure) : HilbertBasis ι 𝕜 E := HilbertBasis.ofRepr <| (hv.isHilbertSum hsp).linearIsometryEquiv -theorem _root_.Orthonormal.linearIsometryEquiv_symm_apply_single_one (h i) : +theorem _root_.Orthonormal.linearIsometryEquiv_symm_apply_single_one [DecidableEq ι] (h i) : (hv.isHilbertSum h).linearIsometryEquiv.symm (lp.single 2 i 1) = v i := by rw [IsHilbertSum.linearIsometryEquiv_symm_apply_single, LinearIsometry.toSpanSingleton_apply, one_smul] @@ -498,6 +506,7 @@ theorem _root_.Orthonormal.linearIsometryEquiv_symm_apply_single_one (h i) : @[simp] protected theorem coe_mk (hsp : ⊤ ≤ (span 𝕜 (Set.range v)).topologicalClosure) : ⇑(HilbertBasis.mk hv hsp) = v := by + classical apply funext <| Orthonormal.linearIsometryEquiv_symm_apply_single_one hv hsp /-- An orthonormal family of vectors whose span has trivial orthogonal complement is a Hilbert diff --git a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean index e641f4975fd6e..69ee742513d76 100644 --- a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean +++ b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean @@ -44,13 +44,11 @@ then the identities from `E` to `E'` and from `E'`to `E` are continuous thanks t `LinearMap.continuous_of_finiteDimensional`. This gives the desired norm equivalence. -/ - universe u v w x noncomputable section -open Set FiniteDimensional TopologicalSpace Filter Asymptotics Classical Topology - NNReal Metric +open Set FiniteDimensional TopologicalSpace Filter Asymptotics Topology NNReal Metric namespace LinearIsometry @@ -153,6 +151,7 @@ theorem ContinuousLinearMap.continuous_det : Continuous fun f : E →L[𝕜] E = by_cases h : ∃ s : Finset E, Nonempty (Basis (↥s) 𝕜 E) · rcases h with ⟨s, ⟨b⟩⟩ haveI : FiniteDimensional 𝕜 E := FiniteDimensional.of_fintype_basis b + classical simp_rw [LinearMap.det_eq_det_toMatrix_of_finset b] refine Continuous.matrix_det ?_ exact @@ -239,6 +238,7 @@ theorem ContinuousLinearMap.isOpen_injective [FiniteDimensional 𝕜 E] : protected theorem LinearIndependent.eventually {ι} [Finite ι] {f : ι → E} (hf : LinearIndependent 𝕜 f) : ∀ᶠ g in 𝓝 f, LinearIndependent 𝕜 g := by cases nonempty_fintype ι + classical simp only [Fintype.linearIndependent_iff'] at hf ⊢ rcases LinearMap.exists_antilipschitzWith _ hf with ⟨K, K0, hK⟩ have : Tendsto (fun g : ι → E => ∑ i, ‖g i - f i‖) (𝓝 f) (𝓝 <| ∑ i, ‖f i - f i‖) := diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean index 1993c366e6d1b..c69364cab074f 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean @@ -30,7 +30,7 @@ suppress_compilation open Bornology open Filter hiding map_smul -open scoped Classical NNReal Topology Uniformity +open scoped NNReal Topology Uniformity -- the `ₗ` subscript variables are for special cases about linear (as opposed to semilinear) maps variable {𝕜 𝕜₂ 𝕜₃ E Eₗ F Fₗ G Gₗ 𝓕 : Type*} diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean index d76f44dd9ab98..8be1705a830fc 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean @@ -19,7 +19,7 @@ suppress_compilation open Bornology open Filter hiding map_smul -open scoped Classical NNReal Topology Uniformity +open scoped NNReal Topology Uniformity -- the `ₗ` subscript variables are for special cases about linear (as opposed to semilinear) maps variable {𝕜 𝕜₂ 𝕜₃ E Eₗ F Fₗ G Gₗ 𝓕 : Type*} diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean index ba20e4a471025..62d8d4d30f47c 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean @@ -17,7 +17,7 @@ suppress_compilation open Bornology Metric Set Real open Filter hiding map_smul -open scoped Classical NNReal Topology Uniformity +open scoped NNReal Topology Uniformity -- the `ₗ` subscript variables are for special cases about linear (as opposed to semilinear) maps variable {𝕜 𝕜₂ 𝕜₃ E Eₗ F Fₗ G Gₗ 𝓕 : Type*} diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean index 455753c23595d..1886ea8a03f89 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean @@ -15,9 +15,8 @@ of multiplication and scalar-multiplication operations in normed algebras and no suppress_compilation - open Metric -open scoped Classical NNReal Topology Uniformity +open scoped NNReal Topology Uniformity variable {𝕜 E : Type*} [NontriviallyNormedField 𝕜] diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean index 664473cd62040..c57a94705a314 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/NormedSpace.lean @@ -18,7 +18,7 @@ suppress_compilation open Bornology open Filter hiding map_smul -open scoped Classical NNReal Topology Uniformity +open scoped NNReal Topology Uniformity -- the `ₗ` subscript variables are for special cases about linear (as opposed to semilinear) maps variable {𝕜 𝕜₂ 𝕜₃ E Eₗ F Fₗ G Gₗ 𝓕 : Type*} diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index 283c31b2e3971..d09579fb63b54 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -27,14 +27,12 @@ Sec. 4.5][HubbardWest-ode], where `norm_le_gronwallBound_of_norm_deriv_right_le` of `K x` and `f x`. -/ +open Metric Set Asymptotics Filter Real +open scoped Topology NNReal variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] -open Metric Set Asymptotics Filter Real - -open scoped Classical Topology NNReal - /-! ### Technical lemmas about `gronwallBound` -/ diff --git a/Mathlib/Analysis/SpecialFunctions/Bernstein.lean b/Mathlib/Analysis/SpecialFunctions/Bernstein.lean index bfc43cb6b00fb..5cfa043e78334 100644 --- a/Mathlib/Analysis/SpecialFunctions/Bernstein.lean +++ b/Mathlib/Analysis/SpecialFunctions/Bernstein.lean @@ -43,10 +43,9 @@ This result proves Weierstrass' theorem that polynomials are dense in `C([0,1], although we defer an abstract statement of this until later. -/ - noncomputable section -open scoped Classical BoundedContinuousFunction unitInterval +open scoped BoundedContinuousFunction unitInterval /-- The Bernstein polynomials, as continuous functions on `[0,1]`. -/ diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index d5dc0f84e7bc8..b5bd9ce06d960 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -17,12 +17,11 @@ limits of `Real.exp` at infinity. exp -/ - noncomputable section open Finset Filter Metric Asymptotics Set Function Bornology -open scoped Classical Topology Nat +open scoped Topology Nat namespace Complex diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index c6e5e2336dd26..4079b96dab467 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -17,12 +17,10 @@ In this file we prove that `Complex.exp` and `Real.exp` are infinitely smooth fu exp, derivative -/ - noncomputable section open Filter Asymptotics Set Function - -open scoped Classical Topology +open scoped Topology /-! ## `Complex.exp` -/ diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index 9c9ee98b23343..02967ef06b371 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -20,8 +20,7 @@ We also prove differentiability and provide derivatives for the power functions noncomputable section -open scoped Classical Real Topology NNReal ENNReal Filter - +open scoped Real Topology NNReal ENNReal open Filter namespace Complex diff --git a/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean index 6bc1053f456c6..20b9a0ce94d94 100644 --- a/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean +++ b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean @@ -23,12 +23,9 @@ cannot have: noncomputable section -open scoped Classical Topology - +open scoped Topology open Polynomial Real Filter Set Function -open scoped Polynomial - /-- `expNegInvGlue` is the real function given by `x ↦ exp (-1/x)` for `x > 0` and `0` for `x ≤ 0`. It is a basic building block to construct smooth partitions of unity. Its main property is that it vanishes for `x ≤ 0`, it is positive for `x > 0`, and the junction between the two diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean index 9293d0b7de6f3..e11d7973f3b2c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Deriv.lean @@ -21,12 +21,10 @@ computed. sin, cos, tan, angle -/ - noncomputable section -open scoped Classical Topology Filter - -open Set Filter +open scoped Topology Filter +open Set namespace Complex diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean index 7043d33bf06eb..36b15f83044f4 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean @@ -12,14 +12,10 @@ import Mathlib.Analysis.SpecialFunctions.Trigonometric.Deriv Derivatives of `arcsin` and `arccos`. -/ - noncomputable section -open scoped Classical Topology Filter - -open Set Filter - -open scoped Real +open scoped Topology Filter Real +open Set namespace Real diff --git a/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean b/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean index bb2917af1e79c..2521a5e952967 100644 --- a/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid/FreeGroupoid.lean @@ -35,10 +35,7 @@ and finally quotienting by the reducibility relation. -/ - -open Set Classical Function - -attribute [local instance] propDecidable +open Set Function namespace CategoryTheory diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 2f84e4fb04713..8979d1b534d65 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -18,7 +18,7 @@ hyperbolic sine, hyperbolic cosine, and hyperbolic tangent functions. -/ open CauSeq Finset IsAbsoluteValue -open scoped Classical ComplexConjugate +open scoped ComplexConjugate namespace Complex diff --git a/Mathlib/FieldTheory/AbelRuffini.lean b/Mathlib/FieldTheory/AbelRuffini.lean index 2fb1bfe62c908..ab3b3dba2d9ba 100644 --- a/Mathlib/FieldTheory/AbelRuffini.lean +++ b/Mathlib/FieldTheory/AbelRuffini.lean @@ -23,7 +23,6 @@ by radicals, then its minimal polynomial has solvable Galois group. that is solvable by radicals has a solvable Galois group. -/ - noncomputable section open Polynomial IntermediateField diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index 25a98255a9528..bb6ec619c92ea 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -29,8 +29,6 @@ For example, `Algebra.adjoin K {x}` might not include `x⁻¹`. open FiniteDimensional Polynomial -open scoped Classical Polynomial - namespace IntermediateField section AdjoinDef @@ -822,6 +820,7 @@ theorem exists_finset_of_mem_adjoin {S : Set E} {x : E} (hx : x ∈ adjoin F S) ∃ T : Finset E, (T : Set E) ⊆ S ∧ x ∈ adjoin F (T : Set E) := by simp_rw [← biSup_adjoin_simple S, ← iSup_subtype''] at hx obtain ⟨s, hx'⟩ := exists_finset_of_mem_iSup hx + classical refine ⟨s.image Subtype.val, by simp, SetLike.le_def.mp ?_ hx'⟩ simp_rw [Finset.coe_image, iSup_le_iff, adjoin_le_iff] rintro _ h _ rfl @@ -1236,6 +1235,7 @@ theorem fg_of_noetherian (S : IntermediateField F E) [IsNoetherian F E] : S.FG : theorem induction_on_adjoin_finset (S : Finset E) (P : IntermediateField F E → Prop) (base : P ⊥) (ih : ∀ (K : IntermediateField F E), ∀ x ∈ S, P K → P (K⟮x⟯.restrictScalars F)) : P (adjoin F S) := by + classical refine Finset.induction_on' S ?_ (fun ha _ _ h => ?_) · simp [base] · rw [Finset.coe_insert, Set.insert_eq, Set.union_comm, ← adjoin_adjoin_left] diff --git a/Mathlib/FieldTheory/Finite/Polynomial.lean b/Mathlib/FieldTheory/Finite/Polynomial.lean index e6e7fa481079a..8207de7d06c14 100644 --- a/Mathlib/FieldTheory/Finite/Polynomial.lean +++ b/Mathlib/FieldTheory/Finite/Polynomial.lean @@ -45,8 +45,6 @@ namespace MvPolynomial noncomputable section -open scoped Classical - open Set LinearMap Submodule variable {K : Type*} {σ : Type*} @@ -73,6 +71,7 @@ theorem degrees_indicator (c : σ → K) : degrees (indicator c) ≤ ∑ s : σ, (Fintype.card K - 1) • {s} := by rw [indicator] refine le_trans (degrees_prod _ _) (Finset.sum_le_sum fun s _ => ?_) + classical refine le_trans (degrees_sub _ _) ?_ rw [degrees_one, ← bot_eq_zero, bot_sup_eq] refine le_trans (degrees_pow _ _) (nsmul_le_nsmul_right ?_ _) @@ -82,6 +81,7 @@ theorem degrees_indicator (c : σ → K) : theorem indicator_mem_restrictDegree (c : σ → K) : indicator c ∈ restrictDegree σ K (Fintype.card K - 1) := by + classical rw [mem_restrictDegree_iff_sup, indicator] intro n refine le_trans (Multiset.count_le_of_le _ <| degrees_indicator _) (le_of_eq ?_) @@ -126,6 +126,7 @@ variable [Field K] [Fintype K] [Finite σ] theorem map_restrict_dom_evalₗ : (restrictDegree σ K (Fintype.card K - 1)).map (evalₗ K σ) = ⊤ := by cases nonempty_fintype σ refine top_unique (SetLike.le_def.2 fun e _ => mem_map.2 ?_) + classical refine ⟨∑ n : σ → K, e n • indicator n, ?_, ?_⟩ · exact sum_mem fun c _ => smul_mem _ _ (indicator_mem_restrictDegree _) · ext n @@ -146,8 +147,7 @@ end MvPolynomial namespace MvPolynomial -open scoped Classical Cardinal - +open scoped Cardinal open LinearMap Submodule universe u @@ -177,6 +177,8 @@ section CommRing variable [CommRing K] +-- TODO: would be nice to replace this by suitable decidability assumptions +open Classical in noncomputable instance decidableRestrictDegree (m : ℕ) : DecidablePred (· ∈ { n : σ →₀ ℕ | ∀ i, n i ≤ m }) := by simp only [Set.mem_setOf_eq]; infer_instance @@ -185,6 +187,7 @@ end CommRing variable [Field K] +open Classical in theorem rank_R [Fintype σ] : Module.rank K (R σ K) = Fintype.card (σ → K) := calc Module.rank K (R σ K) = @@ -205,11 +208,13 @@ theorem rank_R [Fintype σ] : Module.rank K (R σ K) = Fintype.card (σ → K) : instance [Finite σ] : FiniteDimensional K (R σ K) := by cases nonempty_fintype σ + classical exact IsNoetherian.iff_fg.1 (IsNoetherian.iff_rank_lt_aleph0.mpr <| by simpa only [rank_R] using Cardinal.nat_lt_aleph0 (Fintype.card (σ → K))) +open Classical in theorem finrank_R [Fintype σ] : FiniteDimensional.finrank K (R σ K) = Fintype.card (σ → K) := FiniteDimensional.finrank_eq_of_rank_eq (rank_R σ K) @@ -222,6 +227,7 @@ theorem range_evalᵢ [Finite σ] : range (evalᵢ σ K) = ⊤ := by theorem ker_evalₗ [Finite σ] : ker (evalᵢ σ K) = ⊥ := by cases nonempty_fintype σ refine (ker_eq_bot_iff_range_eq_top_of_finrank_eq_finrank ?_).mpr (range_evalᵢ σ K) + classical rw [FiniteDimensional.finrank_fintype_fun_eq_card, finrank_R] theorem eq_zero_of_eval_eq_zero [Finite σ] (p : MvPolynomial σ K) (h : ∀ v : σ → K, eval v p = 0) diff --git a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index 41a375698ada0..662deae08d25c 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -26,13 +26,10 @@ In this file we construct the algebraic closure of a field algebraic closure, algebraically closed -/ - universe u v w noncomputable section -open scoped Classical Polynomial - open Polynomial variable (k : Type u) [Field k] @@ -54,6 +51,7 @@ indeterminate. -/ def spanEval : Ideal (MvPolynomial (MonicIrreducible k) k) := Ideal.span <| Set.range <| evalXSelf k +open Classical in /-- Given a finset of monic irreducible polynomials, construct an algebra homomorphism to the splitting field of the product of the polynomials sending each indeterminate `x_f` represented by the polynomial `f` in the finset to a root of `f`. -/ diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index f710eead22c3f..f50d91d7f2c11 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -41,11 +41,8 @@ algebraic closure, algebraically closed -/ - universe u v w -open scoped Classical Polynomial - open Polynomial variable (k : Type u) [Field k] @@ -127,6 +124,7 @@ theorem of_exists_root (H : ∀ p : k[X], p.Monic → Irreducible p → ∃ x, p refine ⟨fun p ↦ Or.inr ?_⟩ intro q hq _ have : Irreducible (q * C (leadingCoeff q)⁻¹) := by + classical rw [← coe_normUnit_of_ne_zero hq.ne_zero] exact (associated_normalize _).irreducible hq obtain ⟨x, hx⟩ := H (q * C (leadingCoeff q)⁻¹) (monic_mul_leadingCoeff_inv hq.ne_zero) this diff --git a/Mathlib/FieldTheory/Minpoly/IsIntegrallyClosed.lean b/Mathlib/FieldTheory/Minpoly/IsIntegrallyClosed.lean index 7da3e377b531d..5abd2d46ef12c 100644 --- a/Mathlib/FieldTheory/Minpoly/IsIntegrallyClosed.lean +++ b/Mathlib/FieldTheory/Minpoly/IsIntegrallyClosed.lean @@ -26,9 +26,6 @@ This file specializes the theory of minpoly to the case of an algebra over a GCD -/ - -open scoped Classical Polynomial - open Polynomial Set Function minpoly namespace minpoly diff --git a/Mathlib/FieldTheory/PrimitiveElement.lean b/Mathlib/FieldTheory/PrimitiveElement.lean index 544521dd8fdca..74874763ddba2 100644 --- a/Mathlib/FieldTheory/PrimitiveElement.lean +++ b/Mathlib/FieldTheory/PrimitiveElement.lean @@ -34,7 +34,6 @@ exists_adjoin_simple_eq_top -/ - noncomputable section open FiniteDimensional Polynomial IntermediateField @@ -81,6 +80,7 @@ theorem primitive_element_inf_aux_exists_c (f g : F[X]) : classical let sf := (f.map ϕ).roots let sg := (g.map ϕ).roots + classical let s := (sf.bind fun α' => sg.map fun β' => -(α' - α) / (β' - β)).toFinset let s' := s.preimage ϕ fun x _ y _ h => ϕ.injective h obtain ⟨c, hc⟩ := Infinite.exists_not_mem_finset s' @@ -116,6 +116,7 @@ theorem primitive_element_inf_aux [Algebra.IsSeparable F E] : ∃ γ : E, F⟮α have α_in_Fαβ : α ∈ F⟮α, β⟯ := subset_adjoin F {α, β} (Set.mem_insert α {β}) have β_in_Fαβ : β ∈ F⟮α, β⟯ := subset_adjoin F {α, β} (Set.mem_insert_of_mem α rfl) exact F⟮α, β⟯.add_mem α_in_Fαβ (F⟮α, β⟯.smul_mem β_in_Fαβ) + classical let p := EuclideanDomain.gcd ((f.map (algebraMap F F⟮γ⟯)).comp (C (AdjoinSimple.gen F γ) - (C ↑c : F⟮γ⟯[X]) * X)) (g.map (algebraMap F F⟮γ⟯)) let h := EuclideanDomain.gcd ((f.map ιFE).comp (C γ - C (ιFE c) * X)) (g.map ιFE) diff --git a/Mathlib/FieldTheory/SeparableClosure.lean b/Mathlib/FieldTheory/SeparableClosure.lean index 044959bd5502b..35439789c6f2a 100644 --- a/Mathlib/FieldTheory/SeparableClosure.lean +++ b/Mathlib/FieldTheory/SeparableClosure.lean @@ -59,8 +59,6 @@ separable degree, degree, separable closure -/ -open scoped Polynomial - open FiniteDimensional Polynomial IntermediateField Field noncomputable section diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index 60df02f64b2ed..e9e5e31f32d7d 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -118,8 +118,6 @@ separable degree, degree, polynomial -/ -open scoped Classical Polynomial - open FiniteDimensional Polynomial IntermediateField Field noncomputable section @@ -261,6 +259,7 @@ namespace Polynomial variable {F E} variable (f : F[X]) +open Classical in /-- The separable degree `Polynomial.natSepDegree` of a polynomial is a natural number, defined to be the number of distinct roots of it over its splitting field. This is similar to `Polynomial.natDegree` but not to `Polynomial.degree`, namely, the separable @@ -271,6 +270,7 @@ def natSepDegree : ℕ := (f.aroots f.SplittingField).toFinset.card theorem natSepDegree_le_natDegree : f.natSepDegree ≤ f.natDegree := by have := f.map (algebraMap F f.SplittingField) |>.card_roots' rw [← aroots_def, natDegree_map] at this + classical exact (f.aroots f.SplittingField).toFinset_card_le.trans this @[simp] @@ -300,6 +300,7 @@ theorem natSepDegree_one : (1 : F[X]).natSepDegree = 0 := by theorem natSepDegree_ne_zero (h : f.natDegree ≠ 0) : f.natSepDegree ≠ 0 := by rw [natSepDegree, ne_eq, Finset.card_eq_zero, ← ne_eq, ← Finset.nonempty_iff_ne_empty] use rootOfSplits _ (SplittingField.splits f) (ne_of_apply_ne _ h) + classical rw [Multiset.mem_toFinset, mem_aroots] exact ⟨ne_of_apply_ne _ h, map_rootOfSplits _ (SplittingField.splits f) (ne_of_apply_ne _ h)⟩ @@ -315,6 +316,7 @@ theorem natSepDegree_ne_zero_iff : f.natSepDegree ≠ 0 ↔ f.natDegree ≠ 0 := it is separable. -/ theorem natSepDegree_eq_natDegree_iff (hf : f ≠ 0) : f.natSepDegree = f.natDegree ↔ f.Separable := by + classical simp_rw [← card_rootSet_eq_natDegree_iff_of_splits hf (SplittingField.splits f), rootSet_def, Finset.coe_sort_coe, Fintype.card_coe] rfl @@ -331,8 +333,9 @@ theorem Separable.natSepDegree_eq_natDegree (h : f.Separable) : /-- If a polynomial splits over `E`, then its separable degree is equal to the number of distinct roots of it over `E`. -/ -theorem natSepDegree_eq_of_splits (h : f.Splits (algebraMap F E)) : +theorem natSepDegree_eq_of_splits [DecidableEq E] (h : f.Splits (algebraMap F E)) : f.natSepDegree = (f.aroots E).toFinset.card := by + classical rw [aroots, ← (SplittingField.lift f h).comp_algebraMap, ← map_map, roots_map _ ((splits_id_iff_splits _).mpr <| SplittingField.splits f), Multiset.toFinset_map, Finset.card_image_of_injective _ (RingHom.injective _), natSepDegree] @@ -340,27 +343,31 @@ theorem natSepDegree_eq_of_splits (h : f.Splits (algebraMap F E)) : variable (E) in /-- The separable degree of a polynomial is equal to the number of distinct roots of it over any algebraically closed field. -/ -theorem natSepDegree_eq_of_isAlgClosed [IsAlgClosed E] : +theorem natSepDegree_eq_of_isAlgClosed [DecidableEq E] [IsAlgClosed E] : f.natSepDegree = (f.aroots E).toFinset.card := natSepDegree_eq_of_splits f (IsAlgClosed.splits_codomain f) variable (E) in theorem natSepDegree_map : (f.map (algebraMap F E)).natSepDegree = f.natSepDegree := by + classical simp_rw [natSepDegree_eq_of_isAlgClosed (AlgebraicClosure E), aroots_def, map_map, ← IsScalarTower.algebraMap_eq] @[simp] theorem natSepDegree_C_mul {x : F} (hx : x ≠ 0) : (C x * f).natSepDegree = f.natSepDegree := by + classical simp only [natSepDegree_eq_of_isAlgClosed (AlgebraicClosure F), aroots_C_mul _ hx] @[simp] theorem natSepDegree_smul_nonzero {x : F} (hx : x ≠ 0) : (x • f).natSepDegree = f.natSepDegree := by + classical simp only [natSepDegree_eq_of_isAlgClosed (AlgebraicClosure F), aroots_smul_nonzero _ hx] @[simp] theorem natSepDegree_pow {n : ℕ} : (f ^ n).natSepDegree = if n = 0 then 0 else f.natSepDegree := by + classical simp only [natSepDegree_eq_of_isAlgClosed (AlgebraicClosure F), aroots_pow] by_cases h : n = 0 · simp only [h, zero_smul, Multiset.toFinset_zero, Finset.card_empty, ite_true] @@ -384,6 +391,7 @@ theorem natSepDegree_mul (g : F[X]) : (f * g).natSepDegree ≤ f.natSepDegree + g.natSepDegree := by by_cases h : f * g = 0 · simp only [h, natSepDegree_zero, zero_le] + classical simp_rw [natSepDegree_eq_of_isAlgClosed (AlgebraicClosure F), aroots_mul h, Multiset.toFinset_add] exact Finset.card_union_le _ _ @@ -403,6 +411,7 @@ theorem natSepDegree_mul_eq_iff (g : F[X]) : rintro (⟨-, h⟩ | ⟨x, -, h⟩) · exact ⟨0, by rw [h, map_zero]⟩ exact ⟨x, h⟩ + classical simp_rw [natSepDegree_eq_of_isAlgClosed (AlgebraicClosure F), aroots_mul h, Multiset.toFinset_add, Finset.card_union_eq_card_add_card, Finset.disjoint_iff_ne, Multiset.mem_toFinset, mem_aroots] rw [mul_eq_zero, not_or] at h @@ -425,6 +434,7 @@ theorem natSepDegree_mul_of_isCoprime (g : F[X]) (hc : IsCoprime f g) : theorem natSepDegree_le_of_dvd (g : F[X]) (h1 : f ∣ g) (h2 : g ≠ 0) : f.natSepDegree ≤ g.natSepDegree := by + classical simp_rw [natSepDegree_eq_of_isAlgClosed (AlgebraicClosure F)] exact Finset.card_le_card <| Multiset.toFinset_subset.mpr <| Multiset.Le.subset <| roots.le_of_dvd (map_ne_zero h2) <| map_dvd _ h1 @@ -436,6 +446,7 @@ theorem natSepDegree_expand (q : ℕ) [hF : ExpChar F q] {n : ℕ} : cases' hF with _ _ hprime _ · simp only [one_pow, expand_one] haveI := Fact.mk hprime + classical simpa only [natSepDegree_eq_of_isAlgClosed (AlgebraicClosure F), aroots_def, map_expand, Fintype.card_coe] using Fintype.card_eq.2 ⟨(f.map (algebraMap F (AlgebraicClosure F))).rootsExpandPowEquivRoots q n⟩ @@ -513,6 +524,7 @@ some non-zero natural number `m` and some element `y` of `F`. -/ theorem eq_X_sub_C_pow_of_natSepDegree_eq_one_of_splits (hm : f.Monic) (hs : f.Splits (RingHom.id F)) (h : f.natSepDegree = 1) : ∃ (m : ℕ) (y : F), m ≠ 0 ∧ f = (X - C y) ^ m := by + classical have h1 := eq_prod_roots_of_monic_of_splits_id hm hs have h2 := (natSepDegree_eq_of_splits f hs).symm rw [h, aroots_def, Algebra.id.map_eq_id, map_id, Multiset.toFinset_card_eq_one_iff] at h2 @@ -551,6 +563,7 @@ theorem eq_X_pow_char_pow_sub_C_pow_of_natSepDegree_eq_one (q : ℕ) [ExpChar F Nat.pos_of_ne_zero <| (natSepDegree_ne_zero_iff _).2 hI.natDegree_pos.ne' obtain ⟨n, y, H, hp⟩ := hM.eq_X_pow_char_pow_sub_C_of_natSepDegree_eq_one_of_irreducible q hI hD have hF := multiplicity_finite_of_degree_pos_of_monic (degree_pos_of_irreducible hI) hM hm.ne_zero + classical have hne := (multiplicity.pos_of_dvd hF hf).ne' refine ⟨_, n, y, hne, H, ?_⟩ obtain ⟨c, hf, H⟩ := multiplicity.exists_eq_pow_mul_and_not_dvd hF @@ -641,6 +654,7 @@ theorem finSepDegree_adjoin_simple_eq_natSepDegree {α : E} (halg : IsAlgebraic finSepDegree F F⟮α⟯ = (minpoly F α).natSepDegree := by have : finSepDegree F F⟮α⟯ = _ := Nat.card_congr (algHomAdjoinIntegralEquiv F (K := AlgebraicClosure F⟮α⟯) halg.isIntegral) + classical rw [this, Nat.card_eq_fintype_card, natSepDegree_eq_of_isAlgClosed (E := AlgebraicClosure F⟮α⟯), ← Fintype.card_coe] simp_rw [Multiset.mem_toFinset] diff --git a/Mathlib/FieldTheory/SplittingField/Construction.lean b/Mathlib/FieldTheory/SplittingField/Construction.lean index bb19704f24094..0a11bdfa13be0 100644 --- a/Mathlib/FieldTheory/SplittingField/Construction.lean +++ b/Mathlib/FieldTheory/SplittingField/Construction.lean @@ -28,11 +28,8 @@ actual `SplittingField` will be a quotient of a `MvPolynomial`, it has nice inst -/ - noncomputable section -open scoped Classical Polynomial - universe u v w variable {F : Type u} {K : Type v} {L : Type w} @@ -45,6 +42,7 @@ open Polynomial section SplittingField +open Classical in /-- Non-computably choose an irreducible factor from a polynomial. -/ def factor (f : K[X]) : K[X] := if H : ∃ g, Irreducible g ∧ g ∣ f then Classical.choose H else X @@ -188,6 +186,7 @@ theorem adjoin_rootSet (n : ℕ) : have hndf : f.natDegree ≠ 0 := by intro h; rw [h] at hfn; cases hfn have hfn0 : f ≠ 0 := by intro h; rw [h] at hndf; exact hndf rfl have hmf0 : map (algebraMap K (SplittingFieldAux n.succ f)) f ≠ 0 := map_ne_zero hfn0 + classical rw [rootSet_def, aroots_def] rw [algebraMap_succ, ← map_map, ← X_sub_C_mul_removeFactor _ hndf, Polynomial.map_mul] at hmf0 ⊢ -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 diff --git a/Mathlib/FieldTheory/SplittingField/IsSplittingField.lean b/Mathlib/FieldTheory/SplittingField/IsSplittingField.lean index 9ab6d7389fb49..4ca9cadac0e55 100644 --- a/Mathlib/FieldTheory/SplittingField/IsSplittingField.lean +++ b/Mathlib/FieldTheory/SplittingField/IsSplittingField.lean @@ -27,11 +27,8 @@ if it is the smallest field extension of `K` such that `f` splits. -/ - noncomputable section -open scoped Classical Polynomial - universe u v w variable {F : Type u} (K : Type v) (L : Type w) @@ -77,6 +74,7 @@ theorem splits_iff (f : K[X]) [IsSplittingField K L f] : rw [eq_bot_iff, ← adjoin_rootSet L f, rootSet, aroots, roots_map (algebraMap K L) h, Algebra.adjoin_le_iff] intro y hy + classical rw [Multiset.toFinset_map, Finset.mem_coe, Finset.mem_image] at hy obtain ⟨x : K, -, hxy : algebraMap K L x = y⟩ := hy rw [← hxy] @@ -91,6 +89,7 @@ theorem mul (f g : F[X]) (hf : f ≠ 0) (hg : g ≠ 0) [IsSplittingField F K f] ⟨(IsScalarTower.algebraMap_eq F K L).symm ▸ splits_mul _ (splits_comp_of_splits _ _ (splits K f)) ((splits_map_iff _ _).1 (splits L <| g.map <| algebraMap F K)), by + classical rw [rootSet, aroots_mul (mul_ne_zero hf hg), Multiset.toFinset_add, Finset.coe_union, Algebra.adjoin_union_eq_adjoin_adjoin, aroots_def, aroots_def, IsScalarTower.algebraMap_eq F K L, ← map_map, @@ -101,6 +100,7 @@ theorem mul (f g : F[X]) (hf : f ≠ 0) (hg : g ≠ 0) [IsSplittingField F K f] end ScalarTower +open Classical in /-- Splitting field of `f` embeds into any field that splits `f`. -/ def lift [Algebra K F] (f : K[X]) [IsSplittingField K L f] (hf : Splits (algebraMap K F) f) : L →ₐ[K] F := @@ -117,8 +117,9 @@ def lift [Algebra K F] (f : K[X]) [IsSplittingField K L f] ⟨IsAlgebraic.isIntegral ⟨f, hf0, this⟩, splits_of_splits_of_dvd _ hf0 hf <| minpoly.dvd _ _ this⟩)) Algebra.toTop -theorem finiteDimensional (f : K[X]) [IsSplittingField K L f] : FiniteDimensional K L := - ⟨@Algebra.top_toSubmodule K L _ _ _ ▸ +theorem finiteDimensional (f : K[X]) [IsSplittingField K L f] : FiniteDimensional K L := by + classical + exact ⟨@Algebra.top_toSubmodule K L _ _ _ ▸ adjoin_rootSet L f ▸ fg_adjoin_of_finite (Finset.finite_toSet _) fun y hy ↦ if hf : f = 0 then by rw [hf, rootSet_zero] at hy; cases hy else IsAlgebraic.isIntegral ⟨f, hf, (mem_rootSet'.mp hy).2⟩⟩ @@ -145,6 +146,7 @@ variable {K L} [Field K] [Field L] [Algebra K L] {p : K[X]} {F : IntermediateFie theorem IntermediateField.splits_of_splits (h : p.Splits (algebraMap K L)) (hF : ∀ x ∈ p.rootSet L, x ∈ F) : p.Splits (algebraMap K F) := by + classical simp_rw [← F.fieldRange_val, rootSet_def, Finset.mem_coe, Multiset.mem_toFinset] at hF exact splits_of_comp _ F.val.toRingHom h hF diff --git a/Mathlib/Geometry/Manifold/BumpFunction.lean b/Mathlib/Geometry/Manifold/BumpFunction.lean index 8d0691a799c0d..753de631d7746 100644 --- a/Mathlib/Geometry/Manifold/BumpFunction.lean +++ b/Mathlib/Geometry/Manifold/BumpFunction.lean @@ -36,7 +36,7 @@ variable {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensi open Function Filter FiniteDimensional Set Metric -open scoped Topology Manifold Classical Filter +open scoped Topology Manifold noncomputable section diff --git a/Mathlib/LinearAlgebra/TensorProduct/Subalgebra.lean b/Mathlib/LinearAlgebra/TensorProduct/Subalgebra.lean index fa52e52ee6477..5e2c363f1116d 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Subalgebra.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Subalgebra.lean @@ -33,7 +33,7 @@ mainly used in the definition of linearly disjointness. -/ -open scoped Classical TensorProduct +open scoped TensorProduct open FiniteDimensional diff --git a/Mathlib/LinearAlgebra/TensorProduct/Vanishing.lean b/Mathlib/LinearAlgebra/TensorProduct/Vanishing.lean index 07d09f1673d29..968e2a986f444 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Vanishing.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Vanishing.lean @@ -60,7 +60,7 @@ variable (R : Type u) [CommRing R] variable {M : Type u} [AddCommGroup M] [Module R M] variable {N : Type u} [AddCommGroup N] [Module R N] -open Classical DirectSum LinearMap Function Submodule +open DirectSum LinearMap Function Submodule namespace TensorProduct @@ -140,6 +140,7 @@ theorem vanishesTrivially_of_sum_tmul_eq_zero (hm : Submodule.span R (Set.range use fun ⟨⟨_, yj⟩, _⟩ ↦ yj constructor · intro i + classical apply_fun finsuppScalarLeft R N ι at hkn apply_fun (· i) at hkn symm at hkn diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index 620a010e28008..dde83153317ce 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -26,7 +26,7 @@ import Mathlib.MeasureTheory.Constructions.BorelSpace.Order open Set Filter MeasureTheory MeasurableSpace -open scoped Classical Topology NNReal ENNReal MeasureTheory +open scoped Topology NNReal ENNReal universe u v w x y @@ -303,6 +303,7 @@ lemma aemeasurable_of_tendsto' {ι : Type*} {f : ι → α → ℝ≥0∞} {g : set p : α → (ℕ → ℝ≥0∞) → Prop := fun x f' ↦ Tendsto f' atTop (𝓝 (g x)) have hp : ∀ᵐ x ∂μ, p x fun n ↦ f (v n) x := by filter_upwards [hlim] with x hx using hx.comp hv + classical set aeSeqLim := fun x ↦ ite (x ∈ aeSeqSet h'f p) (g x) (⟨f (v 0) x⟩ : Nonempty ℝ≥0∞).some refine ⟨aeSeqLim, measurable_of_tendsto' atTop (aeSeq.measurable h'f p) (tendsto_pi_nhds.mpr fun x ↦ ?_), ?_⟩ diff --git a/Mathlib/MeasureTheory/Function/UnifTight.lean b/Mathlib/MeasureTheory/Function/UnifTight.lean index b20d135ad8f9f..cc50e8c1de965 100644 --- a/Mathlib/MeasureTheory/Function/UnifTight.lean +++ b/Mathlib/MeasureTheory/Function/UnifTight.lean @@ -37,15 +37,12 @@ is also proved later in the file. uniform integrable, uniformly tight, Vitali convergence theorem -/ - namespace MeasureTheory -open Classical Set Filter Topology MeasureTheory NNReal ENNReal +open Set Filter Topology MeasureTheory NNReal ENNReal variable {α β ι : Type*} {m : MeasurableSpace α} {μ : Measure α} [NormedAddCommGroup β] - - section UnifTight /- This follows closely the `UnifIntegrable` section @@ -53,7 +50,6 @@ from `Mathlib.MeasureTheory.Functions.UniformIntegrable`.-/ variable {f g : ι → α → β} {p : ℝ≥0∞} - /-- A sequence of functions `f` is uniformly tight in `L^p` if for all `ε > 0`, there exists some measurable set `s` with finite measure such that the Lp-norm of `f i` restricted to `sᶜ` is smaller than `ε` for all `i`. -/ diff --git a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean index d371725adeeeb..ed2483621633c 100644 --- a/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean +++ b/Mathlib/RingTheory/DedekindDomain/FiniteAdeleRing.lean @@ -358,7 +358,8 @@ instance : CoeFun (FiniteAdeleRing R K) (fun _ ↦ ∀ (v : HeightOneSpectrum R), adicCompletion K v) where coe a v := a.1 v -open scoped algebraMap in +open scoped algebraMap -- coercion from R to `FiniteAdeleRing R K` + variable {R K} in lemma exists_finiteIntegralAdele_iff (a : FiniteAdeleRing R K) : (∃ c : R_hat R K, a = c) ↔ ∀ (v : HeightOneSpectrum R), a v ∈ adicCompletionIntegers K v := @@ -366,9 +367,7 @@ lemma exists_finiteIntegralAdele_iff (a : FiniteAdeleRing R K) : (∃ c : R_hat section Topology -open Classical nonZeroDivisors Multiplicative Additive IsDedekindDomain.HeightOneSpectrum - -open scoped algebraMap -- coercion from R to FiniteAdeleRing R K +open nonZeroDivisors open scoped DiscreteValuation variable {R K} in @@ -393,8 +392,6 @@ lemma mul_nonZeroDivisor_mem_finiteIntegralAdeles (a : FiniteAdeleRing R K) : rw [← mul_assoc] exact mul_mem (h v (a v)) <| coe_mem_adicCompletionIntegers _ _ -open scoped Pointwise - theorem submodulesRingBasis : SubmodulesRingBasis (fun (r : R⁰) ↦ Submodule.span (R_hat R K) {((r : R) : FiniteAdeleRing R K)}) where inter i j := ⟨i * j, by diff --git a/Mathlib/RingTheory/MvPolynomial/Localization.lean b/Mathlib/RingTheory/MvPolynomial/Localization.lean index 1aa0fb1a03909..6bb865ab1f60a 100644 --- a/Mathlib/RingTheory/MvPolynomial/Localization.lean +++ b/Mathlib/RingTheory/MvPolynomial/Localization.lean @@ -23,8 +23,6 @@ In this file we show some results connecting multivariate polynomial rings and l -/ -open Classical - variable {σ R : Type*} [CommRing R] (M : Submonoid R) variable (S : Type*) [CommRing S] [Algebra R S] @@ -63,6 +61,7 @@ instance isLocalization : IsLocalization (M.map <| C (σ := σ)) simp_rw [algebraMap_def, MvPolynomial.ext_iff, coeff_map] at h choose c hc using (fun m ↦ IsLocalization.exists_of_eq (M := M) (h m)) simp only [Subtype.exists, Submonoid.mem_map, exists_prop, exists_exists_and_eq_and] + classical refine ⟨Finset.prod (p.support ∪ q.support) (fun m ↦ c m), ?_, ?_⟩ · exact M.prod_mem (fun m _ ↦ (c m).property) · ext m From d167cb0fe0e8a34993b9567920a631441a5c00c4 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 5 Aug 2024 16:02:55 +0000 Subject: [PATCH 038/359] feat: add `mul_self_tsub_mul_self` (#15494) --- Mathlib/Algebra/Order/Ring/Canonical.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/Algebra/Order/Ring/Canonical.lean b/Mathlib/Algebra/Order/Ring/Canonical.lean index 750d5c7e2a2ff..432f79a5c6202 100644 --- a/Mathlib/Algebra/Order/Ring/Canonical.lean +++ b/Mathlib/Algebra/Order/Ring/Canonical.lean @@ -116,4 +116,15 @@ theorem tsub_mul (a b c : α) : (a - b) * c = a * c - b * c := lemma mul_tsub_one (a b : α) : a * (b - 1) = a * b - a := by rw [mul_tsub, mul_one] lemma tsub_one_mul (a b : α) : (a - 1) * b = a * b - b := by rw [tsub_mul, one_mul] +/-- The `tsub` version of `mul_self_sub_mul_self`. Notably, this holds for `Nat` and `NNReal`. -/ +theorem mul_self_tsub_mul_self (a b : α) : a * a - b * b = (a + b) * (a - b) := by + rw [mul_tsub, add_mul, add_mul, tsub_add_eq_tsub_tsub, mul_comm b a, add_tsub_cancel_right] + +/-- The `tsub` version of `sq_sub_sq`. Notably, this holds for `Nat` and `NNReal`. -/ +theorem sq_tsub_sq (a b : α) : a ^ 2 - b ^ 2 = (a + b) * (a - b) := by + rw [sq, sq, mul_self_tsub_mul_self] + +theorem mul_self_tsub_one (a : α) : a * a - 1 = (a + 1) * (a - 1) := by + rw [← mul_self_tsub_mul_self, mul_one] + end Sub From 124f27e89c4f233bc23e4c6418f3c632b7edc6a1 Mon Sep 17 00:00:00 2001 From: Peter Nelson <71660771+apnelson1@users.noreply.github.com> Date: Mon, 5 Aug 2024 16:52:46 +0000 Subject: [PATCH 039/359] refactor(Order/Minimal): change minimality to a predicate (#14721) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR changes the notion of minimality from being a set with an unbundled relation `r` (i.e. having type signature `minimals (r : α → α → Prop) (s : Set α) : Set α`) to being a predicate with a bundled order relation in a typeclass: `Minimal [LE α] (P : α → Prop) : α → Prop`. It also does the same for maximality. This is hopefully an improvement, for a few reasons : * ergonomic API with dot notation (eg `Minimal.eq_of_le`, `Minimal.not_lt`), * no need to specify a particular relation explicitly (no use of `minimals`/`maximals` for a relation that was not defeq to `LE.le` currently appears anywhere in mathlib), * the potential in future to unify ad hoc notions of min/maximality appearing in various locations in mathlib (eg `zorn_preorder`, `Sylow.is_maximal'` to name a couple of random examples). Since this change removes the definitions `minimals` and `maximals` and all their API, it required a total overhaul of the contents of `Order.Minimal`, and necessarily required updating all uses of `maximals`/`minimals` elsewhere. There were not all too many of these; this PR is designed to be the smallest it could be while making this change. The changes required outside `Order.Minimal` were: - redefinition of `Ideal.minimalPrimes` in `RingTheory/Ideal/MinimalPrime` to use `Minimal` rather than `minimals`, with the necessary minor changes to proofs that used the old definition, including in `AlgebraicGeometry/PrimeSpectrum/Basic`. - redefinition of `irreducibleComponents` in `Topology/Irreducible` to use `Minimal` rather than `minimals`, with necessary minor changes to proofs that used the old definition. - updating of the maximality used in the definition of a matroid in `Data/Matroid/Basic` from `maximals` to `Maximal`, with the necessary changes to everywhere in `Data/Matroid/*` that a matroid was specified using the definition directly. - the necessary changes to two lemmas in `Data/Finset/Lattice` that used `minimals`/`maximals` in the statements. I attempted to clean up the API in `Order.Minimal` to work smoothly with the new definitions, but numerous judgment calls were needed; I removed a few lemmas about `maximals/minimals` that were used nowhere, and seemed unlikely to be useful with the new definition. Comment is welcome! --- .../PrimeSpectrum/Basic.lean | 8 +- Mathlib/Data/Finset/Lattice.lean | 24 +- Mathlib/Data/Matroid/Basic.lean | 94 +- Mathlib/Data/Matroid/Constructions.lean | 26 +- Mathlib/Data/Matroid/Dual.lean | 63 +- Mathlib/Data/Matroid/IndepAxioms.lean | 117 +-- Mathlib/Data/Matroid/Map.lean | 47 +- Mathlib/Data/Matroid/Restrict.lean | 26 +- Mathlib/Data/Matroid/Sum.lean | 14 +- Mathlib/Order/Minimal.lean | 992 +++++++++++------- Mathlib/RingTheory/Ideal/MinimalPrime.lean | 10 +- Mathlib/Topology/Irreducible.lean | 4 +- 12 files changed, 794 insertions(+), 631 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index b0d4fb892fa7a..91c08ba82ce51 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -611,7 +611,8 @@ protected def minimalPrimes.equivIrreducibleComponents : let e : {p : Ideal R | p.IsPrime ∧ ⊥ ≤ p} ≃o PrimeSpectrum R := ⟨⟨fun x ↦ ⟨x.1, x.2.1⟩, fun x ↦ ⟨x.1, x.2, bot_le⟩, fun _ ↦ rfl, fun _ ↦ rfl⟩, Iff.rfl⟩ rw [irreducibleComponents_eq_maximals_closed] - exact OrderIso.minimalsIsoMaximals (e.trans ((PrimeSpectrum.pointsEquivIrreducibleCloseds R).trans + exact OrderIso.setOfMinimalIsoSetOfMaximal + (e.trans ((PrimeSpectrum.pointsEquivIrreducibleCloseds R).trans (TopologicalSpace.IrreducibleCloseds.orderIsoSubtype' (PrimeSpectrum R)).dual)) namespace PrimeSpectrum @@ -620,8 +621,9 @@ lemma vanishingIdeal_irreducibleComponents : vanishingIdeal '' (irreducibleComponents <| PrimeSpectrum R) = minimalPrimes R := by rw [irreducibleComponents_eq_maximals_closed, minimalPrimes_eq_minimals, - ← minimals_swap, ← vanishingIdeal_isClosed_isIrreducible, image_minimals_of_rel_iff_rel] - exact fun s t hs _ ↦ vanishingIdeal_anti_mono_iff hs.1 + image_antitone_setOf_maximal (fun s t hs _ ↦ (vanishingIdeal_anti_mono_iff hs.1).symm), + ← funext (@Set.mem_setOf_eq _ · Ideal.IsPrime), ← vanishingIdeal_isClosed_isIrreducible] + rfl lemma zeroLocus_minimalPrimes : zeroLocus ∘ (↑) '' minimalPrimes R = diff --git a/Mathlib/Data/Finset/Lattice.lean b/Mathlib/Data/Finset/Lattice.lean index 57b07070f1188..16acc98aceb59 100644 --- a/Mathlib/Data/Finset/Lattice.lean +++ b/Mathlib/Data/Finset/Lattice.lean @@ -1800,19 +1800,17 @@ section minimal variable [DecidableEq α] {P : Finset α → Prop} {s : Finset α} -theorem mem_maximals_iff_forall_insert (hP : ∀ ⦃s t⦄, P t → s ⊆ t → P s) : - s ∈ maximals (· ⊆ ·) {t | P t} ↔ P s ∧ ∀ x ∉ s, ¬ P (insert x s) := by - simp only [mem_maximals_iff, and_congr_right_iff, Set.mem_setOf_eq] - refine fun _ ↦ ⟨fun h x hx hxs ↦ hx ?_, fun h t ht hst ↦ hst.antisymm fun x hxt ↦ ?_⟩ - · rw [h hxs (subset_insert _ _)]; exact mem_insert_self x s - exact by_contra fun hxs ↦ h x hxs (hP ht (insert_subset hxt hst)) - -theorem mem_minimals_iff_forall_erase (hP : ∀ ⦃s t⦄, P s → s ⊆ t → P t) : - s ∈ minimals (· ⊆ ·) {t | P t} ↔ P s ∧ ∀ x ∈ s, ¬ P (s.erase x) := by - simp only [mem_minimals_iff, Set.mem_setOf_eq, and_congr_right_iff] - refine fun _ ↦ ⟨fun h x hx hxs ↦ ?_, fun h t ht hst ↦ Eq.symm <| hst.antisymm (fun x hxs ↦ ?_)⟩ - · rw [(h hxs (erase_subset x s))] at hx; simp at hx - exact by_contra fun hxt ↦ h x hxs (hP ht <| subset_erase.2 ⟨hst, hxt⟩) +theorem maximal_iff_forall_insert (hP : ∀ ⦃s t⦄, P t → s ⊆ t → P s) : + Maximal P s ↔ P s ∧ ∀ x ∉ s, ¬ P (insert x s) := by + simp only [Maximal, and_congr_right_iff] + exact fun _ ↦ ⟨fun h x hxs hx ↦ hxs <| h hx (subset_insert _ _) (mem_insert_self x s), + fun h t ht hst x hxt ↦ by_contra fun hxs ↦ h x hxs (hP ht (insert_subset hxt hst))⟩ + +theorem minimal_iff_forall_diff_singleton (hP : ∀ ⦃s t⦄, P t → t ⊆ s → P s) : + Minimal P s ↔ P s ∧ ∀ x ∈ s, ¬ P (s.erase x) where + mp h := ⟨h.prop, fun x hxs hx ↦ by simpa using h.le_of_le hx (erase_subset _ _) hxs⟩ + mpr h := ⟨h.1, fun t ht hts x hxs ↦ by_contra fun hxt ↦ + h.2 x hxs <| hP ht (subset_erase.2 ⟨hts, hxt⟩)⟩ end minimal diff --git a/Mathlib/Data/Matroid/Basic.lean b/Mathlib/Data/Matroid/Basic.lean index 3167eb8359205..22af93d6f230f 100644 --- a/Mathlib/Data/Matroid/Basic.lean +++ b/Mathlib/Data/Matroid/Basic.lean @@ -171,7 +171,7 @@ def Matroid.ExchangeProperty {α : Type _} (P : Set α → Prop) : Prop := /-- A set `X` has the maximal subset property for a predicate `P` if every subset of `X` satisfying `P` is contained in a maximal subset of `X` satisfying `P`. -/ def Matroid.ExistsMaximalSubsetProperty {α : Type _} (P : Set α → Prop) (X : Set α) : Prop := - ∀ I, P I → I ⊆ X → (maximals (· ⊆ ·) {Y | P Y ∧ I ⊆ Y ∧ Y ⊆ X}).Nonempty + ∀ I, P I → I ⊆ X → ∃ J, I ⊆ J ∧ Maximal (fun K ↦ P K ∧ K ⊆ X) J /-- A `Matroid α` is a ground set `E` of type `Set α`, and a nonempty collection of its subsets satisfying the exchange property and the maximal subset property. Each such set is called a @@ -461,9 +461,9 @@ theorem eq_of_base_iff_base_forall {M₁ M₂ : Matroid α} (hE : M₁.E = M₂. fun hB ↦ (h <| hB.subset_ground.trans_eq hE.symm).2 hB⟩ ext <;> simp [hE, M₁.indep_iff', M₂.indep_iff', h'] -theorem base_compl_iff_mem_maximals_disjoint_base (hB : B ⊆ M.E := by aesop_mat) : - M.Base (M.E \ B) ↔ B ∈ maximals (· ⊆ ·) {I | I ⊆ M.E ∧ ∃ B, M.Base B ∧ Disjoint I B} := by - simp_rw [mem_maximals_setOf_iff, and_iff_right hB, and_imp, forall_exists_index] +theorem base_compl_iff_maximal_disjoint_base (hB : B ⊆ M.E := by aesop_mat) : + M.Base (M.E \ B) ↔ Maximal (fun I ↦ I ⊆ M.E ∧ ∃ B, M.Base B ∧ Disjoint I B) B := by + simp_rw [maximal_iff, and_iff_right hB, and_imp, forall_exists_index] refine ⟨fun h ↦ ⟨⟨_, h, disjoint_sdiff_right⟩, fun I hI B' ⟨hB', hIB'⟩ hBI ↦ hBI.antisymm ?_⟩, fun ⟨⟨B', hB', hBB'⟩,h⟩ ↦ ?_⟩ · rw [hB'.eq_of_subset_base h, ← subset_compl_iff_disjoint_right, diff_eq, compl_inter, @@ -569,16 +569,15 @@ theorem Base.eq_of_subset_indep (hB : M.Base B) (hI : M.Indep I) (hBI : B ⊆ I) let ⟨B', hB', hB'I⟩ := hI.exists_base_superset hBI.antisymm (by rwa [hB.eq_of_subset_base hB' (hBI.trans hB'I)]) -theorem base_iff_maximal_indep : M.Base B ↔ M.Indep B ∧ ∀ I, M.Indep I → B ⊆ I → B = I := by - refine ⟨fun h ↦ ⟨h.indep, fun _ ↦ h.eq_of_subset_indep ⟩, fun ⟨h, h'⟩ ↦ ?_⟩ +theorem base_iff_maximal_indep : M.Base B ↔ Maximal M.Indep B := by + rw [maximal_subset_iff] + refine ⟨fun h ↦ ⟨h.indep, fun _ ↦ h.eq_of_subset_indep⟩, fun ⟨h, h'⟩ ↦ ?_⟩ obtain ⟨B', hB', hBB'⟩ := h.exists_base_superset - rwa [h' _ hB'.indep hBB'] + rwa [h' hB'.indep hBB'] -theorem setOf_base_eq_maximals_setOf_indep : {B | M.Base B} = maximals (· ⊆ ·) {I | M.Indep I} := by - ext B; rw [mem_maximals_setOf_iff, mem_setOf, base_iff_maximal_indep] - -theorem Indep.base_of_maximal (hI : M.Indep I) (h : ∀ J, M.Indep J → I ⊆ J → I = J) : M.Base I := - base_iff_maximal_indep.mpr ⟨hI,h⟩ +theorem Indep.base_of_maximal (hI : M.Indep I) (h : ∀ ⦃J⦄, M.Indep J → I ⊆ J → I = J) : + M.Base I := by + rwa [base_iff_maximal_indep, maximal_subset_iff, and_iff_right hI] theorem Base.dep_of_ssubset (hB : M.Base B) (h : B ⊂ X) (hX : X ⊆ M.E := by aesop_mat) : M.Dep X := ⟨fun hX ↦ h.ne (hB.eq_of_subset_indep hX h.subset), hX⟩ @@ -637,11 +636,10 @@ theorem Indep.exists_insert_of_not_base (hI : M.Indep I) (hI' : ¬M.Base I) (hB /-- This is the same as `Indep.exists_insert_of_not_base`, but phrased so that it is defeq to the augmentation axiom for independent sets. -/ -theorem Indep.exists_insert_of_not_mem_maximals (M : Matroid α) ⦃I B : Set α⦄ (hI : M.Indep I) - (hInotmax : I ∉ maximals (· ⊆ ·) {I | M.Indep I}) (hB : B ∈ maximals (· ⊆ ·) {I | M.Indep I}) : +theorem Indep.exists_insert_of_not_maximal (M : Matroid α) ⦃I B : Set α⦄ (hI : M.Indep I) + (hInotmax : ¬ Maximal M.Indep I) (hB : Maximal M.Indep B) : ∃ x ∈ B \ I, M.Indep (insert x I) := by - simp only [mem_maximals_iff, mem_setOf_eq, not_and, not_forall, exists_prop, - exists_and_left, iff_true_intro hI, true_imp_iff] at hB hInotmax + simp only [maximal_subset_iff, hI, not_and, not_forall, exists_prop, true_imp_iff] at hB hInotmax refine hI.exists_insert_of_not_base (fun hIb ↦ ?_) ?_ · obtain ⟨I', hII', hI', hne⟩ := hInotmax exact hne <| hIb.eq_of_subset_indep hII' hI' @@ -664,9 +662,12 @@ theorem Base.exists_insert_of_ssubset (hB : M.Base B) (hIB : I ⊂ B) (hB' : M.B theorem eq_of_indep_iff_indep_forall {M₁ M₂ : Matroid α} (hE : M₁.E = M₂.E) (h : ∀ I, I ⊆ M₁.E → (M₁.Indep I ↔ M₂.Indep I)) : M₁ = M₂ := - let h' : ∀ I, M₁.Indep I ↔ M₂.Indep I := fun I ↦ - (em (I ⊆ M₁.E)).elim (h I) (fun h' ↦ iff_of_false (fun hi ↦ h' (hi.subset_ground)) - (fun hi ↦ h' (hi.subset_ground.trans_eq hE.symm))) + have h' : M₁.Indep = M₂.Indep := by + ext I + by_cases hI : I ⊆ M₁.E + · rwa [h] + exact iff_of_false (fun hi ↦ hI hi.subset_ground) + (fun hi ↦ hI (hi.subset_ground.trans_eq hE.symm)) eq_of_base_iff_base_forall hE (fun B _ ↦ by simp_rw [base_iff_maximal_indep, h']) theorem eq_iff_indep_iff_indep_forall {M₁ M₂ : Matroid α} : @@ -709,12 +710,12 @@ section Basis /-- A Basis for a set `X ⊆ M.E` is a maximal independent subset of `X` (Often in the literature, the word 'Basis' is used to refer to what we call a 'Base'). -/ def Basis (M : Matroid α) (I X : Set α) : Prop := - I ∈ maximals (· ⊆ ·) {A | M.Indep A ∧ A ⊆ X} ∧ X ⊆ M.E + Maximal (fun A ↦ M.Indep A ∧ A ⊆ X) I ∧ X ⊆ M.E /-- A `Basis'` is a basis without the requirement that `X ⊆ M.E`. This is convenient for some API building, especially when working with rank and closure. -/ def Basis' (M : Matroid α) (I X : Set α) : Prop := - I ∈ maximals (· ⊆ ·) {A | M.Indep A ∧ A ⊆ X} + Maximal (fun A ↦ M.Indep A ∧ A ⊆ X) I variable {B I J X Y : Set α} {e : α} @@ -736,9 +737,6 @@ theorem Basis'.basis (hI : M.Basis' I X) (hX : X ⊆ M.E := by aesop_mat) : M.Ba theorem Basis'.subset (hI : M.Basis' I X) : I ⊆ X := hI.1.2 -theorem setOf_basis_eq (M : Matroid α) (hX : X ⊆ M.E := by aesop_mat) : - {I | M.Basis I X} = maximals (· ⊆ ·) ({I | M.Indep I} ∩ Iic X) := by - ext I; simp [Matroid.Basis, maximals, iff_true_intro hX] @[aesop unsafe 15% (rule_sets := [Matroid])] theorem Basis.subset_ground (hI : M.Basis I X) : X ⊆ M.E := @@ -759,20 +757,18 @@ theorem Basis.eq_of_subset_indep (hI : M.Basis I X) (hJ : M.Indep J) (hIJ : I theorem Basis.Finite (hI : M.Basis I X) [FiniteRk M] : I.Finite := hI.indep.finite theorem basis_iff' : - M.Basis I X ↔ (M.Indep I ∧ I ⊆ X ∧ ∀ J, M.Indep J → I ⊆ J → J ⊆ X → I = J) ∧ X ⊆ M.E := by - simp [Basis, mem_maximals_setOf_iff, and_assoc, and_congr_left_iff, and_imp, - and_congr_left_iff, and_congr_right_iff, @Imp.swap (_ ⊆ X)] + M.Basis I X ↔ (M.Indep I ∧ I ⊆ X ∧ ∀ ⦃J⦄, M.Indep J → I ⊆ J → J ⊆ X → I = J) ∧ X ⊆ M.E := by + rw [Basis, maximal_subset_iff] + tauto theorem basis_iff (hX : X ⊆ M.E := by aesop_mat) : M.Basis I X ↔ (M.Indep I ∧ I ⊆ X ∧ ∀ J, M.Indep J → I ⊆ J → J ⊆ X → I = J) := by rw [basis_iff', and_iff_left hX] theorem basis'_iff_basis_inter_ground : M.Basis' I X ↔ M.Basis I (X ∩ M.E) := by - rw [Basis', Basis, and_iff_left inter_subset_right] - convert Iff.rfl using 3 - ext I - simp only [subset_inter_iff, mem_setOf_eq, and_congr_right_iff, and_iff_left_iff_imp] - exact fun h _ ↦ h.subset_ground + rw [Basis', Basis, and_iff_left inter_subset_right, maximal_iff_maximal_of_imp_of_forall] + · exact fun I hI ↦ ⟨hI.1, hI.2.trans inter_subset_left⟩ + exact fun I hI ↦ ⟨I, rfl.le, hI.1, subset_inter hI.2 hI.1.subset_ground⟩ theorem basis'_iff_basis (hX : X ⊆ M.E := by aesop_mat) : M.Basis' I X ↔ M.Basis I X := by rw [basis'_iff_basis_inter_ground, inter_eq_self_of_subset_left hX] @@ -791,14 +787,10 @@ theorem Basis'.insert_not_indep (hI : M.Basis' I X) (he : e ∈ X \ I) : ¬ M.In fun hi ↦ he.2 <| insert_eq_self.1 <| Eq.symm <| hI.eq_of_subset_indep hi (subset_insert _ _) (insert_subset he.1 hI.subset) -theorem basis_iff_mem_maximals (hX : X ⊆ M.E := by aesop_mat) : - M.Basis I X ↔ I ∈ maximals (· ⊆ ·) {I | M.Indep I ∧ I ⊆ X} := by +theorem basis_iff_maximal (hX : X ⊆ M.E := by aesop_mat) : + M.Basis I X ↔ Maximal (fun I ↦ M.Indep I ∧ I ⊆ X) I := by rw [Basis, and_iff_left hX] -theorem basis_iff_mem_maximals_Prop (hX : X ⊆ M.E := by aesop_mat) : - M.Basis I X ↔ I ∈ maximals (· ⊆ ·) (fun I ↦ M.Indep I ∧ I ⊆ X) := - basis_iff_mem_maximals - theorem Indep.basis_of_maximal_subset (hI : M.Indep I) (hIX : I ⊆ X) (hmax : ∀ ⦃J⦄, M.Indep J → I ⊆ J → J ⊆ X → J ⊆ I) (hX : X ⊆ M.E := by aesop_mat) : M.Basis I X := by @@ -840,10 +832,8 @@ theorem Basis.not_basis_of_ssubset (hI : M.Basis I X) (hJI : J ⊂ I) : ¬ M.Bas theorem Indep.subset_basis_of_subset (hI : M.Indep I) (hIX : I ⊆ X) (hX : X ⊆ M.E := by aesop_mat) : ∃ J, M.Basis J X ∧ I ⊆ J := by - obtain ⟨J, ⟨(hJ : M.Indep J),hIJ,hJX⟩, hJmax⟩ := M.maximality X hX I hI hIX - use J - rw [and_iff_left hIJ, basis_iff, and_iff_right hJ, and_iff_right hJX] - exact fun K hK hJK hKX ↦ hJK.antisymm (hJmax ⟨hK, hIJ.trans hJK, hKX⟩ hJK) + obtain ⟨J, hJ, hJmax⟩ := M.maximality X hX I hI hIX + exact ⟨J, ⟨hJmax, hX⟩, hJ⟩ theorem Indep.subset_basis'_of_subset (hI : M.Indep I) (hIX : I ⊆ X) : ∃ J, M.Basis' J X ∧ I ⊆ J := by @@ -887,23 +877,21 @@ theorem Basis.exists_base (hI : M.Basis I X) : ∃ B, M.Base B ∧ I = B ∩ X : inter_subset_right])⟩ @[simp] theorem basis_ground_iff : M.Basis B M.E ↔ M.Base B := by - rw [base_iff_maximal_indep, basis_iff', and_assoc, and_congr_right] - rw [and_iff_left (rfl.subset : M.E ⊆ M.E)] - exact fun h ↦ ⟨fun h' I hI hBI ↦ h'.2 _ hI hBI hI.subset_ground, - fun h' ↦ ⟨h.subset_ground,fun J hJ hBJ _ ↦ h' J hJ hBJ⟩⟩ + rw [Basis, and_iff_left rfl.subset, base_iff_maximal_indep, + maximal_and_iff_right_of_imp (fun _ h ↦ h.subset_ground), + and_iff_left_of_imp (fun h ↦ h.1.subset_ground)] theorem Base.basis_ground (hB : M.Base B) : M.Basis B M.E := basis_ground_iff.mpr hB theorem Indep.basis_iff_forall_insert_dep (hI : M.Indep I) (hIX : I ⊆ X) : M.Basis I X ↔ ∀ e ∈ X \ I, M.Dep (insert e I) := by - rw [basis_iff', and_iff_right hIX, and_iff_right hI] - refine ⟨fun h e he ↦ ⟨fun hi ↦ he.2 ?_, insert_subset (h.2 he.1) hI.subset_ground⟩, - fun h ↦ ⟨fun J hJ hIJ hJX ↦ hIJ.antisymm (fun e heJ ↦ by_contra (fun heI ↦ ?_)), ?_⟩⟩ - · exact (h.1 _ hi (subset_insert _ _) (insert_subset he.1 hIX)).symm.subset (mem_insert e I) - · exact (h e ⟨hJX heJ, heI⟩).not_indep (hJ.subset (insert_subset heJ hIJ)) - rw [← diff_union_of_subset hIX, union_subset_iff, and_iff_left hI.subset_ground] - exact fun e he ↦ (h e he).subset_ground (mem_insert _ _) + rw [Basis, maximal_iff_forall_insert (fun I J hI hIJ ↦ ⟨hI.1.subset hIJ, hIJ.trans hI.2⟩)] + simp only [hI, hIX, and_self, insert_subset_iff, and_true, not_and, true_and, mem_diff, and_imp, + Dep, hI.subset_ground] + exact ⟨fun h e heX heI ↦ ⟨fun hi ↦ h.1 e heI hi heX, h.2 heX⟩, + fun h ↦ ⟨fun e heI hi heX ↦ (h e heX heI).1 hi, + fun e heX ↦ (em (e ∈ I)).elim (fun h ↦ hI.subset_ground h) fun heI ↦ (h _ heX heI).2 ⟩⟩ theorem Indep.basis_of_forall_insert (hI : M.Indep I) (hIX : I ⊆ X) (he : ∀ e ∈ X \ I, M.Dep (insert e I)) : M.Basis I X := diff --git a/Mathlib/Data/Matroid/Constructions.lean b/Mathlib/Data/Matroid/Constructions.lean index 6d45423bd5d7d..d9be13a22679b 100644 --- a/Mathlib/Data/Matroid/Constructions.lean +++ b/Mathlib/Data/Matroid/Constructions.lean @@ -45,7 +45,7 @@ def emptyOn (α : Type*) : Matroid α where indep_iff' := by simp [subset_empty_iff] exists_base := ⟨∅, rfl⟩ base_exchange := by rintro _ _ rfl; simp - maximality := by rintro _ _ _ rfl -; exact ⟨∅, by simp [mem_maximals_iff]⟩ + maximality := by rintro _ _ _ rfl -; exact ⟨∅, by simp [Maximal]⟩ subset_ground := by simp @[simp] theorem emptyOn_ground : (emptyOn α).E = ∅ := rfl @@ -97,8 +97,7 @@ theorem eq_loopyOn_iff : M = loopyOn E ↔ M.E = E ∧ ∀ X ⊆ M.E, M.Indep X refine ⟨fun h I hI ↦ (h I hI).1, fun h I hIE ↦ ⟨h I hIE, by rintro rfl; simp⟩⟩ @[simp] theorem loopyOn_base_iff : (loopyOn E).Base B ↔ B = ∅ := by - simp only [base_iff_maximal_indep, loopyOn_indep_iff, forall_eq, and_iff_left_iff_imp] - exact fun h _ ↦ h + simp [Maximal, base_iff_maximal_indep] @[simp] theorem loopyOn_basis_iff : (loopyOn E).Basis I X ↔ I = ∅ ∧ X ⊆ E := ⟨fun h ↦ ⟨loopyOn_indep_iff.mp h.indep, h.subset_ground⟩, @@ -116,9 +115,10 @@ theorem Finite.loopyOn_finite (hE : E.Finite) : Matroid.Finite (loopyOn E) := exact fun _ h _ ↦ h theorem empty_base_iff : M.Base ∅ ↔ M = loopyOn M.E := by - simp only [base_iff_maximal_indep, empty_indep, empty_subset, eq_comm (a := ∅), true_implies, - true_and, eq_iff_indep_iff_indep_forall, loopyOn_ground, loopyOn_indep_iff] - exact ⟨fun h I _ ↦ ⟨h I, by rintro rfl; simp⟩, fun h I hI ↦ (h I hI.subset_ground).1 hI⟩ + simp only [base_iff_maximal_indep, Maximal, empty_indep, le_eq_subset, empty_subset, + subset_empty_iff, true_implies, true_and, eq_iff_indep_iff_indep_forall, loopyOn_ground, + loopyOn_indep_iff] + exact ⟨fun h I _ ↦ ⟨@h _, fun hI ↦ by simp [hI]⟩, fun h I hI ↦ (h I hI.subset_ground).1 hI⟩ theorem eq_loopyOn_or_rkPos (M : Matroid α) : M = loopyOn M.E ∨ RkPos M := by rw [← empty_base_iff, rkPos_iff_empty_not_base]; apply em @@ -208,15 +208,13 @@ theorem uniqueBaseOn_indep_iff (hIE : I ⊆ E) : (uniqueBaseOn I E).Indep J ↔ rw [uniqueBaseOn, restrict_indep_iff, freeOn_indep_iff, and_iff_left_iff_imp] exact fun h ↦ h.trans hIE -theorem uniqueBaseOn_basis_iff (hI : I ⊆ E) (hX : X ⊆ E) : - (uniqueBaseOn I E).Basis J X ↔ J = X ∩ I := by - rw [basis_iff_mem_maximals] - simp_rw [uniqueBaseOn_indep_iff', ← subset_inter_iff, ← le_eq_subset, Iic_def, maximals_Iic, - mem_singleton_iff, inter_eq_self_of_subset_left hI, inter_comm I] +theorem uniqueBaseOn_basis_iff (hX : X ⊆ E) : (uniqueBaseOn I E).Basis J X ↔ J = X ∩ I := by + rw [basis_iff_maximal] + exact maximal_iff_eq (by simp [inter_subset_left.trans hX]) + (by simp (config := {contextual := true})) -theorem uniqueBaseOn_inter_basis (hI : I ⊆ E) (hX : X ⊆ E) : - (uniqueBaseOn I E).Basis (X ∩ I) X := by - rw [uniqueBaseOn_basis_iff hI hX] +theorem uniqueBaseOn_inter_basis (hX : X ⊆ E) : (uniqueBaseOn I E).Basis (X ∩ I) X := by + rw [uniqueBaseOn_basis_iff hX] @[simp] theorem uniqueBaseOn_dual_eq (I E : Set α) : (uniqueBaseOn I E)✶ = uniqueBaseOn (E \ I) E := by diff --git a/Mathlib/Data/Matroid/Dual.lean b/Mathlib/Data/Matroid/Dual.lean index 21a3f579dc710..3d79fbd6d5b3c 100644 --- a/Mathlib/Data/Matroid/Dual.lean +++ b/Mathlib/Data/Matroid/Dual.lean @@ -48,10 +48,10 @@ section dual indep_aug := by rintro I X ⟨hIE, B, hB, hIB⟩ hI_not_max hX_max have hXE := hX_max.1.1 - have hB' := (base_compl_iff_mem_maximals_disjoint_base hXE).mpr hX_max + have hB' := (base_compl_iff_maximal_disjoint_base hXE).mpr hX_max set B' := M.E \ X with hX - have hI := (not_iff_not.mpr (base_compl_iff_mem_maximals_disjoint_base)).mpr hI_not_max + have hI := (not_iff_not.mpr (base_compl_iff_maximal_disjoint_base)).mpr hI_not_max obtain ⟨B'', hB'', hB''₁, hB''₂⟩ := (hB'.indep.diff I).exists_base_subset_union_base hB rw [← compl_subset_compl, ← hIB.sdiff_eq_right, ← union_diff_distrib, diff_eq, compl_inter, compl_compl, union_subset_iff, compl_subset_compl] at hB''₂ @@ -67,47 +67,40 @@ section dual rw [← union_singleton, disjoint_union_left, disjoint_singleton_left, and_iff_left heB''] exact disjoint_of_subset_left hB''₂.2 disjoint_compl_left indep_maximal := by - rintro X - I'⟨hI'E, B, hB, hI'B⟩ hI'X + rintro X - I' ⟨hI'E, B, hB, hI'B⟩ hI'X obtain ⟨I, hI⟩ := M.exists_basis (M.E \ X) obtain ⟨B', hB', hIB', hB'IB⟩ := hI.indep.exists_base_subset_union_base hB - refine ⟨(X \ B') ∩ M.E, - ⟨?_, subset_inter (subset_diff.mpr ?_) hI'E, inter_subset_left.trans - diff_subset⟩, ?_⟩ - · simp only [inter_subset_right, true_and] - exact ⟨B', hB', disjoint_of_subset_left inter_subset_left disjoint_sdiff_left⟩ - · rw [and_iff_right hI'X] - refine disjoint_of_subset_right hB'IB ?_ - rw [disjoint_union_right, and_iff_left hI'B] - exact disjoint_of_subset hI'X hI.subset disjoint_sdiff_right - simp only [mem_setOf_eq, subset_inter_iff, and_imp, forall_exists_index] - intros J hJE B'' hB'' hdj _ hJX hssJ - rw [and_iff_left hJE] - rw [diff_eq, inter_right_comm, ← diff_eq, diff_subset_iff] at hssJ - have hI' : (B'' ∩ X) ∪ (B' \ X) ⊆ B' := by - rw [union_subset_iff, and_iff_left diff_subset, - ← inter_eq_self_of_subset_left hB''.subset_ground, inter_right_comm, inter_assoc] + obtain rfl : I = B' \ X := hI.eq_of_subset_indep (hB'.indep.diff _) + (subset_diff.2 ⟨hIB', (subset_diff.1 hI.subset).2⟩) + (diff_subset_diff_left hB'.subset_ground) + simp_rw [maximal_subset_iff'] + refine ⟨(X \ B') ∩ M.E, ?_, ⟨⟨inter_subset_right, ?_⟩, ?_⟩, ?_⟩ + · rw [subset_inter_iff, and_iff_left hI'E, subset_diff, and_iff_right hI'X] + exact Disjoint.mono_right hB'IB <| disjoint_union_right.2 + ⟨disjoint_sdiff_right.mono_left hI'X , hI'B⟩ + · exact ⟨B', hB', (disjoint_sdiff_left (t := X)).mono_left inter_subset_left⟩ + · exact inter_subset_left.trans diff_subset + simp only [subset_inter_iff, subset_diff, and_imp, forall_exists_index] + refine fun J hJE B'' hB'' hdj hJX hXJ ↦ ⟨⟨hJX, ?_⟩, hJE⟩ - calc _ ⊆ _ := inter_subset_inter_right _ hssJ - _ ⊆ _ := by rw [inter_union_distrib_left, hdj.symm.inter_eq, union_empty] - _ ⊆ _ := inter_subset_right + have hI' : (B'' ∩ X) ∪ (B' \ X) ⊆ B' := by + rw [union_subset_iff, and_iff_left diff_subset, ← union_diff_cancel hJX, + inter_union_distrib_left, hdj.symm.inter_eq, empty_union, diff_eq, ← inter_assoc, + ← diff_eq, diff_subset_comm, diff_eq, inter_assoc, ← diff_eq, inter_comm] + exact subset_trans (inter_subset_inter_right _ hB''.subset_ground) hXJ obtain ⟨B₁,hB₁,hI'B₁,hB₁I⟩ := (hB'.indep.subset hI').exists_base_subset_union_base hB'' rw [union_comm, ← union_assoc, union_eq_self_of_subset_right inter_subset_left] at hB₁I - have : B₁ = B' := by + obtain rfl : B₁ = B' := by refine hB₁.eq_of_subset_indep hB'.indep (fun e he ↦ ?_) refine (hB₁I he).elim (fun heB'' ↦ ?_) (fun h ↦ h.1) refine (em (e ∈ X)).elim (fun heX ↦ hI' (Or.inl ⟨heB'', heX⟩)) (fun heX ↦ hIB' ?_) - refine hI.mem_of_insert_indep ⟨hB₁.subset_ground he, heX⟩ - (hB₁.indep.subset (insert_subset he ?_)) - refine (subset_union_of_subset_right (subset_diff.mpr ⟨hIB',?_⟩) _).trans hI'B₁ - exact disjoint_of_subset_left hI.subset disjoint_sdiff_left - - subst this - - refine subset_diff.mpr ⟨hJX, by_contra (fun hne ↦ ?_)⟩ - obtain ⟨e, heJ, heB'⟩ := not_disjoint_iff.mp hne + refine hI.mem_of_insert_indep ⟨hB₁.subset_ground he, heX⟩ ?_ + exact hB₁.indep.subset (insert_subset he (subset_union_right.trans hI'B₁)) + by_contra hdj' + obtain ⟨e, heJ, heB'⟩ := not_disjoint_iff.mp hdj' obtain (heB'' | ⟨-,heX⟩ ) := hB₁I heB' · exact hdj.ne_of_mem heJ heB'' rfl exact heX (hJX heJ) @@ -140,9 +133,9 @@ instance dual_nonempty [M.Nonempty] : M✶.Nonempty := ⟨M.ground_nonempty⟩ @[simp] theorem dual_base_iff (hB : B ⊆ M.E := by aesop_mat) : M✶.Base B ↔ M.Base (M.E \ B) := by - rw [base_compl_iff_mem_maximals_disjoint_base, base_iff_maximal_indep, dual_indep_iff_exists', - mem_maximals_setOf_iff] - simp [dual_indep_iff_exists'] + rw [base_compl_iff_maximal_disjoint_base, base_iff_maximal_indep, maximal_subset_iff, + maximal_subset_iff] + simp [dual_indep_iff_exists', hB] theorem dual_base_iff' : M✶.Base B ↔ M.Base (M.E \ B) ∧ B ⊆ M.E := (em (B ⊆ M.E)).elim (fun h ↦ by rw [dual_base_iff, and_iff_left h]) diff --git a/Mathlib/Data/Matroid/IndepAxioms.lean b/Mathlib/Data/Matroid/IndepAxioms.lean index c83aa7323ea20..a71adf7dde620 100644 --- a/Mathlib/Data/Matroid/IndepAxioms.lean +++ b/Mathlib/Data/Matroid/IndepAxioms.lean @@ -25,7 +25,7 @@ and some form of 'augmentation' axiom, which allows one to enlarge a non-maximal This augmentation axiom is still required when there are finiteness assumptions, but is simpler. It just states that if `I` is a finite independent set and `J` is a larger finite independent set, then there exists `e ∈ J \ I` for which `insert e I` is independent. -This is the axiom that appears in all most of the definitions. +This is the axiom that appears in most of the definitions. ## Implementation Details @@ -97,8 +97,8 @@ structure IndepMatroid (α : Type*) where (Indep : Set α → Prop) (indep_empty : Indep ∅) (indep_subset : ∀ ⦃I J⦄, Indep J → I ⊆ J → Indep I) - (indep_aug : ∀⦃I B⦄, Indep I → I ∉ maximals (· ⊆ ·) {I | Indep I} → - B ∈ maximals (· ⊆ ·) {I | Indep I} → ∃ x ∈ B \ I, Indep (insert x I)) + (indep_aug : ∀ ⦃I B⦄, Indep I → ¬ Maximal Indep I → + Maximal Indep B → ∃ x ∈ B \ I, Indep (insert x I)) (indep_maximal : ∀ X, X ⊆ E → ExistsMaximalSubsetProperty Indep X) (subset_ground : ∀ I, Indep I → I ⊆ E) @@ -107,36 +107,28 @@ namespace IndepMatroid /-- An `M : IndepMatroid α` gives a `Matroid α` whose bases are the maximal `M`-independent sets. -/ @[simps] protected def matroid (M : IndepMatroid α) : Matroid α where E := M.E - Base := (· ∈ maximals (· ⊆ ·) {I | M.Indep I}) + Base := Maximal M.Indep Indep := M.Indep indep_iff' := by refine fun I ↦ ⟨fun h ↦ ?_, fun ⟨B, ⟨h, _⟩, hIB'⟩ ↦ M.indep_subset h hIB'⟩ - obtain ⟨B, hB⟩ := M.indep_maximal M.E Subset.rfl I h (M.subset_ground I h) - simp only [mem_maximals_iff, mem_setOf_eq, and_imp] at hB ⊢ - exact ⟨B, ⟨hB.1.1,fun J hJ hBJ ↦ hB.2 hJ (hB.1.2.1.trans hBJ) (M.subset_ground J hJ) hBJ⟩, - hB.1.2.1⟩ + obtain ⟨J, hIJ, hmax⟩ := M.indep_maximal M.E rfl.subset I h (M.subset_ground I h) + rw [maximal_and_iff_right_of_imp M.subset_ground] at hmax + exact ⟨J, hmax.1, hIJ⟩ exists_base := by - obtain ⟨B, ⟨hB,-,-⟩, hB₁⟩ := - M.indep_maximal M.E rfl.subset ∅ M.indep_empty (empty_subset _) - exact ⟨B, ⟨hB, fun B' hB' h' ↦ hB₁ ⟨hB', empty_subset _,M.subset_ground B' hB'⟩ h'⟩⟩ - base_exchange := by - rintro B B' ⟨hB, hBmax⟩ ⟨hB',hB'max⟩ e he - have hnotmax : B \ {e} ∉ maximals (· ⊆ ·) {I | M.Indep I} := by - simp only [mem_maximals_setOf_iff, diff_singleton_subset_iff, not_and, not_forall, - exists_prop, exists_and_left] - exact fun _ ↦ ⟨B, hB, subset_insert _ _, by simpa using he.1⟩ - - obtain ⟨f,hf,hfB⟩ := M.indep_aug (M.indep_subset hB diff_subset) hnotmax ⟨hB',hB'max⟩ - simp only [mem_diff, mem_singleton_iff, not_and, not_not] at hf - - have hfB' : f ∉ B := by (intro hfB; obtain rfl := hf.2 hfB; exact he.2 hf.1) - - refine ⟨f, ⟨hf.1, hfB'⟩, by_contra (fun hnot ↦ ?_)⟩ - obtain ⟨x,hxB, hind⟩ := M.indep_aug hfB hnot ⟨hB, hBmax⟩ - simp only [mem_diff, mem_insert_iff, mem_singleton_iff, not_or, not_and, not_not] at hxB - obtain rfl := hxB.2.2 hxB.1 - rw [insert_comm, insert_diff_singleton, insert_eq_of_mem he.1] at hind - exact not_mem_subset (hBmax hind (subset_insert _ _)) hfB' (mem_insert _ _) + obtain ⟨B, -, hB⟩ := M.indep_maximal M.E rfl.subset ∅ M.indep_empty <| empty_subset _ + rw [maximal_and_iff_right_of_imp M.subset_ground] at hB + exact ⟨B, hB.1⟩ + base_exchange B B' hB hB' e he := by + have hnotmax : ¬ Maximal M.Indep (B \ {e}) := + fun h ↦ h.not_prop_of_ssuperset (diff_singleton_sSubset.2 he.1) hB.prop + obtain ⟨f, hf, hfB⟩ := M.indep_aug (M.indep_subset hB.prop diff_subset) hnotmax hB' + replace hf := show f ∈ B' \ B by simpa [show f ≠ e by rintro rfl; exact he.2 hf.1] using hf + refine ⟨f, hf, by_contra fun hnot ↦ ?_⟩ + obtain ⟨x, hxB, hind⟩ := M.indep_aug hfB hnot hB + obtain ⟨-, rfl⟩ : _ ∧ x = e := by simpa [hxB.1] using hxB + refine hB.not_prop_of_ssuperset ?_ hind + rw [insert_comm, insert_diff_singleton, insert_eq_of_mem he.1] + exact ssubset_insert hf.2 maximality := M.indep_maximal subset_ground B hB := M.subset_ground B hB.1 @@ -167,23 +159,18 @@ namespace IndepMatroid (indep_subset := indep_subset) (indep_aug := by intro I B hI hImax hBmax - simp only [mem_maximals_iff, mem_setOf_eq, not_and, not_forall, exists_prop, - exists_and_left, iff_true_intro hI, true_imp_iff] at hImax hBmax - obtain ⟨I', hI', hII', hne⟩ := hImax - obtain ⟨e, heI', heI⟩ := exists_of_ssubset (hII'.ssubset_of_ne hne) - have hins : Indep (insert e I) := indep_subset hI' (insert_subset heI' hII') - obtain (heB | heB) := em (e ∈ B) + obtain ⟨e, heI, hins⟩ := exists_insert_of_not_maximal indep_subset hI hImax + by_cases heB : e ∈ B · exact ⟨e, ⟨heB, heI⟩, hins⟩ by_contra hcon; push_neg at hcon - have heBdep : ¬Indep (insert e B) := - fun hi ↦ heB <| insert_eq_self.1 (hBmax.2 hi (subset_insert _ _)).symm + have heBdep := hBmax.not_prop_of_ssuperset (ssubset_insert heB) -- There is a finite subset `B₀` of `B` so that `B₀ + e` is dependent - obtain ⟨B₀, hB₀B, hB₀fin, hB₀e⟩ := htofin B e hBmax.1 heBdep + obtain ⟨B₀, hB₀B, hB₀fin, hB₀e⟩ := htofin B e hBmax.1 heBdep have hB₀ := indep_subset hBmax.1 hB₀B - -- There is a finite subset `I₀` of `I` so that `I₀` doesn't extend into `B₀` + -- `I` has a finite subset `I₀` that doesn't extend into `B₀` have hexI₀ : ∃ I₀, I₀ ⊆ I ∧ I₀.Finite ∧ ∀ x, x ∈ B₀ \ I₀ → ¬Indep (insert x I₀) := by have hchoose : ∀ (b : ↑(B₀ \ I)), ∃ Ib, Ib ⊆ I ∧ Ib.Finite ∧ ¬Indep (insert (b : α) Ib) := by rintro ⟨b, hb⟩; exact htofin I b hI (hcon b ⟨hB₀B hb.1, hb.2⟩) @@ -232,8 +219,9 @@ namespace IndepMatroid (indep_maximal := by rintro X - I hI hIX have hzorn := zorn_subset_nonempty {Y | Indep Y ∧ I ⊆ Y ∧ Y ⊆ X} ?_ I ⟨hI, Subset.rfl, hIX⟩ - · obtain ⟨J, hJ, -, hJmax⟩ := hzorn - exact ⟨J, hJ, fun K hK hJK ↦ (hJmax K hK hJK).subset⟩ + · obtain ⟨J, ⟨hJi, hIJ, hJX⟩, -, hJmax⟩ := hzorn + simp_rw [maximal_subset_iff] + exact ⟨J, hIJ, ⟨hJi, hJX⟩, fun K hK hJK ↦ (hJmax K ⟨hK.1, hIJ.trans hJK, hK.2⟩ hJK).symm⟩ refine fun Is hIs hchain ⟨K, hK⟩ ↦ ⟨⋃₀ Is, ⟨?_,?_,?_⟩, fun _ ↦ subset_sUnion_of_mem⟩ · refine indep_compact _ fun J hJ hJfin ↦ ?_ @@ -273,20 +261,22 @@ theorem _root_.Matroid.existsMaximalSubsetProperty_of_bdd {P : Set α → Prop} rintro x ⟨Y, ⟨hY,-,-⟩, rfl⟩ obtain ⟨n₀, heq, hle⟩ := hP Y hY rwa [ncard_def, heq, ENat.toNat_coe] - -- have := (hP Y hY).2 - obtain ⟨Y, hY, hY'⟩ := Finite.exists_maximal_wrt' ncard _ hfin ⟨I, hI, rfl.subset, hIX⟩ - refine ⟨Y, hY, fun J ⟨hJ, hIJ, hJX⟩ (hYJ : Y ⊆ J) ↦ (?_ : J ⊆ Y)⟩ - have hJfin := finite_of_encard_le_coe (hP J hJ) - refine (eq_of_subset_of_ncard_le hYJ ?_ hJfin).symm.subset - rw [hY' J ⟨hJ, hIJ, hJX⟩ (ncard_le_ncard hYJ hJfin)] + obtain ⟨Y, ⟨hY, hIY, hYX⟩, hY'⟩ := + Finite.exists_maximal_wrt' ncard _ hfin ⟨I, hI, rfl.subset, hIX⟩ + + refine ⟨Y, hIY, ⟨hY, hYX⟩, fun K ⟨hPK, hKX⟩ hYK ↦ ?_⟩ + have hKfin : K.Finite := finite_of_encard_le_coe (hP K hPK) + + refine (eq_of_subset_of_ncard_le hYK ?_ hKfin).symm.subset + rw [hY' K ⟨hPK, hIY.trans hYK, hKX⟩ (ncard_le_ncard hYK hKfin)] /-- If there is an absolute upper bound on the size of an independent set, then the maximality axiom isn't needed to define a matroid by independent sets. -/ @[simps E] protected def ofBdd (E : Set α) (Indep : Set α → Prop) (indep_empty : Indep ∅) (indep_subset : ∀ ⦃I J⦄, Indep J → I ⊆ J → Indep I) - (indep_aug : ∀⦃I B⦄, Indep I → I ∉ maximals (· ⊆ ·) {I | Indep I} → - B ∈ maximals (· ⊆ ·) {I | Indep I} → ∃ x ∈ B \ I, Indep (insert x I)) + (indep_aug : ∀⦃I B⦄, Indep I → ¬ Maximal Indep I → Maximal Indep B → + ∃ x ∈ B \ I, Indep (insert x I)) (subset_ground : ∀ I, Indep I → I ⊆ E) (indep_bdd : ∃ (n : ℕ), ∀ I, Indep I → I.encard ≤ n ) : IndepMatroid α where E := E @@ -324,20 +314,19 @@ protected def ofBddAugment (E : Set α) (Indep : Set α → Prop) (indep_empty := indep_empty) (indep_subset := indep_subset) (indep_aug := by - simp_rw [mem_maximals_setOf_iff, not_and, not_forall, exists_prop, mem_diff, - and_imp, and_assoc] - rintro I B hI hImax hB hBmax - obtain ⟨J, hJ, hIJ, hne⟩ := hImax hI - obtain ⟨n, h_bdd⟩ := indep_bdd - - have hlt : I.encard < J.encard := - (finite_of_encard_le_coe (h_bdd J hJ)).encard_lt_encard (hIJ.ssubset_of_ne hne) - have hle : J.encard ≤ B.encard := by - refine le_of_not_lt (fun hlt' ↦ ?_) - obtain ⟨e, he⟩ := indep_aug hB hJ hlt' - rw [hBmax he.2.2 (subset_insert _ _)] at he - exact he.2.1 (mem_insert _ _) - exact indep_aug hI hB (hlt.trans_le hle)) + rintro I B hI hImax hBmax + suffices hcard : I.encard < B.encard by + obtain ⟨e, heB, heI, hi⟩ := indep_aug hI hBmax.prop hcard + exact ⟨e, ⟨heB, heI⟩, hi⟩ + refine lt_of_not_le fun hle ↦ ?_ + obtain ⟨x, hxnot, hxI⟩ := exists_insert_of_not_maximal indep_subset hI hImax + have hlt : B.encard < (insert x I).encard := by + rwa [encard_insert_of_not_mem hxnot, ← not_le, ENat.add_one_le_iff, not_lt] + rw [encard_ne_top_iff] + obtain ⟨n, hn⟩ := indep_bdd + exact finite_of_encard_le_coe (hn _ hI) + obtain ⟨y, -, hyB, hi⟩ := indep_aug hBmax.prop hxI hlt + exact hBmax.not_prop_of_ssuperset (ssubset_insert hyB) hi) (indep_bdd := indep_bdd) (subset_ground := subset_ground) @[simp] theorem ofBddAugment_E (E : Set α) Indep indep_empty indep_subset indep_aug @@ -444,7 +433,7 @@ namespace Matroid IndepMatroid.mk (E := E) (Indep := Indep) (indep_empty := by obtain ⟨M, -, rfl⟩ := hex; exact M.empty_indep) (indep_subset := by obtain ⟨M, -, rfl⟩ := hex; exact fun I J hJ hIJ ↦ hJ.subset hIJ) - (indep_aug := by obtain ⟨M, -, rfl⟩ := hex; exact Indep.exists_insert_of_not_mem_maximals M) + (indep_aug := by obtain ⟨M, -, rfl⟩ := hex; exact Indep.exists_insert_of_not_maximal M) (indep_maximal := by obtain ⟨M, rfl, rfl⟩ := hex; exact M.existsMaximalSubsetProperty_indep) (subset_ground := by obtain ⟨M, rfl, rfl⟩ := hex; exact fun I ↦ Indep.subset_ground) diff --git a/Mathlib/Data/Matroid/Map.lean b/Mathlib/Data/Matroid/Map.lean index 74b9dcb61b27b..b231376efdada 100644 --- a/Mathlib/Data/Matroid/Map.lean +++ b/Mathlib/Data/Matroid/Map.lean @@ -119,47 +119,43 @@ def comap (N : Matroid β) (f : α → β) : Matroid α := indep_subset := fun I J h hIJ ↦ ⟨h.1.subset (image_subset _ hIJ), InjOn.mono hIJ h.2⟩ indep_aug := by rintro I B ⟨hI, hIinj⟩ hImax hBmax - simp only [mem_maximals_iff, mem_setOf_eq, hI, hIinj, and_self, and_imp, - true_and, not_forall, exists_prop, exists_and_left] at hImax hBmax - - obtain ⟨I', hI', hI'inj, hII', hne⟩ := hImax + obtain ⟨I', hII', hI', hI'inj⟩ := (not_maximal_subset_iff ⟨hI, hIinj⟩).1 hImax have h₁ : ¬(N ↾ range f).Base (f '' I) := by - refine fun hB ↦ hne ?_ - have h_im := hB.eq_of_subset_indep (by simpa) (image_subset _ hII') - rwa [hI'inj.image_eq_image_iff hII' Subset.rfl] at h_im + refine fun hB ↦ hII'.ne ?_ + have h_im := hB.eq_of_subset_indep (by simpa) (image_subset _ hII'.subset) + rwa [hI'inj.image_eq_image_iff hII'.subset Subset.rfl] at h_im have h₂ : (N ↾ range f).Base (f '' B) := by refine Indep.base_of_forall_insert (by simpa using hBmax.1.1) ?_ rintro _ ⟨⟨e, heB, rfl⟩, hfe⟩ hi rw [restrict_indep_iff, ← image_insert_eq] at hi have hinj : InjOn f (insert e B) := by - rw [injOn_insert (fun heB ↦ hfe (mem_image_of_mem f heB))]; exact ⟨hBmax.1.2, hfe⟩ - rw [hBmax.2 hi.1 hinj <| subset_insert _ _] at hfe; simp at hfe + rw [injOn_insert (fun heB ↦ hfe (mem_image_of_mem f heB))] + exact ⟨hBmax.1.2, hfe⟩ + refine hBmax.not_prop_of_ssuperset (t := insert e B) (ssubset_insert ?_) ⟨hi.1, hinj⟩ + exact fun heB ↦ hfe <| mem_image_of_mem f heB obtain ⟨_, ⟨⟨e, he, rfl⟩, he'⟩, hei⟩ := Indep.exists_insert_of_not_base (by simpa) h₁ h₂ have heI : e ∉ I := fun heI ↦ he' (mem_image_of_mem f heI) rw [← image_insert_eq, restrict_indep_iff] at hei exact ⟨e, ⟨he, heI⟩, hei.1, (injOn_insert heI).2 ⟨hIinj, he'⟩⟩ - indep_maximal := by + indep_maximal := by rintro X - I ⟨hI, hIinj⟩ hIX obtain ⟨J, hJ⟩ := (N ↾ range f).existsMaximalSubsetProperty_indep (f '' X) (by simp) (f '' I) (by simpa) (image_subset _ hIX) - simp only [restrict_indep_iff, image_subset_iff, mem_maximals_iff, mem_setOf_eq, - and_imp] at hJ + simp only [restrict_indep_iff, image_subset_iff, maximal_subset_iff, mem_setOf_eq, and_imp, + and_assoc] at hJ ⊢ + obtain ⟨hIJ, hJ, hJf, hJX, hJmax⟩ := hJ obtain ⟨J₀, hIJ₀, hJ₀X, hbj⟩ := hIinj.bijOn_image.exists_extend_of_subset hIX - (image_subset f hJ.1.2.1) (image_subset_iff.2 <| preimage_mono hJ.1.2.2) - - have him := hbj.image_eq; rw [image_preimage_eq_of_subset hJ.1.1.2] at him; subst him - use J₀ - simp only [mem_maximals_iff, mem_setOf_eq, hJ.1.1.1, hbj.injOn, and_self, hIJ₀, - hJ₀X, and_imp, true_and] - intro K hK hinj hIK hKX hJ₀K - rw [← hinj.image_eq_image_iff hJ₀K Subset.rfl, hJ.2 hK (image_subset_range _ _) - (fun e he ↦ ⟨e, hIK he, rfl⟩) (image_subset _ hKX) (image_subset _ hJ₀K)] + (image_subset f hIJ) (image_subset_iff.2 <| preimage_mono hJX) + obtain rfl : f '' J₀ = J := by rw [← image_preimage_eq_of_subset hJf, hbj.image_eq] + refine ⟨J₀, hIJ₀, hJ, hbj.injOn, hJ₀X, fun K hK hKinj hKX hJ₀K ↦ ?_⟩ + rw [← hKinj.image_eq_image_iff hJ₀K Subset.rfl, hJmax hK (image_subset_range _ _) + (image_subset f hKX) (image_subset f hJ₀K)] subset_ground := fun I hI e heI ↦ hI.1.subset_ground ⟨e, heI, rfl⟩ } @[simp] lemma comap_indep_iff : (N.comap f).Indep I ↔ N.Indep (f '' I) ∧ InjOn f I := Iff.rfl @@ -375,12 +371,13 @@ lemma map_image_indep_iff {hf} {I : Set α} (hI : I ⊆ M.E) : @[simp] lemma map_base_iff (M : Matroid α) (f : α → β) (hf) {B : Set β} : (M.map f hf).Base B ↔ ∃ B₀, M.Base B₀ ∧ B = f '' B₀ := by rw [base_iff_maximal_indep] - refine ⟨fun ⟨h, h'⟩ ↦ ?_, ?_⟩ - · obtain ⟨B₀, hB₀, hbij⟩ := h.exists_bijOn_of_map + refine ⟨fun h ↦ ?_, ?_⟩ + · obtain ⟨B₀, hB₀, hbij⟩ := h.prop.exists_bijOn_of_map refine ⟨B₀, hB₀.base_of_maximal fun J hJ hB₀J ↦ ?_, hbij.image_eq.symm⟩ - rw [← hf.image_eq_image_iff hB₀.subset_ground hJ.subset_ground] - exact hbij.image_eq ▸ h' _ (hJ.map f hf) (hbij.image_eq ▸ image_subset f hB₀J) + rw [← hf.image_eq_image_iff hB₀.subset_ground hJ.subset_ground, hbij.image_eq] + exact h.eq_of_subset (hJ.map f hf) (hbij.image_eq ▸ image_subset f hB₀J) rintro ⟨B, hB, rfl⟩ + rw [maximal_subset_iff] refine ⟨hB.indep.map f hf, fun I hI hBI ↦ ?_⟩ obtain ⟨I₀, hI₀, hbij⟩ := hI.exists_bijOn_of_map rw [← hbij.image_eq, hf.image_subset_image_iff hB.subset_ground hI₀.subset_ground] at hBI diff --git a/Mathlib/Data/Matroid/Restrict.lean b/Mathlib/Data/Matroid/Restrict.lean index 949d84729c5d5..46d88f70b7cf7 100644 --- a/Mathlib/Data/Matroid/Restrict.lean +++ b/Mathlib/Data/Matroid/Restrict.lean @@ -105,11 +105,12 @@ section restrict exact ⟨e, ⟨⟨(hBIB' he.1.1).elim (fun h ↦ (he.2 h).elim) id,he.1.2⟩, he.2⟩, hI'.indep.subset (insert_subset he.1 hIJ), insert_subset he.1.2.1 hIY⟩ indep_maximal := by - rintro A hAX I ⟨hI, _⟩ hIA - obtain ⟨J, hJ, hIJ⟩ := hI.subset_basis'_of_subset hIA; use J - rw [mem_maximals_setOf_iff, and_iff_left hJ.subset, and_iff_left hIJ, - and_iff_right ⟨hJ.indep, hJ.subset.trans hAX⟩] - exact fun K ⟨⟨hK, _⟩, _, hKA⟩ hJK ↦ hJ.eq_of_subset_indep hK hJK hKA + rintro A hAR I ⟨hI, _⟩ hIA + obtain ⟨J, hJ, hIJ⟩ := hI.subset_basis'_of_subset hIA + use J + simp only [hIJ, and_assoc, maximal_subset_iff, hJ.indep, hJ.subset, and_imp, true_and, + hJ.subset.trans hAR] + exact fun K hK _ hKA hJK ↦ hJ.eq_of_subset_indep hK hJK hKA subset_ground I := And.right /-- Change the ground set of a matroid to some `R : Set α`. The independent sets of the restriction @@ -150,17 +151,13 @@ theorem restrict_restrict_eq {R₁ R₂ : Set α} (M : Matroid α) (hR : R₂ @[simp] theorem base_restrict_iff (hX : X ⊆ M.E := by aesop_mat) : (M ↾ X).Base I ↔ M.Basis I X := by - simp_rw [base_iff_maximal_indep, basis_iff', restrict_indep_iff, and_iff_left hX, and_assoc] - aesop + simp_rw [base_iff_maximal_indep, Basis, and_iff_left hX, maximal_iff, restrict_indep_iff] theorem base_restrict_iff' : (M ↾ X).Base I ↔ M.Basis' I X := by - simp_rw [Basis', base_iff_maximal_indep, mem_maximals_setOf_iff, restrict_indep_iff] + simp_rw [base_iff_maximal_indep, Basis', maximal_iff, restrict_indep_iff] -theorem Basis.restrict_base (h : M.Basis I X) : (M ↾ X).Base I := by - rw [basis_iff'] at h - simp_rw [base_iff_maximal_indep, restrict_indep_iff, and_imp, and_assoc, and_iff_right h.1.1, - and_iff_right h.1.2.1] - exact fun J hJ hJX hIJ ↦ h.1.2.2 _ hJ hIJ hJX +theorem Basis.restrict_base (h : M.Basis I X) : (M ↾ X).Base I := + (base_restrict_iff h.subset_ground).2 h instance restrict_finiteRk [M.FiniteRk] (R : Set α) : (M ↾ R).FiniteRk := let ⟨_, hB⟩ := (M ↾ R).exists_base @@ -180,7 +177,8 @@ theorem Basis.basis_restrict_of_subset (hI : M.Basis I X) (hXY : X ⊆ Y) : (M rwa [← base_restrict_iff, M.restrict_restrict_eq hXY, base_restrict_iff] theorem basis'_restrict_iff : (M ↾ R).Basis' I X ↔ M.Basis' I (X ∩ R) ∧ I ⊆ R := by - simp_rw [Basis', mem_maximals_setOf_iff, restrict_indep_iff, subset_inter_iff, and_imp]; tauto + simp_rw [Basis', maximal_iff, restrict_indep_iff, subset_inter_iff, and_imp] + tauto theorem basis_restrict_iff' : (M ↾ R).Basis I X ↔ M.Basis I (X ∩ M.E) ∧ X ⊆ R := by rw [basis_iff_basis'_subset_ground, basis'_restrict_iff, restrict_ground_eq, and_congr_left_iff, diff --git a/Mathlib/Data/Matroid/Sum.lean b/Mathlib/Data/Matroid/Sum.lean index 3c335b4c6da39..77f18e0275913 100644 --- a/Mathlib/Data/Matroid/Sum.lean +++ b/Mathlib/Data/Matroid/Sum.lean @@ -85,16 +85,16 @@ protected def sigma (M : (i : ι) → Matroid (α i)) : Matroid ((i : ι) × α fun i ↦ (hI i).subset_basis'_of_subset (preimage_mono (f := Sigma.mk i) hIX) use univ.sigma Js - simp only [mem_maximals_setOf_iff, mem_univ, mk_preimage_sigma, and_imp] - refine ⟨⟨fun i ↦ (hJs i).1.indep, ⟨?_, ?_⟩⟩, fun S hS _ hSX h ↦ h.antisymm ?_⟩ + simp only [maximal_subset_iff', mem_univ, mk_preimage_sigma, le_eq_subset, and_imp] + refine ⟨?_, ⟨fun i ↦ (hJs i).1.indep, ?_⟩, fun S hS hSX hJS ↦ ?_⟩ · rw [← univ_sigma_preimage_mk I] exact sigma_mono rfl.subset fun i ↦ (hJs i).2 · rw [← univ_sigma_preimage_mk X] exact sigma_mono rfl.subset fun i ↦ (hJs i).1.subset rw [← univ_sigma_preimage_mk S] refine sigma_mono rfl.subset fun i ↦ ?_ - rw [sigma_subset_iff] at h - rw [(hJs i).1.eq_of_subset_indep (hS i) (h <| mem_univ i)] + rw [sigma_subset_iff] at hJS + rw [(hJs i).1.eq_of_subset_indep (hS i) (hJS <| mem_univ i)] exact preimage_mono hSX subset_ground B hB := by @@ -111,12 +111,12 @@ protected def sigma (M : (i : ι) → Matroid (α i)) : Matroid ((i : ι) × α @[simp] lemma sigma_basis_iff {I X} : (Matroid.sigma M).Basis I X ↔ ∀ i, (M i).Basis (Sigma.mk i ⁻¹' I) (Sigma.mk i ⁻¹' X) := by - simp only [Basis, sigma_indep_iff, mem_maximals_iff, mem_setOf_eq, and_imp, and_assoc, - sigma_ground_eq, forall_and, and_congr_right_iff] + simp only [Basis, sigma_indep_iff, maximal_subset_iff, and_imp, and_assoc, sigma_ground_eq, + forall_and, and_congr_right_iff] refine fun hI ↦ ⟨fun ⟨hIX, h, h'⟩ ↦ ⟨fun i ↦ preimage_mono hIX, fun i I₀ hI₀ hI₀X hII₀ ↦ ?_, ?_⟩, fun ⟨hIX, h', h''⟩ ↦ ⟨?_, ?_, ?_⟩⟩ · refine hII₀.antisymm ?_ - specialize h (y := I ∪ Sigma.mk i '' I₀) + specialize h (t := I ∪ Sigma.mk i '' I₀) simp only [preimage_union, union_subset_iff, hIX, image_subset_iff, hI₀X, and_self, subset_union_left, true_implies] at h rw [h, preimage_union, sigma_mk_preimage_image_eq_self] diff --git a/Mathlib/Order/Minimal.lean b/Mathlib/Order/Minimal.lean index 800a0acdc6f28..eec9b8aea31ce 100644 --- a/Mathlib/Order/Minimal.lean +++ b/Mathlib/Order/Minimal.lean @@ -1,472 +1,672 @@ /- Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies +Authors: Yaël Dillies, Peter Nelson -/ import Mathlib.Order.Antichain import Mathlib.Order.UpperLower.Basic import Mathlib.Order.Interval.Set.Basic -import Mathlib.Order.RelIso.Set /-! -# Minimal/maximal elements of a set +# Minimality and Maximality -This file defines minimal and maximal of a set with respect to an arbitrary relation. +This file defines minimality and maximality of an element with respect to a predicate `P` on +an ordered type `α`. ## Main declarations -* `maximals r s`: Maximal elements of `s` with respect to `r`. -* `minimals r s`: Minimal elements of `s` with respect to `r`. +* `Minimal P x`: `x` is minimal satisfying `P`. +* `Maximal P x`: `x` is maximal satisfying `P`. + +## Implementation Details + +This file underwent a refactor from a version where minimality and maximality were defined using +sets rather than predicates, and with an unbundled order relation rather than a `LE` instance. + +A side effect is that it has become less straightforward to state that something is minimal +with respect to a relation that is *not* defeq to the default `LE`. +One possible way would be with a type synonym, +and another would be with an ad hoc `LE` instance and `@` notation. +This was not an issue in practice anywhere in mathlib at the time of the refactor, +but it may be worth re-examining this to make it easier in the future; see the TODO below. ## TODO -Do we need a `Finset` version? +* In the linearly ordered case, versions of lemmas like `minimal_mem_image` will hold with + `MonotoneOn`/`AntitoneOn` assumptions rather than the stronger `x ≤ y ↔ f x ≤ f y` assumptions. + +* `Set.maximal_iff_forall_insert` and `Set.minimal_iff_forall_diff_singleton` will generalize to + lemmas about covering in the case of an `IsStronglyAtomic`/`IsStronglyCoatomic` order. + +* `Finset` versions of the lemmas about sets. + +* API to allow for easily expressing min/maximality with respect to an arbitrary non-`LE` relation. + -/ -open Function Set +open Set OrderDual -variable {α β : Type*} (r r₁ r₂ : α → α → Prop) (s t : Set α) (a b : α) +variable {α : Type*} {P Q : α → Prop} {a x y : α} -/-- Turns a set into an antichain by keeping only the "maximal" elements. -/ -def maximals : Set α := - { a ∈ s | ∀ ⦃b⦄, b ∈ s → r a b → r b a } +section LE -/-- Turns a set into an antichain by keeping only the "minimal" elements. -/ -def minimals : Set α := - { a ∈ s | ∀ ⦃b⦄, b ∈ s → r b a → r a b } +variable [LE α] -theorem maximals_subset : maximals r s ⊆ s := - sep_subset _ _ +/-- `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 -theorem minimals_subset : minimals r s ⊆ s := - sep_subset _ _ +/-- `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 -@[simp] -theorem maximals_empty : maximals r ∅ = ∅ := - sep_empty _ +lemma Minimal.prop (h : Minimal P x) : P x := + h.1 -@[simp] -theorem minimals_empty : minimals r ∅ = ∅ := - sep_empty _ +lemma Maximal.prop (h : Maximal P x) : P x := + h.1 -@[simp] -theorem maximals_singleton : maximals r {a} = {a} := - (maximals_subset _ _).antisymm <| - singleton_subset_iff.2 <| - ⟨rfl, by - rintro b (rfl : b = a) - exact id⟩ +lemma Minimal.le_of_le (h : Minimal P x) (hy : P y) (hle : y ≤ x) : x ≤ y := + h.2 hy hle -@[simp] -theorem minimals_singleton : minimals r {a} = {a} := - maximals_singleton _ _ +lemma Maximal.le_of_ge (h : Maximal P x) (hy : P y) (hge : x ≤ y) : y ≤ x := + h.2 hy hge -theorem maximals_swap : maximals (swap r) s = minimals r s := - rfl +@[simp] theorem minimal_toDual : Minimal (fun x ↦ P (ofDual x)) (toDual x) ↔ Maximal P x := + Iff.rfl -theorem minimals_swap : minimals (swap r) s = maximals r s := - rfl +alias ⟨Minimal.of_dual, Minimal.dual⟩ := minimal_toDual -variable {r s t a b} -variable {x : α} +@[simp] theorem maximal_toDual : Maximal (fun x ↦ P (ofDual x)) (toDual x) ↔ Minimal P x := + Iff.rfl -/-- This theorem can't be used to rewrite without specifying `rlt`, since `rlt` would have to be - guessed. See `mem_minimals_iff_forall_ssubset_not_mem` and `mem_minimals_iff_forall_lt_not_mem` - for `⊆` and `≤` versions. -/ -theorem mem_minimals_iff_forall_lt_not_mem' - (rlt : α → α → Prop) [IsNonstrictStrictOrder α r rlt] : - x ∈ minimals r s ↔ x ∈ s ∧ ∀ ⦃y⦄, rlt y x → y ∉ s := by - simp [minimals, right_iff_left_not_left_of r rlt, not_imp_not, imp.swap (a := _ ∈ _)] +alias ⟨Maximal.of_dual, Maximal.dual⟩ := maximal_toDual -theorem mem_maximals_iff_forall_lt_not_mem' - (rlt : α → α → Prop) [IsNonstrictStrictOrder α r rlt] : - x ∈ maximals r s ↔ x ∈ s ∧ ∀ ⦃y⦄, rlt x y → y ∉ s := by - simp [maximals, right_iff_left_not_left_of r rlt, not_imp_not, imp.swap (a := _ ∈ _)] +@[simp] theorem minimal_false : ¬ Minimal (fun _ ↦ False) x := by + simp [Minimal] -section IsAntisymm +@[simp] theorem maximal_false : ¬ Maximal (fun _ ↦ False) x := by + simp [Maximal] -variable [IsAntisymm α r] +@[simp] theorem minimal_true : Minimal (fun _ ↦ True) x ↔ IsMin x := by + simp [IsMin, Minimal] -theorem eq_of_mem_maximals (ha : a ∈ maximals r s) (hb : b ∈ s) (h : r a b) : a = b := - antisymm h <| ha.2 hb h +@[simp] theorem maximal_true : Maximal (fun _ ↦ True) x ↔ IsMax x := + minimal_true (α := αᵒᵈ) -theorem eq_of_mem_minimals (ha : a ∈ minimals r s) (hb : b ∈ s) (h : r b a) : a = b := - antisymm (ha.2 hb h) h +@[simp] theorem minimal_subtype {x : Subtype Q} : + Minimal (fun x ↦ P x.1) x ↔ Minimal (P ⊓ Q) x := by + obtain ⟨x, hx⟩ := x + simp only [Minimal, Subtype.forall, Subtype.mk_le_mk, Pi.inf_apply, inf_Prop_eq] + tauto -theorem mem_maximals_iff : x ∈ maximals r s ↔ x ∈ s ∧ ∀ ⦃y⦄, y ∈ s → r x y → x = y := by - simp only [maximals, Set.mem_sep_iff, and_congr_right_iff] - refine fun _ ↦ ⟨fun h y hys hxy ↦ antisymm hxy (h hys hxy), fun h y hys hxy ↦ ?_⟩ - convert hxy <;> rw [h hys hxy] - -theorem mem_maximals_setOf_iff {P : α → Prop} : - x ∈ maximals r (setOf P) ↔ P x ∧ ∀ ⦃y⦄, P y → r x y → x = y := - mem_maximals_iff - -theorem mem_minimals_iff : x ∈ minimals r s ↔ x ∈ s ∧ ∀ ⦃y⦄, y ∈ s → r y x → x = y := - have := IsAntisymm.swap r - mem_maximals_iff - -theorem mem_minimals_setOf_iff {P : α → Prop} : - x ∈ minimals r (setOf P) ↔ P x ∧ ∀ ⦃y⦄, P y → r y x → x = y := - mem_minimals_iff - -theorem minimals_eq_minimals_of_subset_of_forall [IsTrans α r] (hts : t ⊆ s) - (h : ∀ x ∈ s, ∃ y ∈ t, r y x) : minimals r s = minimals r t := by - refine Set.ext fun a ↦ ⟨fun ⟨has, hmin⟩ ↦ ⟨?_,fun b hbt ↦ hmin (hts hbt)⟩, - fun ⟨hat, hmin⟩ ↦ ⟨hts hat, fun b hbs hba ↦ ?_⟩⟩ - · obtain ⟨a', ha', haa'⟩ := h _ has - rwa [antisymm (hmin (hts ha') haa') haa'] - obtain ⟨b', hb't, hb'b⟩ := h b hbs - rwa [antisymm (hmin hb't (Trans.trans hb'b hba)) (Trans.trans hb'b hba)] +@[simp] theorem maximal_subtype {x : Subtype Q} : + Maximal (fun x ↦ P x.1) x ↔ Maximal (P ⊓ Q) x := + minimal_subtype (α := αᵒᵈ) -theorem maximals_eq_maximals_of_subset_of_forall [IsTrans α r] (hts : t ⊆ s) - (h : ∀ x ∈ s, ∃ y ∈ t, r x y) : maximals r s = maximals r t := - have := IsTrans.swap r - have := IsAntisymm.swap r - minimals_eq_minimals_of_subset_of_forall hts h - -variable (r s) - -theorem maximals_antichain : IsAntichain r (maximals r s) := fun _a ha _b hb hab h => - hab <| eq_of_mem_maximals ha hb.1 h - -theorem minimals_antichain : IsAntichain r (minimals r s) := - have := IsAntisymm.swap r - (maximals_antichain _ _).swap - -end IsAntisymm - -theorem mem_minimals_iff_forall_ssubset_not_mem {x : Set α} (s : Set (Set α)) : - x ∈ minimals (· ⊆ ·) s ↔ x ∈ s ∧ ∀ ⦃y⦄, y ⊂ x → y ∉ s := - mem_minimals_iff_forall_lt_not_mem' (· ⊂ ·) - -theorem mem_minimals_iff_forall_lt_not_mem [Preorder α] {x : α} {s : Set α} : - x ∈ minimals (· ≤ ·) s ↔ x ∈ s ∧ ∀ ⦃y⦄, y < x → y ∉ s := - mem_minimals_iff_forall_lt_not_mem' (· < ·) - -theorem mem_maximals_iff_forall_ssubset_not_mem {x : Set α} {s : Set (Set α)} : - x ∈ maximals (· ⊆ ·) s ↔ x ∈ s ∧ ∀ ⦃y⦄, x ⊂ y → y ∉ s := - mem_maximals_iff_forall_lt_not_mem' (· ⊂ ·) - -theorem mem_maximals_iff_forall_lt_not_mem [Preorder α] {x : α} {s : Set α} : - x ∈ maximals (· ≤ ·) s ↔ x ∈ s ∧ ∀ ⦃y⦄, x < y → y ∉ s := - mem_maximals_iff_forall_lt_not_mem' (· < ·) - -theorem Set.mem_maximals_iff_forall_insert {P : Set α → Prop} (hP : ∀ ⦃s t⦄, P t → s ⊆ t → P s) : - s ∈ maximals (· ⊆ ·) {t | P t} ↔ P s ∧ ∀ x ∉ s, ¬ P (insert x s) := by - simp only [mem_maximals_iff, mem_setOf_eq, and_congr_right_iff] - refine fun _ ↦ ⟨fun h x hx hxs ↦ hx ?_, fun h t ht hst ↦ hst.antisymm fun x hxt ↦ ?_⟩ - · rw [h hxs (subset_insert _ _)]; apply mem_insert - exact by_contra fun hxs ↦ h x hxs (hP ht (insert_subset hxt hst)) - -theorem Set.mem_minimals_iff_forall_diff_singleton {P : Set α → Prop} - (hP : ∀ ⦃s t⦄, P s → s ⊆ t → P t) : - s ∈ minimals (· ⊆ ·) {t | P t} ↔ P s ∧ ∀ x ∈ s, ¬ P (s \ {x}) := by - simp only [mem_minimals_iff, mem_setOf_eq, and_congr_right_iff] - refine fun _ ↦ ⟨fun h x hx hxs ↦ ?_, fun h t ht hst ↦ Eq.symm <| hst.antisymm (fun x hxs ↦ ?_)⟩ - · rw [(h hxs diff_subset)] at hx; simp at hx - exact by_contra fun hxt ↦ h x hxs (hP ht (subset_diff_singleton hst hxt)) - -theorem maximals_of_symm [IsSymm α r] : maximals r s = s := - sep_eq_self_iff_mem_true.2 fun _ _ _ _ => symm - -theorem minimals_of_symm [IsSymm α r] : minimals r s = s := - sep_eq_self_iff_mem_true.2 fun _ _ _ _ => symm - -theorem maximals_eq_minimals [IsSymm α r] : maximals r s = minimals r s := by - rw [minimals_of_symm, maximals_of_symm] - -variable {r₁ r₂} - --- Porting note (#11215): TODO: use `h.induction_on` -theorem Set.Subsingleton.maximals_eq (h : s.Subsingleton) : maximals r s = s := by - rcases h.eq_empty_or_singleton with (rfl | ⟨x, rfl⟩) - exacts [minimals_empty _, maximals_singleton _ _] - -theorem Set.Subsingleton.minimals_eq (h : s.Subsingleton) : minimals r s = s := - h.maximals_eq - -theorem maximals_mono [IsAntisymm α r₂] (h : ∀ a b, r₁ a b → r₂ a b) : - maximals r₂ s ⊆ maximals r₁ s := fun a ha => - ⟨ha.1, fun b hb hab => by - have := eq_of_mem_maximals ha hb (h _ _ hab) - subst this - exact hab⟩ - -theorem minimals_mono [IsAntisymm α r₂] (h : ∀ a b, r₁ a b → r₂ a b) : - minimals r₂ s ⊆ minimals r₁ s := fun a ha => - ⟨ha.1, fun b hb hab => by - have := eq_of_mem_minimals ha hb (h _ _ hab) - subst this - exact hab⟩ - -theorem maximals_union : maximals r (s ∪ t) ⊆ maximals r s ∪ maximals r t := by - intro a ha - obtain h | h := ha.1 - · exact Or.inl ⟨h, fun b hb => ha.2 <| Or.inl hb⟩ - · exact Or.inr ⟨h, fun b hb => ha.2 <| Or.inr hb⟩ - -theorem minimals_union : minimals r (s ∪ t) ⊆ minimals r s ∪ minimals r t := - maximals_union - -theorem maximals_inter_subset : maximals r s ∩ t ⊆ maximals r (s ∩ t) := fun _a ha => - ⟨⟨ha.1.1, ha.2⟩, fun _b hb => ha.1.2 hb.1⟩ - -theorem minimals_inter_subset : minimals r s ∩ t ⊆ minimals r (s ∩ t) := - maximals_inter_subset - -theorem inter_maximals_subset : s ∩ maximals r t ⊆ maximals r (s ∩ t) := fun _a ha => - ⟨⟨ha.1, ha.2.1⟩, fun _b hb => ha.2.2 hb.2⟩ - -theorem inter_minimals_subset : s ∩ minimals r t ⊆ minimals r (s ∩ t) := - inter_maximals_subset - -theorem IsAntichain.maximals_eq (h : IsAntichain r s) : maximals r s = s := - (maximals_subset _ _).antisymm fun a ha => - ⟨ha, fun b hb hab => by - obtain rfl := h.eq ha hb hab - exact hab⟩ - -theorem IsAntichain.minimals_eq (h : IsAntichain r s) : minimals r s = s := - h.flip.maximals_eq - -@[simp] -theorem maximals_idem : maximals r (maximals r s) = maximals r s := - (maximals_subset _ _).antisymm fun _a ha => ⟨ha, fun _b hb => ha.2 hb.1⟩ - -@[simp] -theorem minimals_idem : minimals r (minimals r s) = minimals r s := - maximals_idem - -/-- If `maximals r s` is included in but *shadows* the antichain `t`, then it is actually -equal to `t`. -/ -theorem IsAntichain.max_maximals (ht : IsAntichain r t) (h : maximals r s ⊆ t) - (hs : ∀ ⦃a⦄, a ∈ t → ∃ b ∈ maximals r s, r b a) : maximals r s = t := by - refine h.antisymm fun a ha => ?_ - obtain ⟨b, hb, hr⟩ := hs ha - rwa [of_not_not fun hab => ht (h hb) ha (Ne.symm hab) hr] - -/-- If `minimals r s` is included in but *shadows* the antichain `t`, then it is actually -equal to `t`. -/ -theorem IsAntichain.max_minimals (ht : IsAntichain r t) (h : minimals r s ⊆ t) - (hs : ∀ ⦃a⦄, a ∈ t → ∃ b ∈ minimals r s, r a b) : minimals r s = t := by - refine h.antisymm fun a ha => ?_ - obtain ⟨b, hb, hr⟩ := hs ha - rwa [of_not_not fun hab => ht ha (h hb) hab hr] - -theorem IsLeast.mem_minimals [Preorder α] (h : IsLeast s a) : a ∈ minimals (· ≤ ·) s := - ⟨h.1, fun _b hb _ => h.2 hb⟩ - -theorem IsGreatest.mem_maximals [Preorder α] (h : IsGreatest s a) : a ∈ maximals (· ≤ ·) s := - ⟨h.1, fun _b hb _ => h.2 hb⟩ - -theorem IsLeast.minimals_eq [PartialOrder α] (h : IsLeast s a) : minimals (· ≤ ·) s = {a} := - eq_singleton_iff_unique_mem.2 ⟨h.mem_minimals, fun _b hb => eq_of_mem_minimals hb h.1 <| h.2 hb.1⟩ - -theorem IsGreatest.maximals_eq [PartialOrder α] (h : IsGreatest s a) : maximals (· ≤ ·) s = {a} := - eq_singleton_iff_unique_mem.2 ⟨h.mem_maximals, fun _b hb => eq_of_mem_maximals hb h.1 <| h.2 hb.1⟩ - -theorem IsAntichain.minimals_upperClosure [PartialOrder α] (hs : IsAntichain (· ≤ ·) s) : - minimals (· ≤ ·) (upperClosure s : Set α) = s := - hs.max_minimals - (fun a ⟨⟨b, hb, hba⟩, _⟩ => by rwa [eq_of_mem_minimals ‹a ∈ _› (subset_upperClosure hb) hba]) - fun a ha => - ⟨a, ⟨subset_upperClosure ha, fun b ⟨c, hc, hcb⟩ hba => by rwa [hs.eq' ha hc (hcb.trans hba)]⟩, - le_rfl⟩ - -theorem IsAntichain.maximals_lowerClosure [PartialOrder α] (hs : IsAntichain (· ≤ ·) s) : - maximals (· ≤ ·) (lowerClosure s : Set α) = s := - hs.to_dual.minimals_upperClosure +theorem maximal_true_subtype {x : Subtype P} : Maximal (fun _ ↦ True) x ↔ Maximal P x := by + obtain ⟨x, hx⟩ := x + simp [Maximal, hx] -section Image +theorem minimal_true_subtype {x : Subtype P} : Minimal (fun _ ↦ True) x ↔ Minimal P x := by + obtain ⟨x, hx⟩ := x + simp [Minimal, hx] + +@[simp] theorem minimal_minimal : Minimal (Minimal P) x ↔ Minimal P x := + ⟨fun h ↦ h.prop, fun h ↦ ⟨h, fun _ hy hyx ↦ h.le_of_le hy.prop hyx⟩⟩ + +@[simp] theorem maximal_maximal : Maximal (Maximal P) x ↔ Maximal P x := + minimal_minimal (α := αᵒᵈ) + +/-- If `P` is down-closed, then minimal elements satisfying `P` are exactly the globally minimal +elements satisfying `P`. -/ +theorem minimal_iff_isMin (hP : ∀ ⦃x y⦄, P y → x ≤ y → P x) : Minimal P x ↔ P x ∧ IsMin x := + ⟨fun h ↦ ⟨h.prop, fun _ h' ↦ h.le_of_le (hP h.prop h') h'⟩, fun h ↦ ⟨h.1, fun _ _ h' ↦ h.2 h'⟩⟩ + +/-- If `P` is up-closed, then maximal elements satisfying `P` are exactly the globally maximal +elements satisfying `P`. -/ +theorem maximal_iff_isMax (hP : ∀ ⦃x y⦄, P y → y ≤ x → P x) : Maximal P x ↔ P x ∧ IsMax x := + ⟨fun h ↦ ⟨h.prop, fun _ h' ↦ h.le_of_ge (hP h.prop h') h'⟩, fun h ↦ ⟨h.1, fun _ _ h' ↦ h.2 h'⟩⟩ + +theorem Minimal.mono (h : Minimal P x) (hle : Q ≤ P) (hQ : Q x) : Minimal Q x := + ⟨hQ, fun y hQy ↦ h.le_of_le (hle y hQy)⟩ + +theorem Maximal.mono (h : Maximal P x) (hle : Q ≤ P) (hQ : Q x) : Maximal Q x := + ⟨hQ, fun y hQy ↦ h.le_of_ge (hle y hQy)⟩ + +theorem Minimal.and_right (h : Minimal P x) (hQ : Q x) : Minimal (fun x ↦ P x ∧ Q x) x := + h.mono (fun _ ↦ And.left) ⟨h.prop, hQ⟩ + +theorem Minimal.and_left (h : Minimal P x) (hQ : Q x) : Minimal (fun x ↦ (Q x ∧ P x)) x := + h.mono (fun _ ↦ And.right) ⟨hQ, h.prop⟩ + +theorem Maximal.and_right (h : Maximal P x) (hQ : Q x) : Maximal (fun x ↦ (P x ∧ Q x)) x := + h.mono (fun _ ↦ And.left) ⟨h.prop, hQ⟩ + +theorem Maximal.and_left (h : Maximal P x) (hQ : Q x) : Maximal (fun x ↦ (Q x ∧ P x)) x := + h.mono (fun _ ↦ And.right) ⟨hQ, h.prop⟩ + +@[simp] theorem minimal_eq_iff : Minimal (· = y) x ↔ x = y := by + simp (config := {contextual := true}) [Minimal] + +@[simp] theorem maximal_eq_iff : Maximal (· = y) x ↔ x = y := by + simp (config := {contextual := true}) [Maximal] + +theorem not_minimal_iff (hx : P x) : ¬ Minimal P x ↔ ∃ y, P y ∧ y ≤ x ∧ ¬ (x ≤ y) := by + simp [Minimal, hx] + +theorem not_maximal_iff (hx : P x) : ¬ Maximal P x ↔ ∃ y, P y ∧ x ≤ y ∧ ¬ (y ≤ x) := + not_minimal_iff (α := αᵒᵈ) hx + +theorem Minimal.or (h : Minimal (fun x ↦ P x ∨ Q x) x) : Minimal P x ∨ Minimal Q x := by + obtain ⟨h | h, hmin⟩ := h + · exact .inl ⟨h, fun y hy hyx ↦ hmin (Or.inl hy) hyx⟩ + exact .inr ⟨h, fun y hy hyx ↦ hmin (Or.inr hy) hyx⟩ + +theorem Maximal.or (h : Maximal (fun x ↦ P x ∨ Q x) x) : Maximal P x ∨ Maximal Q x := + Minimal.or (α := αᵒᵈ) h + +theorem minimal_and_iff_right_of_imp (hPQ : ∀ ⦃x⦄, P x → Q x) : + Minimal (fun x ↦ P x ∧ Q x) x ↔ (Minimal P x) ∧ Q x := by + simp_rw [and_iff_left_of_imp (fun x ↦ hPQ x), iff_self_and] + exact fun h ↦ hPQ h.prop + +theorem minimal_and_iff_left_of_imp (hPQ : ∀ ⦃x⦄, P x → Q x) : + Minimal (fun x ↦ Q x ∧ P x) x ↔ Q x ∧ (Minimal P x) := by + simp_rw [iff_comm, and_comm, minimal_and_iff_right_of_imp hPQ, and_comm] + +theorem maximal_and_iff_right_of_imp (hPQ : ∀ ⦃x⦄, P x → Q x) : + Maximal (fun x ↦ P x ∧ Q x) x ↔ (Maximal P x) ∧ Q x := + minimal_and_iff_right_of_imp (α := αᵒᵈ) hPQ + +theorem maximal_and_iff_left_of_imp (hPQ : ∀ ⦃x⦄, P x → Q x) : + Maximal (fun x ↦ Q x ∧ P x) x ↔ Q x ∧ (Maximal P x) := + minimal_and_iff_left_of_imp (α := αᵒᵈ) hPQ + +end LE + +section Preorder + +variable [Preorder α] + +theorem minimal_iff_forall_lt : Minimal P x ↔ P x ∧ ∀ ⦃y⦄, y < x → ¬ P y := by + simp [Minimal, lt_iff_le_not_le, not_imp_not, imp.swap] + +theorem maximal_iff_forall_gt : Maximal P x ↔ P x ∧ ∀ ⦃y⦄, x < y → ¬ P y := + minimal_iff_forall_lt (α := αᵒᵈ) + +theorem Minimal.not_prop_of_lt (h : Minimal P x) (hlt : y < x) : ¬ P y := + (minimal_iff_forall_lt.1 h).2 hlt + +theorem Maximal.not_prop_of_gt (h : Maximal P x) (hlt : x < y) : ¬ P y := + (maximal_iff_forall_gt.1 h).2 hlt + +theorem Minimal.not_lt (h : Minimal P x) (hy : P y) : ¬ (y < x) := + fun hlt ↦ h.not_prop_of_lt hlt hy + +theorem Maximal.not_gt (h : Maximal P x) (hy : P y) : ¬ (x < y) := + fun hlt ↦ h.not_prop_of_gt hlt hy + +@[simp] theorem minimal_le_iff : Minimal (· ≤ y) x ↔ x ≤ y ∧ IsMin x := + minimal_iff_isMin (fun _ _ h h' ↦ h'.trans h) + +@[simp] theorem maximal_ge_iff : Maximal (y ≤ ·) x ↔ y ≤ x ∧ IsMax x := + minimal_le_iff (α := αᵒᵈ) + +@[simp] theorem minimal_lt_iff : Minimal (· < y) x ↔ x < y ∧ IsMin x := + minimal_iff_isMin (fun _ _ h h' ↦ h'.trans_lt h) + +@[simp] theorem maximal_gt_iff : Maximal (y < ·) x ↔ y < x ∧ IsMax x := + minimal_lt_iff (α := αᵒᵈ) + +theorem not_minimal_iff_exists_lt (hx : P x) : ¬ Minimal P x ↔ ∃ y, y < x ∧ P y := by + simp_rw [not_minimal_iff hx, lt_iff_le_not_le, and_comm] + +alias ⟨exists_lt_of_not_minimal, _⟩ := not_minimal_iff_exists_lt + +theorem not_maximal_iff_exists_gt (hx : P x) : ¬ Maximal P x ↔ ∃ y, x < y ∧ P y := + not_minimal_iff_exists_lt (α := αᵒᵈ) hx + +alias ⟨exists_gt_of_not_maximal, _⟩ := not_maximal_iff_exists_gt + +end Preorder + +section PartialOrder + +variable [PartialOrder α] + +theorem Minimal.eq_of_ge (hx : Minimal P x) (hy : P y) (hge : y ≤ x) : x = y := + (hx.2 hy hge).antisymm hge + +theorem Minimal.eq_of_le (hx : Minimal P x) (hy : P y) (hle : y ≤ x) : y = x := + (hx.eq_of_ge hy hle).symm + +theorem Maximal.eq_of_le (hx : Maximal P x) (hy : P y) (hle : x ≤ y) : x = y := + hle.antisymm <| hx.2 hy hle + +theorem Maximal.eq_of_ge (hx : Maximal P x) (hy : P y) (hge : x ≤ y) : y = x := + (hx.eq_of_le hy hge).symm + +theorem minimal_iff : Minimal P x ↔ P x ∧ ∀ ⦃y⦄, P y → y ≤ x → x = y := + ⟨fun h ↦ ⟨h.1, fun _ ↦ h.eq_of_ge⟩, fun h ↦ ⟨h.1, fun _ hy hle ↦ (h.2 hy hle).le⟩⟩ + +theorem maximal_iff : Maximal P x ↔ P x ∧ ∀ ⦃y⦄, P y → x ≤ y → x = y := + minimal_iff (α := αᵒᵈ) + +theorem minimal_mem_iff {s : Set α} : Minimal (· ∈ s) x ↔ x ∈ s ∧ ∀ ⦃y⦄, y ∈ s → y ≤ x → x = y := + minimal_iff + +theorem maximal_mem_iff {s : Set α} : Maximal (· ∈ s) x ↔ x ∈ s ∧ ∀ ⦃y⦄, y ∈ s → x ≤ y → x = y := + maximal_iff + +/-- If `P y` holds, and everything satisfying `P` is above `y`, then `y` is the unique minimal +element satisfying `P`. -/ +theorem minimal_iff_eq (hy : P y) (hP : ∀ ⦃x⦄, P x → y ≤ x) : Minimal P x ↔ x = y := + ⟨fun h ↦ h.eq_of_ge hy (hP h.prop), by rintro rfl; exact ⟨hy, fun z hz _ ↦ hP hz⟩⟩ + +/-- If `P y` holds, and everything satisfying `P` is below `y`, then `y` is the unique maximal +element satisfying `P`. -/ +theorem maximal_iff_eq (hy : P y) (hP : ∀ ⦃x⦄, P x → x ≤ y) : Maximal P x ↔ x = y := + minimal_iff_eq (α := αᵒᵈ) hy hP -variable {f : α → β} {r : α → α → Prop} {s : β → β → Prop} +@[simp] theorem minimal_ge_iff : Minimal (y ≤ ·) x ↔ x = y := + minimal_iff_eq rfl.le fun _ ↦ id -section +@[simp] theorem maximal_le_iff : Maximal (· ≤ y) x ↔ x = y := + maximal_iff_eq rfl.le fun _ ↦ id -variable {x : Set α} (hf : ∀ ⦃a a'⦄, a ∈ x → a' ∈ x → (r a a' ↔ s (f a) (f a'))) {a : α} +theorem minimal_iff_minimal_of_imp_of_forall (hPQ : ∀ ⦃x⦄, Q x → P x) + (h : ∀ ⦃x⦄, P x → ∃ y, y ≤ x ∧ Q y) : Minimal P x ↔ Minimal Q x := by + refine ⟨fun h' ↦ ⟨?_, fun y hy hyx ↦ h'.le_of_le (hPQ hy) hyx⟩, + fun h' ↦ ⟨hPQ h'.prop, fun y hy hyx ↦ ?_⟩⟩ + · obtain ⟨y, hyx, hy⟩ := h h'.prop + rwa [((h'.le_of_le (hPQ hy)) hyx).antisymm hyx] + obtain ⟨z, hzy, hz⟩ := h hy + exact (h'.le_of_le hz (hzy.trans hyx)).trans hzy -theorem map_mem_minimals (ha : a ∈ minimals r x) : f a ∈ minimals s (f '' x) := - ⟨⟨a, ha.1, rfl⟩, by rintro _ ⟨a', h', rfl⟩; rw [← hf ha.1 h', ← hf h' ha.1]; exact ha.2 h'⟩ +theorem maximal_iff_maximal_of_imp_of_forall (hPQ : ∀ ⦃x⦄, Q x → P x) + (h : ∀ ⦃x⦄, P x → ∃ y, x ≤ y ∧ Q y) : Maximal P x ↔ Maximal Q x := + minimal_iff_minimal_of_imp_of_forall (α := αᵒᵈ) hPQ h -theorem map_mem_maximals (ha : a ∈ maximals r x) : f a ∈ maximals s (f '' x) := - map_mem_minimals (fun _ _ h₁ h₂ ↦ by exact hf h₂ h₁) ha +end PartialOrder -theorem map_mem_minimals_iff (ha : a ∈ x) : f a ∈ minimals s (f '' x) ↔ a ∈ minimals r x := - ⟨fun ⟨_, hmin⟩ ↦ ⟨ha, fun a' h' ↦ by - simpa only [hf h' ha, hf ha h'] using hmin ⟨a', h', rfl⟩⟩, map_mem_minimals hf⟩ +section Subset -theorem map_mem_maximals_iff (ha : a ∈ x) : f a ∈ maximals s (f '' x) ↔ a ∈ maximals r x := - map_mem_minimals_iff (fun _ _ h₁ h₂ ↦ by exact hf h₂ h₁) ha +variable {P : Set α → Prop} {s t : Set α} -theorem image_minimals_of_rel_iff_rel : f '' minimals r x = minimals s (f '' x) := by - ext b; refine ⟨?_, fun h ↦ ?_⟩ - · rintro ⟨a, ha, rfl⟩; exact map_mem_minimals hf ha - · obtain ⟨a, ha, rfl⟩ := h.1; exact ⟨a, (map_mem_minimals_iff hf ha).mp h, rfl⟩ +theorem Minimal.eq_of_superset (h : Minimal P s) (ht : P t) (hts : t ⊆ s) : s = t := + h.eq_of_ge ht hts -theorem image_maximals_of_rel_iff_rel : f '' maximals r x = maximals s (f '' x) := - image_minimals_of_rel_iff_rel fun _ _ h₁ h₂ ↦ hf h₂ h₁ +theorem Maximal.eq_of_subset (h : Maximal P s) (ht : P t) (hst : s ⊆ t) : s = t := + h.eq_of_le ht hst -end +theorem Minimal.eq_of_subset (h : Minimal P s) (ht : P t) (hts : t ⊆ s) : t = s := + h.eq_of_le ht hts -theorem RelEmbedding.image_minimals_eq (f : r ↪r s) (x : Set α) : - f '' minimals r x = minimals s (f '' x) := by - rw [image_minimals_of_rel_iff_rel]; simp [f.map_rel_iff] +theorem Maximal.eq_of_superset (h : Maximal P s) (ht : P t) (hst : s ⊆ t) : t = s := + h.eq_of_ge ht hst -theorem RelEmbedding.image_maximals_eq (f : r ↪r s) (x : Set α) : - f '' maximals r x = maximals s (f '' x) := - f.swap.image_minimals_eq x +theorem minimal_subset_iff : Minimal P s ↔ P s ∧ ∀ ⦃t⦄, P t → t ⊆ s → s = t := + _root_.minimal_iff -section +theorem maximal_subset_iff : Maximal P s ↔ P s ∧ ∀ ⦃t⦄, P t → s ⊆ t → s = t := + _root_.maximal_iff -variable [LE α] [LE β] {s : Set α} {t : Set β} +theorem minimal_subset_iff' : Minimal P s ↔ P s ∧ ∀ ⦃t⦄, P t → t ⊆ s → s ⊆ t := + Iff.rfl -theorem image_minimals_univ : - Subtype.val '' minimals (· ≤ ·) (univ : Set s) = minimals (· ≤ ·) s := by - rw [image_minimals_of_rel_iff_rel, image_univ, Subtype.range_val]; intros; rfl +theorem maximal_subset_iff' : Maximal P s ↔ P s ∧ ∀ ⦃t⦄, P t → s ⊆ t → t ⊆ s := + Iff.rfl -theorem image_maximals_univ : - Subtype.val '' maximals (· ≤ ·) (univ : Set s) = maximals (· ≤ ·) s := - image_minimals_univ (α := αᵒᵈ) +theorem not_minimal_subset_iff (hs : P s) : ¬ Minimal P s ↔ ∃ t, t ⊂ s ∧ P t := + not_minimal_iff_exists_lt hs -nonrec theorem OrderIso.map_mem_minimals (f : s ≃o t) {x : α} - (hx : x ∈ minimals (· ≤ ·) s) : (f ⟨x, hx.1⟩).val ∈ minimals (· ≤ ·) t := by - rw [← image_minimals_univ] at hx - obtain ⟨x, h, rfl⟩ := hx - convert map_mem_minimals (f := Subtype.val ∘ f) (fun _ _ _ _ ↦ f.map_rel_iff.symm) h - rw [image_comp, image_univ, f.range_eq, image_univ, Subtype.range_val] +theorem not_maximal_subset_iff (hs : P s) : ¬ Maximal P s ↔ ∃ t, s ⊂ t ∧ P t := + not_maximal_iff_exists_gt hs -theorem OrderIso.map_mem_maximals (f : s ≃o t) {x : α} - (hx : x ∈ maximals (· ≤ ·) s) : (f ⟨x, hx.1⟩).val ∈ maximals (· ≤ ·) t := - (show OrderDual.ofDual ⁻¹' s ≃o OrderDual.ofDual ⁻¹' t from f.dual).map_mem_minimals hx +theorem Set.minimal_iff_forall_ssubset : Minimal P s ↔ P s ∧ ∀ ⦃t⦄, t ⊂ s → ¬ P t := + minimal_iff_forall_lt + +theorem Minimal.not_prop_of_ssubset (h : Minimal P s) (ht : t ⊂ s) : ¬ P t := + (minimal_iff_forall_lt.1 h).2 ht + +theorem Minimal.not_ssubset (h : Minimal P s) (ht : P t) : ¬ t ⊂ s := + h.not_lt ht + +theorem Maximal.mem_of_prop_insert (h : Maximal P s) (hx : P (insert x s)) : x ∈ s := + h.eq_of_subset hx (subset_insert _ _) ▸ mem_insert .. + +theorem Minimal.not_mem_of_prop_diff_singleton (h : Minimal P s) (hx : P (s \ {x})) : x ∉ s := + fun hxs ↦ ((h.eq_of_superset hx diff_subset).subset hxs).2 rfl + +theorem Set.minimal_iff_forall_diff_singleton (hP : ∀ ⦃s t⦄, P t → t ⊆ s → P s) : + Minimal P s ↔ P s ∧ ∀ x ∈ s, ¬ P (s \ {x}) := + ⟨fun h ↦ ⟨h.1, fun _ hx hP ↦ h.not_mem_of_prop_diff_singleton hP hx⟩, + fun h ↦ ⟨h.1, fun _ ht hts x hxs ↦ by_contra fun hxt ↦ + h.2 x hxs (hP ht <| subset_diff_singleton hts hxt)⟩⟩ + +theorem Set.exists_diff_singleton_of_not_minimal (hP : ∀ ⦃s t⦄, P t → t ⊆ s → P s) (hs : P s) + (h : ¬ Minimal P s) : ∃ x ∈ s, P (s \ {x}) := by + simpa [Set.minimal_iff_forall_diff_singleton hP, hs] using h + +theorem Set.maximal_iff_forall_ssuperset : Maximal P s ↔ P s ∧ ∀ ⦃t⦄, s ⊂ t → ¬ P t := + maximal_iff_forall_gt + +theorem Maximal.not_prop_of_ssuperset (h : Maximal P s) (ht : s ⊂ t) : ¬ P t := + (maximal_iff_forall_gt.1 h).2 ht + +theorem Maximal.not_ssuperset (h : Maximal P s) (ht : P t) : ¬ s ⊂ t := + h.not_gt ht + +theorem Set.maximal_iff_forall_insert (hP : ∀ ⦃s t⦄, P t → s ⊆ t → P s) : + Maximal P s ↔ P s ∧ ∀ x ∉ s, ¬ P (insert x s) := by + simp only [not_imp_not] + exact ⟨fun h ↦ ⟨h.1, fun x ↦ h.mem_of_prop_insert⟩, + fun h ↦ ⟨h.1, fun t ht hst x hxt ↦ h.2 x (hP ht <| insert_subset hxt hst)⟩⟩ + +theorem Set.exists_insert_of_not_maximal (hP : ∀ ⦃s t⦄, P t → s ⊆ t → P s) (hs : P s) + (h : ¬ Maximal P s) : ∃ x ∉ s, P (insert x s) := by + simpa [Set.maximal_iff_forall_insert hP, hs] using h + +/- TODO : generalize `minimal_iff_forall_diff_singleton` and `maximal_iff_forall_insert` +to `IsStronglyCoatomic`/`IsStronglyAtomic` orders. -/ + +end Subset + +section Set + +variable {s t : Set α} +section Preorder + +variable [Preorder α] + +theorem setOf_minimal_subset (s : Set α) : {x | Minimal (· ∈ s) x} ⊆ s := + sep_subset .. + +theorem setOf_maximal_subset (s : Set α) : {x | Minimal (· ∈ s) x} ⊆ s := + sep_subset .. + +theorem Set.Subsingleton.maximal_mem_iff (h : s.Subsingleton) : Maximal (· ∈ s) x ↔ x ∈ s := by + obtain (rfl | ⟨x, rfl⟩) := h.eq_empty_or_singleton <;> simp + +theorem Set.Subsingleton.minimal_mem_iff (h : s.Subsingleton) : Minimal (· ∈ s) x ↔ x ∈ s := by + obtain (rfl | ⟨x, rfl⟩) := h.eq_empty_or_singleton <;> simp + +theorem IsLeast.minimal (h : IsLeast s x) : Minimal (· ∈ s) x := + ⟨h.1, fun _b hb _ ↦ h.2 hb⟩ + +theorem IsGreatest.maximal (h : IsGreatest s x) : Maximal (· ∈ s) x := + ⟨h.1, fun _b hb _ ↦ h.2 hb⟩ + +theorem IsAntichain.minimal_mem_iff (hs : IsAntichain (· ≤ ·) s) : Minimal (· ∈ s) x ↔ x ∈ s := + ⟨fun h ↦ h.prop, fun h ↦ ⟨h, fun _ hys hyx ↦ (hs.eq hys h hyx).symm.le⟩⟩ + +theorem IsAntichain.maximal_mem_iff (hs : IsAntichain (· ≤ ·) s) : Maximal (· ∈ s) x ↔ x ∈ s := + hs.to_dual.minimal_mem_iff + +/-- If `t` is an antichain shadowing and including the set of maximal elements of `s`, +then `t` *is* the set of maximal elements of `s`. -/ +theorem IsAntichain.eq_setOf_maximal (ht : IsAntichain (· ≤ ·) t) + (h : ∀ x, Maximal (· ∈ s) x → x ∈ t) (hs : ∀ a ∈ t, ∃ b, b ≤ a ∧ Maximal (· ∈ s) b) : + {x | Maximal (· ∈ s) x} = t := by + refine Set.ext fun x ↦ ⟨h _, fun hx ↦ ?_⟩ + obtain ⟨y, hyx, hy⟩ := hs x hx + rwa [← ht.eq (h y hy) hx hyx] + +/-- If `t` is an antichain shadowed by and including the set of minimal elements of `s`, +then `t` *is* the set of minimal elements of `s`. -/ +theorem IsAntichain.eq_setOf_minimal (ht : IsAntichain (· ≤ ·) t) + (h : ∀ x, Minimal (· ∈ s) x → x ∈ t) (hs : ∀ a ∈ t, ∃ b, a ≤ b ∧ Minimal (· ∈ s) b) : + {x | Minimal (· ∈ s) x} = t := + ht.to_dual.eq_setOf_maximal h hs + +end Preorder + +section PartialOrder + +variable [PartialOrder α] + +theorem setOf_maximal_antichain (P : α → Prop) : IsAntichain (· ≤ ·) {x | Maximal P x} := + fun _ hx _ ⟨hy, _⟩ hne hle ↦ hne (hle.antisymm <| hx.2 hy hle) + +theorem setOf_minimal_antichain (P : α → Prop) : IsAntichain (· ≤ ·) {x | Minimal P x} := + (setOf_maximal_antichain (α := αᵒᵈ) P).swap + +theorem IsAntichain.minimal_mem_upperClosure_iff_mem (hs : IsAntichain (· ≤ ·) s) : + Minimal (· ∈ upperClosure s) x ↔ x ∈ s := by + simp only [upperClosure, UpperSet.mem_mk, mem_setOf_eq] + refine ⟨fun h ↦ ?_, fun h ↦ ⟨⟨x, h, rfl.le⟩, fun b ⟨a, has, hab⟩ hbx ↦ ?_⟩⟩ + · obtain ⟨a, has, hax⟩ := h.prop + rwa [h.eq_of_ge ⟨a, has, rfl.le⟩ hax] + rwa [← hs.eq has h (hab.trans hbx)] + +theorem IsAntichain.maximal_mem_lowerClosure_iff_mem (hs : IsAntichain (· ≤ ·) s) : + Maximal (· ∈ lowerClosure s) x ↔ x ∈ s := + hs.to_dual.minimal_mem_upperClosure_iff_mem + +theorem IsLeast.minimal_iff (h : IsLeast s a) : Minimal (· ∈ s) x ↔ x = a := + ⟨fun h' ↦ h'.eq_of_ge h.1 (h.2 h'.prop), fun h' ↦ h' ▸ h.minimal⟩ + +theorem IsGreatest.maximal_iff (h : IsGreatest s a) : Maximal (· ∈ s) x ↔ x = a := + ⟨fun h' ↦ h'.eq_of_le h.1 (h.2 h'.prop), fun h' ↦ h' ▸ h.maximal⟩ + +end PartialOrder + +end Set + +section Image + +variable [Preorder α] {β : Type*} [Preorder β] {s : Set α} {t : Set β} +section Function + +variable {f : α → β} + +theorem minimal_mem_image_monotone (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ x ≤ y)) + (hx : Minimal (· ∈ s) x) : Minimal (· ∈ f '' s) (f x) := by + refine ⟨mem_image_of_mem f hx.prop, ?_⟩ + rintro _ ⟨y, hy, rfl⟩ + rw [hf hx.prop hy, hf hy hx.prop] + exact hx.le_of_le hy + +theorem maximal_mem_image_monotone (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ x ≤ y)) + (hx : Maximal (· ∈ s) x) : Maximal (· ∈ f '' s) (f x) := + minimal_mem_image_monotone (α := αᵒᵈ) (β := βᵒᵈ) (s := s) (fun _ _ hx hy ↦ hf hy hx) hx + +theorem minimal_mem_image_monotone_iff (ha : a ∈ s) + (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ x ≤ y)) : + Minimal (· ∈ f '' s) (f a) ↔ Minimal (· ∈ s) a := by + refine ⟨fun h ↦ ⟨ha, fun y hys ↦ ?_⟩, minimal_mem_image_monotone hf⟩ + rw [← hf ha hys, ← hf hys ha] + exact h.le_of_le (mem_image_of_mem f hys) + +theorem maximal_mem_image_monotone_iff (ha : a ∈ s) + (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ x ≤ y)) : + Maximal (· ∈ f '' s) (f a) ↔ Maximal (· ∈ s) a := + minimal_mem_image_monotone_iff (α := αᵒᵈ) (β := βᵒᵈ) (s := s) ha fun _ _ hx hy ↦ hf hy hx + +theorem minimal_mem_image_antitone (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ y ≤ x)) + (hx : Minimal (· ∈ s) x) : Maximal (· ∈ f '' s) (f x) := + minimal_mem_image_monotone (β := βᵒᵈ) (fun _ _ h h' ↦ hf h' h) hx + +theorem maximal_mem_image_antitone (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ y ≤ x)) + (hx : Maximal (· ∈ s) x) : Minimal (· ∈ f '' s) (f x) := + maximal_mem_image_monotone (β := βᵒᵈ) (fun _ _ h h' ↦ hf h' h) hx + +theorem minimal_mem_image_antitone_iff (ha : a ∈ s) + (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ y ≤ x)) : + Minimal (· ∈ f '' s) (f a) ↔ Maximal (· ∈ s) a := + maximal_mem_image_monotone_iff (β := βᵒᵈ) ha (fun _ _ h h' ↦ hf h' h) + +theorem maximal_mem_image_antitone_iff (ha : a ∈ s) + (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ y ≤ x)) : + Maximal (· ∈ f '' s) (f a) ↔ Minimal (· ∈ s) a := + minimal_mem_image_monotone_iff (β := βᵒᵈ) ha (fun _ _ h h' ↦ hf h' h) + +theorem image_monotone_setOf_minimal (hf : ∀ ⦃x y⦄, P x → P y → (f x ≤ f y ↔ x ≤ y)) : + f '' {x | Minimal P x} = {x | Minimal (∃ x₀, P x₀ ∧ f x₀ = ·) x} := by + refine Set.ext fun x ↦ ⟨?_, fun h ↦ ?_⟩ + · rintro ⟨x, (hx : Minimal _ x), rfl⟩ + exact (minimal_mem_image_monotone_iff hx.prop hf).2 hx + obtain ⟨y, hy, rfl⟩ := (mem_setOf_eq ▸ h).prop + exact mem_image_of_mem _ <| (minimal_mem_image_monotone_iff (s := setOf P) hy hf).1 h + +theorem image_monotone_setOf_maximal (hf : ∀ ⦃x y⦄, P x → P y → (f x ≤ f y ↔ x ≤ y)) : + f '' {x | Maximal P x} = {x | Maximal (∃ x₀, P x₀ ∧ f x₀ = ·) x} := + image_monotone_setOf_minimal (α := αᵒᵈ) (β := βᵒᵈ) (fun _ _ hx hy ↦ hf hy hx) + +theorem image_antitone_setOf_minimal (hf : ∀ ⦃x y⦄, P x → P y → (f x ≤ f y ↔ y ≤ x)) : + f '' {x | Minimal P x} = {x | Maximal (∃ x₀, P x₀ ∧ f x₀ = ·) x} := + image_monotone_setOf_minimal (β := βᵒᵈ) (fun _ _ hx hy ↦ hf hy hx) + +theorem image_antitone_setOf_maximal (hf : ∀ ⦃x y⦄, P x → P y → (f x ≤ f y ↔ y ≤ x)) : + f '' {x | Maximal P x} = {x | Minimal (∃ x₀, P x₀ ∧ f x₀ = ·) x} := + image_monotone_setOf_maximal (β := βᵒᵈ) (fun _ _ hx hy ↦ hf hy hx) + +theorem image_monotone_setOf_minimal_mem (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ x ≤ y)) : + f '' {x | Minimal (· ∈ s) x} = {x | Minimal (· ∈ f '' s) x} := + image_monotone_setOf_minimal hf + +theorem image_monotone_setOf_maximal_mem (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ x ≤ y)) : + f '' {x | Maximal (· ∈ s) x} = {x | Maximal (· ∈ f '' s) x} := + image_monotone_setOf_maximal hf + +theorem image_antitone_setOf_minimal_mem (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ y ≤ x)) : + f '' {x | Minimal (· ∈ s) x} = {x | Maximal (· ∈ f '' s) x} := + image_antitone_setOf_minimal hf + +theorem image_antitone_setOf_maximal_mem (hf : ∀ ⦃x y⦄, x ∈ s → y ∈ s → (f x ≤ f y ↔ y ≤ x)) : + f '' {x | Maximal (· ∈ s) x} = {x | Minimal (· ∈ f '' s) x} := + image_antitone_setOf_maximal hf + +end Function + +namespace OrderEmbedding + +variable {f : α ↪o β} {t : Set β} + +theorem minimal_mem_image (f : α ↪o β) (hx : Minimal (· ∈ s) x) : Minimal (· ∈ f '' s) (f x) := + _root_.minimal_mem_image_monotone (by simp [f.le_iff_le]) hx + +theorem maximal_mem_image (f : α ↪o β) (hx : Maximal (· ∈ s) x) : Maximal (· ∈ f '' s) (f x) := + _root_.maximal_mem_image_monotone (by simp [f.le_iff_le]) hx + +theorem minimal_mem_image_iff (ha : a ∈ s) : Minimal (· ∈ f '' s) (f a) ↔ Minimal (· ∈ s) a := + _root_.minimal_mem_image_monotone_iff ha (by simp [f.le_iff_le]) + +theorem maximal_mem_image_iff (ha : a ∈ s) : Maximal (· ∈ f '' s) (f a) ↔ Maximal (· ∈ s) a := + _root_.maximal_mem_image_monotone_iff ha (by simp [f.le_iff_le]) + +theorem minimal_apply_mem_inter_range_iff : + Minimal (· ∈ t ∩ range f) (f x) ↔ Minimal (fun x ↦ f x ∈ t) x := by + refine ⟨fun h ↦ ⟨h.prop.1, fun y hy ↦ ?_⟩, fun h ↦ ⟨⟨h.prop, by simp⟩, ?_⟩⟩ + · rw [← f.le_iff_le, ← f.le_iff_le] + exact h.le_of_le ⟨hy, by simp⟩ + rintro _ ⟨hyt, ⟨y, rfl⟩⟩ + simp_rw [f.le_iff_le] + exact h.le_of_le hyt + +theorem maximal_apply_mem_inter_range_iff : + Maximal (· ∈ t ∩ range f) (f x) ↔ Maximal (fun x ↦ f x ∈ t) x := + f.dual.minimal_apply_mem_inter_range_iff + +theorem minimal_apply_mem_iff (ht : t ⊆ Set.range f) : + Minimal (· ∈ t) (f x) ↔ Minimal (fun x ↦ f x ∈ t) x := by + rw [← f.minimal_apply_mem_inter_range_iff, inter_eq_self_of_subset_left ht] + +theorem maximal_apply_iff (ht : t ⊆ Set.range f) : + Maximal (· ∈ t) (f x) ↔ Maximal (fun x ↦ f x ∈ t) x := + f.dual.minimal_apply_mem_iff ht + +@[simp] theorem image_setOf_minimal : f '' {x | Minimal (· ∈ s) x} = {x | Minimal (· ∈ f '' s) x} := + _root_.image_monotone_setOf_minimal (by simp [f.le_iff_le]) + +@[simp] theorem image_setOf_maximal : f '' {x | Maximal (· ∈ s) x} = {x | Maximal (· ∈ f '' s) x} := + _root_.image_monotone_setOf_maximal (by simp [f.le_iff_le]) + +theorem inter_preimage_setOf_minimal_eq_of_subset (hts : t ⊆ f '' s) : + x ∈ s ∩ f ⁻¹' {y | Minimal (· ∈ t) y} ↔ Minimal (· ∈ s ∩ f ⁻¹' t) x := by + simp_rw [mem_inter_iff, preimage_setOf_eq, mem_setOf_eq, mem_preimage, + f.minimal_apply_mem_iff (hts.trans (image_subset_range _ _)), + minimal_and_iff_left_of_imp (fun _ hx ↦ f.injective.mem_set_image.1 <| hts hx)] + +theorem inter_preimage_setOf_maximal_eq_of_subset (hts : t ⊆ f '' s) : + x ∈ s ∩ f ⁻¹' {y | Maximal (· ∈ t) y} ↔ Maximal (· ∈ s ∩ f ⁻¹' t) x := + f.dual.inter_preimage_setOf_minimal_eq_of_subset hts + +end OrderEmbedding + +namespace OrderIso + +theorem image_setOf_minimal (f : α ≃o β) (P : α → Prop) : + f '' {x | Minimal P x} = {x | Minimal (fun x ↦ P (f.symm x)) x} := by + convert _root_.image_monotone_setOf_minimal (f := f) (by simp [f.le_iff_le]) + aesop + +theorem image_setOf_maximal (f : α ≃o β) (P : α → Prop) : + f '' {x | Maximal P x} = {x | Maximal (fun x ↦ P (f.symm x)) x} := by + convert _root_.image_monotone_setOf_maximal (f := f) (by simp [f.le_iff_le]) + aesop + +theorem map_minimal_mem (f : s ≃o t) (hx : Minimal (· ∈ s) x) : + Minimal (· ∈ t) (f ⟨x, hx.prop⟩) := by + simpa only [show t = range (Subtype.val ∘ f) by simp, mem_univ, minimal_true_subtype, hx, + true_imp_iff, image_univ] using OrderEmbedding.minimal_mem_image + (f.toOrderEmbedding.trans (OrderEmbedding.subtype t)) (s := univ) (x := ⟨x, hx.prop⟩) + +theorem map_maximal_mem (f : s ≃o t) (hx : Maximal (· ∈ s) x) : + Maximal (· ∈ t) (f ⟨x, hx.prop⟩) := by + simpa only [show t = range (Subtype.val ∘ f) by simp, mem_univ, maximal_true_subtype, hx, + true_imp_iff, image_univ] using OrderEmbedding.maximal_mem_image + (f.toOrderEmbedding.trans (OrderEmbedding.subtype t)) (s := univ) (x := ⟨x, hx.prop⟩) /-- If two sets are order isomorphic, their minimals are also order isomorphic. -/ -def OrderIso.mapMinimals (f : s ≃o t) : minimals (· ≤ ·) s ≃o minimals (· ≤ ·) t where - toFun x := ⟨f ⟨x, x.2.1⟩, f.map_mem_minimals x.2⟩ - invFun x := ⟨f.symm ⟨x, x.2.1⟩, f.symm.map_mem_minimals x.2⟩ +def mapSetOfMinimal (f : s ≃o t) : {x | Minimal (· ∈ s) x} ≃o {x | Minimal (· ∈ t) x} where + toFun x := ⟨f ⟨x, x.2.1⟩, f.map_minimal_mem x.2⟩ + invFun x := ⟨f.symm ⟨x, x.2.1⟩, f.symm.map_minimal_mem x.2⟩ left_inv x := Subtype.ext (by apply congr_arg Subtype.val <| f.left_inv ⟨x, x.2.1⟩) right_inv x := Subtype.ext (by apply congr_arg Subtype.val <| f.right_inv ⟨x, x.2.1⟩) map_rel_iff' {_ _} := f.map_rel_iff /-- If two sets are order isomorphic, their maximals are also order isomorphic. -/ -def OrderIso.mapMaximals (f : s ≃o t) : maximals (· ≤ ·) s ≃o maximals (· ≤ ·) t where - toFun x := ⟨f ⟨x, x.2.1⟩, f.map_mem_maximals x.2⟩ - invFun x := ⟨f.symm ⟨x, x.2.1⟩, f.symm.map_mem_maximals x.2⟩ - __ := (show OrderDual.ofDual ⁻¹' s ≃o OrderDual.ofDual ⁻¹' t from f.dual).mapMinimals - -- defeq abuse to fill in the proof fields. - -- If OrderDual ever becomes a structure, just copy the last three lines from OrderIso.mapMinimals - -open OrderDual in -/-- If two sets are antitonically order isomorphic, their minimals are too. -/ -def OrderIso.minimalsIsoMaximals (f : s ≃o tᵒᵈ) : - minimals (· ≤ ·) s ≃o (maximals (· ≤ ·) t)ᵒᵈ where - toFun x := toDual ⟨↑(ofDual (f ⟨x, x.2.1⟩)), (show s ≃o ofDual ⁻¹' t from f).map_mem_minimals x.2⟩ - invFun x := ⟨f.symm (toDual ⟨_, (ofDual x).2.1⟩), - (show ofDual ⁻¹' t ≃o s from f.symm).map_mem_minimals x.2⟩ - __ := (show s ≃o ofDual⁻¹' t from f).mapMinimals - -open OrderDual in -/-- If two sets are antitonically order isomorphic, their minimals are too. -/ -def OrderIso.maximalsIsoMinimals (f : s ≃o tᵒᵈ) : - maximals (· ≤ ·) s ≃o (minimals (· ≤ ·) t)ᵒᵈ where - toFun x := toDual ⟨↑(ofDual (f ⟨x, x.2.1⟩)), (show s ≃o ofDual ⁻¹' t from f).map_mem_maximals x.2⟩ - invFun x := ⟨f.symm (toDual ⟨_, (ofDual x).2.1⟩), - (show ofDual ⁻¹' t ≃o s from f.symm).map_mem_maximals x.2⟩ - __ := (show s ≃o ofDual⁻¹' t from f).mapMaximals - -end - -theorem inter_minimals_preimage_inter_eq_of_rel_iff_rel_on - {x : Set α} (hf : ∀ ⦃a a'⦄, a ∈ x → a' ∈ x → (r a a' ↔ s (f a) (f a'))) (y : Set β) : - x ∩ f ⁻¹' (minimals s ((f '' x) ∩ y)) = minimals r (x ∩ f ⁻¹' y) := by - ext a - simp only [minimals, mem_inter_iff, mem_image, and_imp, forall_exists_index, - forall_apply_eq_imp_iff₂, preimage_setOf_eq, mem_setOf_eq, mem_preimage] - exact ⟨fun ⟨hax,⟨_,hay⟩,h2⟩ ↦ ⟨⟨hax, hay⟩, fun a₁ ha₁ ha₁y ha₁a ↦ - (hf hax ha₁).mpr (h2 _ ha₁ ha₁y ((hf ha₁ hax).mp ha₁a))⟩ , - fun ⟨⟨hax,hay⟩,h⟩ ↦ ⟨hax, ⟨⟨_, hax, rfl⟩, hay⟩, fun a' ha' ha'y hsa' ↦ - (hf hax ha').mp (h ha' ha'y ((hf ha' hax).mpr hsa'))⟩⟩ - -theorem inter_preimage_minimals_eq_of_rel_iff_rel_on_of_subset {x : Set α} {y : Set β} - (hf : ∀ ⦃a a'⦄, a ∈ x → a' ∈ x → (r a a' ↔ s (f a) (f a'))) (hy : y ⊆ f '' x) : - x ∩ f ⁻¹' (minimals s y) = minimals r (x ∩ f ⁻¹' y) := by - rw [← inter_eq_self_of_subset_right hy, inter_minimals_preimage_inter_eq_of_rel_iff_rel_on hf, - preimage_inter, ← inter_assoc, inter_eq_self_of_subset_left (subset_preimage_image f x)] - -theorem RelEmbedding.inter_preimage_minimals_eq (f : r ↪r s) (x : Set α) (y : Set β) : - x ∩ f⁻¹' (minimals s ((f '' x) ∩ y)) = minimals r (x ∩ f ⁻¹' y) := - inter_minimals_preimage_inter_eq_of_rel_iff_rel_on (by simp [f.map_rel_iff]) y - -theorem RelEmbedding.inter_preimage_minimals_eq_of_subset {y : Set β} {x : Set α} - (f : r ↪r s) (h : y ⊆ f '' x) : - x ∩ f ⁻¹' (minimals s y) = minimals r (x ∩ f ⁻¹' y) := by - rw [inter_preimage_minimals_eq_of_rel_iff_rel_on_of_subset _ h]; simp [f.map_rel_iff] - -theorem RelEmbedding.minimals_preimage_eq (f : r ↪r s) (y : Set β) : - minimals r (f ⁻¹' y) = f ⁻¹' minimals s (y ∩ range f) := by - convert (f.inter_preimage_minimals_eq univ y).symm - · simp [univ_inter] - · simp [inter_comm] - -theorem RelIso.minimals_preimage_eq (f : r ≃r s) (y : Set β) : - minimals r (f ⁻¹' y) = f ⁻¹' (minimals s y) := by - convert f.toRelEmbedding.minimals_preimage_eq y; simp - -theorem RelIso.maximals_preimage_eq (f : r ≃r s) (y : Set β) : - maximals r (f ⁻¹' y) = f ⁻¹' (maximals s y) := - f.swap.minimals_preimage_eq y - -theorem inter_maximals_preimage_inter_eq_of_rel_iff_rel_on {x : Set α} - (hf : ∀ ⦃a a'⦄, a ∈ x → a' ∈ x → (r a a' ↔ s (f a) (f a'))) (y : Set β) : - x ∩ f ⁻¹' (maximals s ((f '' x) ∩ y)) = maximals r (x ∩ f ⁻¹' y) := by - apply inter_minimals_preimage_inter_eq_of_rel_iff_rel_on - exact fun _ _ a b ↦ hf b a - -theorem inter_preimage_maximals_eq_of_rel_iff_rel_on_of_subset {y : Set β} {x : Set α} - (hf : ∀ ⦃a a'⦄, a ∈ x → a' ∈ x → (r a a' ↔ s (f a) (f a'))) (hy : y ⊆ f '' x) : - x ∩ f ⁻¹' (maximals s y) = maximals r (x ∩ f ⁻¹' y) := by - apply inter_preimage_minimals_eq_of_rel_iff_rel_on_of_subset _ hy - exact fun _ _ a b ↦ hf b a - -theorem RelEmbedding.inter_preimage_maximals_eq (f : r ↪r s) (x : Set α) (y : Set β) : - x ∩ f⁻¹' (maximals s ((f '' x) ∩ y)) = maximals r (x ∩ f ⁻¹' y) := - inter_minimals_preimage_inter_eq_of_rel_iff_rel_on (by simp [f.map_rel_iff]) y - -theorem RelEmbedding.inter_preimage_maximals_eq_of_subset {y : Set β} {x : Set α} - (f : r ↪r s) (h : y ⊆ f '' x) : x ∩ f ⁻¹' (maximals s y) = maximals r (x ∩ f ⁻¹' y) := by - rw [inter_preimage_maximals_eq_of_rel_iff_rel_on_of_subset _ h]; simp [f.map_rel_iff] - -theorem RelEmbedding.maximals_preimage_eq (f : r ↪r s) (y : Set β) : - maximals r (f ⁻¹' y) = f ⁻¹' maximals s (y ∩ range f) := by - convert (f.inter_preimage_maximals_eq univ y).symm - · simp [univ_inter] - · simp [inter_comm] +def mapSetOfMaximal (f : s ≃o t) : {x | Maximal (· ∈ s) x} ≃o {x | Maximal (· ∈ t) x} where + toFun x := ⟨f ⟨x, x.2.1⟩, f.map_maximal_mem x.2⟩ + invFun x := ⟨f.symm ⟨x, x.2.1⟩, f.symm.map_maximal_mem x.2⟩ + __ := (show OrderDual.ofDual ⁻¹' s ≃o OrderDual.ofDual ⁻¹' t from f.dual).mapSetOfMinimal + +/-- If two sets are antitonically order isomorphic, their minimals/maximals are too. -/ +def setOfMinimalIsoSetOfMaximal (f : s ≃o tᵒᵈ) : + {x | Minimal (· ∈ s) x} ≃o {x | Maximal (· ∈ t) (ofDual x)} where + toFun x := ⟨(f ⟨x.1, x.2.1⟩).1, ((show s ≃o ofDual ⁻¹' t from f).mapSetOfMinimal x).2⟩ + invFun x := ⟨(f.symm ⟨x.1, x.2.1⟩).1, + ((show ofDual ⁻¹' t ≃o s from f.symm).mapSetOfMinimal x).2⟩ + __ := (show s ≃o ofDual⁻¹' t from f).mapSetOfMinimal + +/-- If two sets are antitonically order isomorphic, their maximals/minimals are too. -/ +def setOfMaximalIsoSetOfMinimal (f : s ≃o tᵒᵈ) : + {x | Maximal (· ∈ s) x} ≃o {x | Minimal (· ∈ t) (ofDual x)} where + toFun x := ⟨(f ⟨x.1, x.2.1⟩).1, ((show s ≃o ofDual ⁻¹' t from f).mapSetOfMaximal x).2⟩ + invFun x := ⟨(f.symm ⟨x.1, x.2.1⟩).1, + ((show ofDual ⁻¹' t ≃o s from f.symm).mapSetOfMaximal x).2⟩ + __ := (show s ≃o ofDual⁻¹' t from f).mapSetOfMaximal + +end OrderIso end Image - section Interval variable [PartialOrder α] {a b : α} -@[simp] theorem maximals_Iic (a : α) : maximals (· ≤ ·) (Iic a) = {a} := - Set.ext fun _ ↦ ⟨fun h ↦ h.1.antisymm (h.2 rfl.le h.1), - fun h ↦ ⟨h.trans_le rfl.le, fun _ hba _ ↦ le_trans hba h.symm.le⟩⟩ - -@[simp] theorem minimals_Ici (a : α) : minimals (· ≤ ·) (Ici a) = {a} := - maximals_Iic (α := αᵒᵈ) a +theorem minimal_mem_Icc (hab : a ≤ b) : Minimal (· ∈ Icc a b) x ↔ x = a := + minimal_iff_eq ⟨rfl.le, hab⟩ (fun _ ↦ And.left) -theorem maximals_Icc (hab : a ≤ b) : maximals (· ≤ ·) (Icc a b) = {b} := - Set.ext fun x ↦ ⟨fun h ↦ h.1.2.antisymm (h.2 ⟨hab, rfl.le⟩ h.1.2), - fun (h : x = b) ↦ ⟨⟨hab.trans_eq h.symm, h.le⟩, fun _ hy _ ↦ hy.2.trans_eq h.symm⟩⟩ +theorem maximal_mem_Icc (hab : a ≤ b) : Maximal (· ∈ Icc a b) x ↔ x = b := + maximal_iff_eq ⟨hab, rfl.le⟩ (fun _ ↦ And.right) -theorem minimals_Icc (hab : a ≤ b) : minimals (· ≤ ·) (Icc a b) = {a} := by - simp_rw [Icc, and_comm (a := (a ≤ _))]; exact maximals_Icc (α := αᵒᵈ) hab +theorem minimal_mem_Ico (hab : a < b) : Minimal (· ∈ Ico a b) x ↔ x = a := + minimal_iff_eq ⟨rfl.le, hab⟩ (fun _ ↦ And.left) -theorem maximals_Ioc (hab : a < b) : maximals (· ≤ ·) (Ioc a b) = {b} := - Set.ext fun x ↦ ⟨fun h ↦ h.1.2.antisymm (h.2 ⟨hab, rfl.le⟩ h.1.2), - fun (h : x = b) ↦ ⟨⟨hab.trans_eq h.symm, h.le⟩, fun _ hy _ ↦ hy.2.trans_eq h.symm⟩⟩ +theorem maximal_mem_Ioc (hab : a < b) : Maximal (· ∈ Ioc a b) x ↔ x = b := + maximal_iff_eq ⟨hab, rfl.le⟩ (fun _ ↦ And.right) -theorem minimals_Ico (hab : a < b) : minimals (· ≤ ·) (Ico a b) = {a} := by - simp_rw [Ico, and_comm (a := _ ≤ _)]; exact maximals_Ioc (α := αᵒᵈ) hab +/- Note : The one-sided interval versions of these lemmas are unnecessary, +since `simp` handles them with `maximal_le_iff` and `minimal_ge_iff`. -/ end Interval diff --git a/Mathlib/RingTheory/Ideal/MinimalPrime.lean b/Mathlib/RingTheory/Ideal/MinimalPrime.lean index 671885ea79228..c4c6347f666fb 100644 --- a/Mathlib/RingTheory/Ideal/MinimalPrime.lean +++ b/Mathlib/RingTheory/Ideal/MinimalPrime.lean @@ -36,7 +36,7 @@ variable {R S : Type*} [CommSemiring R] [CommSemiring S] (I J : Ideal R) /-- `I.minimalPrimes` is the set of ideals that are minimal primes over `I`. -/ protected def Ideal.minimalPrimes : Set (Ideal R) := - minimals (· ≤ ·) { p | p.IsPrime ∧ I ≤ p } + {p | Minimal (fun q ↦ q.IsPrime ∧ I ≤ q) p} variable (R) in /-- `minimalPrimes R` is the set of minimal primes of `R`. @@ -44,8 +44,8 @@ This is defined as `Ideal.minimalPrimes ⊥`. -/ def minimalPrimes : Set (Ideal R) := Ideal.minimalPrimes ⊥ -lemma minimalPrimes_eq_minimals : minimalPrimes R = minimals (· ≤ ·) (setOf Ideal.IsPrime) := - congr_arg (minimals (· ≤ ·)) (by simp) +lemma minimalPrimes_eq_minimals : minimalPrimes R = {x | Minimal Ideal.IsPrime x} := + congr_arg Minimal (by simp) variable {I J} @@ -216,8 +216,8 @@ variable {R : Type*} [CommSemiring R] {I : Ideal R} [hI : I.IsPrime] (hMin : I theorem _root_.IsLocalization.AtPrime.prime_unique_of_minimal {S} [CommSemiring S] [Algebra R S] [IsLocalization.AtPrime S I] {J K : Ideal S} [J.IsPrime] [K.IsPrime] : J = K := haveI : Subsingleton {i : Ideal R // i.IsPrime ∧ i ≤ I} := ⟨fun i₁ i₂ ↦ Subtype.ext <| by - rw [minimalPrimes_eq_minimals] at hMin - rw [← eq_of_mem_minimals hMin i₁.2.1 i₁.2.2, ← eq_of_mem_minimals hMin i₂.2.1 i₂.2.2]⟩ + rw [minimalPrimes_eq_minimals, Set.mem_setOf] at hMin + rw [hMin.eq_of_le i₁.2.1 i₁.2.2, hMin.eq_of_le i₂.2.1 i₂.2.2]⟩ Subtype.ext_iff.mp <| (IsLocalization.AtPrime.orderIsoOfPrime S I).injective (a₁ := ⟨J, ‹_›⟩) (a₂ := ⟨K, ‹_›⟩) (Subsingleton.elim _ _) diff --git a/Mathlib/Topology/Irreducible.lean b/Mathlib/Topology/Irreducible.lean index fe622ea65817c..cf88aaae25d57 100644 --- a/Mathlib/Topology/Irreducible.lean +++ b/Mathlib/Topology/Irreducible.lean @@ -93,7 +93,7 @@ theorem exists_preirreducible (s : Set X) (H : IsPreirreducible s) : /-- The set of irreducible components of a topological space. -/ def irreducibleComponents (X : Type*) [TopologicalSpace X] : Set (Set X) := - maximals (· ≤ ·) { s : Set X | IsIrreducible s } + {s | Maximal IsIrreducible s} theorem isClosed_of_mem_irreducibleComponents (s) (H : s ∈ irreducibleComponents X) : IsClosed s := by @@ -101,7 +101,7 @@ theorem isClosed_of_mem_irreducibleComponents (s) (H : s ∈ irreducibleComponen exact subset_closure.antisymm (H.2 H.1.closure subset_closure) theorem irreducibleComponents_eq_maximals_closed (X : Type*) [TopologicalSpace X] : - irreducibleComponents X = maximals (· ≤ ·) { s : Set X | IsClosed s ∧ IsIrreducible s } := by + irreducibleComponents X = { s | Maximal (fun x ↦ IsClosed x ∧ IsIrreducible x) s} := by ext s constructor · intro H From 2d01c49ac4c2d9765e1b4ed0de7e4cb8789f8158 Mon Sep 17 00:00:00 2001 From: Rida Hamadani Date: Mon, 5 Aug 2024 20:15:14 +0000 Subject: [PATCH 040/359] feat(Combinatorics/SimpleGraph): distance of vertices in a subgraph (#15464) --- Mathlib/Combinatorics/SimpleGraph/Metric.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/Mathlib/Combinatorics/SimpleGraph/Metric.lean b/Mathlib/Combinatorics/SimpleGraph/Metric.lean index 1967ec31109a8..5fac07b81db53 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Metric.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Metric.lean @@ -128,6 +128,15 @@ theorem edist_eq_one_iff_adj : G.edist u v = 1 ↔ G.Adj u v := by exact w.adj_of_length_eq_one <| Nat.cast_eq_one.mp <| h ▸ hw · exact le_antisymm (edist_le h.toWalk) (ENat.one_le_iff_pos.mpr <| edist_pos_of_ne h.ne) +/-- Supergraphs have smaller or equal extended distances to their subgraphs. -/ +theorem edist_le_subgraph_edist {G' : SimpleGraph V} (h : G ≤ G') : + G'.edist u v ≤ G.edist u v := by + by_cases hr : G.Reachable u v + · obtain ⟨_, hw⟩ := hr.exists_walk_length_eq_edist + rw [← hw, ← Walk.length_map (Hom.mapSpanningSubgraphs h)] + apply edist_le + · exact edist_eq_top_of_not_reachable hr ▸ le_top + end edist section dist @@ -234,6 +243,13 @@ lemma Connected.exists_path_of_dist (hconn : G.Connected) (u v : V) : obtain ⟨p, h⟩ := hconn.exists_walk_length_eq_dist u v exact ⟨p, p.isPath_of_length_eq_dist h, h⟩ +/-- Supergraphs have smaller or equal distances to their subgraphs. -/ +theorem dist_le_subgraph_dist {G' : SimpleGraph V} (h : G ≤ G') (hr : G.Reachable u v) : + G'.dist u v ≤ G.dist u v := by + obtain ⟨_, hw⟩ := hr.exists_walk_length_eq_dist + rw [← hw, ← Walk.length_map (Hom.mapSpanningSubgraphs h)] + apply dist_le + end dist end SimpleGraph From 95197cb84e872a92da03d95c232e86b238eccdbb Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 5 Aug 2024 21:07:09 +0000 Subject: [PATCH 041/359] chore: generalize tower instances for Module.End (#14991) This now matches the generality of `Module.End.instAlgebra`, meaning that the generality is available to `M/R/S` not just `A/R/S`. [Context](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Quadratic.20forms.20with.20values.20in.20a.20larger.20ring/near/452999484) on Zulip. --- Mathlib/Algebra/Algebra/Equiv.lean | 10 ++++++---- Mathlib/Algebra/Module/Equiv/Basic.lean | 10 ++++++---- Mathlib/Algebra/Module/LinearMap/End.lean | 14 ++++++++------ 3 files changed, 20 insertions(+), 14 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index 244dfed080662..6b75f47d33700 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -665,11 +665,13 @@ protected theorem smul_def (f : A₁ ≃ₐ[R] A₁) (a : A₁) : f • a = f a instance apply_faithfulSMul : FaithfulSMul (A₁ ≃ₐ[R] A₁) A₁ := ⟨AlgEquiv.ext⟩ -instance apply_smulCommClass : SMulCommClass R (A₁ ≃ₐ[R] A₁) A₁ where - smul_comm r e a := (map_smul e r a).symm +instance apply_smulCommClass {S} [SMul S R] [SMul S A₁] [IsScalarTower S R A₁] : + SMulCommClass S (A₁ ≃ₐ[R] A₁) A₁ where + smul_comm r e a := (e.toLinearEquiv.map_smul_of_tower r a).symm -instance apply_smulCommClass' : SMulCommClass (A₁ ≃ₐ[R] A₁) R A₁ where - smul_comm e r a := map_smul e r a +instance apply_smulCommClass' {S} [SMul S R] [SMul S A₁] [IsScalarTower S R A₁] : + SMulCommClass (A₁ ≃ₐ[R] A₁) S A₁ := + SMulCommClass.symm _ _ _ instance : MulDistribMulAction (A₁ ≃ₐ[R] A₁) A₁ˣ where smul := fun f => Units.map f diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index 89315cfde15fe..41c09bf110bf0 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -124,11 +124,13 @@ protected theorem smul_def (f : M ≃ₗ[R] M) (a : M) : f • a = f a := instance apply_faithfulSMul : FaithfulSMul (M ≃ₗ[R] M) M := ⟨@fun _ _ ↦ LinearEquiv.ext⟩ -instance apply_smulCommClass : SMulCommClass R (M ≃ₗ[R] M) M where - smul_comm r e m := (e.map_smul r m).symm +instance apply_smulCommClass [SMul S R] [SMul S M] [IsScalarTower S R M] : + SMulCommClass S (M ≃ₗ[R] M) M where + smul_comm r e m := (e.map_smul_of_tower r m).symm -instance apply_smulCommClass' : SMulCommClass (M ≃ₗ[R] M) R M where - smul_comm := LinearEquiv.map_smul +instance apply_smulCommClass' [SMul S R] [SMul S M] [IsScalarTower S R M] : + SMulCommClass (M ≃ₗ[R] M) S M := + SMulCommClass.symm _ _ _ end Automorphisms diff --git a/Mathlib/Algebra/Module/LinearMap/End.lean b/Mathlib/Algebra/Module/LinearMap/End.lean index a6d7f5857710f..4d09f38caa016 100644 --- a/Mathlib/Algebra/Module/LinearMap/End.lean +++ b/Mathlib/Algebra/Module/LinearMap/End.lean @@ -204,14 +204,16 @@ protected theorem smul_def (f : Module.End R M) (a : M) : f • a = f a := instance apply_faithfulSMul : FaithfulSMul (Module.End R M) M := ⟨LinearMap.ext⟩ -instance apply_smulCommClass : SMulCommClass R (Module.End R M) M where - smul_comm r e m := (e.map_smul r m).symm +instance apply_smulCommClass [SMul S R] [SMul S M] [IsScalarTower S R M] : + SMulCommClass S (Module.End R M) M where + smul_comm r e m := (e.map_smul_of_tower r m).symm -instance apply_smulCommClass' : SMulCommClass (Module.End R M) R M where - smul_comm := LinearMap.map_smul +instance apply_smulCommClass' [SMul S R] [SMul S M] [IsScalarTower S R M] : + SMulCommClass (Module.End R M) S M := + SMulCommClass.symm _ _ _ -instance apply_isScalarTower {R M : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] : - IsScalarTower R (Module.End R M) M := +instance apply_isScalarTower [Monoid S] [DistribMulAction S M] [SMulCommClass R S M] : + IsScalarTower S (Module.End R M) M := ⟨fun _ _ _ ↦ rfl⟩ end Endomorphisms From 604cb3610ee4c147dd3ef1966d0505c57cabbf0c Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 5 Aug 2024 22:14:34 +0000 Subject: [PATCH 042/359] feat(Analysis/SpecialFunctions/Exp): cexp is uniformlyContinuousOn (#15379) We show that the complex exponential is uniformly continuous on left half planes. Co-authored-by: Chris Birkbeck --- Mathlib/Analysis/SpecialFunctions/Exp.lean | 31 ++++++++++++++++++++++ 1 file changed, 31 insertions(+) diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index b5bd9ce06d960..33170348b5f1f 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -116,6 +116,37 @@ theorem ContinuousOn.cexp (h : ContinuousOn f s) : ContinuousOn (fun y => exp (f theorem Continuous.cexp (h : Continuous f) : Continuous fun y => exp (f y) := continuous_iff_continuousAt.2 fun _ => h.continuousAt.cexp +/-- The complex exponential function is uniformly continuous on left half planes. -/ +lemma UniformlyContinuousOn.cexp (a : ℝ) : UniformContinuousOn exp {x : ℂ | x.re ≤ a} := by + have : Continuous (cexp - 1) := Continuous.sub (Continuous.cexp continuous_id') continuous_one + rw [Metric.uniformContinuousOn_iff, Metric.continuous_iff'] at * + intro ε hε + simp only [gt_iff_lt, Pi.sub_apply, Pi.one_apply, dist_sub_eq_dist_add_right, + sub_add_cancel] at this + have ha : 0 < ε / (2 * Real.exp a) := by positivity + have H := this 0 (ε / (2 * Real.exp a)) ha + rw [Metric.eventually_nhds_iff] at H + obtain ⟨δ, hδ⟩ := H + refine ⟨δ, hδ.1, ?_⟩ + intros x _ y hy hxy + have h3 := hδ.2 (y := x - y) (by simpa only [dist_zero_right, norm_eq_abs] using hxy) + rw [dist_eq_norm, exp_zero] at * + have : cexp x - cexp y = cexp y * (cexp (x - y) - 1) := by + rw [mul_sub_one, ← exp_add] + ring_nf + rw [this, mul_comm] + have hya : ‖cexp y‖ ≤ Real.exp a := by + simp only [norm_eq_abs, abs_exp, Real.exp_le_exp] + exact hy + simp only [gt_iff_lt, dist_zero_right, norm_eq_abs, Set.mem_setOf_eq, norm_mul, + Complex.abs_exp] at * + apply lt_of_le_of_lt (mul_le_mul h3.le hya (Real.exp_nonneg y.re) (le_of_lt ha)) + have hrr : ε / (2 * a.exp) * a.exp = ε / 2 := by + nth_rw 2 [mul_comm] + field_simp [mul_assoc] + rw [hrr] + exact div_two_lt_of_pos hε + end ComplexContinuousExpComp namespace Real From 89f95eeaa70c281c63adab81ccffc2fab7a3b318 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 6 Aug 2024 02:26:11 +0000 Subject: [PATCH 043/359] chore(Perm/List): `List.get` -> `getElem` (#15477) Deprecate `List.get` lemmas, use `getElem` versions in proofs. --- Mathlib/GroupTheory/Perm/Cycle/Basic.lean | 4 +- Mathlib/GroupTheory/Perm/Cycle/Concrete.lean | 6 +- Mathlib/GroupTheory/Perm/List.lean | 59 ++++++++++---------- 3 files changed, 36 insertions(+), 33 deletions(-) diff --git a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean index 3e920ed2d45ec..4cc974bdf7585 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean @@ -867,10 +867,10 @@ theorem Nodup.isCycleOn_formPerm (h : l.Nodup) : l.formPerm.IsCycleOn { a | a ∈ l } := by refine ⟨l.formPerm.bijOn fun _ => List.formPerm_mem_iff_mem, fun a ha b hb => ?_⟩ rw [Set.mem_setOf, ← List.indexOf_lt_length] at ha hb - rw [← List.indexOf_get ha, ← List.indexOf_get hb] + rw [← List.getElem_indexOf ha, ← List.getElem_indexOf hb] refine ⟨l.indexOf b - l.indexOf a, ?_⟩ simp only [sub_eq_neg_add, zpow_add, zpow_neg, Equiv.Perm.inv_eq_iff_eq, zpow_natCast, - Equiv.Perm.coe_mul, List.formPerm_pow_apply_get _ h, Function.comp] + Equiv.Perm.coe_mul, List.formPerm_pow_apply_getElem _ h, Function.comp] rw [add_comm] end diff --git a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean index 509c1209b4f26..9792193750a29 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean @@ -77,10 +77,9 @@ theorem isCycle_formPerm (hl : Nodup l) (hn : 2 ≤ l.length) : IsCycle (formPer · rwa [formPerm_apply_mem_ne_self_iff _ hl _ (mem_cons_self _ _)] · intro w hw have : w ∈ x::y::l := mem_of_formPerm_ne_self _ _ hw - obtain ⟨k, hk⟩ := get_of_mem this + obtain ⟨k, hk, rfl⟩ := getElem_of_mem this use k - rw [← hk] - simp only [zpow_natCast, formPerm_pow_apply_head _ _ hl k, Nat.mod_eq_of_lt k.isLt] + simp only [zpow_natCast, formPerm_pow_apply_head _ _ hl k, Nat.mod_eq_of_lt hk] theorem pairwise_sameCycle_formPerm (hl : Nodup l) (hn : 2 ≤ l.length) : Pairwise l.formPerm.SameCycle l := @@ -110,6 +109,7 @@ theorem cycleType_formPerm (hl : Nodup l) (hn : 2 ≤ l.length) : · simpa using isCycle_formPerm hl hn · simp +set_option linter.deprecated false in theorem formPerm_apply_mem_eq_next (hl : Nodup l) (x : α) (hx : x ∈ l) : formPerm l x = next l x hx := by obtain ⟨k, rfl⟩ := get_of_mem hx diff --git a/Mathlib/GroupTheory/Perm/List.lean b/Mathlib/GroupTheory/Perm/List.lean index 76f10c832bac0..b19007fd40f9a 100644 --- a/Mathlib/GroupTheory/Perm/List.lean +++ b/Mathlib/GroupTheory/Perm/List.lean @@ -136,18 +136,19 @@ theorem formPerm_apply_getLast (x : α) (xs : List α) : @[simp] theorem formPerm_apply_getElem_length (x : α) (xs : List α) : - formPerm (x :: xs) ((x :: xs)[xs.length]) = x := by + formPerm (x :: xs) (x :: xs)[xs.length] = x := by rw [getElem_cons_length _ _ _ rfl, formPerm_apply_getLast] +@[deprecated formPerm_apply_getElem_length (since := "2024-08-03")] theorem formPerm_apply_get_length (x : α) (xs : List α) : - formPerm (x :: xs) ((x :: xs).get (Fin.mk xs.length (by simp))) = x := by - simp [formPerm_apply_getElem_length] + formPerm (x :: xs) ((x :: xs).get (Fin.mk xs.length (by simp))) = x := + formPerm_apply_getElem_length .. set_option linter.deprecated false in -@[simp, deprecated formPerm_apply_get_length (since := "2024-05-30")] +@[deprecated formPerm_apply_getElem_length (since := "2024-05-30")] theorem formPerm_apply_nthLe_length (x : α) (xs : List α) : - formPerm (x :: xs) ((x :: xs).nthLe xs.length (by simp)) = x := by - apply formPerm_apply_get_length + formPerm (x :: xs) ((x :: xs).nthLe xs.length (by simp)) = x := + formPerm_apply_getElem_length .. theorem formPerm_apply_head (x y : α) (xs : List α) (h : Nodup (x :: y :: xs)) : formPerm (x :: y :: xs) x = y := by simp [formPerm_apply_of_not_mem h.not_mem] @@ -159,12 +160,13 @@ theorem formPerm_apply_getElem_zero (l : List α) (h : Nodup l) (hl : 1 < l.leng · simp at hl · rw [getElem_cons_zero, formPerm_apply_head _ _ _ h, getElem_cons_succ, getElem_cons_zero] +@[deprecated formPerm_apply_getElem_zero (since := "2024-08-03")] theorem formPerm_apply_get_zero (l : List α) (h : Nodup l) (hl : 1 < l.length) : - formPerm l (l.get (Fin.mk 0 (by omega))) = l.get (Fin.mk 1 hl) := by - simp_all [formPerm_apply_getElem_zero] + formPerm l (l.get (Fin.mk 0 (by omega))) = l.get (Fin.mk 1 hl) := + formPerm_apply_getElem_zero l h hl set_option linter.deprecated false in -@[deprecated formPerm_apply_get_zero (since := "2024-05-30")] +@[deprecated formPerm_apply_getElem_zero (since := "2024-05-30")] theorem formPerm_apply_nthLe_zero (l : List α) (h : Nodup l) (hl : 1 < l.length) : formPerm l (l.nthLe 0 (by omega)) = l.nthLe 1 hl := by apply formPerm_apply_get_zero _ h @@ -176,7 +178,7 @@ theorem formPerm_eq_head_iff_eq_getLast (x y : α) : Iff.trans (by rw [formPerm_apply_getLast]) (formPerm (y :: l)).injective.eq_iff theorem formPerm_apply_lt_getElem (xs : List α) (h : Nodup xs) (n : ℕ) (hn : n + 1 < xs.length) : - formPerm xs (xs[n]'((Nat.lt_succ_self n).trans hn)) = xs[n + 1] := by + formPerm xs xs[n] = xs[n + 1] := by induction' n with n IH generalizing xs · simpa using formPerm_apply_getElem_zero _ h _ · rcases xs with (_ | ⟨x, _ | ⟨y, l⟩⟩) @@ -191,6 +193,7 @@ theorem formPerm_apply_lt_getElem (xs : List α) (h : Nodup xs) (n : ℕ) (hn : rw [← hx, IH] at h simp [getElem_mem] at h +@[deprecated formPerm_apply_lt_getElem (since := "2024-08-03")] theorem formPerm_apply_lt_get (xs : List α) (h : Nodup xs) (n : ℕ) (hn : n + 1 < xs.length) : formPerm xs (xs.get (Fin.mk n ((Nat.lt_succ_self n).trans hn))) = xs.get (Fin.mk (n + 1) hn) := by @@ -202,7 +205,7 @@ theorem formPerm_apply_lt (xs : List α) (h : Nodup xs) (n : ℕ) (hn : n + 1 < formPerm xs (xs.nthLe n ((Nat.lt_succ_self n).trans hn)) = xs.nthLe (n + 1) hn := by apply formPerm_apply_lt_get _ h -theorem formPerm_apply_getElem (xs : List α) (w : Nodup xs) (i : Nat) (h : i < xs.length) : +theorem formPerm_apply_getElem (xs : List α) (w : Nodup xs) (i : ℕ) (h : i < xs.length) : formPerm xs xs[i] = xs[(i + 1) % xs.length]'(Nat.mod_lt _ (i.zero_le.trans_lt h)) := by cases' xs with x xs @@ -216,6 +219,7 @@ theorem formPerm_apply_getElem (xs : List α) (w : Nodup xs) (i : Nat) (h : i < congr rw [Nat.mod_eq_of_lt]; simpa [Nat.succ_eq_add_one] +@[deprecated formPerm_apply_getElem (since := "2024-08-03")] theorem formPerm_apply_get (xs : List α) (h : Nodup xs) (i : Fin xs.length) : formPerm xs (xs.get i) = xs.get ⟨((i.val + 1) % xs.length), (Nat.mod_lt _ (i.val.zero_le.trans_lt i.isLt))⟩ := by @@ -234,8 +238,8 @@ theorem support_formPerm_of_nodup' (l : List α) (h : Nodup l) (h' : ∀ x : α, · exact support_formPerm_le' l · intro x hx simp only [Finset.mem_coe, mem_toFinset] at hx - obtain ⟨⟨n, hn⟩, rfl⟩ := get_of_mem hx - rw [Set.mem_setOf_eq, formPerm_apply_get _ h] + obtain ⟨n, hn, rfl⟩ := getElem_of_mem hx + rw [Set.mem_setOf_eq, formPerm_apply_getElem _ h] intro H rw [nodup_iff_injective_get, Function.Injective] at h specialize h H @@ -255,8 +259,8 @@ theorem formPerm_rotate_one (l : List α) (h : Nodup l) : formPerm (l.rotate 1) have h' : Nodup (l.rotate 1) := by simpa using h ext x by_cases hx : x ∈ l.rotate 1 - · obtain ⟨⟨k, hk⟩, rfl⟩ := get_of_mem hx - rw [formPerm_apply_get _ h', get_rotate l, get_rotate l, formPerm_apply_get _ h] + · obtain ⟨k, hk, rfl⟩ := getElem_of_mem hx + rw [formPerm_apply_getElem _ h', getElem_rotate l, getElem_rotate l, formPerm_apply_getElem _ h] simp · rw [formPerm_apply_of_not_mem hx, formPerm_apply_of_not_mem] simpa using hx @@ -287,7 +291,7 @@ theorem formPerm_reverse : ∀ l : List α, formPerm l.reverse = (formPerm l)⁻ | a::b::l => by simp [formPerm_append_pair, swap_comm, ← formPerm_reverse (b::l)] -theorem formPerm_pow_apply_getElem (l : List α) (w : Nodup l) (n : ℕ) (i : Nat) (h : i < l.length) : +theorem formPerm_pow_apply_getElem (l : List α) (w : Nodup l) (n : ℕ) (i : ℕ) (h : i < l.length) : (formPerm l ^ n) l[i] = l[(i + n) % l.length]'(Nat.mod_lt _ (i.zero_le.trans_lt h)) := by induction' n with n hn @@ -295,6 +299,7 @@ theorem formPerm_pow_apply_getElem (l : List α) (w : Nodup l) (n : ℕ) (i : Na · simp [pow_succ', mul_apply, hn, formPerm_apply_getElem _ w, Nat.succ_eq_add_one, ← Nat.add_assoc] +@[deprecated formPerm_pow_apply_getElem (since := "2024-08-03")] theorem formPerm_pow_apply_get (l : List α) (h : Nodup l) (n : ℕ) (i : Fin l.length) : (formPerm l ^ n) (l.get i) = l.get ⟨((i.val + n) % l.length), (Nat.mod_lt _ (i.val.zero_le.trans_lt i.isLt))⟩ := by @@ -309,8 +314,8 @@ theorem formPerm_pow_apply_nthLe (l : List α) (h : Nodup l) (n k : ℕ) (hk : k theorem formPerm_pow_apply_head (x : α) (l : List α) (h : Nodup (x :: l)) (n : ℕ) : (formPerm (x :: l) ^ n) x = - (x :: l).get ⟨(n % (x :: l).length), (Nat.mod_lt _ (Nat.zero_lt_succ _))⟩ := by - convert formPerm_pow_apply_get _ h n ⟨0, Nat.succ_pos _⟩ + (x :: l)[(n % (x :: l).length)]'(Nat.mod_lt _ (Nat.zero_lt_succ _)) := by + convert formPerm_pow_apply_getElem _ h n 0 (Nat.succ_pos _) simp theorem formPerm_ext_iff {x y x' y' : α} {l l' : List α} (hd : Nodup (x :: y :: l)) @@ -333,27 +338,25 @@ theorem formPerm_ext_iff {x y x' y' : α} {l l' : List α} (hd : Nodup (x :: y : support_formPerm_of_nodup' _ hd' (by simp)] simp only [h] use n - apply List.ext_get + apply List.ext_getElem · rw [length_rotate, hl] · intro k hk hk' - rw [get_rotate] + rw [getElem_rotate] induction' k with k IH · refine Eq.trans ?_ hx' congr simpa using hn - · conv => congr <;> · arg 2; (congr; (simp only [Fin.val_mk]; rw [← Nat.mod_eq_of_lt hk'])) - rw [← formPerm_apply_get _ hd' ⟨k, k.lt_succ_self.trans hk'⟩, - ← IH (k.lt_succ_self.trans hk), ← h, formPerm_apply_get _ hd] - congr 2 - simp only [Fin.val_mk] + · conv => congr <;> · arg 2; (rw [← Nat.mod_eq_of_lt hk']) + rw [← formPerm_apply_getElem _ hd' k (k.lt_succ_self.trans hk'), + ← IH (k.lt_succ_self.trans hk), ← h, formPerm_apply_getElem _ hd] + congr 1 rw [hl, Nat.mod_eq_of_lt hk', add_right_comm] apply Nat.add_mod theorem formPerm_apply_mem_eq_self_iff (hl : Nodup l) (x : α) (hx : x ∈ l) : formPerm l x = x ↔ length l ≤ 1 := by - obtain ⟨⟨k, hk⟩, rfl⟩ := get_of_mem hx - rw [formPerm_apply_get _ hl ⟨k, hk⟩, hl.get_inj_iff, Fin.mk.inj_iff] - simp only [Fin.val_mk] + obtain ⟨k, hk, rfl⟩ := getElem_of_mem hx + rw [formPerm_apply_getElem _ hl k hk, hl.getElem_inj_iff] cases hn : l.length · exact absurd k.zero_le (hk.trans_le hn.le).not_le · rw [hn] at hk From 994b33b7e7e8b63a7d5912033854fded25feaf52 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 6 Aug 2024 04:16:35 +0000 Subject: [PATCH 044/359] chore(*): rename `Cstar*` to `CStar*` (#15495) Moves: - Cstar* -> CStar* --- Mathlib.lean | 32 +++++----- .../{CstarAlgebra => CStarAlgebra}/Basic.lean | 48 +++++++-------- .../ContinuousFunctionalCalculus/Basic.lean | 4 +- .../Instances.lean | 44 +++++++------- .../NonUnital.lean | 6 +- .../ContinuousFunctionalCalculus/Order.lean | 50 ++++++++-------- .../Restrict.lean | 2 +- .../ContinuousFunctionalCalculus/Unique.lean | 2 +- .../ContinuousFunctionalCalculus/Unital.lean | 4 +- .../ContinuousFunctionalCalculus/Unitary.lean | 2 +- .../Exponential.lean | 0 .../GelfandDuality.lean | 14 ++--- .../Matrix.lean | 6 +- .../Module.lean | 58 ++++++++++--------- .../Multiplier.lean | 14 ++--- .../Spectrum.lean | 16 ++--- .../Unitization.lean | 18 +++--- .../Analysis/InnerProductSpace/Adjoint.lean | 2 +- Mathlib/Analysis/Matrix.lean | 2 +- Mathlib/Analysis/Normed/Lp/lpSpace.lean | 6 +- .../NormedSpace/OperatorNorm/Mul.lean | 2 +- Mathlib/Analysis/Quaternion.lean | 2 +- Mathlib/Analysis/RCLike/Basic.lean | 4 +- .../ContinuousFunctionalCalculus/ExpLog.lean | 2 +- Mathlib/Analysis/VonNeumannAlgebra/Basic.lean | 2 +- .../Matrix/HermitianFunctionalCalculus.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Spectrum.lean | 2 +- .../Function/ContinuousMapDense.lean | 2 +- .../Topology/ContinuousFunction/Bounded.lean | 12 ++-- .../Topology/ContinuousFunction/Compact.lean | 8 +-- .../ContinuousFunction/ZeroAtInfty.lean | 8 +-- scripts/noshake.json | 4 +- test/TCSynth.lean | 2 +- 33 files changed, 193 insertions(+), 189 deletions(-) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/Basic.lean (88%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/ContinuousFunctionalCalculus/Basic.lean (99%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/ContinuousFunctionalCalculus/Instances.lean (95%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/ContinuousFunctionalCalculus/NonUnital.lean (99%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/ContinuousFunctionalCalculus/Order.lean (83%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/ContinuousFunctionalCalculus/Restrict.lean (99%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/ContinuousFunctionalCalculus/Unique.lean (99%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/ContinuousFunctionalCalculus/Unital.lean (99%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/ContinuousFunctionalCalculus/Unitary.lean (97%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/Exponential.lean (100%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/GelfandDuality.lean (97%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/Matrix.lean (98%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/Module.lean (87%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/Multiplier.lean (98%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/Spectrum.lean (94%) rename Mathlib/Analysis/{CstarAlgebra => CStarAlgebra}/Unitization.lean (95%) diff --git a/Mathlib.lean b/Mathlib.lean index 2af47ca3954ff..85efd600407a5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -904,6 +904,22 @@ import Mathlib.Analysis.BoxIntegral.Partition.Measure import Mathlib.Analysis.BoxIntegral.Partition.Split import Mathlib.Analysis.BoxIntegral.Partition.SubboxInduction import Mathlib.Analysis.BoxIntegral.Partition.Tagged +import Mathlib.Analysis.CStarAlgebra.Basic +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Basic +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Order +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Restrict +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unique +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unitary +import Mathlib.Analysis.CStarAlgebra.Exponential +import Mathlib.Analysis.CStarAlgebra.GelfandDuality +import Mathlib.Analysis.CStarAlgebra.Matrix +import Mathlib.Analysis.CStarAlgebra.Module +import Mathlib.Analysis.CStarAlgebra.Multiplier +import Mathlib.Analysis.CStarAlgebra.Spectrum +import Mathlib.Analysis.CStarAlgebra.Unitization import Mathlib.Analysis.Calculus.AddTorsor.AffineMap import Mathlib.Analysis.Calculus.AddTorsor.Coord import Mathlib.Analysis.Calculus.BumpFunction.Basic @@ -1068,22 +1084,6 @@ import Mathlib.Analysis.Convex.Strong import Mathlib.Analysis.Convex.Topology import Mathlib.Analysis.Convex.Uniform import Mathlib.Analysis.Convolution -import Mathlib.Analysis.CstarAlgebra.Basic -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Basic -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Instances -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.NonUnital -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Order -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Restrict -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unique -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unitary -import Mathlib.Analysis.CstarAlgebra.Exponential -import Mathlib.Analysis.CstarAlgebra.GelfandDuality -import Mathlib.Analysis.CstarAlgebra.Matrix -import Mathlib.Analysis.CstarAlgebra.Module -import Mathlib.Analysis.CstarAlgebra.Multiplier -import Mathlib.Analysis.CstarAlgebra.Spectrum -import Mathlib.Analysis.CstarAlgebra.Unitization import Mathlib.Analysis.Distribution.AEEqOfIntegralContDiff import Mathlib.Analysis.Distribution.FourierSchwartz import Mathlib.Analysis.Distribution.SchwartzSpace diff --git a/Mathlib/Analysis/CstarAlgebra/Basic.lean b/Mathlib/Analysis/CStarAlgebra/Basic.lean similarity index 88% rename from Mathlib/Analysis/CstarAlgebra/Basic.lean rename to Mathlib/Analysis/CStarAlgebra/Basic.lean index deaff8d8703e5..f0cb7d301a058 100644 --- a/Mathlib/Analysis/CstarAlgebra/Basic.lean +++ b/Mathlib/Analysis/CStarAlgebra/Basic.lean @@ -21,7 +21,7 @@ condition `‖x‖^2 ≤ ‖x⋆ * x‖` for all `x` (which actually implies equ a star algebra, then it is a C⋆-algebra. To get a C⋆-algebra `E` over field `𝕜`, use -`[NormedField 𝕜] [StarRing 𝕜] [NormedRing E] [StarRing E] [CstarRing E] +`[NormedField 𝕜] [StarRing 𝕜] [NormedRing E] [StarRing E] [CStarRing E] [NormedAlgebra 𝕜 E] [StarModule 𝕜 E]`. ## TODO @@ -74,18 +74,20 @@ instance RingHomIsometric.starRingEnd [NormedCommRing E] [StarRing E] [NormedSta /-- A C*-ring is a normed star ring that satisfies the stronger condition `‖x‖ ^ 2 ≤ ‖x⋆ * x‖` for every `x`. Note that this condition actually implies equality, as is shown in `norm_star_mul_self` below. -/ -class CstarRing (E : Type*) [NonUnitalNormedRing E] [StarRing E] : Prop where +class CStarRing (E : Type*) [NonUnitalNormedRing E] [StarRing E] : Prop where norm_mul_self_le : ∀ x : E, ‖x‖ * ‖x‖ ≤ ‖x⋆ * x‖ -instance : CstarRing ℝ where +@[deprecated (since := "2024-08-04")] alias CstarRing := CStarRing + +instance : CStarRing ℝ where norm_mul_self_le x := by simp only [Real.norm_eq_abs, abs_mul_abs_self, star, id, norm_mul, le_refl] -namespace CstarRing +namespace CStarRing section NonUnital -variable [NonUnitalNormedRing E] [StarRing E] [CstarRing E] +variable [NonUnitalNormedRing E] [StarRing E] [CStarRing E] -- see Note [lower instance priority] /-- In a C*-ring, star preserves the norm. -/ @@ -97,14 +99,14 @@ instance (priority := 100) to_normedStarGroup : NormedStarGroup E := · have hnt : 0 < ‖x‖ := norm_pos_iff.mpr htriv have h₁ : ∀ z : E, ‖z⋆ * z‖ ≤ ‖z⋆‖ * ‖z‖ := fun z => norm_mul_le z⋆ z have h₂ : ∀ z : E, 0 < ‖z‖ → ‖z‖ ≤ ‖z⋆‖ := fun z hz => by - rw [← mul_le_mul_right hz]; exact (CstarRing.norm_mul_self_le z).trans (h₁ z) + rw [← mul_le_mul_right hz]; exact (CStarRing.norm_mul_self_le z).trans (h₁ z) have h₃ : ‖x⋆‖ ≤ ‖x‖ := by conv_rhs => rw [← star_star x] exact h₂ x⋆ (gt_of_ge_of_gt (h₂ x hnt) hnt) exact le_antisymm h₃ (h₂ x hnt)⟩ theorem norm_star_mul_self {x : E} : ‖x⋆ * x‖ = ‖x‖ * ‖x‖ := - le_antisymm ((norm_mul_le _ _).trans (by rw [norm_star])) (CstarRing.norm_mul_self_le x) + le_antisymm ((norm_mul_le _ _).trans (by rw [norm_star])) (CStarRing.norm_mul_self_le x) theorem norm_self_mul_star {x : E} : ‖x * x⋆‖ = ‖x‖ * ‖x‖ := by nth_rw 1 [← star_star x] @@ -138,8 +140,8 @@ end NonUnital section ProdPi variable {ι R₁ R₂ : Type*} {R : ι → Type*} -variable [NonUnitalNormedRing R₁] [StarRing R₁] [CstarRing R₁] -variable [NonUnitalNormedRing R₂] [StarRing R₂] [CstarRing R₂] +variable [NonUnitalNormedRing R₁] [StarRing R₁] [CStarRing R₁] +variable [NonUnitalNormedRing R₂] [StarRing R₂] [CStarRing R₂] variable [∀ i, NonUnitalNormedRing (R i)] [∀ i, StarRing (R i)] /-- This instance exists to short circuit type class resolution because of problems with @@ -147,16 +149,16 @@ inference involving Π-types. -/ instance _root_.Pi.starRing' : StarRing (∀ i, R i) := inferInstance -variable [Fintype ι] [∀ i, CstarRing (R i)] +variable [Fintype ι] [∀ i, CStarRing (R i)] -instance _root_.Prod.cstarRing : CstarRing (R₁ × R₂) where +instance _root_.Prod.cstarRing : CStarRing (R₁ × R₂) where norm_mul_self_le x := by dsimp only [norm] simp only [Prod.fst_mul, Prod.fst_star, Prod.snd_mul, Prod.snd_star, norm_star_mul_self, ← sq] rw [le_sup_iff] rcases le_total ‖x.fst‖ ‖x.snd‖ with (h | h) <;> simp [h] -instance _root_.Pi.cstarRing : CstarRing (∀ i, R i) where +instance _root_.Pi.cstarRing : CStarRing (∀ i, R i) where norm_mul_self_le x := by refine le_of_eq (Eq.symm ?_) simp only [norm, Pi.mul_apply, Pi.star_apply, nnnorm_star_mul_self, ← sq] @@ -165,7 +167,7 @@ instance _root_.Pi.cstarRing : CstarRing (∀ i, R i) where (Finset.comp_sup_eq_sup_comp_of_is_total (fun x : NNReal => x ^ 2) (fun x y h => by simpa only [sq] using mul_le_mul' h h) (by simp)).symm -instance _root_.Pi.cstarRing' : CstarRing (ι → R₁) := +instance _root_.Pi.cstarRing' : CStarRing (ι → R₁) := Pi.cstarRing end ProdPi @@ -173,7 +175,7 @@ end ProdPi section Unital -variable [NormedRing E] [StarRing E] [CstarRing E] +variable [NormedRing E] [StarRing E] [CStarRing E] @[simp, nolint simpNF] -- Porting note (#10959): simp cannot prove this theorem norm_one [Nontrivial E] : ‖(1 : E)‖ = 1 := by @@ -185,8 +187,8 @@ instance (priority := 100) [Nontrivial E] : NormOneClass E := ⟨norm_one⟩ theorem norm_coe_unitary [Nontrivial E] (U : unitary E) : ‖(U : E)‖ = 1 := by - rw [← sq_eq_sq (norm_nonneg _) zero_le_one, one_pow 2, sq, ← CstarRing.norm_star_mul_self, - unitary.coe_star_mul_self, CstarRing.norm_one] + rw [← sq_eq_sq (norm_nonneg _) zero_le_one, one_pow 2, sq, ← CStarRing.norm_star_mul_self, + unitary.coe_star_mul_self, CStarRing.norm_one] @[simp] theorem norm_of_mem_unitary [Nontrivial E] {U : E} (hU : U ∈ unitary E) : ‖U‖ = 1 := @@ -226,17 +228,17 @@ theorem norm_mul_mem_unitary (A : E) {U : E} (hU : U ∈ unitary E) : ‖A * U end Unital -end CstarRing +end CStarRing -theorem IsSelfAdjoint.nnnorm_pow_two_pow [NormedRing E] [StarRing E] [CstarRing E] {x : E} +theorem IsSelfAdjoint.nnnorm_pow_two_pow [NormedRing E] [StarRing E] [CStarRing E] {x : E} (hx : IsSelfAdjoint x) (n : ℕ) : ‖x ^ 2 ^ n‖₊ = ‖x‖₊ ^ 2 ^ n := by induction' n with k hk · simp only [pow_zero, pow_one, Nat.zero_eq] · rw [pow_succ', pow_mul', sq] nth_rw 1 [← selfAdjoint.mem_iff.mp hx] - rw [← star_pow, CstarRing.nnnorm_star_mul_self, ← sq, hk, pow_mul'] + rw [← star_pow, CStarRing.nnnorm_star_mul_self, ← sq, hk, pow_mul'] -theorem selfAdjoint.nnnorm_pow_two_pow [NormedRing E] [StarRing E] [CstarRing E] (x : selfAdjoint E) +theorem selfAdjoint.nnnorm_pow_two_pow [NormedRing E] [StarRing E] [CStarRing E] (x : selfAdjoint E) (n : ℕ) : ‖x ^ 2 ^ n‖₊ = ‖x‖₊ ^ 2 ^ n := x.prop.nnnorm_pow_two_pow _ @@ -275,8 +277,8 @@ instance toNormedAlgebra {𝕜 A : Type*} [NormedField 𝕜] [StarRing 𝕜] [Se [NormedAlgebra 𝕜 A] [StarModule 𝕜 A] (S : StarSubalgebra 𝕜 A) : NormedAlgebra 𝕜 S := NormedAlgebra.induced 𝕜 S A S.subtype -instance to_cstarRing {R A} [CommRing R] [StarRing R] [NormedRing A] [StarRing A] [CstarRing A] - [Algebra R A] [StarModule R A] (S : StarSubalgebra R A) : CstarRing S where - norm_mul_self_le x := @CstarRing.norm_mul_self_le A _ _ _ x +instance to_cstarRing {R A} [CommRing R] [StarRing R] [NormedRing A] [StarRing A] [CStarRing A] + [Algebra R A] [StarModule R A] (S : StarSubalgebra R A) : CStarRing S where + norm_mul_self_le x := @CStarRing.norm_mul_self_le A _ _ _ x end StarSubalgebra diff --git a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Basic.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Basic.lean similarity index 99% rename from Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Basic.lean rename to Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Basic.lean index fce7f222738c0..2dc9d1ad67c79 100644 --- a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Basic.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.CstarAlgebra.GelfandDuality +import Mathlib.Analysis.CStarAlgebra.GelfandDuality import Mathlib.Topology.Algebra.StarSubalgebra /-! # Continuous functional calculus @@ -59,7 +59,7 @@ open scoped Pointwise ENNReal NNReal ComplexOrder open WeakDual WeakDual.CharacterSpace elementalStarAlgebra variable {A : Type*} [NormedRing A] [NormedAlgebra ℂ A] -variable [StarRing A] [CstarRing A] [StarModule ℂ A] +variable [StarRing A] [CStarRing A] [StarModule ℂ A] instance {R A : Type*} [CommRing R] [StarRing R] [NormedRing A] [Algebra R A] [StarRing A] [ContinuousStar A] [StarModule R A] (a : A) [IsStarNormal a] : diff --git a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean similarity index 95% rename from Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Instances.lean rename to Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index 6aa55c0bef7e3..4f2c56e68c51b 100644 --- a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -3,11 +3,11 @@ Copyright (c) 2024 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Restrict -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Basic -import Mathlib.Analysis.CstarAlgebra.Spectrum -import Mathlib.Analysis.CstarAlgebra.Unitization -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unique +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Restrict +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Basic +import Mathlib.Analysis.CStarAlgebra.Spectrum +import Mathlib.Analysis.CStarAlgebra.Unitization +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unique /-! # Instances of the continuous functional calculus @@ -23,7 +23,7 @@ import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unique elements in an `ℝ`-algebra with a continuous functional calculus for selfadjoint elements, where every element has compact spectrum, and where nonnegative elements have nonnegative spectrum. In particular, this includes unital C⋆-algebras over `ℝ`. -* `CstarRing.instNonnegSpectrumClass`: In a unital C⋆-algebra over `ℂ` which is also a +* `CStarRing.instNonnegSpectrumClass`: In a unital C⋆-algebra over `ℂ` which is also a `StarOrderedRing`, the spectrum of a nonnegative element is nonnegative. ## Tags @@ -42,7 +42,7 @@ local notation "σ" => spectrum section RCLike -variable {𝕜 A : Type*} [RCLike 𝕜] [NonUnitalNormedRing A] [StarRing A] [CstarRing A] +variable {𝕜 A : Type*} [RCLike 𝕜] [NonUnitalNormedRing A] [StarRing A] [CStarRing A] variable [CompleteSpace A] [NormedSpace 𝕜 A] [IsScalarTower 𝕜 A A] [SMulCommClass 𝕜 A A] variable [StarModule 𝕜 A] {p : A → Prop} {p₁ : Unitization 𝕜 A → Prop} @@ -147,7 +147,7 @@ end RCLike section Normal instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [NormedRing A] [StarRing A] - [CstarRing A] [CompleteSpace A] [NormedAlgebra ℂ A] [StarModule ℂ A] : + [CStarRing A] [CompleteSpace A] [NormedAlgebra ℂ A] [StarModule ℂ A] : ContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop) where predicate_zero := isStarNormal_zero exists_cfc_of_predicate a ha := by @@ -164,7 +164,7 @@ instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [NormedRing A case predicate_hom => exact fun f ↦ ⟨by rw [← map_star]; exact Commute.all (star f) f |>.map _⟩ instance IsStarNormal.instNonUnitalContinuousFunctionalCalculus {A : Type*} [NonUnitalNormedRing A] - [StarRing A] [CstarRing A] [CompleteSpace A] [NormedSpace ℂ A] [IsScalarTower ℂ A A] + [StarRing A] [CStarRing A] [CompleteSpace A] [NormedSpace ℂ A] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] [StarModule ℂ A] : NonUnitalContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop) := RCLike.nonUnitalContinuousFunctionalCalculus Unitization.isStarNormal_inr @@ -355,7 +355,7 @@ section SpectrumRestricts open NNReal ENNReal -variable {A : Type*} [NormedRing A] [StarRing A] [CstarRing A] [CompleteSpace A] +variable {A : Type*} [NormedRing A] [StarRing A] [CStarRing A] [CompleteSpace A] variable [NormedAlgebra ℂ A] [StarModule ℂ A] lemma SpectrumRestricts.nnreal_iff_nnnorm {a : A} {t : ℝ≥0} (ha : IsSelfAdjoint a) (ht : ‖a‖₊ ≤ t) : @@ -476,10 +476,10 @@ end SpectrumRestricts section NonnegSpectrumClass variable {A : Type*} [NormedRing A] [CompleteSpace A] -variable [PartialOrder A] [StarRing A] [StarOrderedRing A] [CstarRing A] +variable [PartialOrder A] [StarRing A] [StarOrderedRing A] [CStarRing A] variable [NormedAlgebra ℂ A] [StarModule ℂ A] -instance CstarRing.instNonnegSpectrumClass : NonnegSpectrumClass ℝ A := +instance CStarRing.instNonnegSpectrumClass : NonnegSpectrumClass ℝ A := .of_spectrum_nonneg fun a ha ↦ by rw [StarOrderedRing.nonneg_iff] at ha induction ha using AddSubmonoid.closure_induction' with @@ -495,7 +495,7 @@ instance CstarRing.instNonnegSpectrumClass : NonnegSpectrumClass ℝ A := exact hx.nnreal_add (.of_nonneg x_mem) (.of_nonneg y_mem) hy open ComplexOrder in -instance CstarRing.instNonnegSpectrumClassComplexUnital : NonnegSpectrumClass ℂ A where +instance CStarRing.instNonnegSpectrumClassComplexUnital : NonnegSpectrumClass ℂ A where quasispectrum_nonneg_of_nonneg a ha x := by rw [mem_quasispectrum_iff] refine (Or.elim · ge_of_eq fun hx ↦ ?_) @@ -506,7 +506,7 @@ end NonnegSpectrumClass section SpectralOrder -variable (A : Type*) [NormedRing A] [CompleteSpace A] [StarRing A] [CstarRing A] +variable (A : Type*) [NormedRing A] [CompleteSpace A] [StarRing A] [CStarRing A] variable [NormedAlgebra ℂ A] [StarModule ℂ A] /-- The partial order on a unital C⋆-algebra defined by `x ≤ y` if and only if `y - x` is @@ -515,7 +515,7 @@ selfadjoint and has nonnegative spectrum. This is not declared as an instance because one may already have a partial order with better definitional properties. However, it can be useful to invoke this as an instance in proofs. -/ @[reducible] -def CstarRing.spectralOrder : PartialOrder A where +def CStarRing.spectralOrder : PartialOrder A where le x y := IsSelfAdjoint (y - x) ∧ SpectrumRestricts (y - x) ContinuousMap.realToNNReal le_refl := by simp only [sub_self, IsSelfAdjoint.zero, true_and, forall_const] @@ -528,9 +528,9 @@ def CstarRing.spectralOrder : PartialOrder A where le_trans x y z hxy hyz := ⟨by simpa using hyz.1.add hxy.1, by simpa using hyz.2.nnreal_add hyz.1 hxy.1 hxy.2⟩ -/-- The `CstarRing.spectralOrder` on a unital C⋆-algebra is a `StarOrderedRing`. -/ -lemma CstarRing.spectralOrderedRing : @StarOrderedRing A _ (CstarRing.spectralOrder A) _ := - let _ := CstarRing.spectralOrder A +/-- The `CStarRing.spectralOrder` on a unital C⋆-algebra is a `StarOrderedRing`. -/ +lemma CStarRing.spectralOrderedRing : @StarOrderedRing A _ (CStarRing.spectralOrder A) _ := + let _ := CStarRing.spectralOrder A { le_iff := by intro x y constructor @@ -559,15 +559,15 @@ end SpectralOrder section NonnegSpectrumClass variable {A : Type*} [NonUnitalNormedRing A] [CompleteSpace A] -variable [PartialOrder A] [StarRing A] [StarOrderedRing A] [CstarRing A] +variable [PartialOrder A] [StarRing A] [StarOrderedRing A] [CStarRing A] variable [NormedSpace ℂ A] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] [StarModule ℂ A] -instance CstarRing.instNonnegSpectrumClass' : NonnegSpectrumClass ℝ A where +instance CStarRing.instNonnegSpectrumClass' : NonnegSpectrumClass ℝ A where quasispectrum_nonneg_of_nonneg a ha := by rw [Unitization.quasispectrum_eq_spectrum_inr' _ ℂ] -- should this actually be an instance on the `Unitization`? (probably scoped) - let _ := CstarRing.spectralOrder (Unitization ℂ A) - have := CstarRing.spectralOrderedRing (Unitization ℂ A) + let _ := CStarRing.spectralOrder (Unitization ℂ A) + have := CStarRing.spectralOrderedRing (Unitization ℂ A) apply spectrum_nonneg_of_nonneg rw [StarOrderedRing.nonneg_iff] at ha ⊢ have := AddSubmonoid.mem_map_of_mem (Unitization.inrNonUnitalStarAlgHom ℂ A) ha diff --git a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean similarity index 99% rename from Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean rename to Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index f9afc68485787..7e7644559f0e2 100644 --- a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -6,7 +6,7 @@ Authors: Jireh Loreaux import Mathlib.Algebra.Algebra.Quasispectrum import Mathlib.Topology.ContinuousFunction.Compact import Mathlib.Topology.ContinuousFunction.ContinuousMapZero -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital import Mathlib.Topology.UniformSpace.CompactConvergence /-! @@ -15,7 +15,7 @@ import Mathlib.Topology.UniformSpace.CompactConvergence This file defines a generic API for the *continuous functional calculus* in *non-unital* algebras which is suitable in a wide range of settings. The design is intended to match as closely as possible that for unital algebras in -`Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital`. Changes to either file should +`Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital`. Changes to either file should be mirrored in its counterpart whenever possible. The underlying reasons for the design decisions in the unital case apply equally in the non-unital case. See the module documentation in that file for more information. @@ -182,7 +182,7 @@ junk value `0`. This is the primary declaration intended for widespread use of the continuous functional calculus for non-unital algebras, and all the API applies to this declaration. For more information, see the -module documentation for `Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital`. -/ +module documentation for `Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital`. -/ noncomputable irreducible_def cfcₙ (f : R → R) (a : A) : A := if h : p a ∧ ContinuousOn f (σₙ R a) ∧ f 0 = 0 then cfcₙHom h.1 ⟨⟨_, h.2.1.restrict⟩, h.2.2⟩ diff --git a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean similarity index 83% rename from Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Order.lean rename to Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index c407b6c5369aa..c16f8ee372b0d 100644 --- a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Instances +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances import Mathlib.Topology.ContinuousFunction.StarOrdered -import Mathlib.Analysis.CstarAlgebra.Unitization +import Mathlib.Analysis.CStarAlgebra.Unitization /-! # Facts about star-ordered rings that depend on the continuous functional calculus @@ -23,9 +23,9 @@ the spectral order. C⋆-algebra. * `mul_star_le_algebraMap_norm_sq` and `star_mul_le_algebraMap_norm_sq`, which give similar statements for `a * star a` and `star a * a`. -* `CstarRing.norm_le_norm_of_nonneg_of_le`: in a non-unital C⋆-algebra, if `0 ≤ a ≤ b`, then +* `CStarRing.norm_le_norm_of_nonneg_of_le`: in a non-unital C⋆-algebra, if `0 ≤ a ≤ b`, then `‖a‖ ≤ ‖b‖`. -* `CstarRing.conjugate_le_norm_smul`: in a non-unital C⋆-algebra, we have that +* `CStarRing.conjugate_le_norm_smul`: in a non-unital C⋆-algebra, we have that `star a * b * a ≤ ‖b‖ • (star a * a)` (and a primed version for the `a * b * star a` case). ## Tags @@ -38,12 +38,12 @@ open scoped NNReal namespace Unitization variable {A : Type*} [NonUnitalNormedRing A] [CompleteSpace A] [Nontrivial A] - [PartialOrder A] [StarRing A] [StarOrderedRing A] [CstarRing A] [NormedSpace ℂ A] [StarModule ℂ A] + [PartialOrder A] [StarRing A] [StarOrderedRing A] [CStarRing A] [NormedSpace ℂ A] [StarModule ℂ A] [SMulCommClass ℂ A A] [IsScalarTower ℂ A A] -instance instPartialOrder : PartialOrder (Unitization ℂ A) := CstarRing.spectralOrder _ +instance instPartialOrder : PartialOrder (Unitization ℂ A) := CStarRing.spectralOrder _ -instance instStarOrderedRing : StarOrderedRing (Unitization ℂ A) := CstarRing.spectralOrderedRing _ +instance instStarOrderedRing : StarOrderedRing (Unitization ℂ A) := CStarRing.spectralOrderedRing _ lemma inr_le_iff (a b : A) (ha : IsSelfAdjoint a := by cfc_tac) (hb : IsSelfAdjoint b := by cfc_tac) : @@ -64,9 +64,9 @@ lemma inr_nonneg_iff {a : A} : 0 ≤ (a : Unitization ℂ A) ↔ 0 ≤ a := by end Unitization -section Cstar_unital +section CStar_unital -variable {A : Type*} [NormedRing A] [StarRing A] [CstarRing A] [CompleteSpace A] +variable {A : Type*} [NormedRing A] [StarRing A] [CStarRing A] [CompleteSpace A] variable [NormedAlgebra ℂ A] [StarModule ℂ A] [PartialOrder A] [StarOrderedRing A] lemma IsSelfAdjoint.le_algebraMap_norm_self {a : A} (ha : IsSelfAdjoint a := by cfc_tac) : @@ -85,50 +85,50 @@ lemma IsSelfAdjoint.neg_algebraMap_norm_le_self {a : A} (ha : IsSelfAdjoint a := exact IsSelfAdjoint.le_algebraMap_norm_self (neg ha) exact neg_le.mp this -lemma CstarRing.mul_star_le_algebraMap_norm_sq {a : A} : a * star a ≤ algebraMap ℝ A (‖a‖ ^ 2) := by +lemma CStarRing.mul_star_le_algebraMap_norm_sq {a : A} : a * star a ≤ algebraMap ℝ A (‖a‖ ^ 2) := by have : a * star a ≤ algebraMap ℝ A ‖a * star a‖ := IsSelfAdjoint.le_algebraMap_norm_self - rwa [CstarRing.norm_self_mul_star, ← pow_two] at this + rwa [CStarRing.norm_self_mul_star, ← pow_two] at this -lemma CstarRing.star_mul_le_algebraMap_norm_sq {a : A} : star a * a ≤ algebraMap ℝ A (‖a‖ ^ 2) := by +lemma CStarRing.star_mul_le_algebraMap_norm_sq {a : A} : star a * a ≤ algebraMap ℝ A (‖a‖ ^ 2) := by have : star a * a ≤ algebraMap ℝ A ‖star a * a‖ := IsSelfAdjoint.le_algebraMap_norm_self - rwa [CstarRing.norm_star_mul_self, ← pow_two] at this + rwa [CStarRing.norm_star_mul_self, ← pow_two] at this lemma IsSelfAdjoint.toReal_spectralRadius_eq_norm {a : A} (ha : IsSelfAdjoint a) : (spectralRadius ℝ a).toReal = ‖a‖ := by simp [ha.spectrumRestricts.spectralRadius_eq, ha.spectralRadius_eq_nnnorm] -lemma CstarRing.norm_or_neg_norm_mem_spectrum [Nontrivial A] {a : A} +lemma CStarRing.norm_or_neg_norm_mem_spectrum [Nontrivial A] {a : A} (ha : IsSelfAdjoint a := by cfc_tac) : ‖a‖ ∈ spectrum ℝ a ∨ -‖a‖ ∈ spectrum ℝ a := by have ha' : SpectrumRestricts a Complex.reCLM := ha.spectrumRestricts rw [← ha.toReal_spectralRadius_eq_norm] exact Real.spectralRadius_mem_spectrum_or (ha'.image ▸ (spectrum.nonempty a).image _) -lemma CstarRing.nnnorm_mem_spectrum_of_nonneg [Nontrivial A] {a : A} (ha : 0 ≤ a := by cfc_tac) : +lemma CStarRing.nnnorm_mem_spectrum_of_nonneg [Nontrivial A] {a : A} (ha : 0 ≤ a := by cfc_tac) : ‖a‖₊ ∈ spectrum ℝ≥0 a := by have : IsSelfAdjoint a := .of_nonneg ha convert NNReal.spectralRadius_mem_spectrum (a := a) ?_ (.nnreal_of_nonneg ha) · simp [this.spectrumRestricts.spectralRadius_eq, this.spectralRadius_eq_nnnorm] · exact this.spectrumRestricts.image ▸ (spectrum.nonempty a).image _ -lemma CstarRing.norm_mem_spectrum_of_nonneg [Nontrivial A] {a : A} (ha : 0 ≤ a := by cfc_tac) : +lemma CStarRing.norm_mem_spectrum_of_nonneg [Nontrivial A] {a : A} (ha : 0 ≤ a := by cfc_tac) : ‖a‖ ∈ spectrum ℝ a := by - simpa using spectrum.algebraMap_mem ℝ <| CstarRing.nnnorm_mem_spectrum_of_nonneg ha + simpa using spectrum.algebraMap_mem ℝ <| CStarRing.nnnorm_mem_spectrum_of_nonneg ha -end Cstar_unital +end CStar_unital -section Cstar_nonunital +section CStar_nonunital variable {A : Type*} [NonUnitalNormedRing A] [CompleteSpace A] [PartialOrder A] [StarRing A] - [StarOrderedRing A] [CstarRing A] [NormedSpace ℂ A] [StarModule ℂ A] + [StarOrderedRing A] [CStarRing A] [NormedSpace ℂ A] [StarModule ℂ A] [SMulCommClass ℂ A A] [IsScalarTower ℂ A A] open ComplexOrder in -instance CstarRing.instNonnegSpectrumClassComplexNonUnital : NonnegSpectrumClass ℂ A where +instance CStarRing.instNonnegSpectrumClassComplexNonUnital : NonnegSpectrumClass ℂ A where quasispectrum_nonneg_of_nonneg a ha x hx := by rw [Unitization.quasispectrum_eq_spectrum_inr' ℂ ℂ a] at hx exact spectrum_nonneg_of_nonneg (Unitization.inr_nonneg_iff.mpr ha) hx -lemma CstarRing.norm_le_norm_of_nonneg_of_le {a b : A} (ha : 0 ≤ a := by cfc_tac) (hab : a ≤ b) : +lemma CStarRing.norm_le_norm_of_nonneg_of_le {a b : A} (ha : 0 ≤ a := by cfc_tac) (hab : a ≤ b) : ‖a‖ ≤ ‖b‖ := by suffices ∀ a b : Unitization ℂ A, 0 ≤ a → a ≤ b → ‖a‖ ≤ ‖b‖ by have hb := ha.trans hab @@ -148,7 +148,7 @@ lemma CstarRing.norm_le_norm_of_nonneg_of_le {a b : A} (ha : 0 ≤ a := by cfc_t rw [cfc_le_iff id (fun _ => ‖b‖) a] at h₂ exact h₂ ‖a‖ <| norm_mem_spectrum_of_nonneg ha -lemma CstarRing.conjugate_le_norm_smul {a b : A} (hb : IsSelfAdjoint b := by cfc_tac) : +lemma CStarRing.conjugate_le_norm_smul {a b : A} (hb : IsSelfAdjoint b := by cfc_tac) : star a * b * a ≤ ‖b‖ • (star a * a) := by suffices ∀ a b : Unitization ℂ A, IsSelfAdjoint b → star a * b * a ≤ ‖b‖ • (star a * a) by rw [← Unitization.inr_le_iff _ _ (by aesop) ((IsSelfAdjoint.all _).smul (.star_mul_self a))] @@ -159,11 +159,11 @@ lemma CstarRing.conjugate_le_norm_smul {a b : A} (hb : IsSelfAdjoint b := by cfc conjugate_le_conjugate hb.le_algebraMap_norm_self _ _ = ‖b‖ • (star a * a) := by simp [Algebra.algebraMap_eq_smul_one] -lemma CstarRing.conjugate_le_norm_smul' {a b : A} (hb : IsSelfAdjoint b := by cfc_tac) : +lemma CStarRing.conjugate_le_norm_smul' {a b : A} (hb : IsSelfAdjoint b := by cfc_tac) : a * b * star a ≤ ‖b‖ • (a * star a) := by have h₁ : a * b * star a = star (star a) * b * star a := by simp have h₂ : a * star a = star (star a) * star a := by simp simp only [h₁, h₂] exact conjugate_le_norm_smul -end Cstar_nonunital +end CStar_nonunital diff --git a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Restrict.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean similarity index 99% rename from Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Restrict.lean rename to Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean index 304e1476b9402..6bd6ffb31d9d7 100644 --- a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Restrict.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Mathlib.Analysis.Normed.Algebra.Spectrum -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.NonUnital +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital /-! # Restriction of the continuous functional calculus to a scalar subring diff --git a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unique.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean similarity index 99% rename from Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unique.lean rename to Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean index 4b6726cf226cb..49ccc1221a710 100644 --- a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unique.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Mathlib.Analysis.Normed.Algebra.Spectrum -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.NonUnital +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital import Mathlib.Topology.ContinuousFunction.StoneWeierstrass /-! diff --git a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean similarity index 99% rename from Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unital.lean rename to Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index a3b48d80d65f3..a824595d91b6c 100644 --- a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -172,7 +172,7 @@ identity to `a`. This is the necessary tool used to establish `cfcHom_comp` and variant `cfc_comp`. This class has instances, which can be found in -`Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unique`, in each of the common cases +`Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unique`, in each of the common cases `ℂ`, `ℝ` and `ℝ≥0` as a consequence of the Stone-Weierstrass theorem. This class is separate from `ContinuousFunctionalCalculus` primarily because we will later use @@ -270,7 +270,7 @@ not continuous on the spectrum of `a`, then `cfc f a` returns the junk value `0` This is the primary declaration intended for widespread use of the continuous functional calculus, and all the API applies to this declaration. For more information, see the module documentation -for `Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital`. -/ +for `Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital`. -/ noncomputable irreducible_def cfc (f : R → R) (a : A) : A := if h : p a ∧ ContinuousOn f (spectrum R a) then cfcHom h.1 ⟨_, h.2.restrict⟩ diff --git a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unitary.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unitary.lean similarity index 97% rename from Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unitary.lean rename to Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unitary.lean index 6e99b1e287921..bda712cdab3b1 100644 --- a/Mathlib/Analysis/CstarAlgebra/ContinuousFunctionalCalculus/Unitary.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unitary.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Mathlib.Tactic.Peel -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital import Mathlib.Analysis.Complex.Basic /-! # Conditions on unitary elements imposed by the continuous functional calculus diff --git a/Mathlib/Analysis/CstarAlgebra/Exponential.lean b/Mathlib/Analysis/CStarAlgebra/Exponential.lean similarity index 100% rename from Mathlib/Analysis/CstarAlgebra/Exponential.lean rename to Mathlib/Analysis/CStarAlgebra/Exponential.lean diff --git a/Mathlib/Analysis/CstarAlgebra/GelfandDuality.lean b/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean similarity index 97% rename from Mathlib/Analysis/CstarAlgebra/GelfandDuality.lean rename to Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean index 0db306dbcac02..c3e2791a52a09 100644 --- a/Mathlib/Analysis/CstarAlgebra/GelfandDuality.lean +++ b/Mathlib/Analysis/CStarAlgebra/GelfandDuality.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.CstarAlgebra.Spectrum +import Mathlib.Analysis.CStarAlgebra.Spectrum import Mathlib.Analysis.Normed.Group.Quotient import Mathlib.Analysis.Normed.Algebra.Basic import Mathlib.Topology.ContinuousFunction.Units @@ -122,10 +122,10 @@ instance [Nontrivial A] : Nonempty (characterSpace ℂ A) := end ComplexBanachAlgebra -section ComplexCstarAlgebra +section ComplexCStarAlgebra variable {A : Type*} [NormedCommRing A] [NormedAlgebra ℂ A] [CompleteSpace A] -variable [StarRing A] [CstarRing A] [StarModule ℂ A] +variable [StarRing A] [CStarRing A] [StarModule ℂ A] theorem gelfandTransform_map_star (a : A) : gelfandTransform ℂ A (star a) = star (gelfandTransform ℂ A a) := @@ -145,7 +145,7 @@ theorem gelfandTransform_isometry : Isometry (gelfandTransform ℂ A) := by unfold spectralRadius; rw [spectrum.gelfandTransform_eq] rw [map_mul, (IsSelfAdjoint.star_mul_self a).spectralRadius_eq_nnnorm, gelfandTransform_map_star, (IsSelfAdjoint.star_mul_self (gelfandTransform ℂ A a)).spectralRadius_eq_nnnorm] at this - simp only [ENNReal.coe_inj, CstarRing.nnnorm_star_mul_self, ← sq] at this + simp only [ENNReal.coe_inj, CStarRing.nnnorm_star_mul_self, ← sq] at this simpa only [Function.comp_apply, NNReal.sqrt_sq] using congr_arg (((↑) : ℝ≥0 → ℝ) ∘ ⇑NNReal.sqrt) this @@ -189,7 +189,7 @@ noncomputable def gelfandStarTransform : A ≃⋆ₐ[ℂ] C(characterSpace ℂ A { gelfandTransform ℂ A with map_star' := fun x => gelfandTransform_map_star x }) (gelfandTransform_bijective A) -end ComplexCstarAlgebra +end ComplexCStarAlgebra section Functoriality @@ -259,8 +259,8 @@ B --- η B ---> C(characterSpace ℂ B, ℂ) ``` -/ theorem gelfandStarTransform_naturality {A B : Type*} [NormedCommRing A] [NormedAlgebra ℂ A] - [CompleteSpace A] [StarRing A] [CstarRing A] [StarModule ℂ A] [NormedCommRing B] - [NormedAlgebra ℂ B] [CompleteSpace B] [StarRing B] [CstarRing B] [StarModule ℂ B] + [CompleteSpace A] [StarRing A] [CStarRing A] [StarModule ℂ A] [NormedCommRing B] + [NormedAlgebra ℂ B] [CompleteSpace B] [StarRing B] [CStarRing B] [StarModule ℂ B] (φ : A →⋆ₐ[ℂ] B) : (gelfandStarTransform B : _ →⋆ₐ[ℂ] _).comp φ = (compContinuousMap φ |>.compStarAlgHom' ℂ ℂ).comp (gelfandStarTransform A : _ →⋆ₐ[ℂ] _) := by diff --git a/Mathlib/Analysis/CstarAlgebra/Matrix.lean b/Mathlib/Analysis/CStarAlgebra/Matrix.lean similarity index 98% rename from Mathlib/Analysis/CstarAlgebra/Matrix.lean rename to Mathlib/Analysis/CStarAlgebra/Matrix.lean index c98c82ec4082f..2a7915775d319 100644 --- a/Mathlib/Analysis/CstarAlgebra/Matrix.lean +++ b/Mathlib/Analysis/CStarAlgebra/Matrix.lean @@ -18,7 +18,7 @@ This transports the operator norm on `EuclideanSpace 𝕜 n →L[𝕜] Euclidean ## Main definitions * `Matrix.instNormedRingL2Op`: the (necessarily unique) normed ring structure on `Matrix n n 𝕜` - which ensure it is a `CstarRing` in `Matrix.instCstarRing`. This is a scoped instance in the + which ensure it is a `CStarRing` in `Matrix.instCStarRing`. This is a scoped instance in the namespace `Matrix.L2OpNorm` in order to avoid choosing a global norm for `Matrix`. ## Main statements @@ -267,10 +267,10 @@ scoped[Matrix.L2OpNorm] attribute [instance] Matrix.instL2OpNormedAlgebra /-- The operator norm on `Matrix n n 𝕜` given by the identification with (continuous) linear endmorphisms of `EuclideanSpace 𝕜 n` makes it into a `L2OpRing`. -/ -lemma instCstarRing : CstarRing (Matrix n n 𝕜) where +lemma instCStarRing : CStarRing (Matrix n n 𝕜) where norm_mul_self_le M := le_of_eq <| Eq.symm <| l2_opNorm_conjTranspose_mul_self M -scoped[Matrix.L2OpNorm] attribute [instance] Matrix.instCstarRing +scoped[Matrix.L2OpNorm] attribute [instance] Matrix.instCStarRing end Matrix diff --git a/Mathlib/Analysis/CstarAlgebra/Module.lean b/Mathlib/Analysis/CStarAlgebra/Module.lean similarity index 87% rename from Mathlib/Analysis/CstarAlgebra/Module.lean rename to Mathlib/Analysis/CStarAlgebra/Module.lean index 8244b2422d6d3..d9abc7498c893 100644 --- a/Mathlib/Analysis/CstarAlgebra/Module.lean +++ b/Mathlib/Analysis/CStarAlgebra/Module.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Order +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Order /-! # Hilbert C⋆-modules @@ -16,19 +16,19 @@ vector space. ## Main declarations -+ `CstarModule`: The class containing the Hilbert C⋆-module structure -+ `CstarModule.normedSpaceCore`: The proof that a Hilbert C⋆-module is a normed vector ++ `CStarModule`: The class containing the Hilbert C⋆-module structure ++ `CStarModule.normedSpaceCore`: The proof that a Hilbert C⋆-module is a normed vector space. This can be used with `NormedAddCommGroup.ofCore` and `NormedSpace.ofCore` to create the relevant instances on a type of interest. -+ `CstarModule.inner_mul_inner_swap_le`: The statement that ++ `CStarModule.inner_mul_inner_swap_le`: The statement that `⟪y, x⟫ * ⟪x, y⟫ ≤ ‖x‖ ^ 2 • ⟪y, y⟫`, which can be viewed as a version of the Cauchy-Schwarz inequality for Hilbert C⋆-modules. -+ `CstarModule.norm_inner_le`, which states that `‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖`, i.e. the ++ `CStarModule.norm_inner_le`, which states that `‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖`, i.e. the Cauchy-Schwarz inequality. ## Implementation notes -The class `CstarModule A E` requires `E` to already have a `Norm E` instance on it, but +The class `CStarModule A E` requires `E` to already have a `Norm E` instance on it, but no other norm-related instances. We then include the fact that this norm agrees with the norm induced by the inner product among the axioms of the class. Furthermore, instead of registering `NormedAddCommGroup E` and `NormedSpace ℂ E` instances (which might already be present on the type, @@ -47,7 +47,7 @@ open scoped ComplexOrder RightActions /-- A *Hilbert C⋆-module* is a complex module `E` endowed with a right `A`-module structure (where `A` is typically a C⋆-algebra) and an inner product `⟪x, y⟫_A` which satisfies the following properties. -/ -class CstarModule (A : outParam <| Type*) (E : Type*) [NonUnitalSemiring A] [StarRing A] +class CStarModule (A : outParam <| Type*) (E : Type*) [NonUnitalSemiring A] [StarRing A] [Module ℂ A] [AddCommGroup E] [Module ℂ E] [PartialOrder A] [SMul Aᵐᵒᵖ E] [Norm A] [Norm E] extends Inner A E where inner_add_right {x} {y} {z} : inner x (y + z) = inner x y + inner x z @@ -58,16 +58,18 @@ class CstarModule (A : outParam <| Type*) (E : Type*) [NonUnitalSemiring A] [Sta star_inner x y : star (inner x y) = inner y x norm_eq_sqrt_norm_inner_self x : ‖x‖ = √‖inner x x‖ -attribute [simp] CstarModule.inner_add_right CstarModule.star_inner - CstarModule.inner_op_smul_right CstarModule.inner_smul_right_complex +attribute [simp] CStarModule.inner_add_right CStarModule.star_inner + CStarModule.inner_op_smul_right CStarModule.inner_smul_right_complex -namespace CstarModule +@[deprecated (since := "2024-08-04")] alias CstarModule := CStarModule + +namespace CStarModule section general variable {A E : Type*} [NonUnitalRing A] [StarRing A] [AddCommGroup E] [Module ℂ A] [Module ℂ E] [PartialOrder A] [SMul Aᵐᵒᵖ E] [StarModule ℂ A] [Norm A] [Norm E] - [CstarModule A E] + [CStarModule A E] local notation "⟪" x ", " y "⟫" => inner (𝕜 := A) x y @@ -133,9 +135,9 @@ end general section norm -variable {A E : Type*} [NonUnitalNormedRing A] [StarRing A] [CstarRing A] [PartialOrder A] +variable {A E : Type*} [NonUnitalNormedRing A] [StarRing A] [CStarRing A] [PartialOrder A] [StarOrderedRing A] [AddCommGroup E] [NormedSpace ℂ A] [Module ℂ E] [SMul Aᵐᵒᵖ E] [Norm E] - [StarModule ℂ A] [CstarModule A E] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] + [StarModule ℂ A] [CStarModule A E] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] local notation "⟪" x ", " y "⟫" => inner (𝕜 := A) x y @@ -151,7 +153,7 @@ protected lemma norm_zero : ‖(0 : E)‖ = 0 := by simp [norm_eq_sqrt_norm_inne lemma norm_zero_iff (x : E) : ‖x‖ = 0 ↔ x = 0 := ⟨fun h => by simpa [norm_eq_sqrt_norm_inner_self, inner_self] using h, - fun h => by simp [norm, h]; rw [CstarModule.norm_zero] ⟩ + fun h => by simp [norm, h]; rw [CStarModule.norm_zero] ⟩ protected lemma norm_nonneg {x : E} : 0 ≤ ‖x‖ := by simp [norm_eq_sqrt_norm_inner_self]; positivity @@ -166,7 +168,7 @@ lemma norm_sq_eq {x : E} : ‖x‖ ^ 2 = ‖⟪x, x⟫‖ := by simp [norm_eq_sq /-- The C⋆-algebra-valued Cauchy-Schwarz inequality for Hilbert C⋆-modules. -/ lemma inner_mul_inner_swap_le [CompleteSpace A] {x y : E} : ⟪y, x⟫ * ⟪x, y⟫ ≤ ‖x‖ ^ 2 • ⟪y, y⟫ := by rcases eq_or_ne x 0 with h|h - · simp [h, CstarModule.norm_zero (E := E)] + · simp [h, CStarModule.norm_zero (E := E)] · have h₁ : ∀ (a : A), (0 : A) ≤ ‖x‖ ^ 2 • (star a * a) - ‖x‖ ^ 2 • (⟪y, x⟫ * a) - ‖x‖ ^ 2 • (star a * ⟪x, y⟫) + ‖x‖ ^ 2 • (‖x‖ ^ 2 • ⟪y, y⟫) := fun a => by @@ -181,7 +183,7 @@ lemma inner_mul_inner_swap_le [CompleteSpace A] {x y : E} : ⟪y, x⟫ * ⟪x, y _ ≤ ‖x‖ ^ 2 • (star a * a) - ‖x‖ ^ 2 • (⟪y, x⟫ * a) - ‖x‖ ^ 2 • (star a * ⟪x, y⟫) + ‖x‖ ^ 2 • (‖x‖ ^ 2 • ⟪y, y⟫) := by gcongr - calc _ ≤ ‖⟪x, x⟫_A‖ • (star a * a) := CstarRing.conjugate_le_norm_smul + calc _ ≤ ‖⟪x, x⟫_A‖ • (star a * a) := CStarRing.conjugate_le_norm_smul _ = (Real.sqrt ‖⟪x, x⟫_A‖) ^ 2 • (star a * a) := by congr have : 0 ≤ ‖⟪x, x⟫_A‖ := by positivity @@ -189,15 +191,15 @@ lemma inner_mul_inner_swap_le [CompleteSpace A] {x y : E} : ⟪y, x⟫ * ⟪x, y _ = ‖x‖ ^ 2 • (star a * a) := by rw [← norm_eq_sqrt_norm_inner_self] specialize h₁ ⟪x, y⟫ simp only [star_inner, sub_self, zero_sub, le_neg_add_iff_add_le, add_zero] at h₁ - rwa [smul_le_smul_iff_of_pos_left (pow_pos (CstarModule.norm_pos h) _)] at h₁ + rwa [smul_le_smul_iff_of_pos_left (pow_pos (CStarModule.norm_pos h) _)] at h₁ variable (E) in /-- The Cauchy-Schwarz inequality for Hilbert C⋆-modules. -/ lemma norm_inner_le [CompleteSpace A] {x y : E} : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := by have := calc ‖⟪x, y⟫‖ ^ 2 = ‖⟪y, x⟫ * ⟪x, y⟫‖ := by - rw [← star_inner x, CstarRing.norm_star_mul_self, pow_two] + rw [← star_inner x, CStarRing.norm_star_mul_self, pow_two] _ ≤ ‖‖x‖^ 2 • ⟪y, y⟫‖ := by - refine CstarRing.norm_le_norm_of_nonneg_of_le ?_ inner_mul_inner_swap_le + refine CStarRing.norm_le_norm_of_nonneg_of_le ?_ inner_mul_inner_swap_le rw [← star_inner x] exact star_mul_self_nonneg ⟪x, y⟫_A _ = ‖x‖ ^ 2 * ‖⟪y, y⟫‖ := by simp [norm_smul] @@ -205,7 +207,7 @@ lemma norm_inner_le [CompleteSpace A] {x y : E} : ‖⟪x, y⟫‖ ≤ ‖x‖ * simp only [norm_eq_sqrt_norm_inner_self, norm_nonneg, Real.sq_sqrt] _ = (‖x‖ * ‖y‖) ^ 2 := by simp only [mul_pow] refine (pow_le_pow_iff_left (R := ℝ) (norm_nonneg ⟪x, y⟫_A) ?_ (by norm_num)).mp this - exact mul_nonneg CstarModule.norm_nonneg CstarModule.norm_nonneg + exact mul_nonneg CStarModule.norm_nonneg CStarModule.norm_nonneg protected lemma norm_triangle [CompleteSpace A] (x y : E) : ‖x + y‖ ≤ ‖x‖ + ‖y‖ := by have h : ‖x + y‖ ^ 2 ≤ (‖x‖ + ‖y‖) ^ 2 := by @@ -218,16 +220,16 @@ protected lemma norm_triangle [CompleteSpace A] (x y : E) : ‖x + y‖ ≤ ‖x _ = ‖x‖ ^ 2 + ‖y‖ * ‖x‖ + ‖x‖ * ‖y‖ + ‖y‖ ^ 2 := by simp [norm_eq_sqrt_norm_inner_self] _ = (‖x‖ + ‖y‖) ^ 2 := by simp only [add_pow_two, add_left_inj]; ring - refine (pow_le_pow_iff_left CstarModule.norm_nonneg ?_ (by norm_num)).mp h - exact add_nonneg CstarModule.norm_nonneg CstarModule.norm_nonneg + refine (pow_le_pow_iff_left CStarModule.norm_nonneg ?_ (by norm_num)).mp h + exact add_nonneg CStarModule.norm_nonneg CStarModule.norm_nonneg /-- This allows us to get `NormedAddCommGroup` and `NormedSpace` instances on `E` via `NormedAddCommGroup.ofCore` and `NormedSpace.ofCore`. -/ lemma normedSpaceCore [CompleteSpace A] : NormedSpace.Core ℂ E where - norm_nonneg x := CstarModule.norm_nonneg + norm_nonneg x := CStarModule.norm_nonneg norm_eq_zero_iff x := norm_zero_iff x norm_smul c x := by simp [norm_eq_sqrt_norm_inner_self, norm_smul, ← mul_assoc] - norm_triangle x y := CstarModule.norm_triangle x y + norm_triangle x y := CStarModule.norm_triangle x y lemma norm_eq_csSup [CompleteSpace A] (v : E) : ‖v‖ = sSup { ‖⟪w, v⟫_A‖ | (w : E) (_ : ‖w‖ ≤ 1) } := by @@ -235,7 +237,7 @@ lemma norm_eq_csSup [CompleteSpace A] (v : E) : let instNS : NormedSpace ℂ E := .ofCore normedSpaceCore refine Eq.symm <| IsGreatest.csSup_eq ⟨⟨‖v‖⁻¹ • v, ?_, ?_⟩, ?_⟩ · simpa only [norm_smul, norm_inv, norm_norm] using inv_mul_le_one_of_le le_rfl (by positivity) - · simp [norm_smul, CstarModule.inner_self_eq_norm_sq, pow_two, ← mul_assoc] + · simp [norm_smul, CStarModule.inner_self_eq_norm_sq, pow_two, ← mul_assoc] · rintro - ⟨w, hw, rfl⟩ calc _ ≤ ‖w‖ * ‖v‖ := norm_inner_le E _ ≤ 1 * ‖v‖ := by gcongr @@ -245,9 +247,9 @@ end norm section NormedAddCommGroup -variable {A E : Type*} [NonUnitalNormedRing A] [StarRing A] [CstarRing A] [PartialOrder A] +variable {A E : Type*} [NonUnitalNormedRing A] [StarRing A] [CStarRing A] [PartialOrder A] [StarOrderedRing A] [NormedSpace ℂ A] [SMul Aᵐᵒᵖ E] [CompleteSpace A] - [NormedAddCommGroup E] [NormedSpace ℂ E] [StarModule ℂ A] [CstarModule A E] [IsScalarTower ℂ A A] + [NormedAddCommGroup E] [NormedSpace ℂ E] [StarModule ℂ A] [CStarModule A E] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] /-- The function `⟨x, y⟩ ↦ ⟪x, y⟫` bundled as a continuous sesquilinear map. -/ @@ -264,4 +266,4 @@ lemma continuous_inner : Continuous (fun x : E × E => ⟪x.1, x.2⟫_A) := by end NormedAddCommGroup -end CstarModule +end CStarModule diff --git a/Mathlib/Analysis/CstarAlgebra/Multiplier.lean b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean similarity index 98% rename from Mathlib/Analysis/CstarAlgebra/Multiplier.lean rename to Mathlib/Analysis/CStarAlgebra/Multiplier.lean index bc423401f3149..06bd224d0b51a 100644 --- a/Mathlib/Analysis/CstarAlgebra/Multiplier.lean +++ b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux, Jon Bannon -/ import Mathlib.Analysis.NormedSpace.OperatorNorm.Completeness -import Mathlib.Analysis.CstarAlgebra.Unitization +import Mathlib.Analysis.CStarAlgebra.Unitization import Mathlib.Analysis.SpecialFunctions.Pow.NNReal /-! @@ -547,7 +547,7 @@ instance [CompleteSpace A] : CompleteSpace 𝓜(𝕜, A) := by continuous_const exact continuous_const.mul ((ContinuousLinearMap.apply 𝕜 A _).continuous.comp continuous_fst) -variable [StarRing A] [CstarRing A] +variable [StarRing A] [CStarRing A] /-- For `a : 𝓜(𝕜, A)`, the norms of `a.fst` and `a.snd` coincide, and hence these also coincide with `‖a‖` which is `max (‖a.fst‖) (‖a.snd‖)`. -/ @@ -569,7 +569,7 @@ theorem norm_fst_eq_snd (a : 𝓜(𝕜, A)) : ‖a.fst‖ = ‖a.snd‖ := by intro b calc ‖a.fst b‖₊ ^ 2 = ‖star (a.fst b) * a.fst b‖₊ := by - simpa only [← sq] using CstarRing.nnnorm_star_mul_self.symm + simpa only [← sq] using CStarRing.nnnorm_star_mul_self.symm _ ≤ ‖a.snd (star (a.fst b))‖₊ * ‖b‖₊ := (a.central (star (a.fst b)) b ▸ nnnorm_mul_le _ _) _ ≤ ‖a.snd‖₊ * ‖a.fst b‖₊ * ‖b‖₊ := nnnorm_star (a.fst b) ▸ mul_le_mul_right' (a.snd.le_opNNNorm _) _ @@ -578,7 +578,7 @@ theorem norm_fst_eq_snd (a : 𝓜(𝕜, A)) : ‖a.fst‖ = ‖a.snd‖ := by intro b calc ‖a.snd b‖₊ ^ 2 = ‖a.snd b * star (a.snd b)‖₊ := by - simpa only [← sq] using CstarRing.nnnorm_self_mul_star.symm + simpa only [← sq] using CStarRing.nnnorm_self_mul_star.symm _ ≤ ‖b‖₊ * ‖a.fst (star (a.snd b))‖₊ := ((a.central b (star (a.snd b))).symm ▸ nnnorm_mul_le _ _) _ = ‖a.fst (star (a.snd b))‖₊ * ‖b‖₊ := mul_comm _ _ @@ -610,10 +610,10 @@ end NontriviallyNormed section DenselyNormed variable {𝕜 A : Type*} [DenselyNormedField 𝕜] [StarRing 𝕜] -variable [NonUnitalNormedRing A] [StarRing A] [CstarRing A] +variable [NonUnitalNormedRing A] [StarRing A] [CStarRing A] variable [NormedSpace 𝕜 A] [SMulCommClass 𝕜 A A] [IsScalarTower 𝕜 A A] [StarModule 𝕜 A] -instance instCstarRing : CstarRing 𝓜(𝕜, A) where +instance instCStarRing : CStarRing 𝓜(𝕜, A) where norm_mul_self_le := fun (a : 𝓜(𝕜, A)) => le_of_eq <| Eq.symm <| congr_arg ((↑) : ℝ≥0 → ℝ) <| show ‖star a * a‖₊ = ‖a‖₊ * ‖a‖₊ by /- The essence of the argument is this: let `a = (L,R)` and recall `‖a‖ = ‖L‖`. @@ -658,7 +658,7 @@ instance instCstarRing : CstarRing 𝓜(𝕜, A) where · refine ⟨‖a‖₊ * ‖a‖₊, ?_⟩ rintro - ⟨y, hy, rfl⟩ exact key (star x) y ((nnnorm_star x).trans_le hx') (mem_closedBall_zero_iff.1 hy) - · simpa only [a.central, star_star, CstarRing.nnnorm_star_mul_self, NNReal.sq_sqrt, ← sq] + · simpa only [a.central, star_star, CStarRing.nnnorm_star_mul_self, NNReal.sq_sqrt, ← sq] using pow_lt_pow_left hxr zero_le' two_ne_zero end DenselyNormed diff --git a/Mathlib/Analysis/CstarAlgebra/Spectrum.lean b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean similarity index 94% rename from Mathlib/Analysis/CstarAlgebra/Spectrum.lean rename to Mathlib/Analysis/CStarAlgebra/Spectrum.lean index 0b1a907133349..924eaed27bea7 100644 --- a/Mathlib/Analysis/CstarAlgebra/Spectrum.lean +++ b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.CstarAlgebra.Basic +import Mathlib.Analysis.CStarAlgebra.Basic import Mathlib.Analysis.Normed.Algebra.Spectrum import Mathlib.Analysis.SpecialFunctions.Exponential import Mathlib.Algebra.Star.StarAlgHom @@ -19,18 +19,18 @@ section open scoped Topology ENNReal -open Filter ENNReal spectrum CstarRing NormedSpace +open Filter ENNReal spectrum CStarRing NormedSpace section UnitarySpectrum -variable {𝕜 : Type*} [NormedField 𝕜] {E : Type*} [NormedRing E] [StarRing E] [CstarRing E] +variable {𝕜 : Type*} [NormedField 𝕜] {E : Type*} [NormedRing E] [StarRing E] [CStarRing E] [NormedAlgebra 𝕜 E] [CompleteSpace E] theorem unitary.spectrum_subset_circle (u : unitary E) : spectrum 𝕜 (u : E) ⊆ Metric.sphere 0 1 := by nontriviality E refine fun k hk => mem_sphere_zero_iff_norm.mpr (le_antisymm ?_ ?_) - · simpa only [CstarRing.norm_coe_unitary u] using norm_le_norm_of_mem hk + · simpa only [CStarRing.norm_coe_unitary u] using norm_le_norm_of_mem hk · rw [← unitary.val_toUnits_apply u] at hk have hnk := ne_zero_of_mem_of_unit hk rw [← inv_inv (unitary.toUnits u), ← spectrum.map_inv, Set.mem_inv] at hk @@ -49,7 +49,7 @@ section ComplexScalars open Complex variable {A : Type*} [NormedRing A] [NormedAlgebra ℂ A] [CompleteSpace A] [StarRing A] - [CstarRing A] + [CStarRing A] local notation "↑ₐ" => algebraMap ℂ A @@ -122,7 +122,7 @@ end ComplexScalars namespace StarAlgHom variable {F A B : Type*} [NormedRing A] [NormedAlgebra ℂ A] [CompleteSpace A] [StarRing A] - [CstarRing A] [NormedRing B] [NormedAlgebra ℂ B] [CompleteSpace B] [StarRing B] [CstarRing B] + [CStarRing A] [NormedRing B] [NormedAlgebra ℂ B] [CompleteSpace B] [StarRing B] [CStarRing B] [FunLike F A B] [AlgHomClass F ℂ A B] [StarAlgHomClass F ℂ A B] (φ : F) /-- A star algebra homomorphism of complex C⋆-algebras is norm contractive. -/ @@ -153,7 +153,7 @@ end StarAlgHom namespace StarAlgEquiv variable {F A B : Type*} [NormedRing A] [NormedAlgebra ℂ A] [CompleteSpace A] [StarRing A] - [CstarRing A] [NormedRing B] [NormedAlgebra ℂ B] [CompleteSpace B] [StarRing B] [CstarRing B] + [CStarRing A] [NormedRing B] [NormedAlgebra ℂ B] [CompleteSpace B] [StarRing B] [CStarRing B] [EquivLike F A B] [NonUnitalAlgEquivClass F ℂ A B] [StarAlgEquivClass F ℂ A B] lemma nnnorm_map (φ : F) (a : A) : ‖φ a‖₊ = ‖a‖₊ := @@ -177,7 +177,7 @@ open ContinuousMap Complex open scoped ComplexStarModule variable {F A : Type*} [NormedRing A] [NormedAlgebra ℂ A] [CompleteSpace A] [StarRing A] - [CstarRing A] [StarModule ℂ A] [FunLike F A ℂ] [hF : AlgHomClass F ℂ A ℂ] + [CStarRing A] [StarModule ℂ A] [FunLike F A ℂ] [hF : AlgHomClass F ℂ A ℂ] /-- This instance is provided instead of `StarAlgHomClass` to avoid type class inference loops. See note [lower instance priority] -/ diff --git a/Mathlib/Analysis/CstarAlgebra/Unitization.lean b/Mathlib/Analysis/CStarAlgebra/Unitization.lean similarity index 95% rename from Mathlib/Analysis/CstarAlgebra/Unitization.lean rename to Mathlib/Analysis/CStarAlgebra/Unitization.lean index 2b9a2688c3699..85446f1873905 100644 --- a/Mathlib/Analysis/CstarAlgebra/Unitization.lean +++ b/Mathlib/Analysis/CStarAlgebra/Unitization.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.CstarAlgebra.Basic +import Mathlib.Analysis.CStarAlgebra.Basic import Mathlib.Analysis.Normed.Algebra.Unitization /-! # The minimal unitization of a C⋆-algebra @@ -49,12 +49,12 @@ lemma isometry_mul_flip : Isometry (mul 𝕜 E).flip := end ContinuousLinearMap -variable [DenselyNormedField 𝕜] [NonUnitalNormedRing E] [StarRing E] [CstarRing E] +variable [DenselyNormedField 𝕜] [NonUnitalNormedRing E] [StarRing E] [CStarRing E] variable [NormedSpace 𝕜 E] [IsScalarTower 𝕜 E E] [SMulCommClass 𝕜 E E] variable (E) /-- A C⋆-algebra over a densely normed field is a regular normed algebra. -/ -instance CstarRing.instRegularNormedAlgebra : RegularNormedAlgebra 𝕜 E where +instance CStarRing.instRegularNormedAlgebra : RegularNormedAlgebra 𝕜 E where isometry_mul' := AddMonoidHomClass.isometry_of_norm (mul 𝕜 E) fun a => NNReal.eq_iff.mp <| show ‖mul 𝕜 E a‖₊ = ‖a‖₊ by rw [← sSup_closed_unit_ball_eq_nnnorm] @@ -71,17 +71,17 @@ instance CstarRing.instRegularNormedAlgebra : RegularNormedAlgebra 𝕜 E where refine ⟨_, ⟨k • star a, ?_, rfl⟩, ?_⟩ · simpa only [mem_closedBall_zero_iff, norm_smul, one_mul, norm_star] using (NNReal.le_inv_iff_mul_le ha.ne').1 (one_mul ‖a‖₊⁻¹ ▸ hk₂.le : ‖k‖₊ ≤ ‖a‖₊⁻¹) - · simp only [map_smul, nnnorm_smul, mul_apply', mul_smul_comm, CstarRing.nnnorm_self_mul_star] + · simp only [map_smul, nnnorm_smul, mul_apply', mul_smul_comm, CStarRing.nnnorm_self_mul_star] rwa [← NNReal.div_lt_iff (mul_pos ha ha).ne', div_eq_mul_inv, mul_inv, ← mul_assoc] section CStarProperty -variable [StarRing 𝕜] [CstarRing 𝕜] [StarModule 𝕜 E] +variable [StarRing 𝕜] [CStarRing 𝕜] [StarModule 𝕜 E] variable {E} -/-- This is the key lemma used to establish the instance `Unitization.instCstarRing` +/-- This is the key lemma used to establish the instance `Unitization.instCStarRing` (i.e., proving that the norm on `Unitization 𝕜 E` satisfies the C⋆-property). We split this one -out so that declaring the `CstarRing` instance doesn't time out. -/ +out so that declaring the `CStarRing` instance doesn't time out. -/ theorem Unitization.norm_splitMul_snd_sq (x : Unitization 𝕜 E) : ‖(Unitization.splitMul 𝕜 E x).snd‖ ^ 2 ≤ ‖(Unitization.splitMul 𝕜 E (star x * x)).snd‖ := by /- The key idea is that we can use `sSup_closed_unit_ball_eq_norm` to make this about @@ -95,7 +95,7 @@ theorem Unitization.norm_splitMul_snd_sq (x : Unitization 𝕜 E) : simp only -- rewrite to a more convenient form; this is where we use the C⋆-property rw [← Real.sqrt_sq (norm_nonneg _), Real.sqrt_le_sqrt_iff (norm_nonneg _), sq, - ← CstarRing.norm_star_mul_self, ContinuousLinearMap.add_apply, star_add, mul_apply', + ← CStarRing.norm_star_mul_self, ContinuousLinearMap.add_apply, star_add, mul_apply', Algebra.algebraMap_eq_smul_one, ContinuousLinearMap.smul_apply, ContinuousLinearMap.one_apply, star_mul, star_smul, add_mul, smul_mul_assoc, ← mul_smul_comm, mul_assoc, ← mul_add, ← sSup_closed_unit_ball_eq_norm] @@ -124,7 +124,7 @@ theorem Unitization.norm_splitMul_snd_sq (x : Unitization 𝕜 E) : variable {𝕜} /-- The norm on `Unitization 𝕜 E` satisfies the C⋆-property -/ -instance Unitization.instCstarRing : CstarRing (Unitization 𝕜 E) where +instance Unitization.instCStarRing : CStarRing (Unitization 𝕜 E) where norm_mul_self_le x := by -- rewrite both sides as a `⊔` simp only [Unitization.norm_def, Prod.norm_def, ← sup_eq_max] diff --git a/Mathlib/Analysis/InnerProductSpace/Adjoint.lean b/Mathlib/Analysis/InnerProductSpace/Adjoint.lean index ca7a178977447..c4a1bcd72f1b5 100644 --- a/Mathlib/Analysis/InnerProductSpace/Adjoint.lean +++ b/Mathlib/Analysis/InnerProductSpace/Adjoint.lean @@ -215,7 +215,7 @@ theorem norm_adjoint_comp_self (A : E →L[𝕜] F) : simp_rw [mul_assoc, Real.sqrt_mul (norm_nonneg _) (‖x‖ * ‖x‖), Real.sqrt_mul_self (norm_nonneg x)] -instance : CstarRing (E →L[𝕜] E) where +instance : CStarRing (E →L[𝕜] E) where norm_mul_self_le x := le_of_eq <| Eq.symm <| norm_adjoint_comp_self x theorem isAdjointPair_inner (A : E →L[𝕜] F) : diff --git a/Mathlib/Analysis/Matrix.lean b/Mathlib/Analysis/Matrix.lean index 56d0adc4ea7c3..af2323fa39d72 100644 --- a/Mathlib/Analysis/Matrix.lean +++ b/Mathlib/Analysis/Matrix.lean @@ -44,7 +44,7 @@ of a matrix. The norm induced by the identification of `Matrix m n 𝕜` with `EuclideanSpace n 𝕜 →L[𝕜] EuclideanSpace m 𝕜` (i.e., the ℓ² operator norm) can be found in -`Analysis.CstarAlgebra.Matrix`. It is separated to avoid extraneous imports in this file. +`Analysis.CStarAlgebra.Matrix`. It is separated to avoid extraneous imports in this file. -/ noncomputable section diff --git a/Mathlib/Analysis/Normed/Lp/lpSpace.lean b/Mathlib/Analysis/Normed/Lp/lpSpace.lean index 58943db12217a..636faeded0845 100644 --- a/Mathlib/Analysis/Normed/Lp/lpSpace.lean +++ b/Mathlib/Analysis/Normed/Lp/lpSpace.lean @@ -32,7 +32,7 @@ The space `lp E p` is the subtype of elements of `∀ i : α, E i` which satisfy a type synonym `PreLp` for `∀ i : α, E i`, and equipped with a `NormedAddCommGroup` structure. Under appropriate conditions, this is also equipped with the instances `lp.normedSpace`, `lp.completeSpace`. For `p=∞`, there is also `lp.inftyNormedRing`, - `lp.inftyNormedAlgebra`, `lp.inftyStarRing` and `lp.inftyCstarRing`. + `lp.inftyNormedAlgebra`, `lp.inftyStarRing` and `lp.inftyCStarRing`. ## Main results @@ -754,11 +754,11 @@ instance inftyStarRing : StarRing (lp B ∞) := { lp.instStarAddMonoid with star_mul := fun _f _g => ext <| star_mul (R := ∀ i, B i) _ _ } -instance inftyCstarRing [∀ i, CstarRing (B i)] : CstarRing (lp B ∞) where +instance inftyCStarRing [∀ i, CStarRing (B i)] : CStarRing (lp B ∞) where norm_mul_self_le f := by rw [← sq, ← Real.le_sqrt (norm_nonneg _) (norm_nonneg _)] refine lp.norm_le_of_forall_le ‖star f * f‖.sqrt_nonneg fun i => ?_ - rw [Real.le_sqrt (norm_nonneg _) (norm_nonneg _), sq, ← CstarRing.norm_star_mul_self] + rw [Real.le_sqrt (norm_nonneg _) (norm_nonneg _), sq, ← CStarRing.norm_star_mul_self] exact lp.norm_apply_le_norm ENNReal.top_ne_zero (star f * f) i end StarRing diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean index 1886ea8a03f89..9eae72e5fb45b 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean @@ -116,7 +116,7 @@ theorem opNorm_mulLeftRight_le : /-- This is a mixin class for non-unital normed algebras which states that the left-regular representation of the algebra on itself is isometric. Every unital normed algebra with `‖1‖ = 1` is a regular normed algebra (see `NormedAlgebra.instRegularNormedAlgebra`). In addition, so is every -C⋆-algebra, non-unital included (see `CstarRing.instRegularNormedAlgebra`), but there are yet other +C⋆-algebra, non-unital included (see `CStarRing.instRegularNormedAlgebra`), but there are yet other examples. Any algebra with an approximate identity (e.g., $$L^1$$) is also regular. This is a useful class because it gives rise to a nice norm on the unitization; in particular it is diff --git a/Mathlib/Analysis/Quaternion.lean b/Mathlib/Analysis/Quaternion.lean index e41480ad1e545..0991d3a5949c9 100644 --- a/Mathlib/Analysis/Quaternion.lean +++ b/Mathlib/Analysis/Quaternion.lean @@ -90,7 +90,7 @@ noncomputable instance : NormedAlgebra ℝ ℍ where norm_smul_le := norm_smul_le toAlgebra := Quaternion.algebra -instance : CstarRing ℍ where +instance : CStarRing ℍ where norm_mul_self_le x := le_of_eq <| Eq.symm <| (norm_mul _ _).trans <| congr_arg (· * ‖x‖) (norm_star x) diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index d7d6e7c20ec0f..55fb23a8c9e54 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ import Mathlib.Algebra.Star.Order -import Mathlib.Analysis.CstarAlgebra.Basic +import Mathlib.Analysis.CStarAlgebra.Basic import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap import Mathlib.Analysis.Normed.Module.Basic import Mathlib.Data.Real.Sqrt @@ -506,7 +506,7 @@ theorem normSq_div (z w : K) : normSq (z / w) = normSq z / normSq w := @[rclike_simps] -- porting note (#10618): was `simp` theorem norm_conj {z : K} : ‖conj z‖ = ‖z‖ := by simp only [← sqrt_normSq_eq_norm, normSq_conj] -instance (priority := 100) : CstarRing K where +instance (priority := 100) : CStarRing K where norm_mul_self_le x := le_of_eq <| ((norm_mul _ _).trans <| congr_arg (· * ‖x‖) norm_conj).symm /-! ### Cast lemmas -/ diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean index 46d637bae1dbd..e781a032ee7fc 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean @@ -6,7 +6,7 @@ Authors: Frédéric Dupuis import Mathlib.Analysis.Normed.Algebra.Spectrum import Mathlib.Analysis.SpecialFunctions.Exponential -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital /-! # The exponential and logarithm based on the continuous functional calculus diff --git a/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean b/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean index 28d385c769837..6486b7bbab466 100644 --- a/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean +++ b/Mathlib/Analysis/VonNeumannAlgebra/Basic.lean @@ -39,7 +39,7 @@ One the other hand, not picking one means that the weak-* topology (which depends on a choice of predual) must be defined using the choice, and we may be unhappy with the resulting opaqueness of the definition. -/ -class WStarAlgebra (M : Type u) [NormedRing M] [StarRing M] [CstarRing M] [Module ℂ M] +class WStarAlgebra (M : Type u) [NormedRing M] [StarRing M] [CStarRing M] [Module ℂ M] [NormedAlgebra ℂ M] [StarModule ℂ M] : Prop where /-- There is a Banach space `X` whose dual is isometrically (conjugate-linearly) isomorphic to the `WStarAlgebra`. -/ diff --git a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean index 4001c12154234..aedad97e0580a 100644 --- a/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean +++ b/Mathlib/LinearAlgebra/Matrix/HermitianFunctionalCalculus.lean @@ -6,7 +6,7 @@ Authors: Jon Bannon, Jireh Loreaux import Mathlib.LinearAlgebra.Matrix.Spectrum import Mathlib.LinearAlgebra.Eigenspace.Matrix -import Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unique +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unique import Mathlib.Topology.ContinuousFunction.Units /-! diff --git a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean index 23e3a7cb7ae7f..d69134ff4820b 100644 --- a/Mathlib/LinearAlgebra/Matrix/Spectrum.lean +++ b/Mathlib/LinearAlgebra/Matrix/Spectrum.lean @@ -7,7 +7,7 @@ import Mathlib.Analysis.InnerProductSpace.Spectrum import Mathlib.Data.Matrix.Rank import Mathlib.LinearAlgebra.Matrix.Diagonal import Mathlib.LinearAlgebra.Matrix.Hermitian -import Mathlib.Analysis.CstarAlgebra.Matrix +import Mathlib.Analysis.CStarAlgebra.Matrix import Mathlib.Topology.Algebra.Module.FiniteDimension /-! # Spectral theory of hermitian matrices diff --git a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean index 6f1eb9b09d61b..bfe6a02c9efa5 100644 --- a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean +++ b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean @@ -125,7 +125,7 @@ theorem exists_continuous_eLpNorm_sub_le_of_closed [μ.OuterRegular] (hp : p ≠ refine (eLpNorm_mono fun x => ?_).trans_lt this by_cases hx : x ∈ v · simp only [hx, abs_of_nonneg (hg_range x).1, (hg_range x).2, Real.norm_eq_abs, - indicator_of_mem, CstarRing.norm_one] + indicator_of_mem, CStarRing.norm_one] · simp only [hgv hx, Pi.zero_apply, Real.norm_eq_abs, abs_zero, abs_nonneg] refine ⟨fun x => g x • c, g.continuous.smul continuous_const, (eLpNorm_mono gc_bd).trans ?_, gc_bd0, diff --git a/Mathlib/Topology/ContinuousFunction/Bounded.lean b/Mathlib/Topology/ContinuousFunction/Bounded.lean index 2a4324e641da9..a76e6a27a8e44 100644 --- a/Mathlib/Topology/ContinuousFunction/Bounded.lean +++ b/Mathlib/Topology/ContinuousFunction/Bounded.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Module.MinimalAxioms import Mathlib.Topology.ContinuousFunction.Algebra import Mathlib.Analysis.Normed.Order.Lattice import Mathlib.Analysis.NormedSpace.OperatorNorm.Basic -import Mathlib.Analysis.CstarAlgebra.Basic +import Mathlib.Analysis.CStarAlgebra.Basic import Mathlib.Analysis.Normed.Operator.ContinuousLinearMap import Mathlib.Topology.Bornology.BoundedOperation @@ -1367,7 +1367,7 @@ instance instStarModule : StarModule 𝕜 (α →ᵇ β) where end NormedAddCommGroup -section CstarRing +section CStarRing variable [TopologicalSpace α] variable [NonUnitalNormedRing β] [StarRing β] @@ -1376,16 +1376,16 @@ instance instStarRing [NormedStarGroup β] : StarRing (α →ᵇ β) where __ := instStarAddMonoid star_mul f g := ext fun x ↦ star_mul (f x) (g x) -variable [CstarRing β] +variable [CStarRing β] -instance instCstarRing : CstarRing (α →ᵇ β) where +instance instCStarRing : CStarRing (α →ᵇ β) where norm_mul_self_le f := by rw [← sq, ← Real.le_sqrt (norm_nonneg _) (norm_nonneg _), norm_le (Real.sqrt_nonneg _)] intro x - rw [Real.le_sqrt (norm_nonneg _) (norm_nonneg _), sq, ← CstarRing.norm_star_mul_self] + rw [Real.le_sqrt (norm_nonneg _) (norm_nonneg _), sq, ← CStarRing.norm_star_mul_self] exact norm_coe_le_norm (star f * f) x -end CstarRing +end CStarRing section NormedLatticeOrderedGroup diff --git a/Mathlib/Topology/ContinuousFunction/Compact.lean b/Mathlib/Topology/ContinuousFunction/Compact.lean index 8b28ab32982a9..24fddb06c48a9 100644 --- a/Mathlib/Topology/ContinuousFunction/Compact.lean +++ b/Mathlib/Topology/ContinuousFunction/Compact.lean @@ -473,19 +473,19 @@ instance [CompactSpace α] : NormedStarGroup C(α, β) where end NormedSpace -section CstarRing +section CStarRing variable {α : Type*} {β : Type*} variable [TopologicalSpace α] [NormedRing β] [StarRing β] -instance [CompactSpace α] [CstarRing β] : CstarRing C(α, β) where +instance [CompactSpace α] [CStarRing β] : CStarRing C(α, β) where norm_mul_self_le f := by rw [← sq, ← Real.le_sqrt (norm_nonneg _) (norm_nonneg _), ContinuousMap.norm_le _ (Real.sqrt_nonneg _)] intro x - rw [Real.le_sqrt (norm_nonneg _) (norm_nonneg _), sq, ← CstarRing.norm_star_mul_self] + rw [Real.le_sqrt (norm_nonneg _) (norm_nonneg _), sq, ← CStarRing.norm_star_mul_self] exact ContinuousMap.norm_coe_le_norm (star f * f) x -end CstarRing +end CStarRing end ContinuousMap diff --git a/Mathlib/Topology/ContinuousFunction/ZeroAtInfty.lean b/Mathlib/Topology/ContinuousFunction/ZeroAtInfty.lean index 40691e8ad7fa6..586b078cd9837 100644 --- a/Mathlib/Topology/ContinuousFunction/ZeroAtInfty.lean +++ b/Mathlib/Topology/ContinuousFunction/ZeroAtInfty.lean @@ -553,12 +553,12 @@ instance instStarRing : StarRing C₀(α, β) := end StarRing -section CstarRing +section CStarRing -instance instCstarRing [NonUnitalNormedRing β] [StarRing β] [CstarRing β] : CstarRing C₀(α, β) where - norm_mul_self_le f := CstarRing.norm_mul_self_le (x := f.toBCF) +instance instCStarRing [NonUnitalNormedRing β] [StarRing β] [CStarRing β] : CStarRing C₀(α, β) where + norm_mul_self_le f := CStarRing.norm_mul_self_le (x := f.toBCF) -end CstarRing +end CStarRing /-! ### C₀ as a functor diff --git a/scripts/noshake.json b/scripts/noshake.json index a4ef16aeb038b..0a86ded76cecb 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -195,8 +195,8 @@ "Mathlib.Topology.Order.LeftRightNhds": ["Mathlib.Data.Set.Pointwise.Basic"], "Mathlib.Topology.Germ": ["Mathlib.Analysis.Normed.Module.Basic"], "Mathlib.Topology.Defs.Basic": ["Mathlib.Tactic.FunProp"], - "Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.NonUnital": - ["Mathlib.Analysis.CstarAlgebra.ContinuousFunctionalCalculus.Unital"], + "Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital": + ["Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital"], "Mathlib.Topology.Category.UniformSpace": ["Mathlib.CategoryTheory.Monad.Limits"], "Mathlib.Topology.Algebra.Nonarchimedean.AdicTopology": diff --git a/test/TCSynth.lean b/test/TCSynth.lean index dd2933a7cfc96..b29a76f58971b 100644 --- a/test/TCSynth.lean +++ b/test/TCSynth.lean @@ -43,7 +43,7 @@ section /-- info: NonUnitalNonAssocRing.toNonUnitalNonAssocSemiring -/ #guard_msgs in variable {A : Type} [NormedRing A] [NormedAlgebra ℂ A] [StarRing A] - [CstarRing A] [StarModule ℂ A] (x : A) in + [CStarRing A] [StarModule ℂ A] (x : A) in #synth NonUnitalNonAssocSemiring (elementalStarAlgebra ℂ x) end From 123eb213bb6485a5f5d662ff6ebd128cdfdc774e Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Tue, 6 Aug 2024 05:02:33 +0000 Subject: [PATCH 045/359] feat(CategoryTheory): opposite functors preserving finite products (#15518) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `F` preserves finite coproducts, then `F.op` preserves finite products, and the seven other variants for `leftOp`, `unop` etc. Also remove some lean3-isms from docstrings --- .../Limits/Preserves/Opposites.lean | 112 +++++++++++++----- 1 file changed, 82 insertions(+), 30 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean b/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean index a5682bfea6a92..d56a29ad3be6d 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Opposites.lean @@ -9,7 +9,7 @@ import Mathlib.CategoryTheory.Limits.Preserves.Finite /-! # Limit preservation properties of `functor.op` and related constructions -We formulate conditions about `F` which imply that `F.op`, `F.unop`, `F.left_op` and `F.right_op` +We formulate conditions about `F` which imply that `F.op`, `F.unop`, `F.leftOp` and `F.rightOp` preserve certain (co)limits. ## Future work @@ -28,26 +28,24 @@ open CategoryTheory namespace CategoryTheory.Limits -section - variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] variable {J : Type w} [Category.{w'} J] -/-- If `F : C ⥤ D` preserves colimits of `K.left_op : Jᵒᵖ ⥤ C`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves +/-- If `F : C ⥤ D` preserves colimits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves limits of `K : J ⥤ Cᵒᵖ`. -/ def preservesLimitOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ D) [PreservesColimit K.leftOp F] : PreservesLimit K F.op where preserves {_} hc := isLimitConeRightOpOfCocone _ (isColimitOfPreserves F (isColimitCoconeLeftOpOfCone _ hc)) -/-- If `F : C ⥤ Dᵒᵖ` preserves colimits of `K.left_op : Jᵒᵖ ⥤ C`, then `F.left_op : Cᵒᵖ ⥤ D` +/-- If `F : C ⥤ Dᵒᵖ` preserves colimits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves limits of `K : J ⥤ Cᵒᵖ`. -/ def preservesLimitLeftOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ Dᵒᵖ) [PreservesColimit K.leftOp F] : PreservesLimit K F.leftOp where preserves {_} hc := isLimitConeUnopOfCocone _ (isColimitOfPreserves F (isColimitCoconeLeftOpOfCone _ hc)) -/-- If `F : Cᵒᵖ ⥤ D` preserves colimits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.right_op : C ⥤ Dᵒᵖ` preserves +/-- If `F : Cᵒᵖ ⥤ D` preserves colimits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves limits of `K : J ⥤ C`. -/ def preservesLimitRightOp (K : J ⥤ C) (F : Cᵒᵖ ⥤ D) [PreservesColimit K.op F] : PreservesLimit K F.rightOp where @@ -61,21 +59,21 @@ def preservesLimitUnop (K : J ⥤ C) (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimit preserves {_} hc := isLimitConeUnopOfCocone _ (isColimitOfPreserves F hc.op) -/-- If `F : C ⥤ D` preserves limits of `K.left_op : Jᵒᵖ ⥤ C`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves +/-- If `F : C ⥤ D` preserves limits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves colimits of `K : J ⥤ Cᵒᵖ`. -/ def preservesColimitOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ D) [PreservesLimit K.leftOp F] : PreservesColimit K F.op where preserves {_} hc := isColimitCoconeRightOpOfCone _ (isLimitOfPreserves F (isLimitConeLeftOpOfCocone _ hc)) -/-- If `F : C ⥤ Dᵒᵖ` preserves limits of `K.left_op : Jᵒᵖ ⥤ C`, then `F.left_op : Cᵒᵖ ⥤ D` preserves +/-- If `F : C ⥤ Dᵒᵖ` preserves limits of `K.leftOp : Jᵒᵖ ⥤ C`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits of `K : J ⥤ Cᵒᵖ`. -/ def preservesColimitLeftOp (K : J ⥤ Cᵒᵖ) (F : C ⥤ Dᵒᵖ) [PreservesLimit K.leftOp F] : PreservesColimit K F.leftOp where preserves {_} hc := isColimitCoconeUnopOfCone _ (isLimitOfPreserves F (isLimitConeLeftOpOfCocone _ hc)) -/-- If `F : Cᵒᵖ ⥤ D` preserves limits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.right_op : C ⥤ Dᵒᵖ` preserves +/-- If `F : Cᵒᵖ ⥤ D` preserves limits of `K.op : Jᵒᵖ ⥤ Cᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits of `K : J ⥤ C`. -/ def preservesColimitRightOp (K : J ⥤ C) (F : Cᵒᵖ ⥤ D) [PreservesLimit K.op F] : PreservesColimit K F.rightOp where @@ -98,12 +96,12 @@ variable (J) def preservesLimitsOfShapeOp (F : C ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F] : PreservesLimitsOfShape J F.op where preservesLimit {K} := preservesLimitOp K F -/-- If `F : C ⥤ Dᵒᵖ` preserves colimits of shape `Jᵒᵖ`, then `F.left_op : Cᵒᵖ ⥤ D` preserves limits +/-- If `F : C ⥤ Dᵒᵖ` preserves colimits of shape `Jᵒᵖ`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves limits of shape `J`. -/ def preservesLimitsOfShapeLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimitsOfShape Jᵒᵖ F] : PreservesLimitsOfShape J F.leftOp where preservesLimit {K} := preservesLimitLeftOp K F -/-- If `F : Cᵒᵖ ⥤ D` preserves colimits of shape `Jᵒᵖ`, then `F.right_op : C ⥤ Dᵒᵖ` preserves limits +/-- If `F : Cᵒᵖ ⥤ D` preserves colimits of shape `Jᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves limits of shape `J`. -/ def preservesLimitsOfShapeRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimitsOfShape Jᵒᵖ F] : PreservesLimitsOfShape J F.rightOp where preservesLimit {K} := preservesLimitRightOp K F @@ -118,12 +116,12 @@ def preservesLimitsOfShapeUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimitsOfSha def preservesColimitsOfShapeOp (F : C ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F] : PreservesColimitsOfShape J F.op where preservesColimit {K} := preservesColimitOp K F -/-- If `F : C ⥤ Dᵒᵖ` preserves limits of shape `Jᵒᵖ`, then `F.left_op : Cᵒᵖ ⥤ D` preserves colimits +/-- If `F : C ⥤ Dᵒᵖ` preserves limits of shape `Jᵒᵖ`, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits of shape `J`. -/ def preservesColimitsOfShapeLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimitsOfShape Jᵒᵖ F] : PreservesColimitsOfShape J F.leftOp where preservesColimit {K} := preservesColimitLeftOp K F -/-- If `F : Cᵒᵖ ⥤ D` preserves limits of shape `Jᵒᵖ`, then `F.right_op : C ⥤ Dᵒᵖ` preserves colimits +/-- If `F : Cᵒᵖ ⥤ D` preserves limits of shape `Jᵒᵖ`, then `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits of shape `J`. -/ def preservesColimitsOfShapeRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimitsOfShape Jᵒᵖ F] : PreservesColimitsOfShape J F.rightOp where preservesColimit {K} := preservesColimitRightOp K F @@ -139,11 +137,11 @@ end def preservesLimitsOp (F : C ⥤ D) [PreservesColimits F] : PreservesLimits F.op where preservesLimitsOfShape {_} _ := preservesLimitsOfShapeOp _ _ -/-- If `F : C ⥤ Dᵒᵖ` preserves colimits, then `F.left_op : Cᵒᵖ ⥤ D` preserves limits. -/ +/-- If `F : C ⥤ Dᵒᵖ` preserves colimits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves limits. -/ def preservesLimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesColimits F] : PreservesLimits F.leftOp where preservesLimitsOfShape {_} _ := preservesLimitsOfShapeLeftOp _ _ -/-- If `F : Cᵒᵖ ⥤ D` preserves colimits, then `F.right_op : C ⥤ Dᵒᵖ` preserves limits. -/ +/-- If `F : Cᵒᵖ ⥤ D` preserves colimits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves limits. -/ def preservesLimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesColimits F] : PreservesLimits F.rightOp where preservesLimitsOfShape {_} _ := preservesLimitsOfShapeRightOp _ _ @@ -155,11 +153,11 @@ def preservesLimitsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesColimits F] : Preser def perservesColimitsOp (F : C ⥤ D) [PreservesLimits F] : PreservesColimits F.op where preservesColimitsOfShape {_} _ := preservesColimitsOfShapeOp _ _ -/-- If `F : C ⥤ Dᵒᵖ` preserves limits, then `F.left_op : Cᵒᵖ ⥤ D` preserves colimits. -/ +/-- If `F : C ⥤ Dᵒᵖ` preserves limits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves colimits. -/ def preservesColimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesLimits F] : PreservesColimits F.leftOp where preservesColimitsOfShape {_} _ := preservesColimitsOfShapeLeftOp _ _ -/-- If `F : Cᵒᵖ ⥤ D` preserves limits, then `F.right_op : C ⥤ Dᵒᵖ` preserves colimits. -/ +/-- If `F : Cᵒᵖ ⥤ D` preserves limits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves colimits. -/ def preservesColimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimits F] : PreservesColimits F.rightOp where preservesColimitsOfShape {_} _ := preservesColimitsOfShapeRightOp _ _ @@ -167,27 +165,19 @@ def preservesColimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesLimits F] : Preserves def preservesColimitsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesLimits F] : PreservesColimits F.unop where preservesColimitsOfShape {_} _ := preservesColimitsOfShapeUnop _ _ -end - -section - --- Preservation of finite (colimits) is only defined when the morphisms of C and D live in the same --- universe. -variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₁} D] - /-- If `F : C ⥤ D` preserves finite colimits, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite limits. -/ def preservesFiniteLimitsOp (F : C ⥤ D) [PreservesFiniteColimits F] : PreservesFiniteLimits F.op where preservesFiniteLimits J (_ : SmallCategory J) _ := preservesLimitsOfShapeOp J F -/-- If `F : C ⥤ Dᵒᵖ` preserves finite colimits, then `F.left_op : Cᵒᵖ ⥤ D` preserves finite +/-- If `F : C ⥤ Dᵒᵖ` preserves finite colimits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite limits. -/ def preservesFiniteLimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteColimits F] : PreservesFiniteLimits F.leftOp where preservesFiniteLimits J (_ : SmallCategory J) _ := preservesLimitsOfShapeLeftOp J F -/-- If `F : Cᵒᵖ ⥤ D` preserves finite colimits, then `F.right_op : C ⥤ Dᵒᵖ` preserves finite +/-- If `F : Cᵒᵖ ⥤ D` preserves finite colimits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite limits. -/ def preservesFiniteLimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteColimits F] : PreservesFiniteLimits F.rightOp where @@ -205,13 +195,13 @@ def preservesFiniteColimitsOp (F : C ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteColimits F.op where preservesFiniteColimits J (_ : SmallCategory J) _ := preservesColimitsOfShapeOp J F -/-- If `F : C ⥤ Dᵒᵖ` preserves finite limits, then `F.left_op : Cᵒᵖ ⥤ D` preserves finite +/-- If `F : C ⥤ Dᵒᵖ` preserves finite limits, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite colimits. -/ def preservesFiniteColimitsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteLimits F] : PreservesFiniteColimits F.leftOp where preservesFiniteColimits J (_ : SmallCategory J) _ := preservesColimitsOfShapeLeftOp J F -/-- If `F : Cᵒᵖ ⥤ D` preserves finite limits, then `F.right_op : C ⥤ Dᵒᵖ` preserves finite +/-- If `F : Cᵒᵖ ⥤ D` preserves finite limits, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite colimits. -/ def preservesFiniteColimitsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteLimits F] : PreservesFiniteColimits F.rightOp where @@ -223,6 +213,68 @@ def preservesFiniteColimitsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteLimits PreservesFiniteColimits F.unop where preservesFiniteColimits J (_ : SmallCategory J) _ := preservesColimitsOfShapeUnop J F -end +/-- If `F : C ⥤ D` preserves finite coproducts, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite + products. -/ +def preservesFiniteProductsOp (F : C ⥤ D) [PreservesFiniteCoproducts F] : + PreservesFiniteProducts F.op where + preserves J _ := by + apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeOp + exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + +/-- If `F : C ⥤ Dᵒᵖ` preserves finite coproducts, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite + products. -/ +def preservesFiniteProductsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteCoproducts F] : + PreservesFiniteProducts F.leftOp where + preserves J _ := by + apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeLeftOp + exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + +/-- If `F : Cᵒᵖ ⥤ D` preserves finite coproducts, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite + products. -/ +def preservesFiniteProductsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteCoproducts F] : + PreservesFiniteProducts F.rightOp where + preserves J _ := by + apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeRightOp + exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + +/-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite coproducts, then `F.unop : C ⥤ D` preserves finite + products. -/ +def preservesFiniteProductsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteCoproducts F] : + PreservesFiniteProducts F.unop where + preserves J _ := by + apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeUnop + exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + +/-- If `F : C ⥤ D` preserves finite products, then `F.op : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite + coproducts. -/ +def preservesFiniteCoproductsOp (F : C ⥤ D) [PreservesFiniteProducts F] : + PreservesFiniteCoproducts F.op where + preserves J _ := by + apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeOp + exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + +/-- If `F : C ⥤ Dᵒᵖ` preserves finite products, then `F.leftOp : Cᵒᵖ ⥤ D` preserves finite + coproducts. -/ +def preservesFiniteCoproductsLeftOp (F : C ⥤ Dᵒᵖ) [PreservesFiniteProducts F] : + PreservesFiniteCoproducts F.leftOp where + preserves J _ := by + apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeLeftOp + exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + +/-- If `F : Cᵒᵖ ⥤ D` preserves finite products, then `F.rightOp : C ⥤ Dᵒᵖ` preserves finite + coproducts. -/ +def preservesFiniteCoproductsRightOp (F : Cᵒᵖ ⥤ D) [PreservesFiniteProducts F] : + PreservesFiniteCoproducts F.rightOp where + preserves J _ := by + apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeRightOp + exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ + +/-- If `F : Cᵒᵖ ⥤ Dᵒᵖ` preserves finite products, then `F.unop : C ⥤ D` preserves finite + coproducts. -/ +def preservesFiniteCoproductsUnop (F : Cᵒᵖ ⥤ Dᵒᵖ) [PreservesFiniteProducts F] : + PreservesFiniteCoproducts F.unop where + preserves J _ := by + apply (config := { allowSynthFailures := true }) preservesColimitsOfShapeUnop + exact preservesLimitsOfShapeOfEquiv (Discrete.opposite J).symm _ end CategoryTheory.Limits From 50ebd3bd148af9c10b565cd78192653b8c348363 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 6 Aug 2024 05:13:04 +0000 Subject: [PATCH 046/359] feat(RingTheory/Coalgebra/*): add basics about sweedler notation (#12109) Though the name is called sweedler notation, no new notations are actually introduced. This PR is actually about expanding tensor product. Co-authored-by: Lihan Jiang Co-authored-by: 101damnations <101damnations@github.com> --- Mathlib/RingTheory/Coalgebra/Basic.lean | 38 ++++++++++++++++++++++++- Mathlib/RingTheory/Coalgebra/Hom.lean | 31 ++++++++++++++++++-- Mathlib/RingTheory/HopfAlgebra.lean | 24 ++++++++++++++++ 3 files changed, 90 insertions(+), 3 deletions(-) diff --git a/Mathlib/RingTheory/Coalgebra/Basic.lean b/Mathlib/RingTheory/Coalgebra/Basic.lean index f4cde43839b73..03811b4c741c4 100644 --- a/Mathlib/RingTheory/Coalgebra/Basic.lean +++ b/Mathlib/RingTheory/Coalgebra/Basic.lean @@ -6,7 +6,7 @@ Authors: Ali Ramsey, Eric Wieser import Mathlib.Algebra.Algebra.Bilinear import Mathlib.LinearAlgebra.Finsupp import Mathlib.LinearAlgebra.Prod -import Mathlib.LinearAlgebra.TensorProduct.Basic +import Mathlib.LinearAlgebra.TensorProduct.Finiteness /-! # Coalgebras @@ -38,6 +38,30 @@ class CoalgebraStruct (R : Type u) (A : Type v) /-- The counit of the coalgebra -/ counit : A →ₗ[R] R +/-- +A representation of an element `a` of a coalgebra `A` is a finite sum of pure tensors `∑ xᵢ ⊗ yᵢ` +that is equal to `comul a`. +-/ +structure Coalgebra.Repr (R : Type u) {A : Type v} + [CommSemiring R] [AddCommMonoid A] [Module R A] [CoalgebraStruct R A] (a : A) where + /-- the indexing type of a representation of `comul a` -/ + {ι : Type*} + /-- the finite indexing set of a representation of `comul a` -/ + (index : Finset ι) + /-- the first coordinate of a representation of `comul a` -/ + (left : ι → A) + /-- the second coordinate of a representation of `comul a` -/ + (right : ι → A) + /-- `comul a` is equal to a finite sum of some pure tensors -/ + (eq : ∑ i ∈ index, left i ⊗ₜ[R] right i = CoalgebraStruct.comul a) + +/-- An arbitrarily chosen representation. -/ +def Coalgebra.Repr.arbitrary (R : Type u) {A : Type v} + [CommSemiring R] [AddCommMonoid A] [Module R A] [CoalgebraStruct R A] (a : A) : + Coalgebra.Repr R a where + index := TensorProduct.exists_finset (R := R) (CoalgebraStruct.comul a) |>.choose + eq := TensorProduct.exists_finset (R := R) (CoalgebraStruct.comul a) |>.choose_spec.symm + namespace Coalgebra export CoalgebraStruct (comul counit) end Coalgebra @@ -81,7 +105,19 @@ theorem rTensor_counit_comul (a : A) : counit.rTensor A (comul a) = 1 ⊗ₜ[R] theorem lTensor_counit_comul (a : A) : counit.lTensor A (comul a) = a ⊗ₜ[R] 1 := LinearMap.congr_fun lTensor_counit_comp_comul a +@[simp] +lemma sum_counit_tmul_eq {a : A} (repr : Coalgebra.Repr R a) : + ∑ i ∈ repr.index, counit (R := R) (repr.left i) ⊗ₜ (repr.right i) = 1 ⊗ₜ[R] a := by + simpa [← repr.eq, map_sum] using congr($(rTensor_counit_comp_comul (R := R) (A := A)) a) + +@[simp] +lemma sum_tmul_counit_eq {a : A} (repr : Coalgebra.Repr R a) : + ∑ i ∈ repr.index, (repr.left i) ⊗ₜ counit (R := R) (repr.right i) = a ⊗ₜ[R] 1 := by + simpa [← repr.eq, map_sum] using congr($(lTensor_counit_comp_comul (R := R) (A := A)) a) + + end Coalgebra + section CommSemiring open Coalgebra diff --git a/Mathlib/RingTheory/Coalgebra/Hom.lean b/Mathlib/RingTheory/Coalgebra/Hom.lean index d138b068c895a..75d7d5fd60d0e 100644 --- a/Mathlib/RingTheory/Coalgebra/Hom.lean +++ b/Mathlib/RingTheory/Coalgebra/Hom.lean @@ -260,9 +260,10 @@ end CoalgHom namespace Coalgebra -variable (R : Type u) (A : Type v) +variable (R : Type u) (A : Type v) (B : Type w) -variable [CommSemiring R] [AddCommMonoid A] [Module R A] [Coalgebra R A] +variable [CommSemiring R] [AddCommMonoid A] [AddCommMonoid B] [Module R A] [Module R B] +variable [Coalgebra R A] [Coalgebra R B] /-- The counit of a coalgebra as a `CoalgHom`. -/ def counitCoalgHom : A →ₗc[R] R := @@ -293,4 +294,30 @@ instance subsingleton_to_ring : Subsingleton (A →ₗc[R] R) := @[ext high] theorem ext_to_ring (f g : A →ₗc[R] R) : f = g := Subsingleton.elim _ _ +variable {A B} +/-- +If `φ : A → B` is a coalgebra map and `a = ∑ xᵢ ⊗ yᵢ`, then `φ a = ∑ φ xᵢ ⊗ φ yᵢ` +-/ +@[simps] +def Repr.induced {a : A} (repr : Repr R a) + {F : Type*} [FunLike F A B] [CoalgHomClass F R A B] + (φ : F) : Repr R (φ a) where + index := repr.index + left := φ ∘ repr.left + right := φ ∘ repr.right + eq := (congr($((CoalgHomClass.map_comp_comul φ).symm) a).trans <| + by rw [LinearMap.comp_apply, ← repr.eq, map_sum]; rfl).symm + +@[simp] +lemma sum_tmul_counit_apply_eq + {F : Type*} [FunLike F A B] [CoalgHomClass F R A B] (φ : F) {a : A} (repr : Repr R a) : + ∑ i ∈ repr.index, counit (R := R) (repr.left i) ⊗ₜ φ (repr.right i) = 1 ⊗ₜ[R] φ a := by + simp [← sum_counit_tmul_eq (repr.induced φ)] + +@[simp] +lemma sum_tmul_apply_counit_eq + {F : Type*} [FunLike F A B] [CoalgHomClass F R A B] (φ : F) {a : A} (repr : Repr R a) : + ∑ i ∈ repr.index, φ (repr.left i) ⊗ₜ counit (R := R) (repr.right i) = φ a ⊗ₜ[R] 1 := by + simp [← sum_tmul_counit_eq (repr.induced φ)] + end Coalgebra diff --git a/Mathlib/RingTheory/HopfAlgebra.lean b/Mathlib/RingTheory/HopfAlgebra.lean index 4da3803f187d4..20f1a7ca2575f 100644 --- a/Mathlib/RingTheory/HopfAlgebra.lean +++ b/Mathlib/RingTheory/HopfAlgebra.lean @@ -70,6 +70,30 @@ theorem mul_antipode_lTensor_comul_apply (a : A) : algebraMap R A (Coalgebra.counit a) := LinearMap.congr_fun mul_antipode_lTensor_comul a +open Coalgebra + +@[simp] +lemma sum_antipode_mul_eq {a : A} (repr : Repr R a) : + ∑ i ∈ repr.index, antipode (R := R) (repr.left i) * repr.right i = + algebraMap R A (counit a) := by + simpa [← repr.eq, map_sum] using congr($(mul_antipode_rTensor_comul (R := R)) a) + +@[simp] +lemma sum_mul_antipode_eq {a : A} (repr : Repr R a) : + ∑ i ∈ repr.index, repr.left i * antipode (R := R) (repr.right i) = + algebraMap R A (counit a) := by + simpa [← repr.eq, map_sum] using congr($(mul_antipode_lTensor_comul (R := R)) a) + +lemma sum_antipode_mul_eq_smul {a : A} (repr : Repr R a) : + ∑ i ∈ repr.index, antipode (R := R) (repr.left i) * repr.right i = + counit (R := R) a • 1 := by + rw [sum_antipode_mul_eq, Algebra.smul_def, mul_one] + +lemma sum_mul_antipode_eq_smul {a : A} (repr : Repr R a) : + ∑ i ∈ repr.index, repr.left i * antipode (R := R) (repr.right i) = + counit (R := R) a • 1 := by + rw [sum_mul_antipode_eq, Algebra.smul_def, mul_one] + end HopfAlgebra namespace CommSemiring From bff8114b962ab84f65f7628622a7a408d0cb7b8e Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 6 Aug 2024 05:13:06 +0000 Subject: [PATCH 047/359] feat(CategoryTheory/Limits/ConcreteCategory): prove some results for concrete category with extra algebraic structures (#15315) --- Mathlib.lean | 3 +- Mathlib/Algebra/Category/Grp/Abelian.lean | 2 +- Mathlib/Algebra/Category/Grp/Limits.lean | 2 +- .../Basic.lean} | 0 .../WithAlgebraicStructures.lean | 80 +++++++++++++++++++ .../Limits/Shapes/ConcreteCategory.lean | 2 +- .../Category/TopCat/Limits/Products.lean | 2 +- 7 files changed, 86 insertions(+), 5 deletions(-) rename Mathlib/CategoryTheory/Limits/{ConcreteCategory.lean => ConcreteCategory/Basic.lean} (100%) create mode 100644 Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean diff --git a/Mathlib.lean b/Mathlib.lean index 85efd600407a5..afdb573c1abee 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1506,7 +1506,8 @@ import Mathlib.CategoryTheory.LiftingProperties.Basic import Mathlib.CategoryTheory.Limits.Bicones import Mathlib.CategoryTheory.Limits.ColimitLimit import Mathlib.CategoryTheory.Limits.Comma -import Mathlib.CategoryTheory.Limits.ConcreteCategory +import Mathlib.CategoryTheory.Limits.ConcreteCategory.Basic +import Mathlib.CategoryTheory.Limits.ConcreteCategory.WithAlgebraicStructures import Mathlib.CategoryTheory.Limits.ConeCategory import Mathlib.CategoryTheory.Limits.Cones import Mathlib.CategoryTheory.Limits.Connected diff --git a/Mathlib/Algebra/Category/Grp/Abelian.lean b/Mathlib/Algebra/Category/Grp/Abelian.lean index a4856700ad059..fb193bc6cc250 100644 --- a/Mathlib/Algebra/Category/Grp/Abelian.lean +++ b/Mathlib/Algebra/Category/Grp/Abelian.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.Category.Grp.Limits import Mathlib.Algebra.Category.Grp.ZModuleEquivalence import Mathlib.Algebra.Category.ModuleCat.Abelian import Mathlib.CategoryTheory.Adjunction.Limits -import Mathlib.CategoryTheory.Limits.ConcreteCategory +import Mathlib.CategoryTheory.Limits.ConcreteCategory.Basic /-! # The category of abelian groups is abelian diff --git a/Mathlib/Algebra/Category/Grp/Limits.lean b/Mathlib/Algebra/Category/Grp/Limits.lean index 67b7f6bfa0b86..f605284f69e1f 100644 --- a/Mathlib/Algebra/Category/Grp/Limits.lean +++ b/Mathlib/Algebra/Category/Grp/Limits.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.Category.Grp.ForgetCorepresentable import Mathlib.Algebra.Category.Grp.Preadditive import Mathlib.Algebra.Group.Subgroup.Basic import Mathlib.CategoryTheory.Comma.Over -import Mathlib.CategoryTheory.Limits.ConcreteCategory +import Mathlib.CategoryTheory.Limits.ConcreteCategory.Basic import Mathlib.CategoryTheory.ConcreteCategory.ReflectsIso /-! diff --git a/Mathlib/CategoryTheory/Limits/ConcreteCategory.lean b/Mathlib/CategoryTheory/Limits/ConcreteCategory/Basic.lean similarity index 100% rename from Mathlib/CategoryTheory/Limits/ConcreteCategory.lean rename to Mathlib/CategoryTheory/Limits/ConcreteCategory/Basic.lean diff --git a/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean new file mode 100644 index 0000000000000..8a6c50484bdcd --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean @@ -0,0 +1,80 @@ +/- +Copyright (c) 2024 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ +import Mathlib.CategoryTheory.Limits.ConcreteCategory.Basic +import Mathlib.Algebra.Module.LinearMap.Defs +import Mathlib.Tactic.CategoryTheory.Elementwise + +/-! + +# Colimits in concrete categories with algebraic structures + +Let `C` be a concrete category and `F : J ⥤ C` a filtered diagram in `C`. We discuss some results +about `colimit F` when objects and morphisms in `C` have some algebraic structures. + +## Main results +- `CategoryTheory.Limits.Concrete.colimit_rep_eq_zero`: Let `C` be a category where its objects have + zero elements and morphisms preserve zero. If `x : Fⱼ` is mapped to `0` in the colimit, then + there exists a `i ⟶ j` such that `x` restricted to `i` is already `0`. + +- `CategoryTheory.Limits.Concrete.colimit_no_zero_smul_divisor`: Let `C` be a category where its + objects are `R`-modules and morphisms `R`-linear maps. Let `r : R` be an element without zero + smul divisors for all small sections, i.e. there exists some `j : J` such that for all `j ⟶ i` + and `x : Fᵢ` we have `r • x = 0` implies `x = 0`, then if `r • x = 0` for `x : colimit F`, then + `x = 0`. + +-/ + +universe t w v u r + +open CategoryTheory + +namespace CategoryTheory.Limits.Concrete + +attribute [local instance] ConcreteCategory.instFunLike ConcreteCategory.hasCoeToSort + +variable {C : Type u} [Category.{v} C] [ConcreteCategory.{max t w} C] {J : Type w} [Category.{r} J] + +section zero + +theorem colimit_rep_eq_zero + (F : J ⥤ C) [PreservesColimit F (forget C)] [IsFiltered J] + [∀ c : C, Zero c] [∀ {c c' : C}, ZeroHomClass (c ⟶ c') c c'] [HasColimit F] + (j : J) (x : F.obj j) (hx : colimit.ι F j x = 0) : + ∃ (j' : J) (i : j ⟶ j'), F.map i x = 0 := by + rw [show 0 = colimit.ι F j 0 by simp, colimit_rep_eq_iff_exists] at hx + obtain ⟨j', i, y, g⟩ := hx + exact ⟨j', i, g ▸ by simp⟩ + +end zero + +section module + +/-- +if `r` has no zero smul divisors for all small-enough sections, then `r` has no zero smul divisors +in the colimit. +-/ +lemma colimit_no_zero_smul_divisor + (F : J ⥤ C) [PreservesColimit F (forget C)] [IsFiltered J] [HasColimit F] + (R : Type*) [Semiring R] + [∀ c : C, AddCommMonoid c] [∀ c : C, Module R c] + [∀ {c c' : C}, LinearMapClass (c ⟶ c') R c c'] + (r : R) (H : ∃ (j' : J), ∀ (j : J) (_ : j' ⟶ j), ∀ (c : F.obj j), r • c = 0 → c = 0) + (x : (forget C).obj (colimit F)) (hx : r • x = 0) : x = 0 := by + classical + obtain ⟨j, x, rfl⟩ := Concrete.colimit_exists_rep F x + rw [← map_smul] at hx + obtain ⟨j', i, h⟩ := Concrete.colimit_rep_eq_zero (hx := hx) + obtain ⟨j'', H⟩ := H + simpa [elementwise_of% (colimit.w F), map_zero] using congr(colimit.ι F _ + $(H (IsFiltered.sup {j, j', j''} { ⟨j, j', by simp, by simp, i⟩ }) + (IsFiltered.toSup _ _ $ by simp) + (F.map (IsFiltered.toSup _ _ $ by simp) x) + (by rw [← IsFiltered.toSup_commutes (f := i) (mY := by simp) (mf := by simp), F.map_comp, + comp_apply, ← map_smul, ← map_smul, h, map_zero]))) + +end module + +end CategoryTheory.Limits.Concrete diff --git a/Mathlib/CategoryTheory/Limits/Shapes/ConcreteCategory.lean b/Mathlib/CategoryTheory/Limits/Shapes/ConcreteCategory.lean index c817b361c71e9..d1f39b75f730a 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/ConcreteCategory.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/ConcreteCategory.lean @@ -5,7 +5,7 @@ Authors: Joël Riou, Scott Morrison, Adam Topaz -/ import Mathlib.CategoryTheory.Limits.Preserves.Shapes.BinaryProducts import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Products -import Mathlib.CategoryTheory.Limits.ConcreteCategory +import Mathlib.CategoryTheory.Limits.ConcreteCategory.Basic import Mathlib.CategoryTheory.Limits.Shapes.Types import Mathlib.CategoryTheory.Limits.Shapes.Multiequalizer import Mathlib.CategoryTheory.Limits.Shapes.Kernels diff --git a/Mathlib/Topology/Category/TopCat/Limits/Products.lean b/Mathlib/Topology/Category/TopCat/Limits/Products.lean index 03a5ba398a9f0..43cd610adcbfb 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Products.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Products.lean @@ -6,7 +6,7 @@ Authors: Patrick Massot, Scott Morrison, Mario Carneiro, Andrew Yang import Mathlib.Topology.Category.TopCat.EpiMono import Mathlib.Topology.Category.TopCat.Limits.Basic import Mathlib.CategoryTheory.Limits.Shapes.Products -import Mathlib.CategoryTheory.Limits.ConcreteCategory +import Mathlib.CategoryTheory.Limits.ConcreteCategory.Basic import Mathlib.Data.Set.Subsingleton import Mathlib.Tactic.CategoryTheory.Elementwise From 6fdbc7edc0cd1bb2b6f316cbc51121e3c39b20f8 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 6 Aug 2024 05:13:07 +0000 Subject: [PATCH 048/359] chore: backports for leanprover/lean4#4814 (part 19) (#15434) see https://github.com/leanprover-community/mathlib4/pull/15245 Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib/GroupTheory/PGroup.lean | 6 ++--- Mathlib/RingTheory/Binomial.lean | 6 ++--- .../DiscreteValuationRing/Basic.lean | 8 +++--- Mathlib/RingTheory/Filtration.lean | 6 ++--- .../RingTheory/MvPolynomial/Homogeneous.lean | 12 ++++----- .../Polynomial/Eisenstein/Basic.lean | 27 ++++++++++--------- Mathlib/RingTheory/RootsOfUnity/Basic.lean | 20 +++++++------- Mathlib/Topology/Sequences.lean | 6 ++--- 8 files changed, 46 insertions(+), 45 deletions(-) diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index 2bd5e430961c9..f5d02b151cb0c 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -190,7 +190,7 @@ theorem card_modEq_card_fixedPoints : Nat.card α ≡ Nat.card (fixedPoints G α /-- If a p-group acts on `α` and the cardinality of `α` is not a multiple of `p` then the action has a fixed point. -/ -theorem nonempty_fixed_point_of_prime_not_dvd_card (hpα : ¬p ∣ Nat.card α) : +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 _ _ @@ -315,12 +315,12 @@ theorem disjoint_of_ne (p₁ p₂ : ℕ) [hp₁ : Fact p₁.Prime] [hp₂ : Fact section P2comm -variable [Fact p.Prime] {n : ℕ} (hGpn : Nat.card G = p ^ n) +variable [Fact p.Prime] {n : ℕ} open Subgroup /-- The cardinality of the `center` of a `p`-group is `p ^ k` where `k` is positive. -/ -theorem card_center_eq_prime_pow (hn : 0 < n) : +theorem card_center_eq_prime_pow (hGpn : Nat.card G = p ^ n) (hn : 0 < n) : ∃ k > 0, Nat.card (center G) = p ^ k := by have : Finite G := Nat.finite_of_card_ne_zero (hGpn ▸ pow_ne_zero n (NeZero.ne p)) have hcG := to_subgroup (of_card hGpn) (center G) diff --git a/Mathlib/RingTheory/Binomial.lean b/Mathlib/RingTheory/Binomial.lean index db3f789449c62..5becb23233943 100644 --- a/Mathlib/RingTheory/Binomial.lean +++ b/Mathlib/RingTheory/Binomial.lean @@ -301,7 +301,7 @@ theorem smeval_ascPochhammer_neg_of_lt {n k : ℕ} (h : n < k) : smeval (ascPochhammer ℕ k) (-n : ℤ) = 0 := by rw [show k = n + (k - n - 1) + 1 by omega, smeval_ascPochhammer_neg_add] -theorem smeval_ascPochhammer_nat_cast [NatPowAssoc R] (n k : ℕ) : +theorem smeval_ascPochhammer_nat_cast {R} [NonAssocRing R] [Pow R ℕ] [NatPowAssoc R] (n k : ℕ) : smeval (ascPochhammer ℕ k) (n : R) = smeval (ascPochhammer ℕ k) n := by rw [smeval_at_natCast (ascPochhammer ℕ k) n] @@ -337,8 +337,8 @@ theorem multichoose_succ_neg_natCast [NatPowAssoc R] (n : ℕ) : rw [factorial_nsmul_multichoose_eq_ascPochhammer, smeval_neg_nat, smeval_ascPochhammer_succ_neg n, Int.cast_zero] -theorem smeval_ascPochhammer_int_ofNat [NatPowAssoc R] (r : R) : ∀ n : ℕ, - smeval (ascPochhammer ℤ n) r = smeval (ascPochhammer ℕ n) r +theorem smeval_ascPochhammer_int_ofNat {R} [NonAssocRing R] [Pow R ℕ] [NatPowAssoc R] (r : R) : + ∀ n : ℕ, smeval (ascPochhammer ℤ n) r = smeval (ascPochhammer ℕ n) r | 0 => by simp only [ascPochhammer_zero, smeval_one] | n + 1 => by diff --git a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean index f4fbe9856cf6f..518c5943e0624 100644 --- a/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean +++ b/Mathlib/RingTheory/DiscreteValuationRing/Basic.lean @@ -148,9 +148,10 @@ def HasUnitMulPowIrreducibleFactorization [CommRing R] : Prop := namespace HasUnitMulPowIrreducibleFactorization -variable {R} [CommRing R] (hR : HasUnitMulPowIrreducibleFactorization R) +variable {R} [CommRing R] -theorem unique_irreducible ⦃p q : R⦄ (hp : Irreducible p) (hq : Irreducible q) : +theorem unique_irreducible (hR : HasUnitMulPowIrreducibleFactorization R) + ⦃p q : R⦄ (hp : Irreducible p) (hq : Irreducible q) : Associated p q := by rcases hR with ⟨ϖ, hϖ, hR⟩ suffices ∀ {p : R} (_ : Irreducible p), Associated p ϖ by @@ -178,7 +179,8 @@ variable [IsDomain R] /-- An integral domain in which there is an irreducible element `p` such that every nonzero element is associated to a power of `p` is a unique factorization domain. See `DiscreteValuationRing.ofHasUnitMulPowIrreducibleFactorization`. -/ -theorem toUniqueFactorizationMonoid : UniqueFactorizationMonoid R := +theorem toUniqueFactorizationMonoid (hR : HasUnitMulPowIrreducibleFactorization R) : + UniqueFactorizationMonoid R := let p := Classical.choose hR let spec := Classical.choose_spec hR UniqueFactorizationMonoid.of_exists_prime_factors fun x hx => by diff --git a/Mathlib/RingTheory/Filtration.lean b/Mathlib/RingTheory/Filtration.lean index d28ab3f545103..a73354fbc33c0 100644 --- a/Mathlib/RingTheory/Filtration.lean +++ b/Mathlib/RingTheory/Filtration.lean @@ -187,9 +187,8 @@ theorem _root_.Ideal.stableFiltration_stable (I : Ideal R) (N : Submodule R M) : rw [add_comm, pow_add, mul_smul, pow_one] variable {F F'} -variable (h : F.Stable) -theorem Stable.exists_pow_smul_eq : ∃ n₀, ∀ k, F.N (n₀ + k) = I ^ k • F.N n₀ := by +theorem Stable.exists_pow_smul_eq (h : F.Stable) : ∃ n₀, ∀ k, F.N (n₀ + k) = I ^ k • F.N n₀ := by obtain ⟨n₀, hn⟩ := h use n₀ intro k @@ -198,7 +197,8 @@ theorem Stable.exists_pow_smul_eq : ∃ n₀, ∀ k, F.N (n₀ + k) = I ^ k • · rw [← add_assoc, ← hn, ih, add_comm, pow_add, mul_smul, pow_one] omega -theorem Stable.exists_pow_smul_eq_of_ge : ∃ n₀, ∀ n ≥ n₀, F.N n = I ^ (n - n₀) • F.N n₀ := by +theorem Stable.exists_pow_smul_eq_of_ge (h : F.Stable) : + ∃ n₀, ∀ n ≥ n₀, F.N n = I ^ (n - n₀) • F.N n₀ := by obtain ⟨n₀, hn₀⟩ := h.exists_pow_smul_eq use n₀ intro n hn diff --git a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean index 2cd4d9fdd0bf9..11a64d6d9870b 100644 --- a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean @@ -80,27 +80,25 @@ lemma weightedHomogeneousSubmodule_one (n : ℕ) : variable {σ R} @[simp] -theorem mem_homogeneousSubmodule [CommSemiring R] (n : ℕ) (p : MvPolynomial σ R) : +theorem mem_homogeneousSubmodule (n : ℕ) (p : MvPolynomial σ R) : p ∈ homogeneousSubmodule σ R n ↔ p.IsHomogeneous n := Iff.rfl variable (σ R) /-- While equal, the former has a convenient definitional reduction. -/ -theorem homogeneousSubmodule_eq_finsupp_supported [CommSemiring R] (n : ℕ) : +theorem homogeneousSubmodule_eq_finsupp_supported (n : ℕ) : homogeneousSubmodule σ R n = Finsupp.supported _ R { d | d.degree = n } := by simp_rw [degree_eq_weight_one] exact weightedHomogeneousSubmodule_eq_finsupp_supported R 1 n variable {σ R} -theorem homogeneousSubmodule_mul [CommSemiring R] (m n : ℕ) : +theorem homogeneousSubmodule_mul (m n : ℕ) : homogeneousSubmodule σ R m * homogeneousSubmodule σ R n ≤ homogeneousSubmodule σ R (m + n) := weightedHomogeneousSubmodule_mul 1 m n section -variable [CommSemiring R] - theorem isHomogeneous_monomial {d : σ →₀ ℕ} (r : R) {n : ℕ} (hn : d.degree = n) : IsHomogeneous (monomial d r) n := by rw [degree_eq_weight_one] at hn @@ -143,7 +141,7 @@ end namespace IsHomogeneous -variable [CommSemiring R] [CommSemiring S] {φ ψ : MvPolynomial σ R} {m n : ℕ} +variable [CommSemiring S] {φ ψ : MvPolynomial σ R} {m n : ℕ} theorem coeff_eq_zero (hφ : IsHomogeneous φ n) {d : σ →₀ ℕ} (hd : d.degree ≠ n) : coeff d φ = 0 := by @@ -448,7 +446,7 @@ section HomogeneousComponent open Finset Finsupp -variable [CommSemiring R] (n : ℕ) (φ ψ : MvPolynomial σ R) +variable (n : ℕ) (φ ψ : MvPolynomial σ R) theorem homogeneousComponent_mem : homogeneousComponent n φ ∈ homogeneousSubmodule σ R n := diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean index e78f604bc116c..924247aa7404b 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/Basic.lean @@ -56,10 +56,10 @@ namespace IsWeaklyEisensteinAt section CommSemiring -variable [CommSemiring R] {𝓟 : Ideal R} {f : R[X]} (hf : f.IsWeaklyEisensteinAt 𝓟) +variable [CommSemiring R] {𝓟 : Ideal R} {f : R[X]} - -theorem map {A : Type v} [CommRing A] (φ : R →+* A) : (f.map φ).IsWeaklyEisensteinAt (𝓟.map φ) := by +theorem map (hf : f.IsWeaklyEisensteinAt 𝓟) {A : Type v} [CommRing A] (φ : R →+* A) : + (f.map φ).IsWeaklyEisensteinAt (𝓟.map φ) := by refine (isWeaklyEisensteinAt_iff _ _).2 fun hn => ?_ rw [coeff_map] exact mem_map_of_mem _ (hf.mem (lt_of_lt_of_le hn (natDegree_map_le _ _))) @@ -68,7 +68,7 @@ end CommSemiring section CommRing -variable [CommRing R] {𝓟 : Ideal R} {f : R[X]} (hf : f.IsWeaklyEisensteinAt 𝓟) +variable [CommRing R] {𝓟 : Ideal R} {f : R[X]} variable {S : Type v} [CommRing S] [Algebra R S] section Principal @@ -118,7 +118,8 @@ end Principal -- Porting note: `Ideal.neg_mem_iff` was `neg_mem_iff` on line 142 but Lean was not able to find -- NegMemClass -theorem pow_natDegree_le_of_root_of_monic_mem {x : R} (hroot : IsRoot f x) (hmo : f.Monic) : +theorem pow_natDegree_le_of_root_of_monic_mem (hf : f.IsWeaklyEisensteinAt 𝓟) + {x : R} (hroot : IsRoot f x) (hmo : f.Monic) : ∀ i, f.natDegree ≤ i → x ^ i ∈ 𝓟 := by intro i hi obtain ⟨k, hk⟩ := exists_add_of_le hi @@ -130,8 +131,8 @@ theorem pow_natDegree_le_of_root_of_monic_mem {x : R} (hroot : IsRoot f x) (hmo rw [eq_neg_of_add_eq_zero_left hroot, Ideal.neg_mem_iff] exact Submodule.sum_mem _ fun i _ => mul_mem_right _ _ (hf.mem (Fin.is_lt i)) -theorem pow_natDegree_le_of_aeval_zero_of_monic_mem_map {x : S} (hx : aeval x f = 0) - (hmo : f.Monic) : +theorem pow_natDegree_le_of_aeval_zero_of_monic_mem_map (hf : f.IsWeaklyEisensteinAt 𝓟) + {x : S} (hx : aeval x f = 0) (hmo : f.Monic) : ∀ i, (f.map (algebraMap R S)).natDegree ≤ i → x ^ i ∈ 𝓟.map (algebraMap R S) := by suffices x ^ (f.map (algebraMap R S)).natDegree ∈ 𝓟.map (algebraMap R S) by intro i hi @@ -180,7 +181,7 @@ namespace IsEisensteinAt section CommSemiring -variable [CommSemiring R] {𝓟 : Ideal R} {f : R[X]} (hf : f.IsEisensteinAt 𝓟) +variable [CommSemiring R] {𝓟 : Ideal R} {f : R[X]} theorem _root_.Polynomial.Monic.leadingCoeff_not_mem (hf : f.Monic) (h : 𝓟 ≠ ⊤) : ¬f.leadingCoeff ∈ 𝓟 := hf.leadingCoeff.symm ▸ (Ideal.ne_top_iff_one _).1 h @@ -192,10 +193,10 @@ theorem _root_.Polynomial.Monic.isEisensteinAt_of_mem_of_not_mem (hf : f.Monic) mem := fun hn => hmem hn not_mem := hnot_mem } -theorem isWeaklyEisensteinAt : IsWeaklyEisensteinAt f 𝓟 := +theorem isWeaklyEisensteinAt (hf : f.IsEisensteinAt 𝓟) : IsWeaklyEisensteinAt f 𝓟 := ⟨fun h => hf.mem h⟩ -theorem coeff_mem {n : ℕ} (hn : n ≠ f.natDegree) : f.coeff n ∈ 𝓟 := by +theorem coeff_mem (hf : f.IsEisensteinAt 𝓟) {n : ℕ} (hn : n ≠ f.natDegree) : f.coeff n ∈ 𝓟 := by cases' ne_iff_lt_or_gt.1 hn with h₁ h₂ · exact hf.mem h₁ · rw [coeff_eq_zero_of_natDegree_lt h₂] @@ -205,12 +206,12 @@ end CommSemiring section IsDomain -variable [CommRing R] [IsDomain R] {𝓟 : Ideal R} {f : R[X]} (hf : f.IsEisensteinAt 𝓟) +variable [CommRing R] [IsDomain R] {𝓟 : Ideal R} {f : R[X]} /-- If a primitive `f` satisfies `f.IsEisensteinAt 𝓟`, where `𝓟.IsPrime`, then `f` is irreducible. -/ -theorem irreducible (hprime : 𝓟.IsPrime) (hu : f.IsPrimitive) (hfd0 : 0 < f.natDegree) : - Irreducible f := +theorem irreducible (hf : f.IsEisensteinAt 𝓟) (hprime : 𝓟.IsPrime) (hu : f.IsPrimitive) + (hfd0 : 0 < f.natDegree) : Irreducible f := irreducible_of_eisenstein_criterion hprime hf.leading (fun _ hn => hf.mem (coe_lt_degree.1 hn)) (natDegree_pos_iff_degree_pos.1 hfd0) hf.not_mem hu diff --git a/Mathlib/RingTheory/RootsOfUnity/Basic.lean b/Mathlib/RingTheory/RootsOfUnity/Basic.lean index 578188521f670..fdfc31cb9f3fc 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Basic.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Basic.lean @@ -306,13 +306,13 @@ theorem mk_of_lt (ζ : M) (hk : 0 < k) (h1 : ζ ^ k = 1) (h : ∀ l : ℕ, 0 < l section CommMonoid -variable {ζ : M} {f : F} (h : IsPrimitiveRoot ζ k) +variable {ζ : M} {f : F} @[nontriviality] theorem of_subsingleton [Subsingleton M] (x : M) : IsPrimitiveRoot x 1 := ⟨Subsingleton.elim _ _, fun _ _ => one_dvd _⟩ -theorem pow_eq_one_iff_dvd (l : ℕ) : ζ ^ l = 1 ↔ k ∣ l := +theorem pow_eq_one_iff_dvd (h : IsPrimitiveRoot ζ k) (l : ℕ) : ζ ^ l = 1 ↔ k ∣ l := ⟨h.dvd_of_pow_eq_one l, by rintro ⟨i, rfl⟩; simp only [pow_mul, h.pow_eq_one, one_pow, PNat.mul_coe]⟩ @@ -320,10 +320,10 @@ theorem isUnit (h : IsPrimitiveRoot ζ k) (h0 : 0 < k) : IsUnit ζ := by apply isUnit_of_mul_eq_one ζ (ζ ^ (k - 1)) rw [← pow_succ', tsub_add_cancel_of_le h0.nat_succ_le, h.pow_eq_one] -theorem pow_ne_one_of_pos_of_lt (h0 : 0 < l) (hl : l < k) : ζ ^ l ≠ 1 := +theorem pow_ne_one_of_pos_of_lt (h : IsPrimitiveRoot ζ k) (h0 : 0 < l) (hl : l < k) : ζ ^ l ≠ 1 := mt (Nat.le_of_dvd h0 ∘ h.dvd_of_pow_eq_one _) <| not_le_of_lt hl -theorem ne_one (hk : 1 < k) : ζ ≠ 1 := +theorem ne_one (h : IsPrimitiveRoot ζ k) (hk : 1 < k) : ζ ≠ 1 := h.pow_ne_one_of_pos_of_lt zero_lt_one hk ∘ (pow_one ζ).trans theorem pow_inj (h : IsPrimitiveRoot ζ k) ⦃i j : ℕ⦄ (hi : i < k) (hj : j < k) (H : ζ ^ i = ζ ^ j) : @@ -343,7 +343,6 @@ theorem one : IsPrimitiveRoot (1 : M) 1 := @[simp] theorem one_right_iff : IsPrimitiveRoot ζ 1 ↔ ζ = 1 := by - clear h constructor · intro h; rw [← pow_one ζ, h.pow_eq_one] · rintro rfl; exact one @@ -364,8 +363,8 @@ lemma isUnit_unit {ζ : M} {n} (hn) (hζ : IsPrimitiveRoot ζ n) : lemma isUnit_unit' {ζ : G} {n} (hn) (hζ : IsPrimitiveRoot ζ n) : IsPrimitiveRoot (hζ.isUnit hn).unit' n := coe_units_iff.mp hζ --- Porting note `variable` above already contains `(h : IsPrimitiveRoot ζ k)` -theorem pow_of_coprime (i : ℕ) (hi : i.Coprime k) : IsPrimitiveRoot (ζ ^ i) k := by +theorem pow_of_coprime (h : IsPrimitiveRoot ζ k) (i : ℕ) (hi : i.Coprime k) : + IsPrimitiveRoot (ζ ^ i) k := by by_cases h0 : k = 0 · subst k; simp_all only [pow_one, Nat.coprime_zero_right] rcases h.isUnit (Nat.pos_of_ne_zero h0) with ⟨ζ, rfl⟩ @@ -405,7 +404,7 @@ protected theorem orderOf (ζ : M) : IsPrimitiveRoot ζ (orderOf ζ) := theorem unique {ζ : M} (hk : IsPrimitiveRoot ζ k) (hl : IsPrimitiveRoot ζ l) : k = l := Nat.dvd_antisymm (hk.2 _ hl.1) (hl.2 _ hk.1) -theorem eq_orderOf : k = orderOf ζ := +theorem eq_orderOf (h : IsPrimitiveRoot ζ k) : k = orderOf ζ := h.unique (IsPrimitiveRoot.orderOf ζ) protected theorem iff (hk : 0 < k) : @@ -561,9 +560,10 @@ end DivisionCommMonoid section CommRing -variable [CommRing R] {n : ℕ} (hn : 1 < n) {ζ : R} (hζ : IsPrimitiveRoot ζ n) +variable [CommRing R] {n : ℕ} {ζ : R} -theorem sub_one_ne_zero : ζ - 1 ≠ 0 := sub_ne_zero.mpr <| hζ.ne_one hn +theorem sub_one_ne_zero (hn : 1 < n) (hζ : IsPrimitiveRoot ζ n) : ζ - 1 ≠ 0 := + sub_ne_zero.mpr <| hζ.ne_one hn end CommRing diff --git a/Mathlib/Topology/Sequences.lean b/Mathlib/Topology/Sequences.lean index 701c434c73709..4729f9ccd86bf 100644 --- a/Mathlib/Topology/Sequences.lean +++ b/Mathlib/Topology/Sequences.lean @@ -189,19 +189,19 @@ theorem continuous_iff_seqContinuous [SequentialSpace X] {f : X → Y} : Continuous f ↔ SeqContinuous f := ⟨Continuous.seqContinuous, SeqContinuous.continuous⟩ -theorem SequentialSpace.coinduced [SequentialSpace X] (f : X → Y) : +theorem SequentialSpace.coinduced [SequentialSpace X] {Y} (f : X → Y) : @SequentialSpace Y (.coinduced f ‹_›) := letI : TopologicalSpace Y := .coinduced f ‹_› ⟨fun s hs ↦ isClosed_coinduced.2 (hs.preimage continuous_coinduced_rng.seqContinuous).isClosed⟩ -protected theorem SequentialSpace.iSup {ι : Sort*} {t : ι → TopologicalSpace X} +protected theorem SequentialSpace.iSup {X} {ι : Sort*} {t : ι → TopologicalSpace X} (h : ∀ i, @SequentialSpace X (t i)) : @SequentialSpace X (⨆ i, t i) := by letI : TopologicalSpace X := ⨆ i, t i refine ⟨fun s hs ↦ isClosed_iSup_iff.2 fun i ↦ ?_⟩ letI := t i exact IsSeqClosed.isClosed fun u x hus hux ↦ hs hus <| hux.mono_right <| nhds_mono <| le_iSup _ _ -protected theorem SequentialSpace.sup {t₁ t₂ : TopologicalSpace X} +protected theorem SequentialSpace.sup {X} {t₁ t₂ : TopologicalSpace X} (h₁ : @SequentialSpace X t₁) (h₂ : @SequentialSpace X t₂) : @SequentialSpace X (t₁ ⊔ t₂) := by rw [sup_eq_iSup] From b1409cd63e3857e58c473d1d9cc2433b3bb56c5c Mon Sep 17 00:00:00 2001 From: grunweg Date: Tue, 6 Aug 2024 05:13:08 +0000 Subject: [PATCH 049/359] chore: include bare 'open Classical' in tech debt metrics (#15490) And remove one commented counter: there is now a lint against deprecations without a date, so keeping it is probably not useful. --- scripts/technical-debt-metrics.sh | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/technical-debt-metrics.sh b/scripts/technical-debt-metrics.sh index 349030067d6a5..4e5f937b7aad9 100755 --- a/scripts/technical-debt-metrics.sh +++ b/scripts/technical-debt-metrics.sh @@ -55,9 +55,9 @@ done printf '%s|%s\n' "$(grep -c 'docBlame' scripts/nolints.json)" "documentation nolint entries" printf '%s|%s\n' "$(grep -c 'ERR_NUM_LIN' scripts/style-exceptions.txt)" "large files" +printf '%s|%s\n' "$(git grep "^open .*Classical" | grep -v " in$" -c)" "bare open (scoped) Classical" # We print the number of files, not the number of matches --- hence, the nested grep. printf '%s|%s\n' "$(git grep -c 'autoImplicit true' | grep -c -v 'test')" "non-test files with autoImplicit true" -#printf '%s|%s\n' "$(git grep '@\[.*deprecated' | grep -c -v 'deprecated .*(since := "')" "deprecations without a date" initFiles="$(git ls-files '**/Init/*.lean' | xargs wc -l | sed 's=^ *==')" From cea57e0da27e0976684bad5f5573acc7d93c6dc6 Mon Sep 17 00:00:00 2001 From: Nicolas Rolland Date: Tue, 6 Aug 2024 08:06:33 +0000 Subject: [PATCH 050/359] =?UTF-8?q?feat(CategoryTheory/Adjunction):=20`typ?= =?UTF-8?q?eToCat`=C2=A0is=20left=20adjoint=20to=20`Cat.objects`=20(#15375?= =?UTF-8?q?)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The embedding of `Type` into `Cat`, which views a set as a discrete category, is left adjoint to the functor `Cat.objects : Cat -> Set` mapping a category to its set of objects --- Mathlib.lean | 1 + .../Category/Cat/Adjunction.lean | 54 +++++++++++++++++++ 2 files changed, 55 insertions(+) create mode 100644 Mathlib/CategoryTheory/Category/Cat/Adjunction.lean diff --git a/Mathlib.lean b/Mathlib.lean index afdb573c1abee..1821804ea1a3b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1375,6 +1375,7 @@ import Mathlib.CategoryTheory.CatCommSq import Mathlib.CategoryTheory.Category.Basic import Mathlib.CategoryTheory.Category.Bipointed import Mathlib.CategoryTheory.Category.Cat +import Mathlib.CategoryTheory.Category.Cat.Adjunction import Mathlib.CategoryTheory.Category.Cat.Limit import Mathlib.CategoryTheory.Category.Factorisation import Mathlib.CategoryTheory.Category.GaloisConnection diff --git a/Mathlib/CategoryTheory/Category/Cat/Adjunction.lean b/Mathlib/CategoryTheory/Category/Cat/Adjunction.lean new file mode 100644 index 0000000000000..b4d89ce7a89ce --- /dev/null +++ b/Mathlib/CategoryTheory/Category/Cat/Adjunction.lean @@ -0,0 +1,54 @@ +/- +Copyright (c) 2024 Nicolas Rolland. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Nicolas Rolland +-/ +import Mathlib.CategoryTheory.Category.Cat +import Mathlib.CategoryTheory.Adjunction.Basic +/-! +# Adjunctions related to Cat, the category of categories + +The embedding `typeToCat: Type ⥤ Cat`, mapping a type to the corresponding discrete category, is +left adjoint to the functor `Cat.objects`, which maps a category to its set of objects. + + + +## Notes +All this could be made with 2-functors + +## TODO + +Define the left adjoint `Cat.connectedComponents`, which map +a category to its set of connected components. + +-/ + +universe u +namespace CategoryTheory.Cat + +variable (X : Type u) (C : Cat) + +private def typeToCatObjectsAdjHomEquiv : (typeToCat.obj X ⟶ C) ≃ (X ⟶ Cat.objects.obj C) where + toFun f x := f.obj ⟨x⟩ + invFun := Discrete.functor + left_inv F := Functor.ext (fun _ ↦ rfl) (fun ⟨_⟩ ⟨_⟩ f => by + obtain rfl := Discrete.eq_of_hom f + simp) + right_inv _ := rfl + +private def typeToCatObjectsAdjCounitApp : (Cat.objects ⋙ typeToCat).obj C ⥤ C where + obj := Discrete.as + map := eqToHom ∘ Discrete.eq_of_hom + +/-- `typeToCat : Type ⥤ Cat` is left adjoint to `Cat.objects : Cat ⥤ Type` -/ +def typeToCatObjectsAdj : typeToCat ⊣ Cat.objects where + homEquiv := typeToCatObjectsAdjHomEquiv + unit := { app:= fun _ ↦ Discrete.mk } + counit := { + app := typeToCatObjectsAdjCounitApp + naturality := fun _ _ _ ↦ Functor.hext (fun _ ↦ rfl) + (by intro ⟨_⟩ ⟨_⟩ f + obtain rfl := Discrete.eq_of_hom f + aesop_cat ) } + +end CategoryTheory.Cat From c63fd8a8ff044e2fcce00dff9e373e11650e1008 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Iv=C3=A1n=20Renison?= <85908989+IvanRenison@users.noreply.github.com> Date: Tue, 6 Aug 2024 08:16:57 +0000 Subject: [PATCH 051/359] feat(Combinatorics/SimpleGraph): Add definition of circulant graph and cycle graph (#15501) Co-authored-by: Bhavik Mehta --- Mathlib.lean | 1 + .../Combinatorics/SimpleGraph/Circulant.lean | 136 ++++++++++++++++++ 2 files changed, 137 insertions(+) create mode 100644 Mathlib/Combinatorics/SimpleGraph/Circulant.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1821804ea1a3b..1352ff3ec2672 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1901,6 +1901,7 @@ import Mathlib.Combinatorics.SetFamily.Shatter import Mathlib.Combinatorics.SimpleGraph.Acyclic import Mathlib.Combinatorics.SimpleGraph.AdjMatrix import Mathlib.Combinatorics.SimpleGraph.Basic +import Mathlib.Combinatorics.SimpleGraph.Circulant import Mathlib.Combinatorics.SimpleGraph.Clique import Mathlib.Combinatorics.SimpleGraph.Coloring import Mathlib.Combinatorics.SimpleGraph.ConcreteColorings diff --git a/Mathlib/Combinatorics/SimpleGraph/Circulant.lean b/Mathlib/Combinatorics/SimpleGraph/Circulant.lean new file mode 100644 index 0000000000000..1b44d179c5ede --- /dev/null +++ b/Mathlib/Combinatorics/SimpleGraph/Circulant.lean @@ -0,0 +1,136 @@ +/- +Copyright (c) 2024 Iván Renison, Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Iván Renison, Bhavik Mehta +-/ +import Mathlib.Combinatorics.SimpleGraph.Hasse +import Mathlib.Data.Set.Pointwise.Basic + +/-! +# Definition of circulant graphs + +This file defines and proves several fact about circulant graphs. +A circulant graph over type `G` with jumps `s : Set G` is a graph in which two vertices `u` and `v` +are adjacent if and only if `u - v ∈ s` or `v - u ∈ s`. The elements of `s` are called jumps. + +## Main declarations + +* `SimpleGraph.circulantGraph s`: the circulant graph over `G` with jumps `s`. +* `SimpleGraph.cycleGraph n`: the cycle graph over `Fin n`. +-/ + +namespace SimpleGraph + +/-- Circulant graph over additive group `G` with jumps `s` -/ +@[simps!] +def circulantGraph {G : Type*} [AddGroup G] (s : Set G) : SimpleGraph G := + fromRel (· - · ∈ s) + +variable {G : Type*} [AddGroup G] (s : Set G) + +theorem circulantGraph_eq_erase_zero : circulantGraph s = circulantGraph (s \ {0}) := by + ext (u v : G) + simp only [circulantGraph, fromRel_adj, and_congr_right_iff] + intro (h : u ≠ v) + apply Iff.intro + · intro h1 + cases h1 with + | inl h1 => exact Or.inl ⟨h1, sub_ne_zero_of_ne h⟩ + | inr h1 => exact Or.inr ⟨h1, sub_ne_zero_of_ne h.symm⟩ + · intro h1 + cases h1 with + | inl h1 => exact Or.inl h1.left + | inr h1 => exact Or.inr h1.left + +theorem circulantGraph_eq_symm : circulantGraph s = circulantGraph (s ∪ (-s)) := by + ext (u v : G) + simp only [circulantGraph, fromRel_adj, Set.mem_union, Set.mem_neg, neg_sub, and_congr_right_iff, + iff_self_or] + intro _ h + exact Or.symm h + +instance [DecidableEq G] [DecidablePred (· ∈ s)] : DecidableRel (circulantGraph s).Adj := + fun _ _ => inferInstanceAs (Decidable (_ ∧ _)) + +theorem circulantGraph_adj_translate {s : Set G} {u v d : G} : + (circulantGraph s).Adj (u + d) (v + d) ↔ (circulantGraph s).Adj u v := by simp + +/-- Cycle graph over `Fin n` -/ +def cycleGraph : (n : ℕ) → SimpleGraph (Fin n) + | 0 => ⊥ + | _ + 1 => circulantGraph {1} + +instance : (n : ℕ) → DecidableRel (cycleGraph n).Adj + | 0 => fun _ _ => inferInstanceAs (Decidable False) + | _ + 1 => inferInstanceAs (DecidableRel (circulantGraph _).Adj) + +theorem cycleGraph_zero_adj {u v : Fin 0} : ¬(cycleGraph 0).Adj u v := id + +theorem cycleGraph_zero_eq_bot : cycleGraph 0 = ⊥ := Subsingleton.elim _ _ +theorem cycleGraph_one_eq_bot : cycleGraph 1 = ⊥ := Subsingleton.elim _ _ +theorem cycleGraph_zero_eq_top : cycleGraph 0 = ⊤ := Subsingleton.elim _ _ +theorem cycleGraph_one_eq_top : cycleGraph 1 = ⊤ := Subsingleton.elim _ _ + +theorem cycleGraph_two_eq_top : cycleGraph 2 = ⊤ := by + simp only [SimpleGraph.ext_iff, Function.funext_iff] + decide + +theorem cycleGraph_three_eq_top : cycleGraph 3 = ⊤ := by + simp only [SimpleGraph.ext_iff, Function.funext_iff] + decide + +theorem cycleGraph_one_adj {u v : Fin 1} : ¬(cycleGraph 1).Adj u v := by + rw [cycleGraph_one_eq_bot] + exact id + +theorem cycleGraph_adj {n : ℕ} {u v : Fin (n + 2)} : + (cycleGraph (n + 2)).Adj u v ↔ u - v = 1 ∨ v - u = 1 := by + simp only [cycleGraph, circulantGraph_adj, Set.mem_singleton_iff, and_iff_right_iff_imp] + intro _ _ + simp_all + +theorem cycleGraph_adj' {n : ℕ} {u v : Fin n} : + (cycleGraph n).Adj u v ↔ (u - v).val = 1 ∨ (v - u).val = 1 := by + match n with + | 0 => exact u.elim0 + | 1 => simp [cycleGraph_one_adj] + | n + 2 => simp [cycleGraph_adj, Fin.ext_iff] + +theorem cycleGraph_neighborSet {n : ℕ} {v : Fin (n + 2)} : + (cycleGraph (n + 2)).neighborSet v = {v - 1, v + 1} := by + ext w + simp only [mem_neighborSet, Set.mem_insert_iff, Set.mem_singleton_iff] + rw [cycleGraph_adj, sub_eq_iff_eq_add', sub_eq_iff_eq_add', eq_sub_iff_add_eq, eq_comm] + +theorem cycleGraph_neighborFinset {n : ℕ} {v : Fin (n + 2)} : + (cycleGraph (n + 2)).neighborFinset v = {v - 1, v + 1} := by + simp [neighborFinset, cycleGraph_neighborSet] + +theorem cycleGraph_degree_two_le {n : ℕ} {v : Fin (n + 2)} : + (cycleGraph (n + 2)).degree v = Finset.card {v - 1, v + 1} := by + rw [SimpleGraph.degree, cycleGraph_neighborFinset] + +theorem cycleGraph_degree_three_le {n : ℕ} {v : Fin (n + 3)} : + (cycleGraph (n + 3)).degree v = 2 := by + rw [cycleGraph_degree_two_le, Finset.card_pair] + simp only [ne_eq, sub_eq_iff_eq_add, add_assoc v, self_eq_add_right] + exact ne_of_beq_false rfl + +theorem pathGraph_le_cycleGraph {n : ℕ} : pathGraph n ≤ cycleGraph n := by + match n with + | 0 | 1 => simp + | n + 2 => + intro u v h + rw [pathGraph_adj] at h + rw [cycleGraph_adj'] + cases h with + | inl h | inr h => + simp [Fin.coe_sub_iff_le.mpr (Nat.lt_of_succ_le h.le).le, Nat.eq_sub_of_add_eq' h] + +theorem cycleGraph_preconnected {n : ℕ} : (cycleGraph n).Preconnected := + (pathGraph_preconnected n).mono pathGraph_le_cycleGraph + +theorem cycleGraph_connected {n : ℕ} : (cycleGraph (n + 1)).Connected := + (pathGraph_connected n).mono pathGraph_le_cycleGraph + +end SimpleGraph From 26fa8d86c3e1a5e69165db1d3767abe6d1265273 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 6 Aug 2024 08:16:58 +0000 Subject: [PATCH 052/359] chore: tidy various files (#15516) --- .../Algebra/Module/LinearMap/Polynomial.lean | 3 ++- Mathlib/Algebra/Module/Zlattice/Basic.lean | 3 +-- Mathlib/Algebra/Module/Zlattice/Covolume.lean | 6 ++--- Mathlib/Algebra/MvPolynomial/Degrees.lean | 10 ++++---- .../Algebra/Order/Group/Unbundled/Abs.lean | 1 - .../Order/GroupWithZero/Canonical.lean | 6 ++--- Mathlib/Algebra/Order/Interval/Basic.lean | 4 ++-- .../PrimeSpectrum/Basic.lean | 1 - Mathlib/Analysis/Analytic/Within.lean | 2 -- Mathlib/Analysis/LocallyConvex/Barrelled.lean | 2 +- .../Sites/Coherent/Equivalence.lean | 4 ++-- .../Sites/MayerVietorisSquare.lean | 2 +- .../Additive/ErdosGinzburgZiv.lean | 10 ++++---- Mathlib/Combinatorics/SimpleGraph/Metric.lean | 12 +++++----- Mathlib/Data/Matrix/Basis.lean | 23 ++++++------------ Mathlib/Data/Num/Lemmas.lean | 3 ++- Mathlib/Data/Set/Finite.lean | 2 +- Mathlib/Data/TypeVec.lean | 6 ++--- .../Geometry/RingedSpace/OpenImmersion.lean | 7 ++---- Mathlib/GroupTheory/CosetCover.lean | 10 ++++---- Mathlib/GroupTheory/PGroup.lean | 11 +++------ Mathlib/GroupTheory/Perm/Cycle/Concrete.lean | 7 ++---- .../GroupTheory/SpecificGroups/Cyclic.lean | 1 - Mathlib/LinearAlgebra/Matrix/ToLin.lean | 3 +-- .../Measure/Haar/InnerProductSpace.lean | 4 ++-- .../Probability/ConditionalProbability.lean | 3 +-- Mathlib/Probability/Kernel/Basic.lean | 5 ++-- Mathlib/RingTheory/Finiteness.lean | 10 ++++---- Mathlib/RingTheory/Flat/Stability.lean | 4 ++-- .../RingTheory/Kaehler/CotangentComplex.lean | 9 ++++--- .../RingTheory/MvPolynomial/Homogeneous.lean | 7 +++--- .../MvPolynomial/WeightedHomogeneous.lean | 24 +++++++------------ Mathlib/RingTheory/PrimeSpectrum.lean | 2 +- Mathlib/SetTheory/Ordinal/Notation.lean | 4 ++-- 34 files changed, 84 insertions(+), 127 deletions(-) diff --git a/Mathlib/Algebra/Module/LinearMap/Polynomial.lean b/Mathlib/Algebra/Module/LinearMap/Polynomial.lean index 67d94c358dafa..4f40705b3285b 100644 --- a/Mathlib/Algebra/Module/LinearMap/Polynomial.lean +++ b/Mathlib/Algebra/Module/LinearMap/Polynomial.lean @@ -280,7 +280,8 @@ open LinearMap in lemma polyCharpolyAux_map_eq_toMatrix_charpoly (x : L) : (polyCharpolyAux φ b bₘ).map (MvPolynomial.eval (b.repr x)) = (toMatrix bₘ bₘ (φ x)).charpoly := by - erw [polyCharpolyAux, Polynomial.map_map, MvPolynomial.comp_eval₂Hom, charpoly.univ_map_eval₂Hom] + rw [polyCharpolyAux, Polynomial.map_map, ← MvPolynomial.eval₂Hom_C_eq_bind₁, + MvPolynomial.comp_eval₂Hom, charpoly.univ_map_eval₂Hom] congr ext rw [of_apply, Function.curry_apply, toMvPolynomial_eval_eq_apply, LinearEquiv.symm_apply_apply] diff --git a/Mathlib/Algebra/Module/Zlattice/Basic.lean b/Mathlib/Algebra/Module/Zlattice/Basic.lean index 4025c1293c57a..d4968d0b69674 100644 --- a/Mathlib/Algebra/Module/Zlattice/Basic.lean +++ b/Mathlib/Algebra/Module/Zlattice/Basic.lean @@ -599,8 +599,7 @@ theorem Zlattice.isAddFundamentalDomain {E : Type*} [NormedAddCommGroup E] [Norm instance instCountable_of_discrete_addSubgroup {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] (L : AddSubgroup E) [DiscreteTopology L] [IsZlattice ℝ L] : Countable L := by - rw [← (Module.Free.chooseBasis ℤ L).ofZlatticeBasis_span ℝ] - change Countable (span ℤ (Set.range (Basis.ofZlatticeBasis ℝ L _))) + simp_rw [← (Module.Free.chooseBasis ℤ L).ofZlatticeBasis_span ℝ, mem_toAddSubgroup] infer_instance end Zlattice diff --git a/Mathlib/Algebra/Module/Zlattice/Covolume.lean b/Mathlib/Algebra/Module/Zlattice/Covolume.lean index b922c5592c6eb..8a3b7297b4d77 100644 --- a/Mathlib/Algebra/Module/Zlattice/Covolume.lean +++ b/Mathlib/Algebra/Module/Zlattice/Covolume.lean @@ -68,9 +68,9 @@ theorem covolume_eq_det_mul_measure {ι : Type*} [Fintype ι] [DecidableEq ι] ( (b₀ : Basis ι ℝ E) : covolume L μ = |b₀.det ((↑) ∘ b)| * (μ (Zspan.fundamentalDomain b₀)).toReal := by rw [covolume_eq_measure_fundamentalDomain L μ (isAddFundamentalDomain b μ), - Zspan.measure_fundamentalDomain _ _ b₀, measure_congr - (Zspan.fundamentalDomain_ae_parallelepiped b₀ μ), ENNReal.toReal_mul, ENNReal.toReal_ofReal - (by positivity)] + Zspan.measure_fundamentalDomain _ _ b₀, + measure_congr (Zspan.fundamentalDomain_ae_parallelepiped b₀ μ), ENNReal.toReal_mul, + ENNReal.toReal_ofReal (by positivity)] congr ext exact b.ofZlatticeBasis_apply ℝ L _ diff --git a/Mathlib/Algebra/MvPolynomial/Degrees.lean b/Mathlib/Algebra/MvPolynomial/Degrees.lean index a47dbbb4d9829..388b7941515bf 100644 --- a/Mathlib/Algebra/MvPolynomial/Degrees.lean +++ b/Mathlib/Algebra/MvPolynomial/Degrees.lean @@ -402,8 +402,7 @@ theorem totalDegree_monomial_le (s : σ →₀ ℕ) (c : R) : if hc : c = 0 then simp only [hc, map_zero, totalDegree_zero, zero_le] else - rw [totalDegree_monomial _ hc] - exact le_rfl + rw [totalDegree_monomial _ hc, Function.id_def] @[simp] theorem totalDegree_X_pow [Nontrivial R] (s : σ) (n : ℕ) : @@ -411,9 +410,9 @@ theorem totalDegree_X_pow [Nontrivial R] (s : σ) (n : ℕ) : theorem totalDegree_list_prod : ∀ s : List (MvPolynomial σ R), s.prod.totalDegree ≤ (s.map MvPolynomial.totalDegree).sum - | [] => by rw [@List.prod_nil (MvPolynomial σ R) _, totalDegree_one]; rfl + | [] => by rw [List.prod_nil, totalDegree_one, List.map_nil, List.sum_nil] | p::ps => by - rw [@List.prod_cons (MvPolynomial σ R) _, List.map, List.sum_cons] + rw [List.prod_cons, List.map, List.sum_cons] exact le_trans (totalDegree_mul _ _) (add_le_add_left (totalDegree_list_prod ps) _) theorem totalDegree_multiset_prod (s : Multiset (MvPolynomial σ R)) : @@ -425,8 +424,7 @@ theorem totalDegree_multiset_prod (s : Multiset (MvPolynomial σ R)) : theorem totalDegree_finset_prod {ι : Type*} (s : Finset ι) (f : ι → MvPolynomial σ R) : (s.prod f).totalDegree ≤ ∑ i ∈ s, (f i).totalDegree := by refine le_trans (totalDegree_multiset_prod _) ?_ - rw [Multiset.map_map] - rfl + simp only [Multiset.map_map, comp_apply, Finset.sum_map_val, le_refl] theorem totalDegree_finset_sum {ι : Type*} (s : Finset ι) (f : ι → MvPolynomial σ R) : (s.sum f).totalDegree ≤ Finset.sup s fun i => (f i).totalDegree := by diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean index 4c6afac4c5ba7..270be2e5a791c 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean @@ -5,7 +5,6 @@ Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import Mathlib.Algebra.Group.Even import Mathlib.Algebra.Order.Group.Lattice --- import Mathlib.Algebra.Order.Group.Defs /-! # Absolute values in ordered groups diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index bbd29978dc36c..fa19028c7999f 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -309,7 +309,6 @@ instance covariantClass_mul_le [Mul α] [CovariantClass α α (· * ·) (· ≤ rw [← coe_mul _ c, ← coe_mul, coe_le_coe] exact mul_le_mul_left' hbc' _ --- Porting note: same issue as `covariantClass_mul_le` protected lemma covariantClass_add_le [AddZeroClass α] [CovariantClass α α (· + ·) (· ≤ ·)] (h : ∀ a : α, 0 ≤ a) : CovariantClass (WithZero α) (WithZero α) (· + ·) (· ≤ ·) := by refine ⟨fun a b c hbc => ?_⟩ @@ -321,9 +320,8 @@ protected lemma covariantClass_add_le [AddZeroClass α] [CovariantClass α α ( · rw [add_zero] · rw [← coe_add, coe_le_coe] exact le_add_of_nonneg_right (h _) - · rcases WithBot.coe_le_iff.1 hbc with ⟨c, rfl, hbc'⟩ - refine le_trans ?_ (le_of_eq <| coe_add _ _) - rw [← coe_add, coe_le_coe] + · rcases WithZero.coe_le_iff.1 hbc with ⟨c, rfl, hbc'⟩ + rw [← coe_add, ← coe_add _ c, coe_le_coe] exact add_le_add_left hbc' _ instance existsAddOfLE [Add α] [ExistsAddOfLE α] : ExistsAddOfLE (WithZero α) := diff --git a/Mathlib/Algebra/Order/Interval/Basic.lean b/Mathlib/Algebra/Order/Interval/Basic.lean index 2103932e456c4..005379a7fa950 100644 --- a/Mathlib/Algebra/Order/Interval/Basic.lean +++ b/Mathlib/Algebra/Order/Interval/Basic.lean @@ -187,8 +187,8 @@ variable [Monoid α] [Preorder α] @[to_additive existing] instance NonemptyInterval.hasPow - [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] : - Pow (NonemptyInterval α) ℕ := + [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] : + Pow (NonemptyInterval α) ℕ := ⟨fun s n => ⟨s.toProd ^ n, pow_le_pow_left' s.fst_le_snd _⟩⟩ namespace NonemptyInterval diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 91c08ba82ce51..a9bd9774ad781 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -532,7 +532,6 @@ section Order ## The specialization order We endow `PrimeSpectrum R` with a partial order, where `x ≤ y` if and only if `y ∈ closure {x}`. -This instance was defined in `RingTheory/PrimeSpectrum/Basic`. -/ theorem le_iff_mem_closure (x y : PrimeSpectrum R) : diff --git a/Mathlib/Analysis/Analytic/Within.lean b/Mathlib/Analysis/Analytic/Within.lean index 630e2ae62bc1d..a35cbce9538cf 100644 --- a/Mathlib/Analysis/Analytic/Within.lean +++ b/Mathlib/Analysis/Analytic/Within.lean @@ -30,8 +30,6 @@ open Topology Filter ENNReal open Set Filter -variable {α : Type*} - variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] variable {E F G H : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] [NormedAddCommGroup F] diff --git a/Mathlib/Analysis/LocallyConvex/Barrelled.lean b/Mathlib/Analysis/LocallyConvex/Barrelled.lean index c92b0399ae1c7..ea40aa6ca9fd7 100644 --- a/Mathlib/Analysis/LocallyConvex/Barrelled.lean +++ b/Mathlib/Analysis/LocallyConvex/Barrelled.lean @@ -15,7 +15,7 @@ Banach-Steinhaus theorem for maps from a barrelled space to a space equipped wit of seminorms generating the topology (i.e `WithSeminorms q` for some family of seminorms `q`). The more standard Banach-Steinhaus theorem for normed spaces is then deduced from that in -`Analysis.Normed.Operator.BanachSteinhaus`. +`Mathlib.Analysis.Normed.Operator.BanachSteinhaus`. ## Main definitions diff --git a/Mathlib/CategoryTheory/Sites/Coherent/Equivalence.lean b/Mathlib/CategoryTheory/Sites/Coherent/Equivalence.lean index 45db74b8177aa..f5e91a85cbb1e 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/Equivalence.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/Equivalence.lean @@ -38,7 +38,7 @@ instance : haveI := precoherent e e.inverse.IsDenseSubsite (coherentTopology D) (coherentTopology C) where functorPushforward_mem_iff := by rw [coherentTopology.eq_induced e.inverse] - rfl + simp only [Functor.mem_inducedTopology_sieves_iff, implies_true] variable (A : Type*) [Category A] @@ -86,7 +86,7 @@ instance : haveI := preregular e e.inverse.IsDenseSubsite (regularTopology D) (regularTopology C) where functorPushforward_mem_iff := by rw [regularTopology.eq_induced e.inverse] - rfl + simp only [Functor.mem_inducedTopology_sieves_iff, implies_true] variable (A : Type*) [Category A] diff --git a/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean b/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean index b89884b96132b..f6b461edc0b5d 100644 --- a/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean +++ b/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean @@ -19,7 +19,7 @@ where `X₁` is the intersection of `X₂` and `X₃`, and `H^q` are the cohomology groups with values in an abelian sheaf. In this file, we introduce a structure -`GrothendieckTopology.MayerVietorisSquare` which extends `Squace C`, +`GrothendieckTopology.MayerVietorisSquare` which extends `Square C`, and asserts properties which shall imply the existence of long exact Mayer-Vietoris sequences in sheaf cohomology (TODO). We require that the map `X₁ ⟶ X₃` is a monomorphism and diff --git a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean index 09fd0c79b7e07..6359c2a1ca925 100644 --- a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean +++ b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean @@ -110,13 +110,13 @@ theorem Int.erdos_ginzburg_ziv (a : ι → ℤ) (hs : 2 * n - 1 ≤ s.card) : classical -- Do induction on the prime factorisation of `n`. Note that we will apply the induction -- hypothesis with `ι := Finset ι`, so we need to generalise. - induction n using Nat.prime_composite_induction generalizing ι + induction n using Nat.prime_composite_induction generalizing ι with -- When `n := 0`, we can set `t := ∅`. - case zero => exact ⟨∅, by simp⟩ + | zero => exact ⟨∅, by simp⟩ -- When `n := 1`, we can take `t` to be any subset of `s` of size `2 * n - 1`. - case one => simpa using exists_subset_card_eq hs + | one => simpa using exists_subset_card_eq hs -- When `n := p` is prime, we use the prime case `Int.erdos_ginzburg_ziv_prime`. - case prime p hp => + | prime p hp => haveI := Fact.mk hp obtain ⟨t, hts, ht⟩ := exists_subset_card_eq hs obtain ⟨u, hut, hu⟩ := Int.erdos_ginzburg_ziv_prime a ht @@ -124,7 +124,7 @@ theorem Int.erdos_ginzburg_ziv (a : ι → ℤ) (hs : 2 * n - 1 ≤ s.card) : -- When `n := m * n` is composite, we pick (by induction hypothesis on `n`) `2 * m - 1` sets of -- size `n` and sums divisible by `n`. Then by induction hypothesis (on `m`) we can pick `m` of -- these sets whose sum is divisible by `m * n`. - case composite m hm ihm n hn ihn => + | composite m hm ihm n hn ihn => -- First, show that it is enough to have those `2 * m - 1` sets. suffices ∀ k ≤ 2 * m - 1, ∃ 𝒜 : Finset (Finset ι), 𝒜.card = k ∧ (𝒜 : Set (Finset ι)).Pairwise _root_.Disjoint ∧ diff --git a/Mathlib/Combinatorics/SimpleGraph/Metric.lean b/Mathlib/Combinatorics/SimpleGraph/Metric.lean index 5fac07b81db53..9b21f7d272b05 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Metric.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Metric.lean @@ -90,12 +90,12 @@ lemma exists_walk_of_edist_ne_top (h : G.edist u v ≠ ⊤) : (reachable_of_edist_ne_top h).exists_walk_length_eq_edist protected theorem edist_triangle : G.edist u w ≤ G.edist u v + G.edist v w := by - rcases eq_or_ne (G.edist u v) ⊤ with huv | huv - case inl => simp [huv] - case inr => - rcases eq_or_ne (G.edist v w) ⊤ with hvw | hvw - case inl => simp [hvw] - case inr => + cases eq_or_ne (G.edist u v) ⊤ with + | inl huv => simp [huv] + | inr huv => + cases eq_or_ne (G.edist v w) ⊤ with + | inl hvw => simp [hvw] + | inr hvw => obtain ⟨p, hp⟩ := exists_walk_of_edist_ne_top huv obtain ⟨q, hq⟩ := exists_walk_of_edist_ne_top hvw rw [← hp, ← hq, ← Nat.cast_add, ← Walk.length_append] diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index a1614750e15a4..104482a97f690 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -60,17 +60,12 @@ theorem matrix_eq_sum_std_basis [Fintype m] [Fintype n] (x : Matrix m n α) : x = ∑ i : m, ∑ j : n, stdBasisMatrix i j (x i j) := by ext i j; symm iterate 2 rw [Finset.sum_apply] - -- Porting note: was `convert` - refine (Fintype.sum_eq_single i ?_).trans ?_; swap - · -- Porting note: `simp` seems unwilling to apply `Fintype.sum_apply` - simp (config := { unfoldPartialApp := true }) only [stdBasisMatrix] - rw [Fintype.sum_apply, Fintype.sum_apply] - simp + convert (Fintype.sum_eq_single i ?_).trans ?_; swap + · -- Porting note(#12717): `simp` seems unwilling to apply `Fintype.sum_apply` + simp (config := { unfoldPartialApp := true }) [stdBasisMatrix, (Fintype.sum_apply)] · intro j' hj' - -- Porting note: `simp` seems unwilling to apply `Fintype.sum_apply` - simp (config := { unfoldPartialApp := true }) only [stdBasisMatrix] - rw [Fintype.sum_apply, Fintype.sum_apply] - simp [hj'] + -- Porting note(#12717): `simp` seems unwilling to apply `Fintype.sum_apply` + simp (config := { unfoldPartialApp := true }) [stdBasisMatrix, (Fintype.sum_apply), hj'] -- TODO: tie this up with the `Basis` machinery of linear algebra -- this is not completely trivial because we are indexing by two types, instead of one @@ -79,12 +74,8 @@ theorem std_basis_eq_basis_mul_basis (i : m) (j : n) : stdBasisMatrix i j (1 : α) = vecMulVec (fun i' => ite (i = i') 1 0) fun j' => ite (j = j') 1 0 := by ext i' j' - -- Porting note: was `norm_num [std_basis_matrix, vec_mul_vec]` though there are no numerals - -- involved. - simp only [stdBasisMatrix, vecMulVec, mul_ite, mul_one, mul_zero, of_apply] - -- Porting note: added next line - simp_rw [@and_comm (i = i')] - exact ite_and _ _ _ _ + -- Porting note: lean3 didn't apply `mul_ite`. + simp [-mul_ite, stdBasisMatrix, vecMulVec, ite_and] -- todo: the old proof used fintypes, I don't know `Finsupp` but this feels generalizable @[elab_as_elim] diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index 07e422e8a1a7f..ac22db399e145 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -1414,7 +1414,8 @@ theorem divMod_to_nat (d n : PosNum) : revert IH; cases' divMod d n with q r; intro IH simp only [divMod] at IH ⊢ apply divMod_to_nat_aux - · simp; rw [← two_mul, ← two_mul, mul_left_comm, ← mul_add, ← IH.1] + · simp only [Num.cast_bit0, cast_bit0] + rw [← two_mul, ← two_mul, mul_left_comm, ← mul_add, ← IH.1] · simpa using IH.2 @[simp] diff --git a/Mathlib/Data/Set/Finite.lean b/Mathlib/Data/Set/Finite.lean index 1f80edc2409d5..933b71188dd78 100644 --- a/Mathlib/Data/Set/Finite.lean +++ b/Mathlib/Data/Set/Finite.lean @@ -768,7 +768,7 @@ theorem Finite.of_preimage (h : (f ⁻¹' s).Finite) (hf : Surjective f) : s.Fin theorem Finite.preimage (I : Set.InjOn f (f ⁻¹' s)) (h : s.Finite) : (f ⁻¹' s).Finite := (h.subset (image_preimage_subset f s)).of_finite_image I -theorem Finite.preimage' (h : s.Finite) (hf : ∀ b ∈ s, (f ⁻¹' {b}).Finite) : +theorem Finite.preimage' (h : s.Finite) (hf : ∀ b ∈ s, (f ⁻¹' {b}).Finite) : (f ⁻¹' s).Finite := by rw [← Set.biUnion_preimage_singleton] exact Set.Finite.biUnion h hf diff --git a/Mathlib/Data/TypeVec.lean b/Mathlib/Data/TypeVec.lean index 1a377872c5e33..104d68020ce33 100644 --- a/Mathlib/Data/TypeVec.lean +++ b/Mathlib/Data/TypeVec.lean @@ -657,10 +657,8 @@ theorem subtypeVal_diagSub {α : TypeVec n} : subtypeVal (repeatEq α) ⊚ diagS ext i x induction' i with _ _ _ i_ih · simp [comp, diagSub, subtypeVal, prod.diag] - · simp only [prod.diag] - simp? [comp, diagSub, subtypeVal, prod.diag] at * says - simp only [comp, subtypeVal, diagSub] at * - apply @i_ih (drop _) + · simp only [comp, subtypeVal, diagSub, prod.diag] at * + apply i_ih @[simp] theorem toSubtype_of_subtype {α : TypeVec n} (p : α ⟹ «repeat» n Prop) : diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index edd9564421caf..bd741b1f3e58f 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -156,11 +156,8 @@ instance comp {Z : PresheafedSpace C} (g : Y ⟶ Z) [hg : IsOpenImmersion g] : · exact c_iso' g ((opensFunctor f).obj U) (by ext; simp) · apply c_iso' f U ext1 - dsimp only [Opens.map_coe, IsOpenMap.functor_obj_coe, comp_base] - -- Porting note: slightly more hand holding here: `g ∘ f` and `fun x => g (f x)` - erw [coe_comp, show g.base ∘ f.base = fun x => g.base (f.base x) from rfl, - ← Set.image_image g.base f.base, Set.preimage_image_eq _ hg.base_open.inj] - -- now `erw` after #13170 + dsimp only [Opens.map_coe, IsOpenMap.functor_obj_coe, comp_base, TopCat.coe_comp] + rw [Set.image_comp, Set.preimage_image_eq _ hg.base_open.inj] /-- For an open immersion `f : X ⟶ Y` and an open set `U ⊆ X`, we have the map `X(U) ⟶ Y(U)`. -/ noncomputable def invApp (U : Opens X) : diff --git a/Mathlib/GroupTheory/CosetCover.lean b/Mathlib/GroupTheory/CosetCover.lean index 3e09c58c89a09..b5c8e0f4f04ee 100644 --- a/Mathlib/GroupTheory/CosetCover.lean +++ b/Mathlib/GroupTheory/CosetCover.lean @@ -149,7 +149,7 @@ theorem exists_finiteIndex_of_leftCoset_cover_aux [DecidableEq (Subgroup G)] let κ := ↥(s.filter (H · ≠ H j)) × Option ↥(s.filter (H · = H j)) let f : κ → G | ⟨k₁, some k₂⟩ => g k₂ * x⁻¹ * g k₁ - | ⟨k₁, none⟩ => g k₁ + | ⟨k₁, none⟩ => g k₁ let K (k : κ) : Subgroup G := H k.1.val have hK' (k : κ) : K k ∈ (s.image H).erase (H j) := by have := Finset.mem_filter.mp k.1.property @@ -332,10 +332,10 @@ theorem exists_index_le_card_of_leftCoset_cover : ∃ i ∈ s, (H i).FiniteIndex ∧ (H i).index ≤ s.card := by by_contra! h apply (one_le_sum_inv_index_of_leftCoset_cover hcovers).not_lt - by_cases hs : s = ∅ - · simp only [hs, Finset.sum_empty, show (0 : ℚ) < 1 from rfl] - rw [← ne_eq, ← Finset.nonempty_iff_ne_empty] at hs - have hs' : 0 < s.card := Nat.pos_of_ne_zero (Finset.card_ne_zero.mpr hs) + cases s.eq_empty_or_nonempty with + | inl hs => simp only [hs, Finset.sum_empty, zero_lt_one] + | inr hs => + have hs' : 0 < s.card := hs.card_pos have hlt : ∀ i ∈ s, ((H i).index : ℚ)⁻¹ < (s.card : ℚ)⁻¹ := fun i hi ↦ by cases eq_or_ne (H i).index 0 with | inl hindex => diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index f5d02b151cb0c..6a0ff7eea9fcc 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -176,14 +176,9 @@ theorem card_modEq_card_fixedPoints : Nat.card α ≡ Nat.card (fixedPoints G α rw [key, mem_fixedPoints_iff_card_orbit_eq_one.mp a.2]) obtain ⟨k, hk⟩ := hG.card_orbit b rw [Nat.card_eq_fintype_card] at hk - have : k = 0 := - Nat.le_zero.1 - (Nat.le_of_lt_succ - (lt_of_not_ge - (mt (pow_dvd_pow p) - (by - rwa [pow_one, ← hk, ← Nat.modEq_zero_iff_dvd, ← ZMod.eq_iff_modEq_nat, ← key, - Nat.cast_zero])))) + have : k = 0 := by + contrapose! hb + simp [-Quotient.eq'', key, hk, hb] exact ⟨⟨b, mem_fixedPoints_iff_card_orbit_eq_one.2 <| by rw [hk, this, pow_zero]⟩, Finset.mem_univ _, ne_of_eq_of_ne Nat.cast_one one_ne_zero, rfl⟩ diff --git a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean index 9792193750a29..924cd62e58286 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean @@ -179,11 +179,8 @@ nonrec theorem formPerm_eq_formPerm_iff {α : Type*} [DecidableEq α] {s s' : Cy revert s s' intro s s' apply @Quotient.inductionOn₂' _ _ _ _ _ s s' - intro l l' - -- Porting note: was `simpa using formPerm_eq_formPerm_iff` - simp only [mk''_eq_coe, nodup_coe_iff, formPerm_coe, coe_eq_coe, length_coe] - intro hs hs' - constructor <;> intro h <;> simp_all only [formPerm_eq_formPerm_iff] + intro l l' hl hl' + simpa using formPerm_eq_formPerm_iff hl hl' end Cycle diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index f32c588e807ba..749712cd51ac2 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -450,7 +450,6 @@ theorem card_orderOf_eq_totient_aux₂ {d : ℕ} (hd : d ∣ Fintype.card α) : apply card_orderOf_eq_totient_aux₁ hn hd by_contra h0 -- Must qualify `Finset.card_eq_zero` because of leanprover/lean4#2849 - -- Must specify the argument `α` to avoid mathlib4#10830 simp_rw [not_lt, Nat.le_zero, Finset.card_eq_zero] at h0 apply lt_irrefl c calc diff --git a/Mathlib/LinearAlgebra/Matrix/ToLin.lean b/Mathlib/LinearAlgebra/Matrix/ToLin.lean index fcce5c87672f1..d0cc4f7aa9d4b 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLin.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLin.lean @@ -954,12 +954,11 @@ attribute [simp] linearMap_repr_apply lemma linearMap_apply (ij : ι₂ × ι₁) : (b₁.linearMap b₂ ij) = (Matrix.toLin b₁ b₂) (Matrix.stdBasis R ι₂ ι₁ ij) := by - erw [linearMap_repr_symm_apply, Finsupp.total_single, one_smul] + simp [linearMap] lemma linearMap_apply_apply (ij : ι₂ × ι₁) (k : ι₁) : (b₁.linearMap b₂ ij) (b₁ k) = if ij.2 = k then b₂ ij.1 else 0 := by have := Classical.decEq ι₂ - rcases ij with ⟨i, j⟩ rw [linearMap_apply, Matrix.stdBasis_eq_stdBasisMatrix, Matrix.toLin_self] dsimp only [Matrix.stdBasisMatrix] simp_rw [ite_smul, one_smul, zero_smul, ite_and, Finset.sum_ite_eq, Finset.mem_univ, if_true] diff --git a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean index e6af597345ae0..cf60664ba60af 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean @@ -115,8 +115,8 @@ theorem PiLp.volume_preserving_equiv_symm : MeasurePreserving (WithLp.equiv 2 ( lemma volume_euclideanSpace_eq_dirac [IsEmpty ι] : (volume : Measure (EuclideanSpace ℝ ι)) = Measure.dirac 0 := by rw [← ((EuclideanSpace.volume_preserving_measurableEquiv ι).symm).map_eq, - volume_pi_eq_dirac 0, map_dirac (MeasurableEquiv.measurable _)] - rfl + volume_pi_eq_dirac 0, map_dirac (MeasurableEquiv.measurable _), + EuclideanSpace.coe_measurableEquiv_symm, WithLp.equiv_symm_zero] end PiLp diff --git a/Mathlib/Probability/ConditionalProbability.lean b/Mathlib/Probability/ConditionalProbability.lean index de801087d5074..7a9f4513f7f3a 100644 --- a/Mathlib/Probability/ConditionalProbability.lean +++ b/Mathlib/Probability/ConditionalProbability.lean @@ -224,8 +224,7 @@ lemma sum_meas_smul_cond_fiber {X : Ω → α} (hX : Measurable X) (μ : Measure Pi.smul_apply, smul_eq_mul] simp_rw [mul_comm (μ _), cond_mul_eq_inter _ (hX (.singleton _))] _ = _ := by - have : ⋃ x ∈ Finset.univ, X ⁻¹' {x} ∩ E = E := by - simp only [Finset.mem_univ, iUnion_true]; ext _; simp + have : ⋃ x ∈ Finset.univ, X ⁻¹' {x} ∩ E = E := by ext; simp rw [← measure_biUnion_finset _ fun _ _ ↦ (hX (.singleton _)).inter hE, this] aesop (add simp [PairwiseDisjoint, Set.Pairwise, Function.onFun, disjoint_left]) diff --git a/Mathlib/Probability/Kernel/Basic.lean b/Mathlib/Probability/Kernel/Basic.lean index bb952329d7bd0..87c0e64295061 100644 --- a/Mathlib/Probability/Kernel/Basic.lean +++ b/Mathlib/Probability/Kernel/Basic.lean @@ -215,10 +215,9 @@ protected theorem measurable_coe (κ : Kernel α β) {s : Set β} (hs : Measurab (Measure.measurable_coe hs).comp κ.measurable lemma apply_congr_of_mem_measurableAtom (κ : Kernel α β) {y' y : α} (hy' : y' ∈ measurableAtom y) : - κ y' = κ y := by + κ y' = κ y := by ext s hs - exact mem_of_mem_measurableAtom hy' - (κ.measurable_coe hs (measurableSet_singleton (κ y s))) rfl + exact mem_of_mem_measurableAtom hy' (κ.measurable_coe hs (measurableSet_singleton (κ y s))) rfl lemma IsFiniteKernel.integrable (μ : Measure α) [IsFiniteMeasure μ] (κ : Kernel α β) [IsFiniteKernel κ] {s : Set β} (hs : MeasurableSet s) : diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index 374abb864e2cc..c0a3dcf78bc80 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -626,11 +626,11 @@ theorem Module.End.isNilpotent_iff_of_finite {R M : Type*} [CommSemiring R] [Add use Finset.sup S g ext m have hm : m ∈ Submodule.span R S := by simp [hS] - induction hm using Submodule.span_induction' - · next x hx => exact LinearMap.pow_map_zero_of_le (Finset.le_sup hx) (hg x) - · simp - · simp_all - · simp_all + induction hm using Submodule.span_induction' with + | mem x hx => exact LinearMap.pow_map_zero_of_le (Finset.le_sup hx) (hg x) + | zero => simp + | add => simp_all + | smul => simp_all variable {R} diff --git a/Mathlib/RingTheory/Flat/Stability.lean b/Mathlib/RingTheory/Flat/Stability.lean index c1d4983e9f581..27ba15c69ed04 100644 --- a/Mathlib/RingTheory/Flat/Stability.lean +++ b/Mathlib/RingTheory/Flat/Stability.lean @@ -157,8 +157,8 @@ variable {R : Type u} {M Mp : Type*} (Rp : Type v) instance localizedModule [Module.Flat R M] (S : Submonoid R) : Module.Flat (Localization S) (LocalizedModule S M) := by - fapply Module.Flat.isBaseChange (R := R) (M := M) (S := Localization S) - exact LocalizedModule.mkLinearMap S M + apply Module.Flat.isBaseChange (R := R) (S := Localization S) + (f := LocalizedModule.mkLinearMap S M) rw [← isLocalizedModule_iff_isBaseChange S] exact localizedModuleIsLocalizedModule S diff --git a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean index 44936bdc10f1c..23c177a98c8b0 100644 --- a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean +++ b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean @@ -52,8 +52,7 @@ lemma cotangentSpaceBasis_repr_tmul (r x i) : P.cotangentSpaceBasis.repr (r ⊗ₜ .D _ _ x) i = r * aeval P.val (pderiv i x) := by classical simp only [cotangentSpaceBasis, Basis.baseChange_repr_tmul, mvPolynomialBasis_repr_apply, - Algebra.smul_def, mul_comm r] - rfl + Algebra.smul_def, mul_comm r, algebraMap_apply] lemma cotangentSpaceBasis_repr_one_tmul (x i) : P.cotangentSpaceBasis.repr (1 ⊗ₜ .D _ _ x) i = aeval P.val (pderiv i x) := by @@ -130,8 +129,8 @@ lemma map_id : CotangentSpace.map (.id P) = LinearMap.id := by apply P.cotangentSpaceBasis.ext intro i - simp only [cotangentSpaceBasis_apply, map_tmul, _root_.map_one, Hom.toAlgHom_X] - rfl + simp only [cotangentSpaceBasis_apply, map_tmul, _root_.map_one, Hom.toAlgHom_X, Hom.id_val, + LinearMap.id_coe, id_eq] lemma map_comp (f : Hom P P') (g : Hom P' P'') : CotangentSpace.map (g.comp f) = @@ -226,7 +225,7 @@ def Hom.sub (f g : Hom P P') : P.CotangentSpace →ₗ[S] P'.Cotangent := by simp only [LinearMap.sub_apply, AlgHom.toLinearMap_apply, _root_.map_one, sub_self, Submodule.zero_mem] · intro x y - ext; + ext simp only [LinearMap.coe_comp, LinearMap.coe_restrictScalars, Function.comp_apply, Cotangent.val_mk, Cotangent.val_add, Cotangent.val_smul''', ← map_smul, ← map_add, Ideal.toCotangent_eq, AddSubmonoid.coe_add, Submodule.coe_toAddSubmonoid, diff --git a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean index 11a64d6d9870b..f9a0ed2868cf3 100644 --- a/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/Homogeneous.lean @@ -241,16 +241,15 @@ lemma totalDegree_le (hφ : IsHomogeneous φ n) : φ.totalDegree ≤ n := by apply Finset.sup_le intro d hd rw [mem_support_iff] at hd - rw [Finsupp.sum, ← hφ hd, weight_apply] - simp only [Pi.one_apply, smul_eq_mul, mul_one] - exact Nat.le.refl + simp_rw [Finsupp.sum, ← hφ hd, weight_apply, Pi.one_apply, smul_eq_mul, mul_one, Finsupp.sum, + le_rfl] theorem totalDegree (hφ : IsHomogeneous φ n) (h : φ ≠ 0) : totalDegree φ = n := by apply le_antisymm hφ.totalDegree_le obtain ⟨d, hd⟩ : ∃ d, coeff d φ ≠ 0 := exists_coeff_ne_zero h simp only [← hφ hd, MvPolynomial.totalDegree, Finsupp.sum] replace hd := Finsupp.mem_support_iff.mpr hd - simp only [weight_apply,Pi.one_apply, smul_eq_mul, mul_one] + simp only [weight_apply, Pi.one_apply, smul_eq_mul, mul_one] -- Porting note: Original proof did not define `f` exact Finset.le_sup (f := fun s ↦ ∑ x ∈ s.support, s x) hd diff --git a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean index 380d5c0dc3152..c906097c8e36e 100644 --- a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean @@ -399,7 +399,7 @@ theorem IsWeightedHomogeneous.weightedHomogeneousComponent_same {m : M} {p : MvP · split_ifs · rfl rw [zero_coeff] - · rw [hp zero_coeff, if_pos]; rfl + · rw [hp zero_coeff, if_pos rfl] theorem IsWeightedHomogeneous.weightedHomogeneousComponent_ne {m : M} (n : M) {p : MvPolynomial σ R} (hp : IsWeightedHomogeneous w p m) : @@ -409,9 +409,7 @@ theorem IsWeightedHomogeneous.weightedHomogeneousComponent_ne {m : M} (n : M) ext x rw [coeff_weightedHomogeneousComponent] by_cases zero_coeff : coeff x p = 0 - · split_ifs - · rw [zero_coeff]; rw [coeff_zero] - · rw [coeff_zero] + · simp [zero_coeff] · rw [if_neg] · rw [coeff_zero] · rw [hp zero_coeff]; exact Ne.symm hn @@ -424,8 +422,8 @@ theorem weightedHomogeneousComponent_of_mem [DecidableEq M] {m n : M} ext x rw [coeff_weightedHomogeneousComponent] by_cases zero_coeff : coeff x p = 0 - · split_ifs - all_goals simp only [zero_coeff, coeff_zero] + · split_ifs <;> + simp only [zero_coeff, coeff_zero] · rw [h zero_coeff] simp only [show n = m ↔ m = n from eq_comm] split_ifs with h1 @@ -439,8 +437,7 @@ theorem weightedHomogeneousComponent_of_isWeightedHomogeneous_same ext x rw [coeff_weightedHomogeneousComponent] by_cases zero_coeff : coeff x p = 0 - · split_ifs - rfl; rw [zero_coeff] + · simp [zero_coeff] · rw [hp zero_coeff, if_pos rfl] theorem weightedHomogeneousComponent_of_isWeightedHomogeneous_ne @@ -450,7 +447,7 @@ theorem weightedHomogeneousComponent_of_isWeightedHomogeneous_ne ext x rw [coeff_weightedHomogeneousComponent] by_cases zero_coeff : coeff x p = 0 - · split_ifs <;> simp only [zero_coeff, coeff_zero] + · simp [zero_coeff] · rw [if_neg (by simp only [hp zero_coeff, hn.symm, not_false_eq_true]), coeff_zero] variable (R w) @@ -520,13 +517,8 @@ def NonTorsionWeight (w : σ → M) := ∀ n x, n • w x = (0 : M) → n = 0 theorem nonTorsionWeight_of [NoZeroSMulDivisors ℕ M] (hw : ∀ i : σ, w i ≠ 0) : - NonTorsionWeight w := by - intro n x - rw [smul_eq_zero] - intro hnx - cases' hnx with hn hx - · exact hn - · exact absurd hx (hw x) + NonTorsionWeight w := + fun _ x hnx => (smul_eq_zero_iff_left (hw x)).mp hnx end CanonicallyOrderedAddCommMonoid diff --git a/Mathlib/RingTheory/PrimeSpectrum.lean b/Mathlib/RingTheory/PrimeSpectrum.lean index e672235438cc2..e42c803e3eef2 100644 --- a/Mathlib/RingTheory/PrimeSpectrum.lean +++ b/Mathlib/RingTheory/PrimeSpectrum.lean @@ -371,7 +371,7 @@ section Order We endow `PrimeSpectrum R` with a partial order induced from the ideal lattice. This is exactly the specialization order. -See the corresponding section at `AlgebraicGeometry/PrimeSpectrum/Basic`. +See the corresponding section at `Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic`. -/ instance : PartialOrder (PrimeSpectrum R) := diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index fb15a7d13a20d..9806ceba908c9 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -557,8 +557,8 @@ theorem repr_mul : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ * o₂) = rep apply add_absorp h₁.snd'.repr_lt simpa using (Ordinal.mul_le_mul_iff_left <| opow_pos _ omega_pos).2 (natCast_le.2 n₁.2) by_cases e0 : e₂ = 0 - · simp [e0, mul] - cases' Nat.exists_eq_succ_of_ne_zero n₂.ne_zero with x xe + · cases' Nat.exists_eq_succ_of_ne_zero n₂.ne_zero with x xe + simp only [e0, repr, PNat.mul_coe, natCast_mul, opow_zero, one_mul] simp only [xe, h₂.zero_of_zero e0, repr, add_zero] rw [natCast_succ x, add_mul_succ _ ao, mul_assoc] · simp only [repr] From 5bc522295333b0c252cbd79b758ca3bc38766943 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 6 Aug 2024 08:16:59 +0000 Subject: [PATCH 053/359] fix: put `Nat.primeCounting` in a narrower scope (#15519) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Without this, it is not possible to mix `!` notation for factorial and `π` notation for `Real.pi`. More precisely; while the following works in `Mathlib/Analysis/SpecialFunctions/Stirling.lean` ```lean lemma factorial_isEquivalent_stirling : (fun n ↦ n ! : ℕ → ℝ) ~[atTop] fun n ↦ Real.sqrt (2 * n * π) * (n / exp 1) ^ n := by ``` it produces an error about ambiguous syntax in a standalone file with `import Mathlib`: ```lean import Mathlib open scoped Real Nat Asymptotics open Real lemma factorial_isEquivalent_stirling : (fun n ↦ n ! : ℕ → ℝ) ~[atTop] fun n ↦ Real.sqrt (2 * n * π) * (n / exp 1) ^ n := by ``` Follows on from #13473 --- Mathlib/NumberTheory/PrimeCounting.lean | 17 +++++++++++------ 1 file changed, 11 insertions(+), 6 deletions(-) diff --git a/Mathlib/NumberTheory/PrimeCounting.lean b/Mathlib/NumberTheory/PrimeCounting.lean index 4c8ceb8fd2f5f..e49c213d69742 100644 --- a/Mathlib/NumberTheory/PrimeCounting.lean +++ b/Mathlib/NumberTheory/PrimeCounting.lean @@ -27,8 +27,8 @@ are not prime, and so only at most `φ(k)/k` fraction of the numbers from `k` to ## Notation -We use the standard notation `π` to represent the prime counting function (and `π'` to represent -the reindexed version). +Wtih `open scoped Nat.Prime`, we use the standard notation `π` to represent the prime counting +function (and `π'` to represent the reindexed version). -/ @@ -39,17 +39,22 @@ open Finset /-- A variant of the traditional prime counting function which gives the number of primes *strictly* less than the input. More convenient for avoiding off-by-one errors. --/ + +With `open scoped Nat.Prime`, this has notation `π'`. -/ def primeCounting' : ℕ → ℕ := Nat.count Prime -/-- The prime counting function: Returns the number of primes less than or equal to the input. -/ +/-- The prime counting function: Returns the number of primes less than or equal to the input. + +With `open scoped Nat.Prime`, this has notation `π`. -/ def primeCounting (n : ℕ) : ℕ := primeCounting' (n + 1) -@[inherit_doc] scoped notation "π" => Nat.primeCounting +@[inherit_doc] scoped[Nat.Prime] notation "π" => Nat.primeCounting + +@[inherit_doc] scoped[Nat.Prime] notation "π'" => Nat.primeCounting' -@[inherit_doc] scoped notation "π'" => Nat.primeCounting' +open scoped Nat.Prime theorem monotone_primeCounting' : Monotone primeCounting' := count_monotone Prime From 1d4f2c16a08a0d2dc2fa7d846fefbcc212ffd46a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 6 Aug 2024 08:17:00 +0000 Subject: [PATCH 054/359] fix: undo the accidental revert of #11877 in #11807 (#15523) A bad merge in #11807 causes the change from #11877 (removing `import ProofWidgets`) to be reverted. --- Mathlib/Tactic/Common.lean | 5 ----- Mathlib/Tactic/Widget/StringDiagram.lean | 1 + 2 files changed, 1 insertion(+), 5 deletions(-) diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index f318dca852025..53c1b1a4c238f 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -11,11 +11,6 @@ import Qq -- Tools for analysing imports, like `#find_home`, `#minimize_imports`, ... import ImportGraph.Imports --- Currently we don't need to import all of ProofWidgets, --- but without this, if you don't run `lake build ProofWidgets` then `make test` will fail. --- Hopefully `lake` will be able to handle tests later. -import ProofWidgets - -- Import common Batteries tactics and commands import Batteries.Tactic.Where diff --git a/Mathlib/Tactic/Widget/StringDiagram.lean b/Mathlib/Tactic/Widget/StringDiagram.lean index b896832c3fa27..56a321a914335 100644 --- a/Mathlib/Tactic/Widget/StringDiagram.lean +++ b/Mathlib/Tactic/Widget/StringDiagram.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yuma Mizuno -/ import ProofWidgets.Component.PenroseDiagram +import ProofWidgets.Component.Panel.Basic import ProofWidgets.Presentation.Expr import Mathlib.Tactic.CategoryTheory.Monoidal From b1b3aa6389057db0dd190387b02a443a5280f5b6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 6 Aug 2024 08:17:02 +0000 Subject: [PATCH 055/359] =?UTF-8?q?fix:=20`Int.cast=5Fsmul=5Feq=5Fnsmul`?= =?UTF-8?q?=20=E2=86=92=20`Int.cast=5Fsmul=5Feq=5Fzsmul`=20(#15528)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is a recently-introduced typo. The lemma is about `ℤ` not `ℕ`. --- Mathlib/Algebra/Lie/Sl2.lean | 4 ++-- Mathlib/Algebra/Lie/Weights/RootSystem.lean | 6 +++--- Mathlib/Algebra/Module/Defs.lean | 12 ++++++------ Mathlib/Algebra/Module/Rat.lean | 2 +- Mathlib/Algebra/Module/Zlattice/Basic.lean | 8 ++++---- Mathlib/Algebra/Star/CHSH.lean | 2 +- Mathlib/Analysis/Distribution/SchwartzSpace.lean | 2 +- Mathlib/Analysis/NormedSpace/Connected.lean | 4 ++-- Mathlib/Data/ZMod/IntUnitsPower.lean | 2 +- Mathlib/Data/ZMod/Module.lean | 2 +- .../TensorProduct/Graded/External.lean | 14 +++++++------- .../TensorProduct/Graded/Internal.lean | 10 +++++----- .../GroupCohomology/Resolution.lean | 2 +- Mathlib/RingTheory/Derivation/Basic.lean | 2 +- 14 files changed, 36 insertions(+), 36 deletions(-) diff --git a/Mathlib/Algebra/Lie/Sl2.lean b/Mathlib/Algebra/Lie/Sl2.lean index 799f311607911..765625e1d8442 100644 --- a/Mathlib/Algebra/Lie/Sl2.lean +++ b/Mathlib/Algebra/Lie/Sl2.lean @@ -154,8 +154,8 @@ lemma pow_toEnd_f_ne_zero_of_eq_nat · next i IH => have : ((i + 1) * (n - i) : ℤ) • (toEnd R L M f ^ i) m = 0 := by have := congr_arg (⁅e, ·⁆) H - simpa [← Int.cast_smul_eq_nsmul R, P.lie_e_pow_succ_toEnd_f, hn] using this - rw [← Int.cast_smul_eq_nsmul R, smul_eq_zero, Int.cast_eq_zero, mul_eq_zero, sub_eq_zero, + simpa [← Int.cast_smul_eq_zsmul R, P.lie_e_pow_succ_toEnd_f, hn] using this + rw [← Int.cast_smul_eq_zsmul R, smul_eq_zero, Int.cast_eq_zero, mul_eq_zero, sub_eq_zero, Nat.cast_inj, ← @Nat.cast_one ℤ, ← Nat.cast_add, Nat.cast_eq_zero] at this simp only [add_eq_zero, one_ne_zero, and_false, false_or] at this exact (hi.trans_eq (this.resolve_right (IH (i.le_succ.trans hi)))).not_lt i.lt_succ_self diff --git a/Mathlib/Algebra/Lie/Weights/RootSystem.lean b/Mathlib/Algebra/Lie/Weights/RootSystem.lean index 184ab5adcd0f0..1054e9d82da97 100644 --- a/Mathlib/Algebra/Lie/Weights/RootSystem.lean +++ b/Mathlib/Algebra/Lie/Weights/RootSystem.lean @@ -313,7 +313,7 @@ lemma eq_neg_one_or_eq_zero_or_eq_one_of_eq_smul (k : K) (h : (β : H → K) = k mul_eq_mul_left_iff, OfNat.ofNat_ne_zero, or_false] at H rw [← Int.cast_natCast, ← Int.cast_natCast (chainTopCoeff α β), ← Int.cast_sub] at H have := (rootSpace_zsmul_add_ne_bot_iff_mem α 0 hα (n - chainTopCoeff α β)).mp - (by rw [← Int.cast_smul_eq_nsmul K, ← H, ← h, Weight.coe_zero, add_zero]; exact β.2) + (by rw [← Int.cast_smul_eq_zsmul K, ← H, ← h, Weight.coe_zero, add_zero]; exact β.2) rw [chainTopCoeff_zero_right α hα, chainBotCoeff_zero_right α hα, Nat.cast_one] at this set k' : ℤ := n - chainTopCoeff α β subst H @@ -330,7 +330,7 @@ lemma eq_neg_one_or_eq_zero_or_eq_one_of_eq_smul (k : K) (h : (β : H → K) = k · simp only [tsub_le_iff_right, le_add_iff_nonneg_right, Nat.cast_nonneg, neg_sub, true_and] rw [← Nat.cast_add, chainBotCoeff_add_chainTopCoeff, hn] omega - rw [h, hk, ← Int.cast_smul_eq_nsmul K, ← add_smul] at this + rw [h, hk, ← Int.cast_smul_eq_zsmul K, ← add_smul] at this simp only [Int.cast_sub, Int.cast_natCast, sub_add_sub_cancel', add_sub_cancel_left, ne_eq] at this cases this (rootSpace_one_div_two_smul α hα) @@ -352,7 +352,7 @@ def reflectRoot (α β : Weight K H L) : Weight K H L where by_cases hα : α.IsZero · simpa [hα.eq] using β.weightSpace_ne_bot rw [sub_eq_neg_add, apply_coroot_eq_cast α β, ← neg_smul, ← Int.cast_neg, - Int.cast_smul_eq_nsmul, rootSpace_zsmul_add_ne_bot_iff α β hα] + Int.cast_smul_eq_zsmul, rootSpace_zsmul_add_ne_bot_iff α β hα] omega lemma reflectRoot_isNonZero (α β : Weight K H L) (hβ : β.IsNonZero) : diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index db4b297d90ae7..86d77ca8b94dd 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -370,24 +370,24 @@ section variable (R) /-- `zsmul` is equal to any other module structure via a cast. -/ -lemma Int.cast_smul_eq_nsmul (n : ℤ) (b : M) : (n : R) • b = n • b := +lemma Int.cast_smul_eq_zsmul (n : ℤ) (b : M) : (n : R) • b = n • b := have : ((smulAddHom R M).flip b).comp (Int.castAddHom R) = (smulAddHom ℤ M).flip b := by apply AddMonoidHom.ext_int simp DFunLike.congr_fun this n -@[deprecated (since := "2024-07-23")] alias intCast_smul := Int.cast_smul_eq_nsmul +@[deprecated (since := "2024-07-23")] alias intCast_smul := Int.cast_smul_eq_zsmul /-- `zsmul` is equal to any other module structure via a cast. -/ -@[deprecated Int.cast_smul_eq_nsmul (since := "2024-07-23")] -theorem zsmul_eq_smul_cast (n : ℤ) (b : M) : n • b = (n : R) • b := (Int.cast_smul_eq_nsmul ..).symm +@[deprecated Int.cast_smul_eq_zsmul (since := "2024-07-23")] +theorem zsmul_eq_smul_cast (n : ℤ) (b : M) : n • b = (n : R) • b := (Int.cast_smul_eq_zsmul ..).symm end /-- Convert back any exotic `ℤ`-smul to the canonical instance. This should not be needed since in mathlib all `AddCommGroup`s should normally have exactly one `ℤ`-module structure by design. -/ theorem int_smul_eq_zsmul (h : Module ℤ M) (n : ℤ) (x : M) : @SMul.smul ℤ M h.toSMul n x = n • x := - Int.cast_smul_eq_nsmul .. + Int.cast_smul_eq_zsmul .. /-- All `ℤ`-module structures are equal. Not an instance since in mathlib all `AddCommGroup` should normally have exactly one `ℤ`-module structure by design. -/ @@ -400,7 +400,7 @@ end AddCommGroup theorem map_intCast_smul [AddCommGroup M] [AddCommGroup M₂] {F : Type*} [FunLike F M M₂] [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [Ring R] [Ring S] [Module R M] [Module S M₂] (x : ℤ) (a : M) : - f ((x : R) • a) = (x : S) • f a := by simp only [Int.cast_smul_eq_nsmul, map_zsmul] + f ((x : R) • a) = (x : S) • f a := by simp only [Int.cast_smul_eq_zsmul, map_zsmul] theorem map_natCast_smul [AddCommMonoid M] [AddCommMonoid M₂] {F : Type*} [FunLike F M M₂] [AddMonoidHomClass F M M₂] (f : F) (R S : Type*) [Semiring R] [Semiring S] [Module R M] diff --git a/Mathlib/Algebra/Module/Rat.lean b/Mathlib/Algebra/Module/Rat.lean index 8c134bd1eb1a9..fa255570c0b86 100644 --- a/Mathlib/Algebra/Module/Rat.lean +++ b/Mathlib/Algebra/Module/Rat.lean @@ -60,6 +60,6 @@ instance SMulCommClass.rat' {R : Type u} {M : Type v} [Semiring R] [AddCommGroup instance (priority := 100) RatModule.noZeroSMulDivisors [AddCommGroup M] [Module ℚ M] : NoZeroSMulDivisors ℤ M := ⟨fun {k} {x : M} h => by - simpa only [← Int.cast_smul_eq_nsmul ℚ k x, smul_eq_zero, Rat.zero_iff_num_zero] using h⟩ + simpa only [← Int.cast_smul_eq_zsmul ℚ k x, smul_eq_zero, Rat.zero_iff_num_zero] using h⟩ -- Porting note: old proof was: --⟨fun {k x} h => by simpa [zsmul_eq_smul_cast ℚ k x] using h⟩ diff --git a/Mathlib/Algebra/Module/Zlattice/Basic.lean b/Mathlib/Algebra/Module/Zlattice/Basic.lean index d4968d0b69674..c26f3cf544512 100644 --- a/Mathlib/Algebra/Module/Zlattice/Basic.lean +++ b/Mathlib/Algebra/Module/Zlattice/Basic.lean @@ -94,13 +94,13 @@ def ceil (m : E) : span ℤ (Set.range b) := ∑ i, ⌈b.repr m i⌉ • b.restr @[simp] theorem repr_floor_apply (m : E) (i : ι) : b.repr (floor b m) i = ⌊b.repr m i⌋ := by - classical simp only [floor, ← Int.cast_smul_eq_nsmul K, b.repr.map_smul, Finsupp.single_apply, + classical simp only [floor, ← Int.cast_smul_eq_zsmul K, b.repr.map_smul, Finsupp.single_apply, Finset.sum_apply', Basis.repr_self, Finsupp.smul_single', mul_one, Finset.sum_ite_eq', coe_sum, Finset.mem_univ, if_true, coe_smul_of_tower, Basis.restrictScalars_apply, map_sum] @[simp] theorem repr_ceil_apply (m : E) (i : ι) : b.repr (ceil b m) i = ⌈b.repr m i⌉ := by - classical simp only [ceil, ← Int.cast_smul_eq_nsmul K, b.repr.map_smul, Finsupp.single_apply, + classical simp only [ceil, ← Int.cast_smul_eq_zsmul K, b.repr.map_smul, Finsupp.single_apply, Finset.sum_apply', Basis.repr_self, Finsupp.smul_single', mul_one, Finset.sum_ite_eq', coe_sum, Finset.mem_univ, if_true, coe_smul_of_tower, Basis.restrictScalars_apply, map_sum] @@ -542,8 +542,8 @@ theorem Zlattice.rank [hs : IsZlattice K L] : finrank ℤ L = finrank K E := by rwa [Ne, add_eq_zero_iff_eq_neg.not, neg_inj, Rat.coe_int_inj, ← Ne] apply (smul_mem_iff _ h_nz).mp refine span_subset_span ℤ ℚ _ ?_ - rwa [add_smul, neg_smul, SetLike.mem_coe, ← Zspan.fract_eq_fract, Int.cast_smul_eq_nsmul ℚ, - Int.cast_smul_eq_nsmul ℚ] + rwa [add_smul, neg_smul, SetLike.mem_coe, ← Zspan.fract_eq_fract, Int.cast_smul_eq_zsmul ℚ, + Int.cast_smul_eq_zsmul ℚ] · -- To prove that `finrank K E ≤ finrank ℤ L`, we use the fact `b` generates `E` over `K` -- and thus `finrank K E ≤ card b = finrank ℤ L` rw [← topEquiv.finrank_eq, ← h_spanE] diff --git a/Mathlib/Algebra/Star/CHSH.lean b/Mathlib/Algebra/Star/CHSH.lean index 6adb8f4bf21a8..8247bb1bd6655 100644 --- a/Mathlib/Algebra/Star/CHSH.lean +++ b/Mathlib/Algebra/Star/CHSH.lean @@ -176,7 +176,7 @@ theorem tsirelson_inequality [OrderedRing R] [StarRing R] [StarOrderedRing R] [A A₀ * B₀ + A₀ * B₁ + A₁ * B₀ - A₁ * B₁ ≤ √2 ^ 3 • (1 : R) := by -- abel will create `ℤ` multiplication. We will `simp` them away to `ℝ` multiplication. have M : ∀ (m : ℤ) (a : ℝ) (x : R), m • a • x = ((m : ℝ) * a) • x := fun m a x => by - rw [← Int.cast_smul_eq_nsmul ℝ, ← mul_smul] + rw [← Int.cast_smul_eq_zsmul ℝ, ← mul_smul] let P := (√2)⁻¹ • (A₁ + A₀) - B₀ let Q := (√2)⁻¹ • (A₁ - A₀) + B₁ have w : √2 ^ 3 • (1 : R) - A₀ * B₀ - A₀ * B₁ - A₁ * B₀ + A₁ * B₁ = (√2)⁻¹ • (P ^ 2 + Q ^ 2) := by diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 85133c9920df7..5b93a0b901333 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -263,7 +263,7 @@ instance instZSMul : SMul ℤ 𝓢(E, F) := ⟨fun c f => { toFun := c • (f : E → F) smooth' := (f.smooth _).const_smul c - decay' := by simpa [← Int.cast_smul_eq_nsmul ℝ] using ((c : ℝ) • f).decay' }⟩ + decay' := by simpa [← Int.cast_smul_eq_zsmul ℝ] using ((c : ℝ) • f).decay' }⟩ end SMul diff --git a/Mathlib/Analysis/NormedSpace/Connected.lean b/Mathlib/Analysis/NormedSpace/Connected.lean index 4137feae466c0..ded791be4e517 100644 --- a/Mathlib/Analysis/NormedSpace/Connected.lean +++ b/Mathlib/Analysis/NormedSpace/Connected.lean @@ -54,11 +54,11 @@ theorem Set.Countable.isPathConnected_compl_of_one_lt_rank have Ia : c - x = a := by simp only [c, x, smul_add, smul_sub] abel_nf - simp [← Int.cast_smul_eq_nsmul ℝ 2] + simp [← Int.cast_smul_eq_zsmul ℝ 2] have Ib : c + x = b := by simp only [c, x, smul_add, smul_sub] abel_nf - simp [← Int.cast_smul_eq_nsmul ℝ 2] + simp [← Int.cast_smul_eq_zsmul ℝ 2] have x_ne_zero : x ≠ 0 := by simpa [x] using sub_ne_zero.2 hab.symm obtain ⟨y, hy⟩ : ∃ y, LinearIndependent ℝ ![x, y] := exists_linearIndependent_pair_of_one_lt_rank h x_ne_zero diff --git a/Mathlib/Data/ZMod/IntUnitsPower.lean b/Mathlib/Data/ZMod/IntUnitsPower.lean index a3e51f8b1b144..4d6bd6c1d8b3d 100644 --- a/Mathlib/Data/ZMod/IntUnitsPower.lean +++ b/Mathlib/Data/ZMod/IntUnitsPower.lean @@ -107,6 +107,6 @@ lemma uzpow_neg (s : ℤˣ) (x : R) : s ^ (-x) = (s ^ x)⁻¹ := @[norm_cast] lemma uzpow_intCast (u : ℤˣ) (z : ℤ) : u ^ (z : R) = u ^ z := by change Additive.toMul ((z : R) • Additive.ofMul u) = _ - rw [Int.cast_smul_eq_nsmul, toMul_zsmul, toMul_ofMul] + rw [Int.cast_smul_eq_zsmul, toMul_zsmul, toMul_ofMul] end CommRing diff --git a/Mathlib/Data/ZMod/Module.lean b/Mathlib/Data/ZMod/Module.lean index aebc50352620e..64cb1d86b6160 100644 --- a/Mathlib/Data/ZMod/Module.lean +++ b/Mathlib/Data/ZMod/Module.lean @@ -52,7 +52,7 @@ theorem map_smul (f : F) (c : ZMod n) (x : M) : f (c • x) = c • f x := by exact map_intCast_smul f _ _ (cast c) x theorem smul_mem (hx : x ∈ K) (c : ZMod n) : c • x ∈ K := by - rw [← ZMod.intCast_zmod_cast c, Int.cast_smul_eq_nsmul] + rw [← ZMod.intCast_zmod_cast c, Int.cast_smul_eq_zsmul] exact zsmul_mem hx (cast c) end ZMod diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean index 8bb7be468d23b..b8af1ec61709a 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean @@ -120,8 +120,8 @@ theorem gradedComm_of_tmul_of (i j : ι) (a : 𝒜 i) (b : ℬ j) : dsimp only [LinearEquiv.trans_apply, LinearEquiv.ofLinear_apply] rw [TensorProduct.directSum_lof_tmul_lof, gradedCommAux_lof_tmul, Units.smul_def, -- Note: #8386 specialized `map_smul` to `LinearEquiv.map_smul` to avoid timeouts. - ← Int.cast_smul_eq_nsmul R, LinearEquiv.map_smul, TensorProduct.directSum_symm_lof_tmul, - Int.cast_smul_eq_nsmul, ← Units.smul_def] + ← Int.cast_smul_eq_zsmul R, LinearEquiv.map_smul, TensorProduct.directSum_symm_lof_tmul, + Int.cast_smul_eq_zsmul, ← Units.smul_def] theorem gradedComm_tmul_of_zero (a : ⨁ i, 𝒜 i) (b : ℬ 0) : gradedComm R 𝒜 ℬ (a ⊗ₜ lof R _ ℬ 0 b) = lof R _ ℬ _ b ⊗ₜ a := by @@ -190,7 +190,7 @@ theorem tmul_of_gradedMul_of_tmul (j₁ i₂ : ι) LinearMap.lTensor_tmul] rw [mul_comm j₁ i₂, gradedComm_of_tmul_of] -- the tower smul lemmas elaborate too slowly - rw [Units.smul_def, Units.smul_def, ← Int.cast_smul_eq_nsmul R, ← Int.cast_smul_eq_nsmul R] + rw [Units.smul_def, Units.smul_def, ← Int.cast_smul_eq_zsmul R, ← Int.cast_smul_eq_zsmul R] -- Note: #8386 had to specialize `map_smul` to avoid timeouts. rw [← smul_tmul', LinearEquiv.map_smul, tmul_smul, LinearEquiv.map_smul, LinearMap.map_smul] dsimp @@ -237,11 +237,11 @@ theorem gradedMul_assoc (x y z : DirectSum _ 𝒜 ⊗[R] DirectSum _ ℬ) : exact DFunLike.congr_fun (DFunLike.congr_fun (DFunLike.congr_fun this x) y) z ext ixa xa ixb xb iya ya iyb yb iza za izb zb dsimp [mA] - simp_rw [tmul_of_gradedMul_of_tmul, Units.smul_def, ← Int.cast_smul_eq_nsmul R, + simp_rw [tmul_of_gradedMul_of_tmul, Units.smul_def, ← Int.cast_smul_eq_zsmul R, LinearMap.map_smul₂, LinearMap.map_smul, DirectSum.lof_eq_of, DirectSum.of_mul_of, ← DirectSum.lof_eq_of R, tmul_of_gradedMul_of_tmul, DirectSum.lof_eq_of, ← DirectSum.of_mul_of, ← DirectSum.lof_eq_of R, mul_assoc] - simp_rw [Int.cast_smul_eq_nsmul R, ← Units.smul_def, smul_smul, ← uzpow_add, add_mul, mul_add] + simp_rw [Int.cast_smul_eq_zsmul R, ← Units.smul_def, smul_smul, ← uzpow_add, add_mul, mul_add] congr 2 abel @@ -256,9 +256,9 @@ theorem gradedComm_gradedMul (x y : DirectSum _ 𝒜 ⊗[R] DirectSum _ ℬ) : dsimp rw [gradedComm_of_tmul_of, gradedComm_of_tmul_of, tmul_of_gradedMul_of_tmul] -- Note: #8386 had to specialize `map_smul` to avoid timeouts. - simp_rw [Units.smul_def, ← Int.cast_smul_eq_nsmul R, LinearEquiv.map_smul, LinearMap.map_smul, + simp_rw [Units.smul_def, ← Int.cast_smul_eq_zsmul R, LinearEquiv.map_smul, LinearMap.map_smul, LinearMap.smul_apply] - simp_rw [Int.cast_smul_eq_nsmul R, ← Units.smul_def, DirectSum.lof_eq_of, DirectSum.of_mul_of, + simp_rw [Int.cast_smul_eq_zsmul R, ← Units.smul_def, DirectSum.lof_eq_of, DirectSum.of_mul_of, ← DirectSum.lof_eq_of R, gradedComm_of_tmul_of, tmul_of_gradedMul_of_tmul, smul_smul, DirectSum.lof_eq_of, ← DirectSum.of_mul_of, ← DirectSum.lof_eq_of R] simp_rw [← uzpow_add, mul_add, add_mul, mul_comm i₁ j₂] diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean index 0f33c9bae0d79..4aaefe468fd08 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean @@ -189,8 +189,8 @@ theorem tmul_coe_mul_coe_tmul {j₁ i₂ : ι} (a₁ : A) (b₁ : ℬ j₁) (a simp_rw [lof_eq_of R] rw [LinearEquiv.symm_symm] -- Note: #8386 had to specialize `map_smul` to `LinearEquiv.map_smul` - rw [@Units.smul_def _ _ (_) (_), ← Int.cast_smul_eq_nsmul R, LinearEquiv.map_smul, map_smul, - Int.cast_smul_eq_nsmul R, ← @Units.smul_def _ _ (_) (_)] + rw [@Units.smul_def _ _ (_) (_), ← Int.cast_smul_eq_zsmul R, LinearEquiv.map_smul, map_smul, + Int.cast_smul_eq_zsmul R, ← @Units.smul_def _ _ (_) (_)] rw [congr_symm_tmul] dsimp simp_rw [decompose_symm_mul, decompose_symm_of, Equiv.symm_apply_apply] @@ -315,8 +315,8 @@ def lift (f : A →ₐ[R] C) (g : B →ₐ[R] C) ext a₂ b₂ : 2 dsimp rw [tmul_coe_mul_coe_tmul] - rw [@Units.smul_def _ _ (_) (_), ← Int.cast_smul_eq_nsmul R, map_smul, map_smul, map_smul] - rw [Int.cast_smul_eq_nsmul R, ← @Units.smul_def _ _ (_) (_)] + rw [@Units.smul_def _ _ (_) (_), ← Int.cast_smul_eq_zsmul R, map_smul, map_smul, map_smul] + rw [Int.cast_smul_eq_zsmul R, ← @Units.smul_def _ _ (_) (_)] rw [of_symm_of, map_tmul, LinearMap.mul'_apply] simp_rw [AlgHom.toLinearMap_apply, _root_.map_mul] simp_rw [mul_assoc (f a₁), ← mul_assoc _ _ (g b₂), h_anti_commutes, mul_smul_comm, @@ -376,7 +376,7 @@ lemma auxEquiv_comm (x : 𝒜 ᵍ⊗[R] ℬ) : comm 𝒜 ℬ (a ᵍ⊗ₜ b) = (-1 : ℤˣ)^(j * i) • (b ᵍ⊗ₜ a : ℬ ᵍ⊗[R] 𝒜) := (auxEquiv R ℬ 𝒜).injective <| by simp_rw [auxEquiv_comm, auxEquiv_tmul, decompose_coe, ← lof_eq_of R, gradedComm_of_tmul_of, - @Units.smul_def _ _ (_) (_), ← Int.cast_smul_eq_nsmul R] + @Units.smul_def _ _ (_) (_), ← Int.cast_smul_eq_zsmul R] -- Qualified `map_smul` to avoid a TC timeout #8386 erw [LinearMap.map_smul, auxEquiv_tmul] simp_rw [decompose_coe, lof_eq_of] diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index 808459c99ce3a..e1fdb9f4e2b18 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -520,7 +520,7 @@ theorem d_eq (n : ℕ) : ((groupCohomology.resolution k G).d (n + 1) n).hom = d /- Porting note (#11039): broken proof was simpa [← @intCast_smul k, simplicial_object.δ] -/ simp_rw [alternatingFaceMapComplex_obj_d, AlternatingFaceMapComplex.objD, SimplicialObject.δ, - Functor.comp_map, ← Int.cast_smul_eq_nsmul k ((-1) ^ _ : ℤ), Int.cast_pow, Int.cast_neg, + Functor.comp_map, ← Int.cast_smul_eq_zsmul k ((-1) ^ _ : ℤ), Int.cast_pow, Int.cast_neg, Int.cast_one, Action.sum_hom, Action.smul_hom, Rep.linearization_map_hom] rw [LinearMap.coeFn_sum, Fintype.sum_apply] erw [d_of (k := k) x] diff --git a/Mathlib/RingTheory/Derivation/Basic.lean b/Mathlib/RingTheory/Derivation/Basic.lean index ba01be750abfd..8b4495bf94265 100644 --- a/Mathlib/RingTheory/Derivation/Basic.lean +++ b/Mathlib/RingTheory/Derivation/Basic.lean @@ -423,7 +423,7 @@ lemma leibniz_zpow (a : K) (n : ℤ) : D (a ^ n) = n • a ^ (n - 1) • D a := congr omega · rw [h, zpow_neg, zpow_natCast, leibniz_inv, leibniz_pow, inv_pow, ← pow_mul, ← zpow_natCast, - ← zpow_natCast, ← Nat.cast_smul_eq_nsmul K, ← Int.cast_smul_eq_nsmul K, smul_smul, smul_smul, + ← zpow_natCast, ← Nat.cast_smul_eq_nsmul K, ← Int.cast_smul_eq_zsmul K, smul_smul, smul_smul, smul_smul] trans (-n.natAbs * (a ^ ((n.natAbs - 1 : ℕ) : ℤ) / (a ^ ((n.natAbs * 2 : ℕ) : ℤ)))) • D a · ring_nf From c82a357fc704576d5dd7e7160e14eb6d3db0b6ab Mon Sep 17 00:00:00 2001 From: Gabin Date: Tue, 6 Aug 2024 09:11:28 +0000 Subject: [PATCH 056/359] feat(ModelTheory/Substructures): define equivalences between substructures and prove simple results (#11175) Define partial equivalences between structures. A partial equivalence between structures `M` and `N` is a subobject of `M` and a subobject of `N` along with an equivalence between. This is useful to build an actual equivalence between `M` and `N`. Co-authored-by: Aaron Anderson --- Mathlib.lean | 1 + Mathlib/ModelTheory/PartialEquiv.lean | 274 +++++++++++++++++++++++++ Mathlib/ModelTheory/Substructures.lean | 28 ++- 3 files changed, 287 insertions(+), 16 deletions(-) create mode 100644 Mathlib/ModelTheory/PartialEquiv.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1352ff3ec2672..f1be574c05c38 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3275,6 +3275,7 @@ import Mathlib.ModelTheory.Fraisse import Mathlib.ModelTheory.Graph import Mathlib.ModelTheory.LanguageMap import Mathlib.ModelTheory.Order +import Mathlib.ModelTheory.PartialEquiv import Mathlib.ModelTheory.Quotients import Mathlib.ModelTheory.Satisfiability import Mathlib.ModelTheory.Semantics diff --git a/Mathlib/ModelTheory/PartialEquiv.lean b/Mathlib/ModelTheory/PartialEquiv.lean new file mode 100644 index 0000000000000..41ce920aed2f2 --- /dev/null +++ b/Mathlib/ModelTheory/PartialEquiv.lean @@ -0,0 +1,274 @@ +/- +Copyright (c) 2024 Gabin Kolly. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Aaron Anderson, Gabin Kolly +-/ +import Mathlib.ModelTheory.Substructures + +/-! +# Partial Isomorphisms +This file defines partial isomorphisms between first-order structures. + +## Main Definitions +- `FirstOrder.Language.Substructure.PartialEquiv` is defined so that `L.PartialEquiv M N`, annotated + `M ≃ₚ[L] N`, is the type of equivalences between substructures of `M` and `N`. + +-/ + +universe u v w w' + +namespace FirstOrder + +namespace Language + +variable (L : Language.{u, v}) (M : Type w) (N : Type w') {P : Type*} +variable [L.Structure M] [L.Structure N] [L.Structure P] + +open FirstOrder Structure Substructure + +/-- A partial `L`-equivalence, implemented as an equivalence between substructures. -/ +structure PartialEquiv where + /-- The substructure which is the domain of the equivalence. -/ + dom : L.Substructure M + /-- The substructure which is the codomain of the equivalence. -/ + cod : L.Substructure N + /-- The equivalence between the two subdomains. -/ + toEquiv : dom ≃[L] cod + +@[inherit_doc] +scoped[FirstOrder] notation:25 M " ≃ₚ[" L "] " N => + FirstOrder.Language.PartialEquiv L M N + +variable {L M N} + +namespace PartialEquiv + +noncomputable instance instInhabited_self : Inhabited (M ≃ₚ[L] M) := + ⟨⊤, ⊤, Equiv.refl L (⊤ : L.Substructure M)⟩ + +/-- Maps to the symmetric partial equivalence. -/ +def symm (f : M ≃ₚ[L] N) : N ≃ₚ[L] M where + dom := f.cod + cod := f.dom + toEquiv := f.toEquiv.symm + +@[simp] +theorem symm_symm (f : M ≃ₚ[L] N) : f.symm.symm = f := + rfl + +@[simp] +theorem symm_apply (f : M ≃ₚ[L] N) (x : f.cod) : f.symm.toEquiv x = f.toEquiv.symm x := + rfl + +instance : LE (M ≃ₚ[L] N) := + ⟨fun f g ↦ ∃h : f.dom ≤ g.dom, + (subtype _).comp (g.toEquiv.toEmbedding.comp (Substructure.inclusion h)) = + (subtype _).comp f.toEquiv.toEmbedding⟩ + +theorem le_def (f g : M ≃ₚ[L] N) : f ≤ g ↔ ∃ h : f.dom ≤ g.dom, + (subtype _).comp (g.toEquiv.toEmbedding.comp (Substructure.inclusion h)) = + (subtype _).comp f.toEquiv.toEmbedding := + Iff.rfl + +@[gcongr] theorem dom_le_dom {f g : M ≃ₚ[L] N} : f ≤ g → f.dom ≤ g.dom := fun ⟨le, _⟩ ↦ le + +@[gcongr] theorem cod_le_cod {f g : M ≃ₚ[L] N} : f ≤ g → f.cod ≤ g.cod := by + rintro ⟨_, eq_fun⟩ n hn + let m := f.toEquiv.symm ⟨n, hn⟩ + have : ((subtype _).comp f.toEquiv.toEmbedding) m = n := by simp only [m, Embedding.comp_apply, + Equiv.coe_toEmbedding, Equiv.apply_symm_apply, coeSubtype] + rw [← this, ← eq_fun] + simp only [Embedding.comp_apply, coe_inclusion, Equiv.coe_toEmbedding, coeSubtype, + SetLike.coe_mem] + +theorem subtype_toEquiv_inclusion {f g : M ≃ₚ[L] N} (h : f ≤ g) : + (subtype _).comp (g.toEquiv.toEmbedding.comp (Substructure.inclusion (dom_le_dom h))) = + (subtype _).comp f.toEquiv.toEmbedding := by + let ⟨_, eq⟩ := h; exact eq + +theorem toEquiv_inclusion {f g : M ≃ₚ[L] N} (h : f ≤ g) : + g.toEquiv.toEmbedding.comp (Substructure.inclusion (dom_le_dom h)) = + (Substructure.inclusion (cod_le_cod h)).comp f.toEquiv.toEmbedding := by + rw [← (subtype _).comp_inj, subtype_toEquiv_inclusion h] + rfl + +theorem toEquiv_inclusion_apply {f g : M ≃ₚ[L] N} (h : f ≤ g) (x : f.dom) : + g.toEquiv (Substructure.inclusion (dom_le_dom h) x) = + Substructure.inclusion (cod_le_cod h) (f.toEquiv x) := by + apply (subtype _).injective + change (subtype _).comp (g.toEquiv.toEmbedding.comp (inclusion _)) x = _ + rw [subtype_toEquiv_inclusion h] + rfl + +theorem le_iff {f g : M ≃ₚ[L] N} : f ≤ g ↔ + ∃ dom_le_dom : f.dom ≤ g.dom, + ∃ cod_le_cod : f.cod ≤ g.cod, + ∀ x, inclusion cod_le_cod (f.toEquiv x) = g.toEquiv (inclusion dom_le_dom x) := by + constructor + · exact fun h ↦ ⟨dom_le_dom h, cod_le_cod h, + by intro x; apply (subtype _).inj'; rwa [toEquiv_inclusion_apply]⟩ + · rintro ⟨dom_le_dom, le_cod, h_eq⟩ + rw [le_def] + exact ⟨dom_le_dom, by ext; change subtype _ (g.toEquiv _) = _; rw [← h_eq]; rfl⟩ + +theorem le_trans (f g h : M ≃ₚ[L] N) : f ≤ g → g ≤ h → f ≤ h := by + rintro ⟨le_fg, eq_fg⟩ ⟨le_gh, eq_gh⟩ + refine ⟨le_fg.trans le_gh, ?_⟩ + rw [← eq_fg, ← Embedding.comp_assoc (g := g.toEquiv.toEmbedding), ← eq_gh] + rfl + +private theorem le_refl (f : M ≃ₚ[L] N) : f ≤ f := ⟨le_rfl, rfl⟩ + +private theorem le_antisymm (f g : M ≃ₚ[L] N) (le_fg : f ≤ g) (le_gf : g ≤ f) : f = g := by + let ⟨dom_f, cod_f, equiv_f⟩ := f + cases _root_.le_antisymm (dom_le_dom le_fg) (dom_le_dom le_gf) + cases _root_.le_antisymm (cod_le_cod le_fg) (cod_le_cod le_gf) + convert rfl + exact Equiv.injective_toEmbedding ((subtype _).comp_injective (subtype_toEquiv_inclusion le_fg)) + +instance : PartialOrder (M ≃ₚ[L] N) where + le_refl := le_refl + le_trans := le_trans + le_antisymm := le_antisymm + +@[gcongr] lemma symm_le_symm {f g : M ≃ₚ[L] N} (hfg : f ≤ g) : f.symm ≤ g.symm := by + rw [le_iff] + refine ⟨cod_le_cod hfg, dom_le_dom hfg, ?_⟩ + intro x + apply g.toEquiv.injective + change g.toEquiv (inclusion _ (f.toEquiv.symm x)) = g.toEquiv (g.toEquiv.symm _) + rw [g.toEquiv.apply_symm_apply, (Equiv.apply_symm_apply f.toEquiv x).symm, + f.toEquiv.symm_apply_apply] + exact toEquiv_inclusion_apply hfg _ + +theorem monotone_symm : Monotone (fun (f : M ≃ₚ[L] N) ↦ f.symm) := fun _ _ => symm_le_symm + +theorem symm_le_iff {f : M ≃ₚ[L] N} {g : N ≃ₚ[L] M} : f.symm ≤ g ↔ f ≤ g.symm := + ⟨by intro h; rw [← f.symm_symm]; exact monotone_symm h, + by intro h; rw [← g.symm_symm]; exact monotone_symm h⟩ + +theorem ext {f g : M ≃ₚ[L] N} (h_dom : f.dom = g.dom) : (∀ x : M, ∀ h : x ∈ f.dom, + subtype _ (f.toEquiv ⟨x, h⟩) = subtype _ (g.toEquiv ⟨x, (h_dom ▸ h)⟩)) → f = g := by + intro h + rcases f with ⟨dom_f, cod_f, equiv_f⟩ + cases h_dom + apply le_antisymm <;> (rw [le_def]; use (by rfl); ext ⟨x, hx⟩) + · exact (h x hx).symm + · exact h x hx + +theorem ext_iff {f g : M ≃ₚ[L] N} : f = g ↔ ∃ h_dom : f.dom = g.dom, + ∀ x : M, ∀ h : x ∈ f.dom, + subtype _ (f.toEquiv ⟨x, h⟩) = subtype _ (g.toEquiv ⟨x, (h_dom ▸ h)⟩) := by + constructor + · intro h_eq + rcases f with ⟨dom_f, cod_f, equiv_f⟩ + cases h_eq + exact ⟨rfl, fun _ _ ↦ rfl⟩ + · rintro ⟨h, H⟩; exact ext h H + +theorem monotone_dom : Monotone (fun f : M ≃ₚ[L] N ↦ f.dom) := fun _ _ ↦ dom_le_dom + +theorem monotone_cod : Monotone (fun f : M ≃ₚ[L] N ↦ f.cod) := fun _ _ ↦ cod_le_cod + +/-- Restriction of a partial equivalence to a substructure of the domain. -/ +noncomputable def domRestrict (f : M ≃ₚ[L] N) {A : L.Substructure M} (h : A ≤ f.dom) : + M ≃ₚ[L] N := by + let g := (subtype _).comp (f.toEquiv.toEmbedding.comp (A.inclusion h)) + exact { + dom := A + cod := g.toHom.range + toEquiv := g.equivRange + } + +theorem domRestrict_le (f : M ≃ₚ[L] N) {A : L.Substructure M} (h : A ≤ f.dom) : + f.domRestrict h ≤ f := ⟨h, rfl⟩ + +theorem le_domRestrict (f g : M ≃ₚ[L] N) {A : L.Substructure M} (hf : f.dom ≤ A) + (hg : A ≤ g.dom) (hfg : f ≤ g) : f ≤ g.domRestrict hg := + ⟨hf, by rw [← (subtype_toEquiv_inclusion hfg)]; rfl⟩ + +/-- Restriction of a partial equivalence to a substructure of the codomain. -/ +noncomputable def codRestrict (f : M ≃ₚ[L] N) {A : L.Substructure N} (h : A ≤ f.cod) : + M ≃ₚ[L] N := + (f.symm.domRestrict h).symm + +theorem codRestrict_le (f : M ≃ₚ[L] N) {A : L.Substructure N} (h : A ≤ f.cod) : + codRestrict f h ≤ f := + symm_le_iff.2 (f.symm.domRestrict_le h) + +theorem le_codRestrict (f g : M ≃ₚ[L] N) {A : L.Substructure N} (hf : f.cod ≤ A) + (hg : A ≤ g.cod) (hfg : f ≤ g) : f ≤ g.codRestrict hg := + symm_le_iff.1 (le_domRestrict f.symm g.symm hf hg (monotone_symm hfg)) + +/-- A partial equivalence as an embedding from its domain. -/ +def toEmbedding (f : M ≃ₚ[L] N) : f.dom ↪[L] N := + (subtype _).comp f.toEquiv.toEmbedding + +@[simp] +theorem toEmbedding_apply {f : M ≃ₚ[L] N} (m : f.dom) : + f.toEmbedding m = f.toEquiv m := by + rcases f with ⟨dom, cod, g⟩ + rfl + +/-- Given a partial equivalence which has the whole structure as domain, + returns the corresponding embedding. -/ +def toEmbeddingOfEqTop {f : M ≃ₚ[L] N} (h : f.dom = ⊤) : M ↪[L] N := + (h ▸ f.toEmbedding).comp topEquiv.symm.toEmbedding + +@[simp] +theorem toEmbeddingOfEqTop__apply {f : M ≃ₚ[L] N} (h : f.dom = ⊤) (m : M) : + toEmbeddingOfEqTop h m = f.toEquiv ⟨m, h.symm ▸ mem_top m⟩ := by + rcases f with ⟨dom, cod, g⟩ + cases h + rfl + +/-- Given a partial equivalence which has the whole structure as domain and + as codomain, returns the corresponding equivalence. -/ +def toEquivOfEqTop {f : M ≃ₚ[L] N} (h_dom : f.dom = ⊤) + (h_cod : f.cod = ⊤) : M ≃[L] N := + (topEquiv (M := N)).comp ((h_dom ▸ h_cod ▸ f.toEquiv).comp (topEquiv (M := M)).symm) + +@[simp] +theorem toEquivOfEqTop_toEmbedding {f : M ≃ₚ[L] N} (h_dom : f.dom = ⊤) + (h_cod : f.cod = ⊤) : + (toEquivOfEqTop h_dom h_cod).toEmbedding = toEmbeddingOfEqTop h_dom := by + rcases f with ⟨dom, cod, g⟩ + cases h_dom + cases h_cod + rfl + +end PartialEquiv + +namespace Embedding + +/-- Given an embedding, returns the corresponding partial equivalence with `⊤` as domain. -/ +noncomputable def toPartialEquiv (f : M ↪[L] N) : M ≃ₚ[L] N := + ⟨⊤, f.toHom.range, f.equivRange.comp (Substructure.topEquiv)⟩ + +theorem toPartialEquiv_injective : + Function.Injective (fun f : M ↪[L] N ↦ f.toPartialEquiv) := by + intro _ _ h + ext + rw [PartialEquiv.ext_iff] at h + rcases h with ⟨_, H⟩ + exact H _ (Substructure.mem_top _) + +@[simp] +theorem toEmbedding_toPartialEquiv (f : M ↪[L] N) : + PartialEquiv.toEmbeddingOfEqTop (f := f.toPartialEquiv) rfl = f := + rfl + +@[simp] +theorem toPartialEquiv_toEmbedding {f : M ≃ₚ[L] N} (h : f.dom = ⊤) : + (PartialEquiv.toEmbeddingOfEqTop h).toPartialEquiv = f := by + rcases f with ⟨_, _, _⟩ + cases h + apply PartialEquiv.ext + intro _ _ + rfl; rfl + +end Embedding + +end Language + +end FirstOrder diff --git a/Mathlib/ModelTheory/Substructures.lean b/Mathlib/ModelTheory/Substructures.lean index 1d96f21064925..31517d42b7c94 100644 --- a/Mathlib/ModelTheory/Substructures.lean +++ b/Mathlib/ModelTheory/Substructures.lean @@ -16,27 +16,28 @@ substructures appearing in the algebra library. ## Main Definitions - A `FirstOrder.Language.Substructure` is defined so that `L.Substructure M` is the type of all - substructures of the `L`-structure `M`. + substructures of the `L`-structure `M`. - `FirstOrder.Language.Substructure.closure` is defined so that if `s : Set M`, `closure L s` is - the least substructure of `M` containing `s`. + the least substructure of `M` containing `s`. - `FirstOrder.Language.Substructure.comap` is defined so that `s.comap f` is the preimage of the - substructure `s` under the homomorphism `f`, as a substructure. + substructure `s` under the homomorphism `f`, as a substructure. - `FirstOrder.Language.Substructure.map` is defined so that `s.map f` is the image of the - substructure `s` under the homomorphism `f`, as a substructure. + substructure `s` under the homomorphism `f`, as a substructure. - `FirstOrder.Language.Hom.range` is defined so that `f.range` is the range of the - homomorphism `f`, as a substructure. + homomorphism `f`, as a substructure. - `FirstOrder.Language.Hom.domRestrict` and `FirstOrder.Language.Hom.codRestrict` restrict - the domain and codomain respectively of first-order homomorphisms to substructures. + the domain and codomain respectively of first-order homomorphisms to substructures. - `FirstOrder.Language.Embedding.domRestrict` and `FirstOrder.Language.Embedding.codRestrict` - restrict the domain and codomain respectively of first-order embeddings to substructures. + restrict the domain and codomain respectively of first-order embeddings to substructures. - `FirstOrder.Language.Substructure.inclusion` is the inclusion embedding between substructures. +- `FirstOrder.Language.Substructure.PartialEquiv` is defined so that `PartialEquiv L M N` is + the type of equivalences between substructures of `M` and `N`. ## Main Results - `L.Substructure M` forms a `CompleteLattice`. -/ - universe u v w namespace FirstOrder @@ -350,7 +351,7 @@ instance small_bot : Small.{u} (⊥ : L.Substructure M) := by /-! ### `comap` and `map` --/ + -/ /-- The preimage of a substructure along a homomorphism is a substructure. -/ @@ -601,11 +602,10 @@ theorem closure_induction' (s : Set M) {p : ∀ x, x ∈ closure L s → Prop} end Substructure -namespace LHom - - open Substructure +namespace LHom + variable {L' : Language} [L'.Structure M] /-- Reduces the language of a substructure along a language hom. -/ @@ -682,8 +682,6 @@ end Substructure namespace Hom -open Substructure - /-- The restriction of a first-order hom to a substructure `s ⊆ M` gives a hom `s → N`. -/ @[simps!] def domRestrict (f : M →[L] N) (p : L.Substructure M) : p →[L] N := @@ -772,8 +770,6 @@ end Hom namespace Embedding -open Substructure - /-- The restriction of a first-order embedding to a substructure `s ⊆ M` gives an embedding `s → N`. -/ def domRestrict (f : M ↪[L] N) (p : L.Substructure M) : p ↪[L] N := From 04b5a7b3455c83115ae0db6ec9b1d217aac91365 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 6 Aug 2024 09:11:29 +0000 Subject: [PATCH 057/359] feat(ZPowers): add `zpowers_inv` and `zmultiples_neg` (#15534) --- Mathlib/Algebra/Group/Subgroup/ZPowers.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Algebra/Group/Subgroup/ZPowers.lean b/Mathlib/Algebra/Group/Subgroup/ZPowers.lean index d96ff9c7cc2db..b91454a9b1837 100644 --- a/Mathlib/Algebra/Group/Subgroup/ZPowers.lean +++ b/Mathlib/Algebra/Group/Subgroup/ZPowers.lean @@ -203,6 +203,10 @@ theorem zpowers_ne_bot : zpowers g ≠ ⊥ ↔ g ≠ 1 := theorem zpowers_one_eq_bot : Subgroup.zpowers (1 : G) = ⊥ := Subgroup.zpowers_eq_bot.mpr rfl +@[to_additive (attr := simp)] +theorem zpowers_inv : zpowers g⁻¹ = zpowers g := + eq_of_forall_ge_iff fun _ ↦ by simp only [zpowers_le, inv_mem_iff] + @[to_additive] theorem centralizer_closure (S : Set G) : centralizer (closure S : Set G) = ⨅ g ∈ S, centralizer (zpowers g : Set G) := From 780b09c212da1d3a59b84d431f8c5565fb4b1555 Mon Sep 17 00:00:00 2001 From: Ralf Stephan Date: Tue, 6 Aug 2024 10:53:52 +0000 Subject: [PATCH 058/359] chore(ZMod/Basic): move two lemmas (#15446) Move the lemmas `val_eq_zero` and `val_intCast` up. Necessary for proving some lemmas without using `ModEq`. --- Mathlib/Data/ZMod/Basic.lean | 32 ++++++++++++++++---------------- 1 file changed, 16 insertions(+), 16 deletions(-) diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 0ac9d27dbe010..0f50b93aa8261 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -510,6 +510,13 @@ end UniversalProperty variable {m n : ℕ} +@[simp] +theorem val_eq_zero : ∀ {n : ℕ} (a : ZMod n), a.val = 0 ↔ a = 0 + | 0, a => Int.natAbs_eq_zero + | n + 1, a => by + rw [Fin.ext_iff] + exact Iff.rfl + theorem intCast_eq_intCast_iff (a b : ℤ) (c : ℕ) : (a : ZMod c) = (b : ZMod c) ↔ a ≡ b [ZMOD c] := CharP.intCast_eq_intCast (ZMod c) c @@ -522,6 +529,15 @@ theorem intCast_eq_intCast_iff' (a b : ℤ) (c : ℕ) : (a : ZMod c) = (b : ZMod @[deprecated (since := "2024-04-17")] alias int_cast_eq_int_cast_iff' := intCast_eq_intCast_iff' +theorem val_intCast {n : ℕ} (a : ℤ) [NeZero n] : ↑(a : ZMod n).val = a % n := by + have hle : (0 : ℤ) ≤ ↑(a : ZMod n).val := Int.natCast_nonneg _ + have hlt : ↑(a : ZMod n).val < (n : ℤ) := Int.ofNat_lt.mpr (ZMod.val_lt a) + refine (Int.emod_eq_of_lt hle hlt).symm.trans ?_ + rw [← ZMod.intCast_eq_intCast_iff', Int.cast_natCast, ZMod.natCast_val, ZMod.cast_id] + +@[deprecated (since := "2024-04-17")] +alias val_int_cast := val_intCast + theorem natCast_eq_natCast_iff (a b c : ℕ) : (a : ZMod c) = (b : ZMod c) ↔ a ≡ b [MOD c] := by simpa [Int.natCast_modEq_iff] using ZMod.intCast_eq_intCast_iff a b c @@ -552,15 +568,6 @@ theorem natCast_zmod_eq_zero_iff_dvd (a b : ℕ) : (a : ZMod b) = 0 ↔ b ∣ a @[deprecated (since := "2024-04-17")] alias nat_cast_zmod_eq_zero_iff_dvd := natCast_zmod_eq_zero_iff_dvd -theorem val_intCast {n : ℕ} (a : ℤ) [NeZero n] : ↑(a : ZMod n).val = a % n := by - have hle : (0 : ℤ) ≤ ↑(a : ZMod n).val := Int.natCast_nonneg _ - have hlt : ↑(a : ZMod n).val < (n : ℤ) := Int.ofNat_lt.mpr (ZMod.val_lt a) - refine (Int.emod_eq_of_lt hle hlt).symm.trans ?_ - rw [← ZMod.intCast_eq_intCast_iff', Int.cast_natCast, ZMod.natCast_val, ZMod.cast_id] - -@[deprecated (since := "2024-04-17")] -alias val_int_cast := val_intCast - theorem coe_intCast (a : ℤ) : cast (a : ZMod n) = a % n := by cases n · rw [Int.ofNat_zero, Int.emod_zero, Int.cast_id]; rfl @@ -971,13 +978,6 @@ theorem natAbs_mod_two (a : ℤ) : (a.natAbs : ZMod 2) = a := by · simp only [Int.natAbs_ofNat, Int.cast_natCast, Int.ofNat_eq_coe] · simp only [neg_eq_self_mod_two, Nat.cast_succ, Int.natAbs, Int.cast_negSucc] -@[simp] -theorem val_eq_zero : ∀ {n : ℕ} (a : ZMod n), a.val = 0 ↔ a = 0 - | 0, a => Int.natAbs_eq_zero - | n + 1, a => by - rw [Fin.ext_iff] - exact Iff.rfl - theorem val_ne_zero {n : ℕ} (a : ZMod n) : a.val ≠ 0 ↔ a ≠ 0 := (val_eq_zero a).not From 31cb2527235be23a4394fd83806d797a5345ac37 Mon Sep 17 00:00:00 2001 From: Henrik Lievonen Date: Tue, 6 Aug 2024 14:46:28 +0000 Subject: [PATCH 059/359] doc(Probability/Notation): tell user to open ProbabilityTheory to use these notations (#15542) Previously this requirement was not documented anywhere apart from the source. Co-authored-by: Henrik Lievonen Co-authored-by: Henrik Lievonen --- Mathlib/Probability/Notation.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Probability/Notation.lean b/Mathlib/Probability/Notation.lean index d0dc50ccbd035..dc866c19c326d 100644 --- a/Mathlib/Probability/Notation.lean +++ b/Mathlib/Probability/Notation.lean @@ -24,6 +24,9 @@ We note that the notation `∂P/∂Q` applies to three different cases, namely, `MeasureTheory.ComplexMeasure.rnDeriv`. - `ℙ` is a notation for `volume` on a measured space. + +To use these notations, you need to use `open scoped ProbabilityTheory` +or `open ProbabilityTheory`. -/ From 1d8b37341778bb9a6a042b2e1f261405291c345b Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 6 Aug 2024 16:17:37 +0000 Subject: [PATCH 060/359] chore: update Mathlib dependencies 2024-08-06 (#15553) 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 a00fdb8a546f2..1f02830a3f5a8 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "ae6ea60e9d8bc2d4b37ff02115854da2e1b710d0", + "rev": "8b31db8e20ef5e31877615b1142160ba811b89a5", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", From b3472b6c0b5d1309ed991a78bc050f08543eb190 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 6 Aug 2024 17:15:42 +0000 Subject: [PATCH 061/359] feat(ENat): CompleteLinearOrder (WithBot ENat) (#15558) from the Carleson project. --- Mathlib/Data/ENat/Lattice.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Data/ENat/Lattice.lean b/Mathlib/Data/ENat/Lattice.lean index 297e4191742e3..869003b893a3e 100644 --- a/Mathlib/Data/ENat/Lattice.lean +++ b/Mathlib/Data/ENat/Lattice.lean @@ -22,6 +22,9 @@ open Set noncomputable instance : CompleteLinearOrder ENat := inferInstanceAs (CompleteLinearOrder (WithTop ℕ)) +noncomputable instance : CompleteLinearOrder (WithBot ENat) := + inferInstanceAs (CompleteLinearOrder (WithBot (WithTop ℕ))) + namespace ENat variable {ι : Sort*} {f : ι → ℕ} {s : Set ℕ} From 0a376d7809ad7879fe41e832c90ce7d8257fac84 Mon Sep 17 00:00:00 2001 From: Lorenzo Luccioli Date: Tue, 6 Aug 2024 17:59:06 +0000 Subject: [PATCH 062/359] feat (MeaureTheory): `measure_univ_le_add_compl` (#15467) Add `measure_univ_le_add_compl`. --- Mathlib/MeasureTheory/OuterMeasure/Basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/MeasureTheory/OuterMeasure/Basic.lean b/Mathlib/MeasureTheory/OuterMeasure/Basic.lean index b0758fd89c38a..9ffe980b91080 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Basic.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Basic.lean @@ -80,6 +80,9 @@ theorem measure_iUnion_fintype_le [Fintype ι] (μ : F) (s : ι → Set α) : theorem measure_union_le (s t : Set α) : μ (s ∪ t) ≤ μ s + μ t := by simpa [union_eq_iUnion] using measure_iUnion_fintype_le μ (cond · s t) +lemma measure_univ_le_add_compl (s : Set α) : μ univ ≤ μ s + μ sᶜ := + s.union_compl_self ▸ measure_union_le s sᶜ + theorem measure_le_inter_add_diff (μ : F) (s t : Set α) : μ s ≤ μ (s ∩ t) + μ (s \ t) := by simpa using measure_union_le (s ∩ t) (s \ t) From 37940a3bd47b4339157b920b85db460b64b3074c Mon Sep 17 00:00:00 2001 From: Hongyu Ouyang <96765450+casavaca@users.noreply.github.com> Date: Tue, 6 Aug 2024 17:59:07 +0000 Subject: [PATCH 063/359] feat(Data.List): Add `List.nodup_permutations_iff` (#15512) --- Mathlib/Data/List/Perm.lean | 19 ++++++++++++++++++- 1 file changed, 18 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index 74b4ead49e16c..f804a9590f642 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Data.List.Count import Mathlib.Data.List.Dedup +import Mathlib.Data.List.Duplicate import Mathlib.Data.List.InsertNth import Mathlib.Data.List.Lattice import Mathlib.Data.List.Permutation @@ -632,7 +633,23 @@ theorem nodup_permutations (s : List α) (hs : Nodup s) : Nodup s.permutations : rw [← hx, nthLe_insertNth_of_lt _ _ _ _ ht (ht.trans_le hn)] exact nthLe_mem _ _ _ --- TODO: `nodup s.permutations ↔ nodup s` +lemma permutations_take_two (x y : α) (s : List α) : + (x :: y :: s).permutations.take 2 = [x :: y :: s, y :: x :: s] := by + induction s <;> simp only [take, permutationsAux, permutationsAux.rec, permutationsAux2, id_eq] + +@[simp] +theorem nodup_permutations_iff {s : List α} : Nodup s.permutations ↔ Nodup s := by + refine ⟨?_, nodup_permutations s⟩ + contrapose + rw [← exists_duplicate_iff_not_nodup] + intro ⟨x, hs⟩ + rw [duplicate_iff_sublist] at hs + obtain ⟨l, ht⟩ := List.Sublist.exists_perm_append hs + rw [List.Perm.nodup_iff (List.Perm.permutations ht), ← exists_duplicate_iff_not_nodup] + use x :: x :: l + rw [List.duplicate_iff_sublist, ← permutations_take_two] + exact take_sublist 2 _ + -- TODO: `count s s.permutations = (zipWith count s s.tails).prod` end Permutations From e0ea30194f5f571dec3100d3c57b1c08304c1fd4 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Tue, 6 Aug 2024 19:16:43 +0000 Subject: [PATCH 064/359] chore(Algebra/GeomSum): turn `induction'` into `induction` (#15561) This PR turns the (soon to be deprecated) `induction'` into `induction`. --- Mathlib/Algebra/GeomSum.lean | 45 ++++++++++++++++++------------------ 1 file changed, 23 insertions(+), 22 deletions(-) diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index cd0d0f54101d6..7bcba9c410122 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -93,33 +93,33 @@ protected theorem Commute.geom_sum₂_mul_add {x y : α} (h : Commute x y) (n : (∑ i ∈ range n, (x + y) ^ i * y ^ (n - 1 - i)) * x + y ^ n = (x + y) ^ n := by let f : ℕ → ℕ → α := fun m i : ℕ => (x + y) ^ i * y ^ (m - 1 - i) change (∑ i ∈ range n, (f n) i) * x + y ^ n = (x + y) ^ n - induction' n with n ih - · rw [range_zero, sum_empty, zero_mul, zero_add, pow_zero, pow_zero] - · have f_last : f (n + 1) n = (x + y) ^ n := by + induction n with + | zero => rw [range_zero, sum_empty, zero_mul, zero_add, pow_zero, pow_zero] + | succ n ih => + have f_last : f (n + 1) n = (x + y) ^ n := by dsimp only [f] rw [← tsub_add_eq_tsub_tsub, Nat.add_comm, tsub_self, pow_zero, mul_one] have f_succ : ∀ i, i ∈ range n → f (n + 1) i = y * f n i := fun i hi => by dsimp only [f] have : Commute y ((x + y) ^ i) := (h.symm.add_right (Commute.refl y)).pow_right i - rw [← mul_assoc, this.eq, mul_assoc, ← pow_succ' y (n - 1 - i)] - congr 2 - rw [add_tsub_cancel_right, ← tsub_add_eq_tsub_tsub, add_comm 1 i] + rw [← mul_assoc, this.eq, mul_assoc, ← pow_succ' y (n - 1 - i), add_tsub_cancel_right, + ← tsub_add_eq_tsub_tsub, add_comm 1 i] have : i + 1 + (n - (i + 1)) = n := add_tsub_cancel_of_le (mem_range.mp hi) rw [add_comm (i + 1)] at this rw [← this, add_tsub_cancel_right, add_comm i 1, ← add_assoc, add_tsub_cancel_right] - rw [pow_succ' (x + y), add_mul, sum_range_succ_comm, add_mul, f_last, add_assoc] - rw [(((Commute.refl x).add_right h).pow_right n).eq] - congr 1 - rw [sum_congr rfl f_succ, ← mul_sum, pow_succ' y, mul_assoc, ← mul_add y, ih] + rw [pow_succ' (x + y), add_mul, sum_range_succ_comm, add_mul, f_last, add_assoc, + (((Commute.refl x).add_right h).pow_right n).eq, sum_congr rfl f_succ, ← mul_sum, + pow_succ' y, mul_assoc, ← mul_add y, ih] end Semiring @[simp] theorem neg_one_geom_sum [Ring α] {n : ℕ} : ∑ i ∈ range n, (-1 : α) ^ i = if Even n then 0 else 1 := by - induction' n with k hk - · simp - · simp only [geom_sum_succ', Nat.even_add_one, hk] + induction n with + | zero => simp + | succ k hk => + simp only [geom_sum_succ', Nat.even_add_one, hk] split_ifs with h · rw [h.neg_one_pow, add_zero] · rw [(Nat.odd_iff_not_even.2 h).neg_one_pow, neg_add_self] @@ -428,15 +428,16 @@ theorem geom_sum_pos_and_lt_one [StrictOrderedRing α] (hx : x < 0) (hx' : 0 < x theorem geom_sum_alternating_of_le_neg_one [StrictOrderedRing α] (hx : x + 1 ≤ 0) (n : ℕ) : if Even n then (∑ i ∈ range n, x ^ i) ≤ 0 else 1 ≤ ∑ i ∈ range n, x ^ i := by have hx0 : x ≤ 0 := (le_add_of_nonneg_right zero_le_one).trans hx - induction' n with n ih - · simp only [Nat.zero_eq, range_zero, sum_empty, le_refl, ite_true, even_zero] - simp only [Nat.even_add_one, geom_sum_succ] - split_ifs at ih with h - · rw [if_neg (not_not_intro h), le_add_iff_nonneg_left] - exact mul_nonneg_of_nonpos_of_nonpos hx0 ih - · rw [if_pos h] - refine (add_le_add_right ?_ _).trans hx - simpa only [mul_one] using mul_le_mul_of_nonpos_left ih hx0 + induction n with + | zero => simp only [Nat.zero_eq, range_zero, sum_empty, le_refl, ite_true, even_zero] + | succ n ih => + simp only [Nat.even_add_one, geom_sum_succ] + split_ifs at ih with h + · rw [if_neg (not_not_intro h), le_add_iff_nonneg_left] + exact mul_nonneg_of_nonpos_of_nonpos hx0 ih + · rw [if_pos h] + refine (add_le_add_right ?_ _).trans hx + simpa only [mul_one] using mul_le_mul_of_nonpos_left ih hx0 theorem geom_sum_alternating_of_lt_neg_one [StrictOrderedRing α] (hx : x + 1 < 0) (hn : 1 < n) : if Even n then (∑ i ∈ range n, x ^ i) < 0 else 1 < ∑ i ∈ range n, x ^ i := by From c7cfd15ca2392e919c9355096f9cd8def5cedf75 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 6 Aug 2024 20:43:02 +0000 Subject: [PATCH 065/359] chore: update Mathlib dependencies 2024-08-06 (#15562) 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 1f02830a3f5a8..6fc3bcbf513a7 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "8b31db8e20ef5e31877615b1142160ba811b89a5", + "rev": "8ff1f48157c719e04bebc729f12af55cc22aabe4", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", From 82f82694b951154678cbfbd1f57540974c5b54b3 Mon Sep 17 00:00:00 2001 From: kkytola Date: Tue, 6 Aug 2024 21:08:21 +0000 Subject: [PATCH 066/359] feat: Miscellaneous small ENNReal lemmas. (#13938) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `ENNReal.*`: `.toNNReal_toReal_eq`, `.sub_add_eq_add_sub`, `.zpow_neg`, `.zpow_sub`, `.tendsto_pow_atTop_nhds_top` + instance `MeasurableSMul₂ ℝ≥0 ℝ≥0∞` Co-authored-by: kkytola <“kalle.kytola@aalto.fi”> Co-authored-by: kkytola <39528102+kkytola@users.noreply.github.com> --- Mathlib/Analysis/SpecificLimits/Basic.lean | 15 +++++++++++++++ Mathlib/Data/ENNReal/Basic.lean | 5 +++++ Mathlib/Data/ENNReal/Inv.lean | 8 ++++++++ Mathlib/Data/ENNReal/Operations.lean | 8 ++++++++ .../Constructions/BorelSpace/Real.lean | 3 +++ 5 files changed, 39 insertions(+) diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index e26b6b2cc6dc5..314ebd5e16ce3 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -243,6 +243,21 @@ protected theorem ENNReal.tendsto_pow_atTop_nhds_zero_iff {r : ℝ≥0∞} : norm_cast at h ⊢ exact NNReal.tendsto_pow_atTop_nhds_zero_iff.mp h +@[simp] +protected theorem ENNReal.tendsto_pow_atTop_nhds_top_iff {r : ℝ≥0∞} : + Tendsto (fun n ↦ r^n) atTop (𝓝 ∞) ↔ 1 < r := by + refine ⟨?_, ?_⟩ + · contrapose! + intro r_le_one h_tends + specialize h_tends (Ioi_mem_nhds one_lt_top) + simp only [Filter.mem_map, mem_atTop_sets, ge_iff_le, Set.mem_preimage, Set.mem_Ioi] at h_tends + obtain ⟨n, hn⟩ := h_tends + exact lt_irrefl _ <| lt_of_lt_of_le (hn n le_rfl) <| pow_le_one n (zero_le _) r_le_one + · intro r_gt_one + have obs := @Tendsto.inv ℝ≥0∞ ℕ _ _ _ (fun n ↦ (r⁻¹)^n) atTop 0 + simp only [ENNReal.tendsto_pow_atTop_nhds_zero_iff, inv_zero] at obs + simpa [← ENNReal.inv_pow] using obs <| ENNReal.inv_lt_one.mpr r_gt_one + /-! ### Geometric series-/ diff --git a/Mathlib/Data/ENNReal/Basic.lean b/Mathlib/Data/ENNReal/Basic.lean index 70b9e6b22c79f..4acf0165b1904 100644 --- a/Mathlib/Data/ENNReal/Basic.lean +++ b/Mathlib/Data/ENNReal/Basic.lean @@ -223,6 +223,11 @@ theorem ofReal_eq_coe_nnreal {x : ℝ} (h : 0 ≤ x) : @[simp] theorem toReal_nonneg {a : ℝ≥0∞} : 0 ≤ a.toReal := a.toNNReal.2 +@[norm_cast] theorem coe_toNNReal_eq_toReal (z : ℝ≥0∞) : (z.toNNReal : ℝ) = z.toReal := rfl + +@[simp] theorem toNNReal_toReal_eq (z : ℝ≥0∞) : z.toReal.toNNReal = z.toNNReal := by + ext; simp [coe_toNNReal_eq_toReal] + @[simp] theorem top_toNNReal : ∞.toNNReal = 0 := rfl @[simp] theorem top_toReal : ∞.toReal = 0 := rfl diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 1438bbc61c075..4115d63979f9a 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -586,5 +586,13 @@ protected theorem zpow_add {x : ℝ≥0∞} (hx : x ≠ 0) (h'x : x ≠ ∞) (m replace hx : x ≠ 0 := by simpa only [Ne, coe_eq_zero] using hx simp only [← coe_zpow hx, zpow_add₀ hx, coe_mul] +protected theorem zpow_neg {x : ℝ≥0∞} (x_ne_zero : x ≠ 0) (x_ne_top : x ≠ ⊤) (m : ℤ) : + x ^ (-m) = (x ^ m)⁻¹ := + ENNReal.eq_inv_of_mul_eq_one_left (by simp [← ENNReal.zpow_add x_ne_zero x_ne_top]) + +protected theorem zpow_sub {x : ℝ≥0∞} (x_ne_zero : x ≠ 0) (x_ne_top : x ≠ ⊤) (m n : ℤ) : + x ^ (m - n) = (x ^ m) * (x ^ n)⁻¹ := by + rw [sub_eq_add_neg, ENNReal.zpow_add x_ne_zero x_ne_top, ENNReal.zpow_neg x_ne_zero x_ne_top n] + end Inv end ENNReal diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index 36bfcf5a23a5a..bcc3076cd29db 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -334,6 +334,14 @@ protected theorem add_sub_cancel_left (ha : a ≠ ∞) : a + b - a = b := protected theorem add_sub_cancel_right (hb : b ≠ ∞) : a + b - b = a := (cancel_of_ne hb).add_tsub_cancel_right +protected theorem sub_add_eq_add_sub (hab : b ≤ a) (b_ne_top : b ≠ ∞) : + a - b + c = a + c - b := by + by_cases c_top : c = ∞ + · simpa [c_top] using ENNReal.eq_sub_of_add_eq b_ne_top rfl + refine (sub_eq_of_add_eq b_ne_top ?_).symm + simp only [add_assoc, add_comm c b] + simpa only [← add_assoc] using (add_left_inj c_top).mpr <| tsub_add_cancel_of_le hab + protected theorem lt_add_of_sub_lt_left (h : a ≠ ∞ ∨ b ≠ ∞) : a - b < c → a < b + c := by obtain rfl | hb := eq_or_ne b ∞ · rw [top_add, lt_top_iff_ne_top] diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index dde83153317ce..e47b7dd06d224 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -441,6 +441,9 @@ theorem AEMeasurable.coe_ereal_ennreal {f : α → ℝ≥0∞} {μ : Measure α} namespace NNReal +instance : MeasurableSMul₂ ℝ≥0 ℝ≥0∞ where + measurable_smul := show Measurable fun r : ℝ≥0 × ℝ≥0∞ ↦ (r.1 : ℝ≥0) * r.2 by fun_prop + /-- A limit (over a general filter) of measurable `ℝ≥0` valued functions is measurable. -/ theorem measurable_of_tendsto' {ι} {f : ι → α → ℝ≥0} {g : α → ℝ≥0} (u : Filter ι) [NeBot u] [IsCountablyGenerated u] (hf : ∀ i, Measurable (f i)) (lim : Tendsto f u (𝓝 g)) : From fe874da10d5846d0cddd07fa0dc6078bce5fc999 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Tue, 6 Aug 2024 22:27:23 +0000 Subject: [PATCH 067/359] feat(GroupTheory/Index): `index_eq_zero_iff_infinite` (#15504) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add the trivial, but apparently missing, lemma: ```lean @[to_additive] lemma index_eq_zero_iff_infinite : H.index = 0 ↔ Infinite (G ⧸ H) := by ``` From AperiodicMonotilesLean. --- Mathlib/GroupTheory/Index.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index e7f2ad5c5c1d1..20a4102ae16ee 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -375,6 +375,10 @@ theorem index_ne_zero_of_finite [hH : Finite (G ⧸ H)] : H.index ≠ 0 := by noncomputable def fintypeOfIndexNeZero (hH : H.index ≠ 0) : Fintype (G ⧸ H) := @Fintype.ofFinite _ (Nat.finite_of_card_ne_zero hH) +@[to_additive] +lemma index_eq_zero_iff_infinite : H.index = 0 ↔ Infinite (G ⧸ H) := by + simp [index_eq_card, Nat.card_eq_zero] + @[to_additive one_lt_index_of_ne_top] theorem one_lt_index_of_ne_top [Finite (G ⧸ H)] (hH : H ≠ ⊤) : 1 < H.index := Nat.one_lt_iff_ne_zero_and_ne_one.mpr ⟨index_ne_zero_of_finite, mt index_eq_one.mp hH⟩ From 48415a3ecf8c07b069fa227a6f2e6052f062fe1d Mon Sep 17 00:00:00 2001 From: Geoffrey Irving Date: Tue, 6 Aug 2024 23:04:26 +0000 Subject: [PATCH 068/359] feat(Tactic.Bound): Add a `bound` tactic for inequalities by structure (#10562) `bound` proves inequalities that follow "obviously by structure", assuming nonnegativity or positivity of subterms where necessary. It is built on top of Aesop. It has significant overlap with `positivity` and `gcongr`, but can also switch back and forth between those modes (such as when proving `0 < a * b - a * c` from `0 < a`, `c < b`). Zulip discussion: https://leanprover.zulipchat.com/#narrow/stream/239415-metaprogramming-.2F-tactics/topic/.60bound.60.20mixes.20.60gcongr.60.20and.20.60positivity.60/near/412120380 --- Mathlib.lean | 3 + Mathlib/Algebra/Order/AbsoluteValue.lean | 7 + .../Order/BigOperators/Group/Finset.lean | 3 + Mathlib/Algebra/Order/Field/Basic.lean | 22 +- Mathlib/Algebra/Order/Floor.lean | 2 + Mathlib/Algebra/Order/Ring/Basic.lean | 20 +- Mathlib/Analysis/Analytic/Basic.lean | 4 + Mathlib/Analysis/Normed/Group/Basic.lean | 6 + Mathlib/Analysis/Normed/MulAction.lean | 2 + .../Analysis/SpecialFunctions/Log/Basic.lean | 3 +- .../Analysis/SpecialFunctions/Pow/Real.lean | 15 +- .../SpecialFunctions/Trigonometric/Basic.lean | 2 + Mathlib/Data/Complex/Abs.lean | 2 + Mathlib/Data/Complex/Exponential.lean | 7 +- Mathlib/Data/ENNReal/Real.lean | 5 +- Mathlib/Data/NNReal/Basic.lean | 7 +- Mathlib/Data/Nat/Log.lean | 2 + Mathlib/Data/Real/Sqrt.lean | 6 +- Mathlib/Tactic.lean | 3 + Mathlib/Tactic/Bound.lean | 258 ++++++++++++++++++ Mathlib/Tactic/Bound/Attribute.lean | 139 ++++++++++ Mathlib/Tactic/Bound/Init.lean | 17 ++ Mathlib/Tactic/GCongr/Core.lean | 4 +- Mathlib/Topology/MetricSpace/Pseudo/Defs.lean | 3 + scripts/noshake.json | 12 + test/Bound/attribute.lean | 32 +++ test/Bound/bound.lean | 137 ++++++++++ 27 files changed, 707 insertions(+), 16 deletions(-) create mode 100644 Mathlib/Tactic/Bound.lean create mode 100644 Mathlib/Tactic/Bound/Attribute.lean create mode 100644 Mathlib/Tactic/Bound/Init.lean create mode 100644 test/Bound/attribute.lean create mode 100644 test/Bound/bound.lean diff --git a/Mathlib.lean b/Mathlib.lean index f1be574c05c38..e2132682089cc 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4009,6 +4009,9 @@ import Mathlib.Tactic.ArithMult.Init import Mathlib.Tactic.Attr.Core import Mathlib.Tactic.Attr.Register import Mathlib.Tactic.Basic +import Mathlib.Tactic.Bound +import Mathlib.Tactic.Bound.Attribute +import Mathlib.Tactic.Bound.Init import Mathlib.Tactic.ByContra import Mathlib.Tactic.CC import Mathlib.Tactic.CC.Addition diff --git a/Mathlib/Algebra/Order/AbsoluteValue.lean b/Mathlib/Algebra/Order/AbsoluteValue.lean index 1872f78df7a16..0dcc6acdcfbb4 100644 --- a/Mathlib/Algebra/Order/AbsoluteValue.lean +++ b/Mathlib/Algebra/Order/AbsoluteValue.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Order.Field.Defs import Mathlib.Algebra.Order.Hom.Basic import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Algebra.Regular.Basic +import Mathlib.Tactic.Bound.Attribute /-! # Absolute values @@ -85,6 +86,7 @@ instance : CoeFun (AbsoluteValue R S) fun _ => R → S := theorem coe_toMulHom : ⇑abv.toMulHom = abv := rfl +@[bound] protected theorem nonneg (x : R) : 0 ≤ abv x := abv.nonneg' x @@ -92,6 +94,7 @@ protected theorem nonneg (x : R) : 0 ≤ abv x := protected theorem eq_zero {x : R} : abv x = 0 ↔ x = 0 := abv.eq_zero' x +@[bound] protected theorem add_le (x y : R) : abv (x + y) ≤ abv x + abv y := abv.add_le' x y @@ -184,6 +187,7 @@ section Ring variable {R S : Type*} [Ring R] [OrderedRing S] (abv : AbsoluteValue R S) +@[bound] protected theorem le_sub (a b : R) : abv a - abv b ≤ abv (a - b) := sub_le_iff_le_add.2 <| by simpa using abv.add_le (a - b) b @@ -205,10 +209,12 @@ protected theorem map_neg (a : R) : abv (-a) = abv a := by protected theorem map_sub (a b : R) : abv (a - b) = abv (b - a) := by rw [← neg_sub, abv.map_neg] /-- Bound `abv (a + b)` from below -/ +@[bound] protected theorem le_add (a b : R) : abv a - abv b ≤ abv (a + b) := by simpa only [tsub_le_iff_right, add_neg_cancel_right, abv.map_neg] using abv.add_le (a + b) (-b) /-- Bound `abv (a - b)` from above -/ +@[bound] lemma sub_le_add (a b : R) : abv (a - b) ≤ abv a + abv b := by simpa only [← sub_eq_add_neg, AbsoluteValue.map_neg] using abv.add_le a (-b) @@ -242,6 +248,7 @@ section LinearOrderedCommRing variable {R S : Type*} [Ring R] [LinearOrderedCommRing S] (abv : AbsoluteValue R S) +@[bound] theorem abs_abv_sub_le_abv_sub (a b : R) : abs (abv a - abv b) ≤ abv (a - b) := abs_sub_le_iff.2 ⟨abv.le_sub _ _, by rw [abv.map_sub]; apply abv.le_sub⟩ diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index 4e4552cca0e22..a37fbfc93fced 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Order.BigOperators.Group.Multiset +import Mathlib.Tactic.Bound.Attribute import Mathlib.Tactic.NormNum.Basic import Mathlib.Tactic.Positivity.Core @@ -100,6 +101,8 @@ equal to the corresponding factor `g i` of another finite product, then theorem prod_le_prod' (h : ∀ i ∈ s, f i ≤ g i) : ∏ i ∈ s, f i ≤ ∏ i ∈ s, g i := Multiset.prod_map_le_prod_map f g h +attribute [bound] sum_le_sum + /-- In an ordered additive commutative monoid, if each summand `f i` of one finite sum is less than or equal to the corresponding summand `g i` of another finite sum, then `∑ i ∈ s, f i ≤ ∑ i ∈ s, g i`. -/ diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 77f2b01fc9ddf..edf647b081fba 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.CharZero.Lemmas import Mathlib.Algebra.Order.Field.Defs import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Order.Bounds.OrderIso +import Mathlib.Tactic.Bound.Attribute import Mathlib.Tactic.Positivity.Core import Mathlib.Algebra.Order.Field.Unbundled.Basic @@ -131,12 +132,15 @@ lemma mul_le_of_nonneg_of_le_div (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ b / c) · simpa using hb · rwa [le_div_iff hc] at h +@[bound] theorem div_le_one_of_le (h : a ≤ b) (hb : 0 ≤ b) : a / b ≤ 1 := div_le_of_nonneg_of_le_mul hb zero_le_one <| by rwa [one_mul] +@[bound] lemma mul_inv_le_one_of_le (h : a ≤ b) (hb : 0 ≤ b) : a * b⁻¹ ≤ 1 := by simpa only [← div_eq_mul_inv] using div_le_one_of_le h hb +@[bound] lemma inv_mul_le_one_of_le (h : a ≤ b) (hb : 0 ≤ b) : b⁻¹ * a ≤ 1 := by simpa only [← div_eq_inv_mul] using div_le_one_of_le h hb @@ -144,7 +148,7 @@ lemma inv_mul_le_one_of_le (h : a ≤ b) (hb : 0 ≤ b) : b⁻¹ * a ≤ 1 := by ### Bi-implications of inequalities using inversions -/ -@[gcongr] +@[gcongr, bound] theorem inv_le_inv_of_le (ha : 0 < a) (h : a ≤ b) : b⁻¹ ≤ a⁻¹ := by rwa [← one_div a, le_div_iff' ha, ← div_eq_mul_inv, div_le_iff (ha.trans_le h), one_mul] @@ -188,6 +192,7 @@ theorem inv_lt_one (ha : 1 < a) : a⁻¹ < 1 := by theorem one_lt_inv (h₁ : 0 < a) (h₂ : a < 1) : 1 < a⁻¹ := by rwa [lt_inv (@zero_lt_one α _ _ _ _ _) h₁, inv_one] +@[bound] theorem inv_le_one (ha : 1 ≤ a) : a⁻¹ ≤ 1 := by rwa [inv_le (zero_lt_one.trans_le ha) zero_lt_one, inv_one] @@ -220,12 +225,12 @@ theorem one_le_inv_iff : 1 ≤ a⁻¹ ↔ 0 < a ∧ a ≤ 1 := -/ -@[mono, gcongr] +@[mono, gcongr, bound] lemma div_le_div_of_nonneg_right (hab : a ≤ b) (hc : 0 ≤ c) : a / c ≤ b / c := by rw [div_eq_mul_one_div a c, div_eq_mul_one_div b c] exact mul_le_mul_of_nonneg_right hab (one_div_nonneg (α := α) |>.2 hc) -@[gcongr] +@[gcongr, bound] lemma div_lt_div_of_pos_right (h : a < b) (hc : 0 < c) : a / c < b / c := by rw [div_eq_mul_one_div a c, div_eq_mul_one_div b c] exact mul_lt_mul_of_pos_right h (one_div_pos (α := α) |>.2 hc) @@ -236,7 +241,7 @@ lemma div_le_div_of_nonneg_left (ha : 0 ≤ a) (hc : 0 < c) (h : c ≤ b) : a / rw [div_eq_mul_inv, div_eq_mul_inv] exact mul_le_mul_of_nonneg_left ((inv_le_inv (hc.trans_le h) hc).mpr h) ha -@[gcongr] +@[gcongr, bound] lemma div_lt_div_of_pos_left (ha : 0 < a) (hc : 0 < c) (h : c < b) : a / b < a / c := by simpa only [div_eq_mul_inv, mul_lt_mul_left ha, inv_lt_inv (hc.trans h) hc] @@ -268,7 +273,7 @@ theorem div_lt_div_iff (b0 : 0 < b) (d0 : 0 < d) : a / b < c / d ↔ a * d < c * theorem div_le_div_iff (b0 : 0 < b) (d0 : 0 < d) : a / b ≤ c / d ↔ a * d ≤ c * b := by rw [le_div_iff d0, div_mul_eq_mul_div, div_le_iff b0] -@[mono, gcongr] +@[mono, gcongr, bound] theorem div_le_div (hc : 0 ≤ c) (hac : a ≤ c) (hd : 0 < d) (hbd : d ≤ b) : a / b ≤ c / d := by rw [div_le_div_iff (hd.trans_le hbd) hd] exact mul_le_mul hac hbd hd.le hc @@ -285,12 +290,15 @@ theorem div_lt_div' (hac : a ≤ c) (hbd : d < b) (c0 : 0 < c) (d0 : 0 < d) : a -/ +@[bound] theorem div_le_self (ha : 0 ≤ a) (hb : 1 ≤ b) : a / b ≤ a := by simpa only [div_one] using div_le_div_of_nonneg_left ha zero_lt_one hb +@[bound] theorem div_lt_self (ha : 0 < a) (hb : 1 < b) : a / b < a := by simpa only [div_one] using div_lt_div_of_pos_left ha zero_lt_one hb +@[bound] theorem le_div_self (ha : 0 ≤ a) (hb₀ : 0 < b) (hb₁ : b ≤ 1) : a ≤ a / b := by simpa only [div_one] using div_le_div_of_nonneg_left ha hb₀ hb₁ @@ -310,6 +318,10 @@ theorem le_one_div (ha : 0 < a) (hb : 0 < b) : a ≤ 1 / b ↔ b ≤ 1 / a := by theorem lt_one_div (ha : 0 < a) (hb : 0 < b) : a < 1 / b ↔ b < 1 / a := by simpa using lt_inv ha hb +@[bound] lemma Bound.one_lt_div_of_pos_of_lt (b0 : 0 < b) : b < a → 1 < a / b := (one_lt_div b0).mpr + +@[bound] lemma Bound.div_lt_one_of_pos_of_lt (b0 : 0 < b) : a < b → a / b < 1 := (div_lt_one b0).mpr + /-! ### Relating two divisions, involving `1` -/ diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 52b1d3b770b27..6e2e7f168de95 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -446,6 +446,7 @@ theorem ceil_add_ofNat (ha : 0 ≤ a) (n : ℕ) [n.AtLeastTwo] : ⌈a + (no_index (OfNat.ofNat n))⌉₊ = ⌈a⌉₊ + OfNat.ofNat n := ceil_add_nat ha n +@[bound] theorem ceil_lt_add_one (ha : 0 ≤ a) : (⌈a⌉₊ : α) < a + 1 := lt_ceil.1 <| (Nat.lt_succ_self _).trans_le (ceil_add_one ha).ge @@ -1060,6 +1061,7 @@ theorem ceil_le_floor_add_one (a : α) : ⌈a⌉ ≤ ⌊a⌋ + 1 := by rw [ceil_le, Int.cast_add, Int.cast_one] exact (lt_floor_add_one a).le +@[bound] theorem le_ceil (a : α) : a ≤ ⌈a⌉ := gc_ceil_coe.le_u_l a diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index 278ac4ae53594..a4780f5c89c08 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -6,6 +6,7 @@ Authors: Jeremy Avigad, Robert Y. Lewis import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Algebra.Ring.Parity +import Mathlib.Tactic.Bound.Attribute /-! # Basic lemmas about ordered rings @@ -58,6 +59,7 @@ theorem pow_add_pow_le (hx : 0 ≤ x) (hy : 0 ≤ y) (hn : n ≠ 0) : x ^ n + y rw [pow_succ' _ n] exact mul_le_mul_of_nonneg_left (ih (Nat.succ_ne_zero k)) h2 +@[bound] theorem pow_le_one : ∀ n : ℕ, 0 ≤ a → a ≤ 1 → a ^ n ≤ 1 | 0, _, _ => (pow_zero a).le | n + 1, h₀, h₁ => (pow_succ a n).le.trans (mul_le_one (pow_le_one n h₀ h₁) h₀ h₁) @@ -68,6 +70,7 @@ theorem pow_lt_one (h₀ : 0 ≤ a) (h₁ : a < 1) : ∀ {n : ℕ}, n ≠ 0 → rw [pow_succ'] exact mul_lt_one_of_nonneg_of_lt_one_left h₀ h₁ (pow_le_one _ h₀ h₁.le) +@[bound] theorem one_le_pow_of_one_le (H : 1 ≤ a) : ∀ n : ℕ, 1 ≤ a ^ n | 0 => by rw [pow_zero] | n + 1 => by @@ -86,7 +89,12 @@ theorem pow_le_pow_right (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := pow_r theorem le_self_pow (ha : 1 ≤ a) (h : m ≠ 0) : a ≤ a ^ m := by simpa only [pow_one] using pow_le_pow_right ha <| Nat.pos_iff_ne_zero.2 h -@[mono, gcongr] +/-- The `bound` tactic can't handle `m ≠ 0` goals yet, so we express as `0 < m` -/ +@[bound] +lemma Bound.le_self_pow_of_pos {m : ℕ} (ha : 1 ≤ a) (h : 0 < m) : a ≤ a ^ m := + le_self_pow ha h.ne' + +@[mono, gcongr, bound] theorem pow_le_pow_left {a b : R} (ha : 0 ≤ a) (hab : a ≤ b) : ∀ n, a ^ n ≤ b ^ n | 0 => by simp | n + 1 => by simpa only [pow_succ'] @@ -103,13 +111,21 @@ lemma pow_add_pow_le' (ha : 0 ≤ a) (hb : 0 ≤ b) : a ^ n + b ^ n ≤ 2 * (a + exact add_le_add (pow_le_pow_left ha (le_add_of_nonneg_right hb) _) (pow_le_pow_left hb (le_add_of_nonneg_left ha) _) +/-- `bound` lemma for branching on `1 ≤ a ∨ a ≤ 1` when proving `a ^ n ≤ a ^ m` -/ +@[bound] +lemma Bound.pow_le_pow_right_of_le_one_or_one_le (h : 1 ≤ a ∧ n ≤ m ∨ 0 ≤ a ∧ a ≤ 1 ∧ m ≤ n) : + a ^ n ≤ a ^ m := by + rcases h with ⟨a1, nm⟩ | ⟨a0, a1, mn⟩ + · exact pow_le_pow_right a1 nm + · exact pow_le_pow_of_le_one a0 a1 mn + end OrderedSemiring section StrictOrderedSemiring variable [StrictOrderedSemiring R] {a x y : R} {n m : ℕ} -@[gcongr] +@[gcongr, bound] theorem pow_lt_pow_left (h : x < y) (hx : 0 ≤ x) : ∀ {n : ℕ}, n ≠ 0 → x ^ n < y ^ n | 0, hn => by contradiction | n + 1, _ => by diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index f3b6c32015311..e81266a4b101c 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Star.Order import Mathlib.Analysis.Calculus.FormalMultilinearSeries import Mathlib.Analysis.SpecificLimits.Normed import Mathlib.Logic.Equiv.Fin +import Mathlib.Tactic.Bound.Attribute import Mathlib.Topology.Algebra.InfiniteSum.Module /-! @@ -373,6 +374,9 @@ def HasFPowerSeriesAt (f : E → F) (p : FormalMultilinearSeries 𝕜 E F) (x : def HasFPowerSeriesWithinAt (f : E → F) (p : FormalMultilinearSeries 𝕜 E F) (s : Set E) (x : E) := ∃ r, HasFPowerSeriesWithinOnBall f p s x r +-- Teach the `bound` tactic that power series have positive radius +attribute [bound_forward] HasFPowerSeriesOnBall.r_pos HasFPowerSeriesWithinOnBall.r_pos + variable (𝕜) /-- Given a function `f : E → F`, we say that `f` is analytic at `x` if it admits a convergent power diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index ebe208edc6074..c5a2fba4dda85 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -6,6 +6,7 @@ Authors: Patrick Massot, Johannes Hölzl, Yaël Dillies import Mathlib.Algebra.CharP.Defs import Mathlib.Algebra.Group.Subgroup.Basic import Mathlib.Analysis.Normed.Group.Seminorm +import Mathlib.Tactic.Bound.Attribute import Mathlib.Topology.Metrizable.Uniformity import Mathlib.Topology.Sequences @@ -403,6 +404,8 @@ theorem norm_nonneg' (a : E) : 0 ≤ ‖a‖ := by rw [← dist_one_right] exact dist_nonneg +attribute [bound] norm_nonneg + @[to_additive (attr := simp) abs_norm] theorem abs_norm' (z : E) : |‖z‖| = ‖z‖ := abs_of_nonneg <| norm_nonneg' _ @@ -454,6 +457,8 @@ theorem zero_lt_one_add_norm_sq' (x : E) : 0 < 1 + ‖x‖ ^ 2 := by theorem norm_div_le (a b : E) : ‖a / b‖ ≤ ‖a‖ + ‖b‖ := by simpa [dist_eq_norm_div] using dist_triangle a 1 b +attribute [bound] norm_sub_le + @[to_additive] theorem norm_div_le_of_le {r₁ r₂ : ℝ} (H₁ : ‖a₁‖ ≤ r₁) (H₂ : ‖a₂‖ ≤ r₂) : ‖a₁ / a₂‖ ≤ r₁ + r₂ := (norm_div_le a₁ a₂).trans <| add_le_add H₁ H₂ @@ -960,6 +965,7 @@ theorem norm_multiset_prod_le (m : Multiset E) : ‖m.prod‖ ≤ (m.map fun x = -- Porting note: had to add `ι` here because otherwise the universe order gets switched compared to -- `norm_prod_le` below +@[bound] theorem norm_sum_le {ι E} [SeminormedAddCommGroup E] (s : Finset ι) (f : ι → E) : ‖∑ i ∈ s, f i‖ ≤ ∑ i ∈ s, ‖f i‖ := s.le_sum_of_subadditive norm norm_zero norm_add_le f diff --git a/Mathlib/Analysis/Normed/MulAction.lean b/Mathlib/Analysis/Normed/MulAction.lean index 3f130c93f187e..a7acc637d6fa0 100644 --- a/Mathlib/Analysis/Normed/MulAction.lean +++ b/Mathlib/Analysis/Normed/MulAction.lean @@ -24,9 +24,11 @@ section SeminormedAddGroup variable [SeminormedAddGroup α] [SeminormedAddGroup β] [SMulZeroClass α β] variable [BoundedSMul α β] +@[bound] theorem norm_smul_le (r : α) (x : β) : ‖r • x‖ ≤ ‖r‖ * ‖x‖ := by simpa [smul_zero] using dist_smul_pair r 0 x +@[bound] theorem nnnorm_smul_le (r : α) (x : β) : ‖r • x‖₊ ≤ ‖r‖₊ * ‖x‖₊ := norm_smul_le _ _ diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index b5a0736903d90..3593e3aa57d9e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -119,7 +119,7 @@ theorem log_inv (x : ℝ) : log x⁻¹ = -log x := by theorem log_le_log_iff (h : 0 < x) (h₁ : 0 < y) : log x ≤ log y ↔ x ≤ y := by rw [← exp_le_exp, exp_log h, exp_log h₁] -@[gcongr] +@[gcongr, bound] lemma log_le_log (hx : 0 < x) (hxy : x ≤ y) : log x ≤ log y := (log_le_log_iff hx (hx.trans_le hxy)).2 hxy @@ -165,6 +165,7 @@ theorem log_neg_of_lt_zero (h0 : x < 0) (h1 : -1 < x) : log x < 0 := by theorem log_nonneg_iff (hx : 0 < x) : 0 ≤ log x ↔ 1 ≤ x := by rw [← not_lt, log_neg_iff hx, not_lt] +@[bound] theorem log_nonneg (hx : 1 ≤ x) : 0 ≤ log x := (log_nonneg_iff (zero_lt_one.trans_le hx)).2 hx diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index ccdb7fd337e63..05be81afdde37 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -102,6 +102,7 @@ theorem rpow_def_of_nonpos {x : ℝ} (hx : x ≤ 0) (y : ℝ) : x ^ y = if x = 0 then if y = 0 then 1 else 0 else exp (log x * y) * cos (y * π) := by split_ifs with h <;> simp [rpow_def, *]; exact rpow_def_of_neg (lt_of_le_of_ne hx h) _ +@[bound] theorem rpow_pos_of_pos {x : ℝ} (hx : 0 < x) (y : ℝ) : 0 < x ^ y := by rw [rpow_def_of_pos hx]; apply exp_pos @@ -142,6 +143,7 @@ theorem zero_rpow_le_one (x : ℝ) : (0 : ℝ) ^ x ≤ 1 := by theorem zero_rpow_nonneg (x : ℝ) : 0 ≤ (0 : ℝ) ^ x := by by_cases h : x = 0 <;> simp [h, zero_le_one] +@[bound] theorem rpow_nonneg {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : 0 ≤ x ^ y := by rw [rpow_def_of_nonneg hx]; split_ifs <;> simp only [zero_le_one, le_refl, le_of_lt (exp_pos _)] @@ -150,6 +152,7 @@ theorem abs_rpow_of_nonneg {x y : ℝ} (hx_nonneg : 0 ≤ x) : |x ^ y| = |x| ^ y have h_rpow_nonneg : 0 ≤ x ^ y := Real.rpow_nonneg hx_nonneg _ rw [abs_eq_self.mpr hx_nonneg, abs_eq_self.mpr h_rpow_nonneg] +@[bound] theorem abs_rpow_le_abs_rpow (x y : ℝ) : |x ^ y| ≤ |x| ^ y := by rcases le_or_lt 0 x with hx | hx · rw [abs_rpow_of_nonneg hx] @@ -476,7 +479,7 @@ in `Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean` instead. -/ -/ -@[gcongr] +@[gcongr, bound] theorem rpow_lt_rpow (hx : 0 ≤ x) (hxy : x < y) (hz : 0 < z) : x ^ z < y ^ z := by rw [le_iff_eq_or_lt] at hx; cases' hx with hx hx · rw [← hx, zero_rpow (ne_of_gt hz)] @@ -488,7 +491,7 @@ theorem strictMonoOn_rpow_Ici_of_exponent_pos {r : ℝ} (hr : 0 < r) : StrictMonoOn (fun (x : ℝ) => x ^ r) (Set.Ici 0) := fun _ ha _ _ hab => rpow_lt_rpow ha hab hr -@[gcongr] +@[gcongr, bound] theorem rpow_le_rpow {x y z : ℝ} (h : 0 ≤ x) (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x ^ z ≤ y ^ z := by rcases eq_or_lt_of_le h₁ with (rfl | h₁'); · rfl rcases eq_or_lt_of_le h₂ with (rfl | h₂'); · simp @@ -755,6 +758,14 @@ lemma antitone_rpow_of_base_le_one {b : ℝ} (hb₀ : 0 < b) (hb₁ : b ≤ 1) : case inl => exact (strictAnti_rpow_of_base_lt_one hb₀ hb₁).antitone case inr => intro _ _ _; simp +/-- Guessing rule for the `bound` tactic: when trying to prove `x ^ y ≤ x ^ z`, we can either assume +`1 ≤ x` or `0 < x ≤ 1`. -/ +@[bound] lemma rpow_le_rpow_of_exponent_le_or_ge {x y z : ℝ} + (h : 1 ≤ x ∧ y ≤ z ∨ 0 < x ∧ x ≤ 1 ∧ z ≤ y) : x ^ y ≤ x ^ z := by + rcases h with ⟨x1, yz⟩ | ⟨x0, x1, zy⟩ + · exact Real.rpow_le_rpow_of_exponent_le x1 yz + · exact Real.rpow_le_rpow_of_exponent_ge x0 x1 zy + end Real namespace Complex diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index e5198e9b7d79d..c9617df307a4e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -145,9 +145,11 @@ theorem pi_le_four : π ≤ 4 := π / 2 ≤ 2 := pi_div_two_le_two _ = 4 / 2 := by norm_num) +@[bound] theorem pi_pos : 0 < π := lt_of_lt_of_le (by norm_num) two_le_pi +@[bound] theorem pi_nonneg : 0 ≤ π := pi_pos.le diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 69b4687bda886..3fdd10b7dd853 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -134,11 +134,13 @@ by `simp only [@map_zpow₀]` -/ theorem abs_zpow (z : ℂ) (n : ℤ) : Complex.abs (z ^ n) = Complex.abs z ^ n := map_zpow₀ Complex.abs z n +@[bound] theorem abs_re_le_abs (z : ℂ) : |z.re| ≤ Complex.abs z := Real.abs_le_sqrt <| by rw [normSq_apply, ← sq] exact le_add_of_nonneg_right (mul_self_nonneg _) +@[bound] theorem abs_im_le_abs (z : ℂ) : |z.im| ≤ Complex.abs z := Real.abs_le_sqrt <| by rw [normSq_apply, ← sq, ← sq] diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 8979d1b534d65..7f358613f520b 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Star.Order import Mathlib.Data.Complex.Abs import Mathlib.Data.Complex.BigOperators import Mathlib.Data.Nat.Choose.Sum +import Mathlib.Tactic.Bound.Attribute /-! # Exponential, trigonometric and hyperbolic trigonometric functions @@ -979,11 +980,13 @@ private theorem add_one_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : x + 1 ≤ ex theorem one_le_exp {x : ℝ} (hx : 0 ≤ x) : 1 ≤ exp x := by linarith [add_one_le_exp_of_nonneg hx] +@[bound] theorem exp_pos (x : ℝ) : 0 < exp x := (le_total 0 x).elim (lt_of_lt_of_le zero_lt_one ∘ one_le_exp) fun h => by rw [← neg_neg x, Real.exp_neg] exact inv_pos.2 (lt_of_lt_of_le zero_lt_one (one_le_exp (neg_nonneg.2 h))) +@[bound] lemma exp_nonneg (x : ℝ) : 0 ≤ exp x := x.exp_pos.le @[simp] @@ -1006,7 +1009,7 @@ theorem exp_lt_exp_of_lt {x y : ℝ} (h : x < y) : exp x < exp y := exp_strictMo theorem exp_monotone : Monotone exp := exp_strictMono.monotone -@[gcongr] +@[gcongr, bound] theorem exp_le_exp_of_le {x y : ℝ} (h : x ≤ y) : exp x ≤ exp y := exp_monotone h @[simp] @@ -1031,6 +1034,8 @@ theorem exp_eq_one_iff : exp x = 1 ↔ x = 0 := @[simp] theorem one_lt_exp_iff {x : ℝ} : 1 < exp x ↔ 0 < x := by rw [← exp_zero, exp_lt_exp] +@[bound] private alias ⟨_, Bound.one_lt_exp_of_pos⟩ := one_lt_exp_iff + @[simp] theorem exp_lt_one_iff {x : ℝ} : exp x < 1 ↔ x < 0 := by rw [← exp_zero, exp_lt_exp] diff --git a/Mathlib/Data/ENNReal/Real.lean b/Mathlib/Data/ENNReal/Real.lean index 855b63757b6d2..d1cdba2b5cfd8 100644 --- a/Mathlib/Data/ENNReal/Real.lean +++ b/Mathlib/Data/ENNReal/Real.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ import Mathlib.Data.ENNReal.Inv +import Mathlib.Tactic.Bound.Attribute /-! # Maps between real and extended non-negative real numbers @@ -155,7 +156,7 @@ theorem toReal_pos_iff : 0 < a.toReal ↔ 0 < a ∧ a < ∞ := theorem toReal_pos {a : ℝ≥0∞} (ha₀ : a ≠ 0) (ha_top : a ≠ ∞) : 0 < a.toReal := toReal_pos_iff.mpr ⟨bot_lt_iff_ne_bot.mpr ha₀, lt_top_iff_ne_top.mpr ha_top⟩ -@[gcongr] +@[gcongr, bound] theorem ofReal_le_ofReal {p q : ℝ} (h : p ≤ q) : ENNReal.ofReal p ≤ ENNReal.ofReal q := by simp [ENNReal.ofReal, Real.toNNReal_le_toNNReal h] @@ -191,6 +192,8 @@ theorem ofReal_lt_ofReal_iff_of_nonneg {p q : ℝ} (hp : 0 ≤ p) : @[simp] theorem ofReal_pos {p : ℝ} : 0 < ENNReal.ofReal p ↔ 0 < p := by simp [ENNReal.ofReal] +@[bound] private alias ⟨_, Bound.ofReal_pos_of_pos⟩ := ofReal_pos + @[simp] theorem ofReal_eq_zero {p : ℝ} : ENNReal.ofReal p = 0 ↔ p ≤ 0 := by simp [ENNReal.ofReal] diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index 6741f66bbca51..26fe515738bcf 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -10,6 +10,7 @@ import Mathlib.Algebra.Order.Nonneg.Field import Mathlib.Algebra.Order.Nonneg.Floor import Mathlib.Data.Real.Pointwise import Mathlib.Order.ConditionallyCompleteLattice.Group +import Mathlib.Tactic.Bound.Attribute import Mathlib.Tactic.GCongr.Core import Mathlib.Algebra.Ring.Regular @@ -116,7 +117,7 @@ theorem _root_.Real.toNNReal_of_nonneg {r : ℝ} (hr : 0 ≤ r) : r.toNNReal = theorem _root_.Real.le_coe_toNNReal (r : ℝ) : r ≤ Real.toNNReal r := le_max_left r 0 -theorem coe_nonneg (r : ℝ≥0) : (0 : ℝ) ≤ r := r.2 +@[bound] theorem coe_nonneg (r : ℝ≥0) : (0 : ℝ) ≤ r := r.2 @[simp, norm_cast] theorem coe_mk (a : ℝ) (ha) : toReal ⟨a, ha⟩ = a := rfl @@ -323,8 +324,12 @@ noncomputable example : LinearOrder ℝ≥0 := by infer_instance @[simp, norm_cast] lemma coe_lt_coe : (r₁ : ℝ) < r₂ ↔ r₁ < r₂ := Iff.rfl +@[bound] private alias ⟨_, Bound.coe_lt_coe_of_lt⟩ := coe_lt_coe + @[simp, norm_cast] lemma coe_pos : (0 : ℝ) < r ↔ 0 < r := Iff.rfl +@[bound] private alias ⟨_, Bound.coe_pos_of_pos⟩ := coe_pos + @[simp, norm_cast] lemma one_le_coe : 1 ≤ (r : ℝ) ↔ 1 ≤ r := by rw [← coe_le_coe, coe_one] @[simp, norm_cast] lemma one_lt_coe : 1 < (r : ℝ) ↔ 1 < r := by rw [← coe_lt_coe, coe_one] @[simp, norm_cast] lemma coe_le_one : (r : ℝ) ≤ 1 ↔ r ≤ 1 := by rw [← coe_le_coe, coe_one] diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index 4fd0786ba5f43..fef1cd6a9e339 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -6,6 +6,7 @@ Authors: Simon Hudon, Yaël Dillies import Mathlib.Data.Nat.Defs import Mathlib.Init.Data.Nat.Lemmas import Mathlib.Order.Interval.Set.Basic +import Mathlib.Tactic.Bound.Attribute import Mathlib.Tactic.Monotonicity.Attr /-! @@ -51,6 +52,7 @@ theorem log_of_left_le_one {b : ℕ} (hb : b ≤ 1) (n) : log b n = 0 := theorem log_pos_iff {b n : ℕ} : 0 < log b n ↔ b ≤ n ∧ 1 < b := by rw [Nat.pos_iff_ne_zero, Ne, log_eq_zero_iff, not_or, not_lt, not_le] +@[bound] theorem log_pos {b n : ℕ} (hb : 1 < b) (hbn : b ≤ n) : 0 < log b n := log_pos_iff.2 ⟨hbn, hb⟩ diff --git a/Mathlib/Data/Real/Sqrt.lean b/Mathlib/Data/Real/Sqrt.lean index fc96a3f4372dd..3c3768bf2fd11 100644 --- a/Mathlib/Data/Real/Sqrt.lean +++ b/Mathlib/Data/Real/Sqrt.lean @@ -99,6 +99,8 @@ theorem continuous_sqrt : Continuous sqrt := sqrt.continuous alias ⟨_, sqrt_pos_of_pos⟩ := sqrt_pos +attribute [bound] sqrt_pos_of_pos + end NNReal namespace Real @@ -197,12 +199,12 @@ theorem sqrt_lt_sqrt_iff (hx : 0 ≤ x) : √x < √y ↔ x < y := theorem sqrt_lt_sqrt_iff_of_pos (hy : 0 < y) : √x < √y ↔ x < y := by rw [Real.sqrt, Real.sqrt, NNReal.coe_lt_coe, NNReal.sqrt_lt_sqrt, toNNReal_lt_toNNReal_iff hy] -@[gcongr] +@[gcongr, bound] theorem sqrt_le_sqrt (h : x ≤ y) : √x ≤ √y := by rw [Real.sqrt, Real.sqrt, NNReal.coe_le_coe, NNReal.sqrt_le_sqrt] exact toNNReal_le_toNNReal h -@[gcongr] +@[gcongr, bound] theorem sqrt_lt_sqrt (hx : 0 ≤ x) (h : x < y) : √x < √y := (sqrt_lt_sqrt_iff hx).2 h diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index cbb2ac2a7b14b..c5fdf61f4f3fd 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -9,6 +9,9 @@ import Mathlib.Tactic.ArithMult.Init import Mathlib.Tactic.Attr.Core import Mathlib.Tactic.Attr.Register import Mathlib.Tactic.Basic +import Mathlib.Tactic.Bound +import Mathlib.Tactic.Bound.Attribute +import Mathlib.Tactic.Bound.Init import Mathlib.Tactic.ByContra import Mathlib.Tactic.CC import Mathlib.Tactic.CC.Addition diff --git a/Mathlib/Tactic/Bound.lean b/Mathlib/Tactic/Bound.lean new file mode 100644 index 0000000000000..47591bb967e5d --- /dev/null +++ b/Mathlib/Tactic/Bound.lean @@ -0,0 +1,258 @@ +/- +Copyright (c) 2024 Geoffrey Irving. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Geoffrey Irving +-/ + +import Aesop +import Mathlib.Tactic.Bound.Attribute +import Mathlib.Tactic.Lemma +import Mathlib.Tactic.Linarith.Frontend +import Mathlib.Tactic.NormNum.Core + +/-! +## The `bound` tactic + +`bound` is an `aesop` wrapper that proves inequalities by straightforward recursion on structure, +assuming that intermediate terms are nonnegative or positive as needed. It also has some support +for guessing where it is unclear where to recurse, such as which side of a `min` or `max` to use +as the bound or whether to assume a power is less than or greater than one. + +The functionality of `bound` overlaps with `positivity` and `gcongr`, but can jump back and forth +between `0 ≤ x` and `x ≤ y`-type inequalities. For example, `bound` proves + `0 ≤ c → b ≤ a → 0 ≤ a * c - b * c` +by turning the goal into `b * c ≤ a * c`, then using `mul_le_mul_of_nonneg_right`. `bound` also +uses specialized lemmas for goals of the form `1 ≤ x, 1 < x, x ≤ 1, x < 1`. + +Additional hypotheses can be passed as `bound [h0, h1 n, ...]`. This is equivalent to declaring +them via `have` before calling `bound`. + +See `test/Bound.lean` for tests. + +### Calc usage + +Since `bound` requires the inequality proof to exactly match the structure of the expression, it is +often useful to iterate between `bound` and `rw / simp` using `calc`. Here is an example: + +``` +-- Calc example: A weak lower bound for `z ↦ z^2 + c` +lemma le_sqr_add {c z : ℂ} (cz : abs c ≤ abs z) (z3 : 3 ≤ abs z) : + 2 * abs z ≤ abs (z^2 + c) := by + calc abs (z^2 + c) + _ ≥ abs (z^2) - abs c := by bound + _ ≥ abs (z^2) - abs z := by bound + _ ≥ (abs z - 1) * abs z := by rw [mul_comm, mul_sub_one, ← pow_two, ← abs.map_pow] + _ ≥ 2 * abs z := by bound +``` + +### Aesop rules + +`bound` uses threes types of aesop rules: `apply`, `forward`, and closing `tactic`s. To register a +lemma as an `apply` rule, tag it with `@[bound]`. It will be automatically converted into either a +`norm apply` or `safe apply` rule depending on the number and type of its hypotheses: + +1. Nonnegativity/positivity/nonpositivity/negativity hypotheses get score 1 (those involving `0`). +2. Other inequalities get score 10. +3. Disjunctions `a ∨ b` get score 100, plus the score of `a` and `b`. + +Score `0` lemmas turn into `norm apply` rules, and score `0 < s` lemmas turn into `safe apply s` +rules. The score is roughly lexicographic ordering on the counts of the three type (guessing, +general, involving-zero), and tries to minimize the complexity of hypotheses we have to prove. +See `Mathlib.Tactic.Bound.Attribute` for the full algorithm. + +To register a lemma as a `forward` rule, tag it with `@[bound_forward]`. The most important +builtin forward rule is `le_of_lt`, so that strict inequalities can be used to prove weak +inequalities. Another example is `HasFPowerSeriesOnBall.r_pos`, so that `bound` knows that any +power series present in the context have positive radius of convergence. Custom `@[bound_forward]` +rules that similarly expose inequalities inside structures are often useful. + +### Guessing apply rules + +There are several cases where there are two standard ways to recurse down an inequality, and it is +not obvious which is correct without more information. For example, `a ≤ min b c` is registered as +a `safe apply 4` rule, since we always need to prove `a ≤ b ∧ a ≤ c`. But if we see `min a b ≤ c`, +either `a ≤ c` or `b ≤ c` suffices, and we don't know which. + +In these cases we declare a new lemma with an `∨` hypotheses that covers the two cases. Tagging +it as `@[bound]` will add a +100 penalty to the score, so that it will be used only if necessary. +Aesop will then try both ways by splitting on the resulting `∨` hypothesis. + +Currently the two types of guessing rules are +1. `min` and `max` rules, for both `≤` and `<` +2. `pow` and `rpow` monotonicity rules which branch on `1 ≤ a` or `a ≤ 1`. + +### Closing tactics + +We close numerical goals with `norm_num` and `linarith`. +-/ + +open Lean Elab Meta Term Mathlib.Tactic Syntax +open Lean.Elab.Tactic (liftMetaTactic liftMetaTactic' TacticM getMainGoal) + +namespace Mathlib.Tactic.Bound + +/-! +### `.mpr` lemmas of iff statements for use as Aesop apply rules + +Once Aesop can do general terms directly, we can remove these: + + https://github.com/leanprover-community/aesop/issues/107 +-/ + +lemma mul_lt_mul_left_of_pos_of_lt {α : Type} {a b c : α} [Mul α] [Zero α] [Preorder α] + [PosMulStrictMono α] [PosMulReflectLT α] (a0 : 0 < a) : b < c → a * b < a * c := + (mul_lt_mul_left a0).mpr + +lemma mul_lt_mul_right_of_pos_of_lt {α : Type} {a b c : α} [Mul α] [Zero α] [Preorder α] + [MulPosStrictMono α] [MulPosReflectLT α] (c0 : 0 < c) : a < b → a * c < b * c := + (mul_lt_mul_right c0).mpr + +lemma Nat.cast_pos_of_pos {R : Type} [OrderedSemiring R] [Nontrivial R] {n : ℕ} : + 0 < n → 0 < (n : R) := + Nat.cast_pos.mpr + +lemma Nat.one_le_cast_of_le {α : Type} [AddCommMonoidWithOne α] [PartialOrder α] + [CovariantClass α α (fun (x y : α) => x + y) fun (x y : α) => x ≤ y] [ZeroLEOneClass α] + [CharZero α] {n : ℕ} : 1 ≤ n → 1 ≤ (n : α) := + Nat.one_le_cast.mpr + +/-! +### Apply rules for `bound` + +Most `bound` lemmas are registered in-place where the lemma is declared. These are only the lemmas +that do not require additional imports within this file. +-/ + +-- Reflexivity +attribute [bound] le_refl + +-- 0 ≤, 0 < +attribute [bound] sq_nonneg Nat.cast_nonneg abs_nonneg Nat.zero_lt_succ pow_pos pow_nonneg + sub_nonneg_of_le sub_pos_of_lt inv_nonneg_of_nonneg inv_pos_of_pos tsub_pos_of_lt mul_pos + mul_nonneg div_pos div_nonneg add_nonneg + +-- 1 ≤, ≤ 1 +attribute [bound] Nat.one_le_cast_of_le one_le_mul_of_one_le_of_one_le + +-- ≤ +attribute [bound] le_abs_self neg_abs_le neg_le_neg tsub_le_tsub_right mul_le_mul_of_nonneg_left + mul_le_mul_of_nonneg_right le_add_of_nonneg_right le_add_of_nonneg_left le_mul_of_one_le_right + mul_le_of_le_one_right sub_le_sub add_le_add mul_le_mul + +-- < +attribute [bound] Nat.cast_pos_of_pos neg_lt_neg sub_lt_sub_left sub_lt_sub_right add_lt_add_left + add_lt_add_right mul_lt_mul_left_of_pos_of_lt mul_lt_mul_right_of_pos_of_lt + +-- min and max +attribute [bound] min_le_right min_le_left le_max_left le_max_right le_min max_le lt_min max_lt + +-- Memorize a few constants to avoid going to `norm_num` +attribute [bound] zero_le_one zero_lt_one zero_le_two zero_lt_two + +/-! +### Forward rules for `bound` +-/ + +-- Bound applies `le_of_lt` to all hypotheses +attribute [bound_forward] le_of_lt + +/-! +### Guessing rules: when we don't know how to recurse +-/ + +section Guessing + +variable {α : Type} [LinearOrder α] {a b c : α} + +-- `min` and `max` guessing lemmas +lemma le_max_of_le_left_or_le_right : a ≤ b ∨ a ≤ c → a ≤ max b c := le_max_iff.mpr +lemma lt_max_of_lt_left_or_lt_right : a < b ∨ a < c → a < max b c := lt_max_iff.mpr +lemma min_le_of_left_le_or_right_le : a ≤ c ∨ b ≤ c → min a b ≤ c := min_le_iff.mpr +lemma min_lt_of_left_lt_or_right_lt : a < c ∨ b < c → min a b < c := min_lt_iff.mpr + +-- Register guessing rules +attribute [bound] + -- Which side of the `max` should we use as the lower bound? + le_max_of_le_left_or_le_right + lt_max_of_lt_left_or_lt_right + -- Which side of the `min` should we use as the upper bound? + min_le_of_left_le_or_right_le + min_lt_of_left_lt_or_right_lt + +end Guessing + +/-! +### Closing tactics + +TODO: Kim Morrison noted that we could check for `ℕ` or `ℤ` and try `omega` as well. +-/ + +/-- Close numerical goals with `norm_num` -/ +def boundNormNum : Aesop.RuleTac := + Aesop.SingleRuleTac.toRuleTac fun i => do + let tac := do Mathlib.Meta.NormNum.elabNormNum .missing .missing .missing + let goals ← Lean.Elab.Tactic.run i.goal tac |>.run' + if !goals.isEmpty then failure + return (#[], none, some .hundred) +attribute [aesop unsafe 10% tactic (rule_sets := [Bound])] boundNormNum + +/-- Close numerical and other goals with `linarith` -/ +def boundLinarith : Aesop.RuleTac := + Aesop.SingleRuleTac.toRuleTac fun i => do + Linarith.linarith false [] {} i.goal + return (#[], none, some .hundred) +attribute [aesop unsafe 5% tactic (rule_sets := [Bound])] boundLinarith + +/-! +### `bound` tactic implementation +-/ + +/-- Aesop configuration for `bound` -/ +def boundConfig : Aesop.Options := { + enableSimp := false +} + +end Mathlib.Tactic.Bound + +/-- `bound` tactic for proving inequalities via straightforward recursion on expression structure. + +An example use case is + +``` +-- Calc example: A weak lower bound for `z ↦ z^2 + c` +lemma le_sqr_add {c z : ℂ} (cz : abs c ≤ abs z) (z3 : 3 ≤ abs z) : + 2 * abs z ≤ abs (z^2 + c) := by + calc abs (z^2 + c) + _ ≥ abs (z^2) - abs c := by bound + _ ≥ abs (z^2) - abs z := by bound + _ ≥ (abs z - 1) * abs z := by rw [mul_comm, mul_sub_one, ← pow_two, ← abs.map_pow] + _ ≥ 2 * abs z := by bound +``` + +`bound` is built on top of `aesop`, and uses +1. Apply lemmas registered via the `@[bound]` attribute +2. Forward lemmas registered via the `@[bound_forward]` attribute +3. Local hypotheses from the context +4. Optionally: additional hypotheses provided as `bound [h₀, h₁]` or similar. These are added to the + context as if by `have := hᵢ`. + +The functionality of `bound` overlaps with `positivity` and `gcongr`, but can jump back and forth +between `0 ≤ x` and `x ≤ y`-type inequalities. For example, `bound` proves + `0 ≤ c → b ≤ a → 0 ≤ a * c - b * c` +by turning the goal into `b * c ≤ a * c`, then using `mul_le_mul_of_nonneg_right`. `bound` also +contains lemmas for goals of the form `1 ≤ x, 1 < x, x ≤ 1, x < 1`. Conversely, `gcongr` can prove +inequalities for more types of relations, supports all `positivity` functionality, and is likely +faster since it is more specialized (not built atop `aesop`). -/ +syntax "bound " (" [" term,* "]")? : tactic + +-- Plain `bound` elaboration, with no hypotheses +elab_rules : tactic + | `(tactic| bound) => do + let tac ← `(tactic| aesop (rule_sets := [Bound, -default]) (config := Bound.boundConfig)) + liftMetaTactic fun g ↦ do return (← Lean.Elab.runTactic g tac.raw).1 + +-- Rewrite `bound [h₀, h₁]` into `have := h₀, have := h₁, bound`, and similar +macro_rules + | `(tactic| bound%$tk [$[$ts],*]) => do + let haves ← ts.mapM fun (t : Term) => withRef t `(tactic| have := $t) + `(tactic| ($haves;*; bound%$tk)) diff --git a/Mathlib/Tactic/Bound/Attribute.lean b/Mathlib/Tactic/Bound/Attribute.lean new file mode 100644 index 0000000000000..f2633e11e5b0c --- /dev/null +++ b/Mathlib/Tactic/Bound/Attribute.lean @@ -0,0 +1,139 @@ +/- +Copyright (c) 2024 Geoffrey Irving. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Geoffrey Irving +-/ +import Mathlib.Algebra.Group.ZeroOne +import Mathlib.Tactic.Bound.Init +import Qq + +/-! +# The `bound` attribute + +Any lemma tagged with `@[bound]` is registered as an apply rule for the `bound` tactic, by +converting it to either `norm apply` or `safe apply `. The classification is based +on the number and types of the lemma's hypotheses. +-/ + +open Lean (MetaM) +open Qq + +namespace Mathlib.Tactic.Bound + +initialize Lean.registerTraceClass `bound.attribute + +variable {u : Lean.Level} {α : Q(Type u)} + +/-- Check if an expression is zero -/ +def isZero (e : Q($α)) : MetaM Bool := + match e with + | ~q(@OfNat.ofNat.{u} _ (nat_lit 0) $i) => return true + | _ => return false + +/-- Map the arguments of an inequality expression to a score -/ +def ineqPriority (a b : Q($α)) : MetaM ℕ := do + return if (← isZero a) || (← isZero b) then 1 else 10 + +/-- Map a hypothesis type to a score -/ +partial def hypPriority (hyp : Q(Prop)) : MetaM ℕ := do + match hyp with + -- Conjunctions add scores + | ~q($a ∧ $b) => pure <| (← hypPriority a) + (← hypPriority b) + -- Guessing (disjunction) gets a big penalty + | ~q($a ∨ $b) => pure <| 100 + (← hypPriority a) + (← hypPriority b) + -- Inequalities get score 1 if they contain zero, 10 otherwise + | ~q(@LE.le _ $i $a $b) => ineqPriority a b + | ~q(@LT.lt _ $i $a $b) => ineqPriority a b + | ~q(@GE.ge _ $i $b $a) => ineqPriority a b + | ~q(@GT.gt _ $i $b $a) => ineqPriority a b + -- Assume anything else is non-relevant + | _ => pure 0 + +/-- Map a type to a score -/ +def typePriority (decl : Lean.Name) (type : Lean.Expr) : MetaM ℕ := + Lean.Meta.forallTelescope type fun xs t ↦ do + checkResult t + xs.foldlM (fun (t : ℕ) x ↦ do return t + (← argPriority x)) 0 + where + /-- Score the type of argument `x` -/ + argPriority (x : Lean.Expr) : MetaM ℕ := do + hypPriority (← Lean.Meta.inferType x) + /-- Insist that our conclusion is an inequality -/ + checkResult (t : Q(Prop)) : MetaM Unit := do match t with + | ~q(@LE.le _ $i $a $b) => return () + | ~q(@LT.lt _ $i $a $b) => return () + | ~q(@GE.ge _ $i $b $a) => return () + | ~q(@GT.gt _ $i $b $a) => return () + | _ => throwError (f!"`{decl}` has invalid type `{type}` as a 'bound' lemma: \ + it should be an inequality") + +/-- Map a theorem decl to a score (0 means `norm apply`, `0 <` means `safe apply`) -/ +def declPriority (decl : Lean.Name) : Lean.MetaM ℕ := do + match (← Lean.getEnv).find? decl with + | some info => do + typePriority decl info.type + | none => throwError "unknown declaration {decl}" + +/-- Map a score to either `norm apply` or `safe apply ` -/ +def scoreToConfig (decl : Lean.Name) (score : ℕ) : Aesop.Frontend.RuleConfig := + let (phase, priority) := match score with + | 0 => (Aesop.PhaseName.norm, 0) -- No hypotheses: this rule closes the goal immediately + | s => (Aesop.PhaseName.safe, s) + { term? := some (Lean.mkIdent decl) + phase? := phase + priority? := some (Aesop.Frontend.Priority.int priority) + builder? := some (.regular .apply) + builderOptions := {} + ruleSets := ⟨#[`Bound]⟩ } + +/-- Register a lemma as an `apply` rule for the `bound` tactic. + +A lemma is appropriate for `bound` if it proves an inequality using structurally simpler +inequalities, "recursing" on the structure of the expressions involved, assuming positivity or +nonnegativity where useful. Examples include +1. `gcongr`-like inequalities over `<` and `≤` such as `f x ≤ f y` where `f` is monotone (note that + `gcongr` supports other relations). +2. `mul_le_mul` which proves `a * b ≤ c * d` from `a ≤ c ∧ b ≤ d ∧ 0 ≤ b ∧ 0 ≤ c` +3. Positivity or nonnegativity inequalities such as `sub_nonneg`: `a ≤ b → 0 ≤ b - a` +4. Inequalities involving `1` such as `one_le_div` or `Real.one_le_exp` +5. Disjunctions where the natural recursion branches, such as `a ^ n ≤ a ^ m` when the inequality + for `n,m` depends on whether `1 ≤ a ∨ a ≤ 1`. + +Each `@[bound]` lemma is assigned a score based on the number and complexity of its hypotheses, +and the `aesop` implementation chooses lemmas with lower scores first: +1. Inequality hypotheses involving `0` add 1 to the score. +2. General inequalities add `10`. +3. Disjuctions `a ∨ b` add `100` plus the sum of the scores of `a` and `b`. + +The functionality of `bound` overlaps with `positivity` and `gcongr`, but can jump back and forth +between `0 ≤ x` and `x ≤ y`-type inequalities. For example, `bound` proves + `0 ≤ c → b ≤ a → 0 ≤ a * c - b * c` +by turning the goal into `b * c ≤ a * c`, then using `mul_le_mul_of_nonneg_right`. `bound` also +uses specialized lemmas for goals of the form `1 ≤ x, 1 < x, x ≤ 1, x < 1`. + +See also `@[bound_forward]` which marks a lemma as a forward rule for `bound`: these lemmas are +applied to hypotheses to extract inequalities (e.g. `HasPowerSeriesOnBall.r_pos`). -/ +initialize Lean.registerBuiltinAttribute { + name := `bound + descr := "Register a theorem as an apply rule for the `bound` tactic." + applicationTime := .afterCompilation + add := fun decl stx attrKind => Lean.withRef stx do + let score ← Aesop.runTermElabMAsCoreM <| declPriority decl + trace[bound.attribute] "'{decl}' has score '{score}'" + let context ← Aesop.runMetaMAsCoreM Aesop.ElabM.Context.forAdditionalGlobalRules + let (rule, ruleSets) ← Aesop.runTermElabMAsCoreM <| + (scoreToConfig decl score).buildGlobalRule.run context + for ruleSet in ruleSets do + Aesop.Frontend.addGlobalRule ruleSet rule attrKind (checkNotExists := true) + erase := fun decl => + let ruleFilter := { name := decl, scope := .global, builders := #[], phases := #[] } + Aesop.Frontend.eraseGlobalRules Aesop.RuleSetNameFilter.all ruleFilter (checkExists := true) +} + +/-- Attribute for `forward` rules for the `bound` tactic. + +`@[bound_forward]` lemmas should produce inequalities given other hypotheses that might be in the +context. A typical example is exposing an inequality field of a structure, such as +`HasPowerSeriesOnBall.r_pos`. -/ +macro "bound_forward" : attr => + `(attr|aesop safe forward (rule_sets := [$(Lean.mkIdent `Bound):ident])) diff --git a/Mathlib/Tactic/Bound/Init.lean b/Mathlib/Tactic/Bound/Init.lean new file mode 100644 index 0000000000000..68e5fdae37e04 --- /dev/null +++ b/Mathlib/Tactic/Bound/Init.lean @@ -0,0 +1,17 @@ +/- +Copyright (c) 2024 Geoffrey Irving. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Geoffrey Irving +-/ + +import Aesop.Frontend.Command + +/-! +# Bound Rule Set + +This module defines the `Bound` Aesop rule set which is used by the +`bound` tactic. Aesop rule sets only become visible once the file in which +they're declared is imported, so we must put this declaration into its own file. +-/ + +declare_aesop_rule_sets [Bound] diff --git a/Mathlib/Tactic/GCongr/Core.lean b/Mathlib/Tactic/GCongr/Core.lean index 2d2034d789474..bcb5eeb0a05d9 100644 --- a/Mathlib/Tactic/GCongr/Core.lean +++ b/Mathlib/Tactic/GCongr/Core.lean @@ -149,7 +149,9 @@ The antecedents of such a lemma are classified as generating "main goals" if the `x₁ ≈ x₂` for some "varying argument" pair `x₁`/`x₂` (and a possibly different relation `≈` to `∼`), or more generally of the form `∀ i h h' j h'', f₁ i j ≈ f₂ i j` (say) for some "varying argument" pair `f₁`/`f₂`. (Other antecedents are considered to generate "side goals".) The index of the -"varying argument" pair corresponding to each "main" antecedent is recorded. -/ +"varying argument" pair corresponding to each "main" antecedent is recorded. + +Lemmas involving `<` or `≤` can also be marked `@[bound]` for use in the related `bound` tactic. -/ initialize registerBuiltinAttribute { name := `gcongr descr := "generalized congruence" diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean index 5f7582e2c59ee..282e00a70e43b 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean @@ -4,6 +4,7 @@ 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.Topology.EMetricSpace.Basic +import Mathlib.Tactic.Bound.Attribute /-! ## Pseudo-metric spaces @@ -169,6 +170,7 @@ theorem dist_comm (x y : α) : dist x y = dist y x := theorem edist_dist (x y : α) : edist x y = ENNReal.ofReal (dist x y) := PseudoMetricSpace.edist_dist x y +@[bound] theorem dist_triangle (x y z : α) : dist x z ≤ dist x y + dist y z := PseudoMetricSpace.dist_triangle x y z @@ -231,6 +233,7 @@ theorem abs_dist_sub_le (x y z : α) : |dist x z - dist y z| ≤ dist x y := abs_sub_le_iff.2 ⟨sub_le_iff_le_add.2 (dist_triangle _ _ _), sub_le_iff_le_add.2 (dist_triangle_left _ _ _)⟩ +@[bound] theorem dist_nonneg {x y : α} : 0 ≤ dist x y := dist_nonneg' dist dist_self dist_comm dist_triangle diff --git a/scripts/noshake.json b/scripts/noshake.json index 0a86ded76cecb..c89c42678e386 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -61,6 +61,7 @@ "Mathlib.Tactic.Attr.Core", "Mathlib.Tactic.Attr.Register", "Mathlib.Tactic.Basic", + "Mathlib.Tactic.Bound.Attribute", "Mathlib.Tactic.ByContra", "Mathlib.Tactic.CC", "Mathlib.Tactic.CancelDenoms", @@ -281,6 +282,17 @@ "Mathlib.Tactic.CC.Addition": ["Mathlib.Logic.Basic", "Mathlib.Tactic.CC.Lemmas"], "Mathlib.Tactic.Basic": ["Mathlib.Tactic.Linter.OldObtain"], + "Mathlib.Tactic.Bound": + ["Mathlib.Analysis.Analytic.Basic", + "Mathlib.Algebra.Order.AbsoluteValue", + "Mathlib.Algebra.Order.BigOperators.Group.Finset", + "Mathlib.Algebra.Order.Floor", + "Mathlib.Tactic.Bound.Attribute", + "Mathlib.Tactic.Bound.Init"], + "Mathlib.Tactic.Bound.Attribute": + ["Mathlib.Algebra.Group.ZeroOne", + "Mathlib.Tactic.Bound.Init"], + "Mathlib.Tactic.Bound.Init": ["Aesop.Frontend.Command"], "Mathlib.Tactic.Attr.Register": ["Lean.Meta.Tactic.Simp.SimpTheorems"], "Mathlib.Tactic.ArithMult": ["Mathlib.Tactic.ArithMult.Init"], "Mathlib.RingTheory.PowerSeries.Basic": diff --git a/test/Bound/attribute.lean b/test/Bound/attribute.lean new file mode 100644 index 0000000000000..7cc3be49cf7ad --- /dev/null +++ b/test/Bound/attribute.lean @@ -0,0 +1,32 @@ +/- +Copyright (c) 2024 Geoffrey Irving. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Geoffrey Irving +-/ + +import Mathlib.Algebra.Order.Field.Basic +import Mathlib.Tactic.Bound + +/-! +## Tests for the `@bound` attribute + +Verify that our heuristic for the priority of a declaration produces the expected values. +-/ + +open Mathlib.Tactic.Bound (declPriority) + +/-- info: 0 -/ +#guard_msgs in +#eval declPriority `le_refl + +/-- info: 0 -/ +#guard_msgs in +#eval declPriority `sq_nonneg + +/-- info: 11 -/ +#guard_msgs in +#eval declPriority `Bound.one_lt_div_of_pos_of_lt + +/-- info: 141 -/ +#guard_msgs in +#eval declPriority `Bound.pow_le_pow_right_of_le_one_or_one_le diff --git a/test/Bound/bound.lean b/test/Bound/bound.lean new file mode 100644 index 0000000000000..cdb633309e96a --- /dev/null +++ b/test/Bound/bound.lean @@ -0,0 +1,137 @@ +/- +Copyright (c) 2024 Geoffrey Irving. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Geoffrey Irving +-/ + +import Mathlib.Analysis.SpecialFunctions.Pow.Real +import Mathlib.Tactic.Bound + +/-! +## Tests for the `bound` tactic +-/ + +open Complex (abs) +open scoped NNReal + +-- Tests that work with `bound`, but not `positivity`, `gcongr`, or `norm_num` +section bound_only +variable {a b c x y : ℝ} {z : ℂ} {n : ℕ} +example (h : x < y) : y - x > 0 := by bound +example (h : x < y) : Real.exp (y - x) > 1 := by bound +example (h : x < y) (y0 : 0 < y) : x / y < 1 := by bound +example (f : ℕ → ℝ) (h : ∀ n, f n ≥ 0) : f n ≥ 0 := by bound [h n] +example (x y : ℝ≥0) (h : x < y) : (x : ℝ) < y := by bound +example : dist a c ≤ dist a b + dist b c := by bound +example {α : Type} {s : Finset α} {f g : α → ℂ} : -- An example that requires function inference + ‖s.sum (fun x ↦ f x + g x)‖ ≤ s.sum (fun x ↦ ‖f x + g x‖) := by bound +end bound_only + +-- Calc example: A weak lower bound for `z ← z^2 + c` +example {c z : ℂ} (cz : Complex.abs c ≤ Complex.abs z) (z3 : 3 ≤ Complex.abs z) : + 2 * Complex.abs z ≤ Complex.abs (z^2 + c) := by + calc Complex.abs (z^2 + c) + _ ≥ Complex.abs (z^2) - Complex.abs c := by bound + _ ≥ Complex.abs (z^2) - Complex.abs z := by bound -- gcongr works here, not for the other two + _ ≥ (Complex.abs z - 1) * Complex.abs z := by + rw [mul_comm, mul_sub_one, ← pow_two, ← Complex.abs.map_pow] + _ ≥ 2 * Complex.abs z := by bound + +-- Testing branching functionality. None of these tests work with `positivity` or `bound`. +section guess_tests +variable {a b c : ℝ} {n m : ℕ} +example (h : a ≤ b) : a ≤ max b c := by bound +example (h : a ≤ c) : a ≤ max b c := by bound +example (h : a < b) : a < max b c := by bound +example (h : a < c) : a < max b c := by bound +example (h : a ≤ c) : min a b ≤ c := by bound +example (h : b ≤ c) : min a b ≤ c := by bound +example (h : a < c) : min a b < c := by bound +example (h : b < c) : min a b < c := by bound +example (a1 : 1 ≤ a) (h : m ≤ n) : a^m ≤ a^n := by bound +example (a0 : 0 ≤ a) (a1 : a ≤ 1) (h : n ≤ m) : a^m ≤ a^n := by bound +example (a1 : 1 ≤ a) (h : b ≤ c) : a^b ≤ a^c := by bound +example (a0 : 0 < a) (a1 : a ≤ 1) (h : c ≤ b) : a^b ≤ a^c := by bound +end guess_tests + +section positive_tests +variable {n : ℕ} {x y : ℝ} {u : ℝ≥0} {z : ℂ} +example (h : 0 < x) : x^2 > 0 := by bound +example (h : x > 0) : x^2 > 0 := by bound +example (p : x > 0) (q : y > 0) : x * y > 0 := by bound +example (p : x > 0) (q : y > 0) : x / y > 0 := by bound +example : 0 < 4 := by bound +example : 0 < 7 := by bound +example : 0 < (4 : ℝ) := by bound +example : 0 < (7 : ℝ) := by bound +example : 0 < (1 : ℝ) := by bound +example (h : u > 0) : 0 < (u : ℝ) := by bound +example : 0 < 2^n := by bound +example : 0 < (1 : ℝ)⁻¹ := by bound +end positive_tests + +section nonneg_tests +variable {n : ℕ} {x y : ℝ} {u : ℝ≥0} {z : ℂ} +example : 0 ≤ abs z := by bound +example : abs z ≥ 0 := by bound +example : x^2 ≥ 0 := by bound +example (p : x ≥ 0) (q : y ≥ 0) : x * y ≥ 0 := by bound +example (p : x ≥ 0) (q : y ≥ 0) : x / y ≥ 0 := by bound +example (p : x ≥ 0) (q : y ≥ 0) : x + y ≥ 0 := by bound +example : (n : ℝ) ≥ 0 := by bound +example : 0 ≤ 7 := by bound +example : 0 ≤ (7 : ℝ) := by bound +example : 0 ≤ (1 : ℝ) := by bound +example : 0 ≤ (u : ℝ) := by bound +example : 0 ≤ (0 : ℝ) := by bound +example : 0 ≤ 2^n := by bound +example : 0 ≤ (0 : ℝ)⁻¹ := by bound +end nonneg_tests + +section bound_tests +variable {a b c x y : ℝ} {z : ℂ} {n : ℕ} +example : (1 : ℝ) < 4 := by bound +example : (2 : ℝ) < 4 := by bound +example (n : x ≥ 0) (h : x ≤ y) : x^2 ≤ y^2 := by bound +example (n : x ≥ 0) (h : x ≤ y) : y^2 ≥ x^2 := by bound +example (n : a ≥ 0) (h : x ≤ y) : a * x ≤ a * y := by bound +example (n : a ≥ 0) (h : x ≤ y) : x * a ≤ y * a := by bound +example (bp : b ≥ 0) (xp : x ≥ 0) (ab : a ≤ b) (xy : x ≤ y) : a * x ≤ b * y := by bound +example (h : x ≤ y) : abs z * x ≤ abs z * y := by bound +example (h : x ≤ y) : a + x ≤ a + y := by bound +example (h : x ≤ y) : x + a ≤ y + a := by bound +example (ab : a ≤ b) (xy : x ≤ y) : a + x ≤ b + y := by bound +example (h : x ≥ y) : a - x ≤ a - y := by bound +example (h : x ≤ y) : x - a ≤ y - a := by bound +example (ab : a ≤ b) (xy : x ≥ y) : a - x ≤ b - y := by bound +example (h : x > 0) : x ≥ 0 := by bound +example (hc : c ≥ 0) (h : a ≤ b) : a / c ≤ b / c := by bound +example (ha : a ≥ 0) (hc : c > 0) (h : b ≥ c) : a / b ≤ a / c := by bound +example (x y : ℝ) (x0 : 0 < x) (h : x ≤ y) : x.log ≤ y.log := by bound +end bound_tests + +/-- This broke without appropriate `g.withContext` use in an older implementation of `bound`. +Leaving the test here just in case. -/ +example {s : Set ℂ} (o : IsOpen s) (z) (h : z ∈ s) : ∃ r : ℝ, r > 0 := by + rw [Metric.isOpen_iff] at o + rcases o z h with ⟨t, tp, bs⟩ + exists t/2 + clear o h bs z s + bound + +-- Test various elaboration issues +example {f : ℂ → ℂ} {z w : ℂ} {s r c e : ℝ} + (sc : ∀ {w}, abs (w - z) < s → abs (f w - f z) < e) (wz : abs (w - z) < s) (wr : abs w < r) + (h : ∀ z : ℂ, abs z < r → abs (f z) ≤ c * abs z) : + abs (f z) ≤ c * abs w + e := by + calc abs (f z) = abs (f w - (f w - f z)) := by ring_nf + _ ≤ abs (f w) + abs (f w - f z) := by bound + _ ≤ c * abs w + e := by bound [h w wr, sc wz] + +-- A test that requires reduction to weak head normal form to work (surfaced by `Hartogs.lean`) +example (x y : ℝ) (h : x < y ∧ True) : x ≤ y := by + bound [h.1] + +-- Used to fail with `unknown identifier n`, since I wasn't elaborating [] inside the goal +theorem test_unknown_identifier {f : ℕ → ℝ} (le : ∀ n, f n ≤ n) : ∀ n : ℕ, f n ≤ n := by + intro n; bound [le n] From c2f3c7e48850127da856a7da0f433e51cbeaf9e7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 6 Aug 2024 23:04:27 +0000 Subject: [PATCH 069/359] chore(Order/Filter/AtTopBot): Generalise from semilattices to directed orders (#15252) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Many lemmas requiring `[SemilatticeSup α]` actually only need `[Preorder α] [IsDirected α (· ≤ ·)]`. --- Mathlib/Order/Filter/AtTopBot.lean | 390 +++++++++++------------ Mathlib/Topology/MetricSpace/Cauchy.lean | 4 +- 2 files changed, 193 insertions(+), 201 deletions(-) diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index a13d9f02fb9eb..b0c8240c7c620 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -109,58 +109,6 @@ theorem disjoint_atBot_atTop [PartialOrder α] [Nontrivial α] : theorem disjoint_atTop_atBot [PartialOrder α] [Nontrivial α] : Disjoint (atTop : Filter α) atBot := disjoint_atBot_atTop.symm -theorem hasAntitoneBasis_atTop [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] : - (@atTop α _).HasAntitoneBasis Ici := - .iInf_principal fun _ _ ↦ Ici_subset_Ici.2 - -theorem atTop_basis [Nonempty α] [SemilatticeSup α] : (@atTop α _).HasBasis (fun _ => True) Ici := - hasAntitoneBasis_atTop.1 - -theorem atTop_eq_generate_Ici [SemilatticeSup α] : atTop = generate (range (Ici (α := α))) := by - rcases isEmpty_or_nonempty α with hα|hα - · simp only [eq_iff_true_of_subsingleton] - · simp [(atTop_basis (α := α)).eq_generate, range] - -theorem atTop_basis' [SemilatticeSup α] (a : α) : (@atTop α _).HasBasis (fun x => a ≤ x) Ici := - ⟨fun _ => - (@atTop_basis α ⟨a⟩ _).mem_iff.trans - ⟨fun ⟨x, _, hx⟩ => ⟨x ⊔ a, le_sup_right, fun _y hy => hx (le_trans le_sup_left hy)⟩, - fun ⟨x, _, hx⟩ => ⟨x, trivial, hx⟩⟩⟩ - -theorem atBot_basis [Nonempty α] [SemilatticeInf α] : (@atBot α _).HasBasis (fun _ => True) Iic := - @atTop_basis αᵒᵈ _ _ - -theorem atBot_basis' [SemilatticeInf α] (a : α) : (@atBot α _).HasBasis (fun x => x ≤ a) Iic := - @atTop_basis' αᵒᵈ _ _ - -@[instance] -theorem atTop_neBot [Nonempty α] [SemilatticeSup α] : NeBot (atTop : Filter α) := - atTop_basis.neBot_iff.2 fun _ => nonempty_Ici - -@[instance] -theorem atBot_neBot [Nonempty α] [SemilatticeInf α] : NeBot (atBot : Filter α) := - @atTop_neBot αᵒᵈ _ _ - -@[simp] -theorem mem_atTop_sets [Nonempty α] [SemilatticeSup α] {s : Set α} : - s ∈ (atTop : Filter α) ↔ ∃ a : α, ∀ b ≥ a, b ∈ s := - atTop_basis.mem_iff.trans <| exists_congr fun _ => true_and_iff _ - -@[simp] -theorem mem_atBot_sets [Nonempty α] [SemilatticeInf α] {s : Set α} : - s ∈ (atBot : Filter α) ↔ ∃ a : α, ∀ b ≤ a, b ∈ s := - @mem_atTop_sets αᵒᵈ _ _ _ - -@[simp] -theorem eventually_atTop [SemilatticeSup α] [Nonempty α] {p : α → Prop} : - (∀ᶠ x in atTop, p x) ↔ ∃ a, ∀ b ≥ a, p b := - mem_atTop_sets - -@[simp] -theorem eventually_atBot [SemilatticeInf α] [Nonempty α] {p : α → Prop} : - (∀ᶠ x in atBot, p x) ↔ ∃ a, ∀ b ≤ a, p b := - mem_atBot_sets - theorem eventually_ge_atTop [Preorder α] (a : α) : ∀ᶠ x in atTop, a ≤ x := mem_atTop a @@ -229,26 +177,6 @@ theorem Tendsto.eventually_forall_le_atBot [Preorder β] {l : Filter α} ∀ᶠ x in l, ∀ y, y ≤ f x → p y := by rw [← Filter.eventually_forall_le_atBot] at h_evtl; exact (h_evtl.comap f).filter_mono hf.le_comap -theorem atTop_basis_Ioi [Nonempty α] [SemilatticeSup α] [NoMaxOrder α] : - (@atTop α _).HasBasis (fun _ => True) Ioi := - atTop_basis.to_hasBasis (fun a ha => ⟨a, ha, Ioi_subset_Ici_self⟩) fun a ha => - (exists_gt a).imp fun _b hb => ⟨ha, Ici_subset_Ioi.2 hb⟩ - -lemma atTop_basis_Ioi' [SemilatticeSup α] [NoMaxOrder α] (a : α) : atTop.HasBasis (a < ·) Ioi := - have : Nonempty α := ⟨a⟩ - atTop_basis_Ioi.to_hasBasis (fun b _ ↦ - let ⟨c, hc⟩ := exists_gt (a ⊔ b) - ⟨c, le_sup_left.trans_lt hc, Ioi_subset_Ioi <| le_sup_right.trans hc.le⟩) fun b _ ↦ - ⟨b, trivial, Subset.rfl⟩ - -theorem atTop_countable_basis [Nonempty α] [SemilatticeSup α] [Countable α] : - HasCountableBasis (atTop : Filter α) (fun _ => True) Ici := - { atTop_basis with countable := to_countable _ } - -theorem atBot_countable_basis [Nonempty α] [SemilatticeInf α] [Countable α] : - HasCountableBasis (atBot : Filter α) (fun _ => True) Iic := - { atBot_basis with countable := to_countable _ } - instance (priority := 200) atTop.isCountablyGenerated [Preorder α] [Countable α] : (atTop : Filter <| α).IsCountablyGenerated := isCountablyGenerated_seq _ @@ -287,57 +215,109 @@ theorem tendsto_atBot_pure [PartialOrder α] [OrderBot α] (f : α → β) : Tendsto f atBot (pure <| f ⊥) := @tendsto_atTop_pure αᵒᵈ _ _ _ _ -theorem Eventually.exists_forall_of_atTop [SemilatticeSup α] [Nonempty α] {p : α → Prop} - (h : ∀ᶠ x in atTop, p x) : ∃ a, ∀ b ≥ a, p b := - eventually_atTop.mp h +section IsDirected +variable [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] {p : α → Prop} -theorem Eventually.exists_forall_of_atBot [SemilatticeInf α] [Nonempty α] {p : α → Prop} - (h : ∀ᶠ x in atBot, p x) : ∃ a, ∀ b ≤ a, p b := - eventually_atBot.mp h +theorem hasAntitoneBasis_atTop : (@atTop α _).HasAntitoneBasis Ici := + .iInf_principal fun _ _ ↦ Ici_subset_Ici.2 + +theorem atTop_basis : (@atTop α _).HasBasis (fun _ => True) Ici := hasAntitoneBasis_atTop.1 + +theorem atTop_eq_generate_Ici : atTop = generate (range (Ici (α := α))) := by + rcases isEmpty_or_nonempty α with hα|hα + · simp only [eq_iff_true_of_subsingleton] + · simp [(atTop_basis (α := α)).eq_generate, range] + +theorem atTop_basis' (a : α) : (@atTop α _).HasBasis (fun x => a ≤ x) Ici := by + refine atTop_basis.to_hasBasis (fun b _ ↦ ?_) fun b _ ↦ ⟨b, trivial, Subset.rfl⟩ + obtain ⟨c, hac, hbc⟩ := exists_ge_ge a b + exact ⟨c, hac, Ici_subset_Ici.2 hbc⟩ + +@[instance] +lemma atTop_neBot : NeBot (atTop : Filter α) := atTop_basis.neBot_iff.2 fun _ => nonempty_Ici + +@[simp] lemma mem_atTop_sets {s : Set α} : s ∈ (atTop : Filter α) ↔ ∃ a : α, ∀ b ≥ a, b ∈ s := + atTop_basis.mem_iff.trans <| exists_congr fun _ => true_and_iff _ -lemma exists_eventually_atTop [SemilatticeSup α] [Nonempty α] {r : α → β → Prop} : +@[simp] lemma eventually_atTop : (∀ᶠ x in atTop, p x) ↔ ∃ a, ∀ b ≥ a, p b := mem_atTop_sets + +theorem frequently_atTop : (∃ᶠ x in atTop, p x) ↔ ∀ a, ∃ b ≥ a, p b := + atTop_basis.frequently_iff.trans <| by simp + +alias ⟨Eventually.exists_forall_of_atTop, _⟩ := eventually_atTop +alias ⟨Frequently.forall_exists_of_atTop, _⟩ := frequently_atTop + +lemma exists_eventually_atTop {r : α → β → Prop} : (∃ b, ∀ᶠ a in atTop, r a b) ↔ ∀ᶠ a₀ in atTop, ∃ b, ∀ a ≥ a₀, r a b := by simp_rw [eventually_atTop, ← exists_swap (α := α)] exact exists_congr fun a ↦ .symm <| forall_ge_iff <| Monotone.exists fun _ _ _ hb H n hn ↦ H n (hb.trans hn) -lemma exists_eventually_atBot [SemilatticeInf α] [Nonempty α] {r : α → β → Prop} : +theorem map_atTop_eq {f : α → β} : atTop.map f = ⨅ a, 𝓟 (f '' { a' | a ≤ a' }) := + (atTop_basis.map f).eq_iInf + +lemma atTop_basis_Ioi [NoMaxOrder α] : (@atTop α _).HasBasis (fun _ => True) Ioi := + atTop_basis.to_hasBasis (fun a ha => ⟨a, ha, Ioi_subset_Ici_self⟩) fun a ha => + (exists_gt a).imp fun _b hb => ⟨ha, Ici_subset_Ioi.2 hb⟩ + +lemma atTop_basis_Ioi' [NoMaxOrder α] (a : α) : atTop.HasBasis (a < ·) Ioi := by + have : Nonempty α := ⟨a⟩ + refine atTop_basis_Ioi.to_hasBasis (fun b _ ↦ ?_) fun b _ ↦ ⟨b, trivial, Subset.rfl⟩ + obtain ⟨c, hac, hbc⟩ := exists_ge_ge a b + obtain ⟨d, hcd⟩ := exists_gt c + exact ⟨d, hac.trans_lt hcd, Ioi_subset_Ioi (hbc.trans hcd.le)⟩ + +theorem frequently_atTop' [NoMaxOrder α] : (∃ᶠ x in atTop, p x) ↔ ∀ a, ∃ b > a, p b := + atTop_basis_Ioi.frequently_iff.trans <| by simp + +lemma atTop_countable_basis [Countable α] : + HasCountableBasis (atTop : Filter α) (fun _ => True) Ici := + { atTop_basis with countable := to_countable _ } + +end IsDirected + +section IsCodirected +variable [Nonempty α] [Preorder α] [IsDirected α (· ≥ ·)] {p : α → Prop} + +lemma atBot_basis : (@atBot α _).HasBasis (fun _ => True) Iic := atTop_basis (α := αᵒᵈ) + +lemma atBot_basis' (a : α) : (@atBot α _).HasBasis (fun x => x ≤ a) Iic := atTop_basis' (α := αᵒᵈ) _ + +@[instance] lemma atBot_neBot : NeBot (atBot : Filter α) := atTop_neBot (α := αᵒᵈ) + +@[simp] lemma mem_atBot_sets {s : Set α} : s ∈ (atBot : Filter α) ↔ ∃ a : α, ∀ b ≤ a, b ∈ s := + mem_atTop_sets (α := αᵒᵈ) + +@[simp] lemma eventually_atBot : (∀ᶠ x in atBot, p x) ↔ ∃ a, ∀ b ≤ a, p b := mem_atBot_sets + +theorem frequently_atBot : (∃ᶠ x in atBot, p x) ↔ ∀ a, ∃ b ≤ a, p b := frequently_atTop (α := αᵒᵈ) + +alias ⟨Eventually.exists_forall_of_atBot, _⟩ := eventually_atBot +alias ⟨Frequently.forall_exists_of_atBot, _⟩ := frequently_atBot + +lemma exists_eventually_atBot {r : α → β → Prop} : (∃ b, ∀ᶠ a in atBot, r a b) ↔ ∀ᶠ a₀ in atBot, ∃ b, ∀ a ≤ a₀, r a b := by simp_rw [eventually_atBot, ← exists_swap (α := α)] exact exists_congr fun a ↦ .symm <| forall_le_iff <| Antitone.exists fun _ _ _ hb H n hn ↦ H n (hn.trans hb) -theorem frequently_atTop [SemilatticeSup α] [Nonempty α] {p : α → Prop} : - (∃ᶠ x in atTop, p x) ↔ ∀ a, ∃ b ≥ a, p b := - atTop_basis.frequently_iff.trans <| by simp - -theorem frequently_atBot [SemilatticeInf α] [Nonempty α] {p : α → Prop} : - (∃ᶠ x in atBot, p x) ↔ ∀ a, ∃ b ≤ a, p b := - @frequently_atTop αᵒᵈ _ _ _ - -theorem frequently_atTop' [SemilatticeSup α] [Nonempty α] [NoMaxOrder α] {p : α → Prop} : - (∃ᶠ x in atTop, p x) ↔ ∀ a, ∃ b > a, p b := - atTop_basis_Ioi.frequently_iff.trans <| by simp +theorem map_atBot_eq {f : α → β} : atBot.map f = ⨅ a, 𝓟 (f '' { a' | a' ≤ a }) := + map_atTop_eq (α := αᵒᵈ) -theorem frequently_atBot' [SemilatticeInf α] [Nonempty α] [NoMinOrder α] {p : α → Prop} : - (∃ᶠ x in atBot, p x) ↔ ∀ a, ∃ b < a, p b := - @frequently_atTop' αᵒᵈ _ _ _ _ +lemma atBot_basis_Iio [NoMinOrder α] : (@atBot α _).HasBasis (fun _ => True) Iio := + atTop_basis_Ioi (α := αᵒᵈ) -theorem Frequently.forall_exists_of_atTop [SemilatticeSup α] [Nonempty α] {p : α → Prop} - (h : ∃ᶠ x in atTop, p x) : ∀ a, ∃ b ≥ a, p b := - frequently_atTop.mp h +lemma atBot_basis_Iio' [NoMinOrder α] (a : α) : atBot.HasBasis (· < a) Iio := + atTop_basis_Ioi' (α := αᵒᵈ) a -theorem Frequently.forall_exists_of_atBot [SemilatticeInf α] [Nonempty α] {p : α → Prop} - (h : ∃ᶠ x in atBot, p x) : ∀ a, ∃ b ≤ a, p b := - frequently_atBot.mp h +theorem frequently_atBot' [NoMinOrder α] : (∃ᶠ x in atBot, p x) ↔ ∀ a, ∃ b < a, p b := + frequently_atTop' (α := αᵒᵈ) -theorem map_atTop_eq [Nonempty α] [SemilatticeSup α] {f : α → β} : - atTop.map f = ⨅ a, 𝓟 (f '' { a' | a ≤ a' }) := - (atTop_basis.map f).eq_iInf +lemma atBot_countable_basis [Countable α] : + HasCountableBasis (atBot : Filter α) (fun _ => True) Iic := + { atBot_basis with countable := to_countable _ } -theorem map_atBot_eq [Nonempty α] [SemilatticeInf α] {f : α → β} : - atBot.map f = ⨅ a, 𝓟 (f '' { a' | a' ≤ a }) := - @map_atTop_eq αᵒᵈ _ _ _ _ +end IsCodirected theorem tendsto_atTop [Preorder β] {m : α → β} {f : Filter α} : Tendsto m f atTop ↔ ∀ b, ∀ᶠ a in f, b ≤ m a := by @@ -431,14 +411,6 @@ namespace Filter ### Sequences -/ -theorem inf_map_atTop_neBot_iff [SemilatticeSup α] [Nonempty α] {F : Filter β} {u : α → β} : - NeBot (F ⊓ map u atTop) ↔ ∀ U ∈ F, ∀ N, ∃ n ≥ N, u n ∈ U := by - simp_rw [inf_neBot_iff_frequently_left, frequently_map, frequently_atTop]; rfl - -theorem inf_map_atBot_neBot_iff [SemilatticeInf α] [Nonempty α] {F : Filter β} {u : α → β} : - NeBot (F ⊓ map u atBot) ↔ ∀ U ∈ F, ∀ N, ∃ n ≤ N, u n ∈ U := - @inf_map_atTop_neBot_iff αᵒᵈ _ _ _ _ _ - theorem extraction_of_frequently_atTop' {P : ℕ → Prop} (h : ∀ N, ∃ n > N, P n) : ∃ φ : ℕ → ℕ, StrictMono φ ∧ ∀ n, P (φ n) := by choose u hu hu' using h @@ -486,28 +458,42 @@ theorem Eventually.atTop_of_arithmetic {p : ℕ → Prop} {n : ℕ} (hn : n ≠ rw [ge_iff_le, Nat.le_div_iff_mul_le hn.bot_lt, mul_comm] exact (Finset.le_sup (f := (n * N ·)) (Finset.mem_range.2 hlt)).trans hb -theorem exists_le_of_tendsto_atTop [SemilatticeSup α] [Preorder β] {u : α → β} - (h : Tendsto u atTop atTop) (a : α) (b : β) : ∃ a' ≥ a, b ≤ u a' := by +section IsDirected +variable [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] [Preorder β] {F : Filter β} {u : α → β} + +theorem inf_map_atTop_neBot_iff : NeBot (F ⊓ map u atTop) ↔ ∀ U ∈ F, ∀ N, ∃ n ≥ N, u n ∈ U := by + simp_rw [inf_neBot_iff_frequently_left, frequently_map, frequently_atTop]; rfl + +lemma exists_le_of_tendsto_atTop (h : Tendsto u atTop atTop) (a : α) (b : β) : + ∃ a' ≥ a, b ≤ u a' := by have : Nonempty α := ⟨a⟩ have : ∀ᶠ x in atTop, a ≤ x ∧ b ≤ u x := (eventually_ge_atTop a).and (h.eventually <| eventually_ge_atTop b) exact this.exists -- @[nolint ge_or_gt] -- Porting note: restore attribute -theorem exists_le_of_tendsto_atBot [SemilatticeSup α] [Preorder β] {u : α → β} - (h : Tendsto u atTop atBot) : ∀ a b, ∃ a' ≥ a, u a' ≤ b := - @exists_le_of_tendsto_atTop _ βᵒᵈ _ _ _ h +theorem exists_le_of_tendsto_atBot (h : Tendsto u atTop atBot) : + ∀ a b, ∃ a' ≥ a, u a' ≤ b := exists_le_of_tendsto_atTop (β := βᵒᵈ) h -theorem exists_lt_of_tendsto_atTop [SemilatticeSup α] [Preorder β] [NoMaxOrder β] {u : α → β} - (h : Tendsto u atTop atTop) (a : α) (b : β) : ∃ a' ≥ a, b < u a' := by +theorem exists_lt_of_tendsto_atTop [NoMaxOrder β] (h : Tendsto u atTop atTop) (a : α) (b : β) : + ∃ a' ≥ a, b < u a' := by cases' exists_gt b with b' hb' rcases exists_le_of_tendsto_atTop h a b' with ⟨a', ha', ha''⟩ exact ⟨a', ha', lt_of_lt_of_le hb' ha''⟩ -- @[nolint ge_or_gt] -- Porting note: restore attribute -theorem exists_lt_of_tendsto_atBot [SemilatticeSup α] [Preorder β] [NoMinOrder β] {u : α → β} - (h : Tendsto u atTop atBot) : ∀ a b, ∃ a' ≥ a, u a' < b := - @exists_lt_of_tendsto_atTop _ βᵒᵈ _ _ _ _ h +theorem exists_lt_of_tendsto_atBot [NoMinOrder β] (h : Tendsto u atTop atBot) : + ∀ a b, ∃ a' ≥ a, u a' < b := exists_lt_of_tendsto_atTop (β := βᵒᵈ) h + +end IsDirected + +section IsCodirected +variable [Nonempty α] [Preorder α] [IsDirected α (· ≥ ·)] {F : Filter β} {u : α → β} + +theorem inf_map_atBot_neBot_iff : NeBot (F ⊓ map u atBot) ↔ ∀ U ∈ F, ∀ N, ∃ n ≤ N, u n ∈ U := + inf_map_atTop_neBot_iff (α := αᵒᵈ) + +end IsCodirected /-- If `u` is a sequence which is unbounded above, then after any point, it reaches a value strictly greater than all previous values. @@ -1178,39 +1164,6 @@ end LinearOrderedField open Filter -theorem tendsto_atTop' [Nonempty α] [SemilatticeSup α] {f : α → β} {l : Filter β} : - Tendsto f atTop l ↔ ∀ s ∈ l, ∃ a, ∀ b ≥ a, f b ∈ s := by - simp only [tendsto_def, mem_atTop_sets, mem_preimage] - -theorem tendsto_atBot' [Nonempty α] [SemilatticeInf α] {f : α → β} {l : Filter β} : - Tendsto f atBot l ↔ ∀ s ∈ l, ∃ a, ∀ b ≤ a, f b ∈ s := - @tendsto_atTop' αᵒᵈ _ _ _ _ _ - -theorem tendsto_atTop_principal [Nonempty β] [SemilatticeSup β] {f : β → α} {s : Set α} : - Tendsto f atTop (𝓟 s) ↔ ∃ N, ∀ n ≥ N, f n ∈ s := by - simp_rw [tendsto_iff_comap, comap_principal, le_principal_iff, mem_atTop_sets, mem_preimage] - -theorem tendsto_atBot_principal [Nonempty β] [SemilatticeInf β] {f : β → α} {s : Set α} : - Tendsto f atBot (𝓟 s) ↔ ∃ N, ∀ n ≤ N, f n ∈ s := - @tendsto_atTop_principal _ βᵒᵈ _ _ _ _ - -/-- A function `f` grows to `+∞` independent of an order-preserving embedding `e`. -/ -theorem tendsto_atTop_atTop [Nonempty α] [SemilatticeSup α] [Preorder β] {f : α → β} : - Tendsto f atTop atTop ↔ ∀ b : β, ∃ i : α, ∀ a : α, i ≤ a → b ≤ f a := - Iff.trans tendsto_iInf <| forall_congr' fun _ => tendsto_atTop_principal - -theorem tendsto_atTop_atBot [Nonempty α] [SemilatticeSup α] [Preorder β] {f : α → β} : - Tendsto f atTop atBot ↔ ∀ b : β, ∃ i : α, ∀ a : α, i ≤ a → f a ≤ b := - @tendsto_atTop_atTop α βᵒᵈ _ _ _ f - -theorem tendsto_atBot_atTop [Nonempty α] [SemilatticeInf α] [Preorder β] {f : α → β} : - Tendsto f atBot atTop ↔ ∀ b : β, ∃ i : α, ∀ a : α, a ≤ i → b ≤ f a := - @tendsto_atTop_atTop αᵒᵈ β _ _ _ f - -theorem tendsto_atBot_atBot [Nonempty α] [SemilatticeInf α] [Preorder β] {f : α → β} : - Tendsto f atBot atBot ↔ ∀ b : β, ∃ i : α, ∀ a : α, a ≤ i → f a ≤ b := - @tendsto_atTop_atTop αᵒᵈ βᵒᵈ _ _ _ f - theorem tendsto_atTop_atTop_of_monotone [Preorder α] [Preorder β] {f : α → β} (hf : Monotone f) (h : ∀ b, ∃ a, b ≤ f a) : Tendsto f atTop atTop := tendsto_iInf.2 fun b => @@ -1231,23 +1184,58 @@ theorem tendsto_atBot_atTop_of_antitone [Preorder α] [Preorder β] {f : α → (h : ∀ b, ∃ a, b ≤ f a) : Tendsto f atBot atTop := @tendsto_atBot_atBot_of_monotone _ βᵒᵈ _ _ _ hf h -theorem tendsto_atTop_atTop_iff_of_monotone [Nonempty α] [SemilatticeSup α] [Preorder β] {f : α → β} - (hf : Monotone f) : Tendsto f atTop atTop ↔ ∀ b : β, ∃ a : α, b ≤ f a := +section IsDirected +variable [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] [Preorder β] {f : α → β} {l : Filter β} + +theorem tendsto_atTop' : Tendsto f atTop l ↔ ∀ s ∈ l, ∃ a, ∀ b ≥ a, f b ∈ s := by + simp only [tendsto_def, mem_atTop_sets, mem_preimage] + +theorem tendsto_atTop_principal {s : Set β} : Tendsto f atTop (𝓟 s) ↔ ∃ N, ∀ n ≥ N, f n ∈ s := by + simp_rw [tendsto_iff_comap, comap_principal, le_principal_iff, mem_atTop_sets, mem_preimage] + +/-- A function `f` grows to `+∞` independent of an order-preserving embedding `e`. -/ +theorem tendsto_atTop_atTop : Tendsto f atTop atTop ↔ ∀ b : β, ∃ i : α, ∀ a : α, i ≤ a → b ≤ f a := + tendsto_iInf.trans <| forall_congr' fun _ => tendsto_atTop_principal + +theorem tendsto_atTop_atBot : Tendsto f atTop atBot ↔ ∀ b : β, ∃ i : α, ∀ a : α, i ≤ a → f a ≤ b := + tendsto_atTop_atTop (β := βᵒᵈ) + +theorem tendsto_atTop_atTop_iff_of_monotone (hf : Monotone f) : + Tendsto f atTop atTop ↔ ∀ b : β, ∃ a, b ≤ f a := tendsto_atTop_atTop.trans <| forall_congr' fun _ => exists_congr fun a => ⟨fun h => h a (le_refl a), fun h _a' ha' => le_trans h <| hf ha'⟩ -theorem tendsto_atTop_atBot_iff_of_antitone [Nonempty α] [SemilatticeSup α] [Preorder β] {f : α → β} - (hf : Antitone f) : Tendsto f atTop atBot ↔ ∀ b : β, ∃ a : α, f a ≤ b := - @tendsto_atTop_atTop_iff_of_monotone _ βᵒᵈ _ _ _ _ hf +theorem tendsto_atTop_atBot_iff_of_antitone (hf : Antitone f) : + Tendsto f atTop atBot ↔ ∀ b : β, ∃ a, f a ≤ b := + tendsto_atTop_atTop_iff_of_monotone (β := βᵒᵈ) hf + +end IsDirected + +section IsCodirected +variable [Nonempty α] [Preorder α] [IsDirected α (· ≥ ·)] [Preorder β] {f : α → β} {l : Filter β} + +theorem tendsto_atBot' : Tendsto f atBot l ↔ ∀ s ∈ l, ∃ a, ∀ b ≤ a, f b ∈ s := + tendsto_atTop' (α := αᵒᵈ) + +theorem tendsto_atBot_principal {s : Set β} : Tendsto f atBot (𝓟 s) ↔ ∃ N, ∀ n ≤ N, f n ∈ s := + tendsto_atTop_principal (α := αᵒᵈ) (β := βᵒᵈ) -theorem tendsto_atBot_atBot_iff_of_monotone [Nonempty α] [SemilatticeInf α] [Preorder β] {f : α → β} - (hf : Monotone f) : Tendsto f atBot atBot ↔ ∀ b : β, ∃ a : α, f a ≤ b := +theorem tendsto_atBot_atTop : Tendsto f atBot atTop ↔ ∀ b : β, ∃ i : α, ∀ a : α, a ≤ i → b ≤ f a := + tendsto_atTop_atTop (α := αᵒᵈ) + +theorem tendsto_atBot_atBot : Tendsto f atBot atBot ↔ ∀ b : β, ∃ i : α, ∀ a : α, a ≤ i → f a ≤ b := + tendsto_atTop_atTop (α := αᵒᵈ) (β := βᵒᵈ) + +theorem tendsto_atBot_atBot_iff_of_monotone (hf : Monotone f) : + Tendsto f atBot atBot ↔ ∀ b : β, ∃ a, f a ≤ b := tendsto_atBot_atBot.trans <| forall_congr' fun _ => exists_congr fun a => ⟨fun h => h a (le_refl a), fun h _a' ha' => le_trans (hf ha') h⟩ -theorem tendsto_atBot_atTop_iff_of_antitone [Nonempty α] [SemilatticeInf α] [Preorder β] {f : α → β} - (hf : Antitone f) : Tendsto f atBot atTop ↔ ∀ b : β, ∃ a : α, b ≤ f a := - @tendsto_atBot_atBot_iff_of_monotone _ βᵒᵈ _ _ _ _ hf +theorem tendsto_atBot_atTop_iff_of_antitone (hf : Antitone f) : + Tendsto f atBot atTop ↔ ∀ b : β, ∃ a, b ≤ f a := + tendsto_atBot_atBot_iff_of_monotone (β := βᵒᵈ) hf + +end IsCodirected alias _root_.Monotone.tendsto_atTop_atTop := tendsto_atTop_atTop_of_monotone @@ -1313,7 +1301,6 @@ theorem tendsto_finset_preimage_atTop_atTop {f : α → β} (hf : Function.Injec (Finset.monotone_preimage hf).tendsto_atTop_finset fun x => ⟨{f x}, Finset.mem_preimage.2 <| Finset.mem_singleton_self _⟩ --- Porting note: generalized from `SemilatticeSup` to `Preorder` theorem prod_atTop_atTop_eq [Preorder α] [Preorder β] : (atTop : Filter α) ×ˢ (atTop : Filter β) = (atTop : Filter (α × β)) := by cases isEmpty_or_nonempty α @@ -1322,17 +1309,14 @@ theorem prod_atTop_atTop_eq [Preorder α] [Preorder β] : · subsingleton simpa [atTop, prod_iInf_left, prod_iInf_right, iInf_prod] using iInf_comm --- Porting note: generalized from `SemilatticeSup` to `Preorder` theorem prod_atBot_atBot_eq [Preorder α] [Preorder β] : (atBot : Filter α) ×ˢ (atBot : Filter β) = (atBot : Filter (α × β)) := @prod_atTop_atTop_eq αᵒᵈ βᵒᵈ _ _ --- Porting note: generalized from `SemilatticeSup` to `Preorder` theorem prod_map_atTop_eq {α₁ α₂ β₁ β₂ : Type*} [Preorder β₁] [Preorder β₂] (u₁ : β₁ → α₁) (u₂ : β₂ → α₂) : map u₁ atTop ×ˢ map u₂ atTop = map (Prod.map u₁ u₂) atTop := by rw [prod_map_map_eq, prod_atTop_atTop_eq, Prod.map_def] --- Porting note: generalized from `SemilatticeSup` to `Preorder` theorem prod_map_atBot_eq {α₁ α₂ β₁ β₂ : Type*} [Preorder β₁] [Preorder β₂] (u₁ : β₁ → α₁) (u₂ : β₂ → α₂) : map u₁ atBot ×ˢ map u₂ atBot = map (Prod.map u₁ u₂) atBot := @prod_map_atTop_eq _ _ β₁ᵒᵈ β₂ᵒᵈ _ _ _ _ @@ -1342,60 +1326,60 @@ theorem Tendsto.subseq_mem {F : Filter α} {V : ℕ → Set α} (h : ∀ n, V n extraction_forall_of_eventually' (fun n => tendsto_atTop'.mp hu _ (h n) : ∀ n, ∃ N, ∀ k ≥ N, u k ∈ V n) -theorem tendsto_atBot_diagonal [SemilatticeInf α] : Tendsto (fun a : α => (a, a)) atBot atBot := by +theorem tendsto_atBot_diagonal [Preorder α] : Tendsto (fun a : α => (a, a)) atBot atBot := by rw [← prod_atBot_atBot_eq] exact tendsto_id.prod_mk tendsto_id -theorem tendsto_atTop_diagonal [SemilatticeSup α] : Tendsto (fun a : α => (a, a)) atTop atTop := by +theorem tendsto_atTop_diagonal [Preorder α] : Tendsto (fun a : α => (a, a)) atTop atTop := by rw [← prod_atTop_atTop_eq] exact tendsto_id.prod_mk tendsto_id -theorem Tendsto.prod_map_prod_atBot [SemilatticeInf γ] {F : Filter α} {G : Filter β} {f : α → γ} +theorem Tendsto.prod_map_prod_atBot [Preorder γ] {F : Filter α} {G : Filter β} {f : α → γ} {g : β → γ} (hf : Tendsto f F atBot) (hg : Tendsto g G atBot) : Tendsto (Prod.map f g) (F ×ˢ G) atBot := by rw [← prod_atBot_atBot_eq] exact hf.prod_map hg -theorem Tendsto.prod_map_prod_atTop [SemilatticeSup γ] {F : Filter α} {G : Filter β} {f : α → γ} +theorem Tendsto.prod_map_prod_atTop [Preorder γ] {F : Filter α} {G : Filter β} {f : α → γ} {g : β → γ} (hf : Tendsto f F atTop) (hg : Tendsto g G atTop) : Tendsto (Prod.map f g) (F ×ˢ G) atTop := by rw [← prod_atTop_atTop_eq] exact hf.prod_map hg -theorem Tendsto.prod_atBot [SemilatticeInf α] [SemilatticeInf γ] {f g : α → γ} +theorem Tendsto.prod_atBot [Preorder α] [Preorder γ] {f g : α → γ} (hf : Tendsto f atBot atBot) (hg : Tendsto g atBot atBot) : Tendsto (Prod.map f g) atBot atBot := by rw [← prod_atBot_atBot_eq] exact hf.prod_map_prod_atBot hg -theorem Tendsto.prod_atTop [SemilatticeSup α] [SemilatticeSup γ] {f g : α → γ} +theorem Tendsto.prod_atTop [Preorder α] [Preorder γ] {f g : α → γ} (hf : Tendsto f atTop atTop) (hg : Tendsto g atTop atTop) : Tendsto (Prod.map f g) atTop atTop := by rw [← prod_atTop_atTop_eq] exact hf.prod_map_prod_atTop hg -theorem eventually_atBot_prod_self [SemilatticeInf α] [Nonempty α] {p : α × α → Prop} : - (∀ᶠ x in atBot, p x) ↔ ∃ a, ∀ k l, k ≤ a → l ≤ a → p (k, l) := by +theorem eventually_atBot_prod_self [Nonempty α] [Preorder α] [IsDirected α (· ≥ ·)] + {p : α × α → Prop} : (∀ᶠ x in atBot, p x) ↔ ∃ a, ∀ k l, k ≤ a → l ≤ a → p (k, l) := by simp [← prod_atBot_atBot_eq, (@atBot_basis α _ _).prod_self.eventually_iff] -theorem eventually_atTop_prod_self [SemilatticeSup α] [Nonempty α] {p : α × α → Prop} : - (∀ᶠ x in atTop, p x) ↔ ∃ a, ∀ k l, a ≤ k → a ≤ l → p (k, l) := +theorem eventually_atTop_prod_self [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] + {p : α × α → Prop} : (∀ᶠ x in atTop, p x) ↔ ∃ a, ∀ k l, a ≤ k → a ≤ l → p (k, l) := eventually_atBot_prod_self (α := αᵒᵈ) -theorem eventually_atBot_prod_self' [SemilatticeInf α] [Nonempty α] {p : α × α → Prop} : - (∀ᶠ x in atBot, p x) ↔ ∃ a, ∀ k ≤ a, ∀ l ≤ a, p (k, l) := by +theorem eventually_atBot_prod_self' [Nonempty α] [Preorder α] [IsDirected α (· ≥ ·)] + {p : α × α → Prop} : (∀ᶠ x in atBot, p x) ↔ ∃ a, ∀ k ≤ a, ∀ l ≤ a, p (k, l) := by simp only [eventually_atBot_prod_self, forall_cond_comm] -theorem eventually_atTop_prod_self' [SemilatticeSup α] [Nonempty α] {p : α × α → Prop} : - (∀ᶠ x in atTop, p x) ↔ ∃ a, ∀ k ≥ a, ∀ l ≥ a, p (k, l) := by +theorem eventually_atTop_prod_self' [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] + {p : α × α → Prop} : (∀ᶠ x in atTop, p x) ↔ ∃ a, ∀ k ≥ a, ∀ l ≥ a, p (k, l) := by simp only [eventually_atTop_prod_self, forall_cond_comm] -theorem eventually_atTop_curry [SemilatticeSup α] [SemilatticeSup β] {p : α × β → Prop} +theorem eventually_atTop_curry [Preorder α] [Preorder β] {p : α × β → Prop} (hp : ∀ᶠ x : α × β in Filter.atTop, p x) : ∀ᶠ k in atTop, ∀ᶠ l in atTop, p (k, l) := by rw [← prod_atTop_atTop_eq] at hp exact hp.curry -theorem eventually_atBot_curry [SemilatticeInf α] [SemilatticeInf β] {p : α × β → Prop} +theorem eventually_atBot_curry [Preorder α] [Preorder β] {p : α × β → Prop} (hp : ∀ᶠ x : α × β in Filter.atBot, p x) : ∀ᶠ k in atBot, ∀ᶠ l in atBot, p (k, l) := @eventually_atTop_curry αᵒᵈ βᵒᵈ _ _ _ hp @@ -1563,8 +1547,11 @@ theorem tendsto_atBot_atBot_of_monotone' [Preorder ι] [LinearOrder α] {u : ι (H : ¬BddBelow (range u)) : Tendsto u atBot atBot := @tendsto_atTop_atTop_of_monotone' ιᵒᵈ αᵒᵈ _ _ _ h.dual H -theorem unbounded_of_tendsto_atTop [Nonempty α] [SemilatticeSup α] [Preorder β] [NoMaxOrder β] - {f : α → β} (h : Tendsto f atTop atTop) : ¬BddAbove (range f) := by +section IsDirected +variable [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] [Preorder β] {f : α → β} + +theorem unbounded_of_tendsto_atTop [NoMaxOrder β] (h : Tendsto f atTop atTop) : + ¬BddAbove (range f) := by rintro ⟨M, hM⟩ cases' mem_atTop_sets.mp (h <| Ioi_mem_atTop M) with a ha apply lt_irrefl M @@ -1572,17 +1559,21 @@ theorem unbounded_of_tendsto_atTop [Nonempty α] [SemilatticeSup α] [Preorder M < f a := ha a le_rfl _ ≤ M := hM (Set.mem_range_self a) -theorem unbounded_of_tendsto_atBot [Nonempty α] [SemilatticeSup α] [Preorder β] [NoMinOrder β] - {f : α → β} (h : Tendsto f atTop atBot) : ¬BddBelow (range f) := - @unbounded_of_tendsto_atTop _ βᵒᵈ _ _ _ _ _ h +theorem unbounded_of_tendsto_atBot [NoMinOrder β] (h : Tendsto f atTop atBot) : + ¬BddBelow (range f) := unbounded_of_tendsto_atTop (β := βᵒᵈ) h + +end IsDirected + +section IsCodirected +variable [Nonempty α] [Preorder α] [IsDirected α (· ≥ ·)] [Preorder β] {f : α → β} + +theorem unbounded_of_tendsto_atTop' [NoMaxOrder β] (h : Tendsto f atBot atTop) : + ¬BddAbove (range f) := unbounded_of_tendsto_atTop (α := αᵒᵈ) h -theorem unbounded_of_tendsto_atTop' [Nonempty α] [SemilatticeInf α] [Preorder β] [NoMaxOrder β] - {f : α → β} (h : Tendsto f atBot atTop) : ¬BddAbove (range f) := - @unbounded_of_tendsto_atTop αᵒᵈ _ _ _ _ _ _ h +theorem unbounded_of_tendsto_atBot' [NoMinOrder β] (h : Tendsto f atBot atBot) : + ¬BddBelow (range f) := unbounded_of_tendsto_atTop (α := αᵒᵈ) (β := βᵒᵈ) h -theorem unbounded_of_tendsto_atBot' [Nonempty α] [SemilatticeInf α] [Preorder β] [NoMinOrder β] - {f : α → β} (h : Tendsto f atBot atBot) : ¬BddBelow (range f) := - @unbounded_of_tendsto_atTop αᵒᵈ βᵒᵈ _ _ _ _ _ h +end IsCodirected /-- If a monotone function `u : ι → α` tends to `atTop` along *some* non-trivial filter `l`, then it tends to `atTop` along `atTop`. -/ @@ -1634,7 +1625,8 @@ protected theorem HasAntitoneBasis.tendsto [Preorder ι] {l : Filter α} {s : ι (hl : l.HasAntitoneBasis s) {φ : ι → α} (h : ∀ i : ι, φ i ∈ s i) : Tendsto φ atTop l := fun _t ht => mem_map.2 <| (hl.eventually_subset ht).mono fun i hi => hi (h i) -theorem HasAntitoneBasis.comp_mono [SemilatticeSup ι] [Nonempty ι] [Preorder ι'] {l : Filter α} +theorem HasAntitoneBasis.comp_mono [Nonempty ι] [Preorder ι] [IsDirected ι (· ≤ ·)] [Preorder ι'] + {l : Filter α} {s : ι' → Set α} (hs : l.HasAntitoneBasis s) {φ : ι → ι'} (φ_mono : Monotone φ) (hφ : Tendsto φ atTop atTop) : l.HasAntitoneBasis (s ∘ φ) := ⟨hs.1.to_hasBasis diff --git a/Mathlib/Topology/MetricSpace/Cauchy.lean b/Mathlib/Topology/MetricSpace/Cauchy.lean index 8c227359c648c..f6af5f9a89078 100644 --- a/Mathlib/Topology/MetricSpace/Cauchy.lean +++ b/Mathlib/Topology/MetricSpace/Cauchy.lean @@ -72,8 +72,8 @@ theorem Metric.uniformCauchySeqOn_iff {γ : Type*} {F : β → γ → α} {s : S · intro h ε hε let u := { a : α × α | dist a.fst a.snd < ε } have hu : u ∈ 𝓤 α := Metric.mem_uniformity_dist.mpr ⟨ε, hε, by simp [u]⟩ - rw [← @Filter.eventually_atTop_prod_self' _ _ _ fun m => - ∀ x ∈ s, dist (F m.fst x) (F m.snd x) < ε] + rw [← Filter.eventually_atTop_prod_self' (p := fun m => + ∀ x ∈ s, dist (F m.fst x) (F m.snd x) < ε)] specialize h u hu rw [prod_atTop_atTop_eq] at h exact h.mono fun n h x hx => h x hx From af4fabaf9620392ea80b26d98e6738e381eeeac4 Mon Sep 17 00:00:00 2001 From: Yoh Tanimoto <57562556+yoh-tanimoto@users.noreply.github.com> Date: Wed, 7 Aug 2024 01:32:34 +0000 Subject: [PATCH 070/359] feat(Analysis/InnerProductSpace/Basic): add PreInnerProductSpace (#14024) add the structure `PreInnerProductSpace`, which does not assume `definite` for the inner product, with a proof of the Cauchy-Schwarz inequality using the discriminant. Motivation: Such structures arise very often in applications, and by quotienting the kernel one obtains `InnerProductSpace.Core`. I duplicated most of the proofs from the namespace `InnerProductSpace.Core` and put them under the namespace `PreInnerProductSpace` . Is there a better way? --- Mathlib/Analysis/InnerProductSpace/Basic.lean | 771 ++++++++++++------ Mathlib/LinearAlgebra/Matrix/PosDef.lean | 2 +- scripts/style-exceptions.txt | 2 +- 3 files changed, 542 insertions(+), 233 deletions(-) diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 752bd81941924..8121332042b11 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -5,6 +5,7 @@ Authors: Zhouhang Zhou, Sébastien Gouëzel, Frédéric Dupuis -/ import Mathlib.Algebra.DirectSum.Module import Mathlib.Algebra.Module.LinearMap.Basic +import Mathlib.Algebra.QuadraticDiscriminant import Mathlib.Analysis.Complex.Basic import Mathlib.Analysis.Convex.Uniform import Mathlib.Analysis.Normed.Module.Completion @@ -93,14 +94,15 @@ scoped[ComplexInnerProductSpace] notation "⟪" x ", " y "⟫" => @inner ℂ _ _ end Notations -/-- An inner product space is a vector space with an additional operation called inner product. -The norm could be derived from the inner product, instead we require the existence of a norm and -the fact that `‖x‖^2 = re ⟪x, x⟫` to be able to put instances on `𝕂` or product -spaces. +/-- A (pre) inner product space is a vector space with an additional operation called inner product. +The (semi)norm could be derived from the inner product, instead we require the existence of a +seminorm and the fact that `‖x‖^2 = re ⟪x, x⟫` to be able to put instances on `𝕂` or product spaces. -To construct a norm from an inner product, see `InnerProductSpace.ofCore`. +Note that `NormedSpace` does not assume that `‖x‖=0` implies `x=0` (it is rather a seminorm). + +To construct a seminorm from an inner product, see `PreInnerProductSpace.ofCore`. -/ -class InnerProductSpace (𝕜 : Type*) (E : Type*) [RCLike 𝕜] [NormedAddCommGroup E] extends +class InnerProductSpace (𝕜 : Type*) (E : Type*) [RCLike 𝕜] [SeminormedAddCommGroup E] extends NormedSpace 𝕜 E, Inner 𝕜 E where /-- The inner product induces the norm. -/ norm_sq_eq_inner : ∀ x : E, ‖x‖ ^ 2 = re (inner x x) @@ -129,28 +131,52 @@ Warning: Do not use this `Core` structure if the space you are interested in alr instance defined on it, otherwise this will create a second non-defeq norm instance! -/ - -/-- A structure requiring that a scalar product is positive definite and symmetric, from which one -can construct an `InnerProductSpace` instance in `InnerProductSpace.ofCore`. -/ --- @[nolint HasNonemptyInstance] porting note: I don't think we have this linter anymore -structure InnerProductSpace.Core (𝕜 : Type*) (F : Type*) [RCLike 𝕜] [AddCommGroup F] +/-- A structure requiring that a scalar product is positive semidefinite and symmetric. -/ +structure PreInnerProductSpace.Core (𝕜 : Type*) (F : Type*) [RCLike 𝕜] [AddCommGroup F] [Module 𝕜 F] extends Inner 𝕜 F where /-- The inner product is *hermitian*, taking the `conj` swaps the arguments. -/ conj_symm : ∀ x y, conj (inner y x) = inner x y /-- The inner product is positive (semi)definite. -/ nonneg_re : ∀ x, 0 ≤ re (inner x x) /-- The inner product is positive definite. -/ - definite : ∀ x, inner x x = 0 → x = 0 - /-- The inner product is additive in the first coordinate. -/ add_left : ∀ x y z, inner (x + y) z = inner x z + inner y z /-- The inner product is conjugate linear in the first coordinate. -/ smul_left : ∀ x y r, inner (r • x) y = conj r * inner x y +attribute [class] PreInnerProductSpace.Core + +/-- A structure requiring that a scalar product is positive definite. Some theorems that +require this assumptions are put under section `InnerProductSpace.Core`. -/ +-- @[nolint HasNonemptyInstance] porting note: I don't think we have this linter anymore +structure InnerProductSpace.Core (𝕜 : Type*) (F : Type*) [RCLike 𝕜] [AddCommGroup F] + [Module 𝕜 F] extends PreInnerProductSpace.Core 𝕜 F where + /-- The inner product is positive definite. -/ + definite : ∀ x, inner x x = 0 → x = 0 + /- We set `InnerProductSpace.Core` to be a class as we will use it as such in the construction of the normed space structure that it produces. However, all the instances we will use will be local to this proof. -/ attribute [class] InnerProductSpace.Core +instance (𝕜 : Type*) (F : Type*) [RCLike 𝕜] [AddCommGroup F] + [Module 𝕜 F] [cd : InnerProductSpace.Core 𝕜 F] : PreInnerProductSpace.Core 𝕜 F where + inner := cd.inner + conj_symm := cd.conj_symm + nonneg_re := cd.nonneg_re + add_left := cd.add_left + smul_left := cd.smul_left + +/-- Define `PreInnerProductSpace.Core` from `PreInnerProductSpace`. Defined to reuse lemmas about +`PreInnerProductSpace.Core` for `PreInnerProductSpace`s. Note that the `Seminorm` instance provided +by `PreInnerProductSpace.Core.norm` is propositionally but not definitionally equal to the original +norm. -/ +def PreInnerProductSpace.toCore [SeminormedAddCommGroup E] [c : InnerProductSpace 𝕜 E] : + PreInnerProductSpace.Core 𝕜 E := + { c with + nonneg_re := fun x => by + rw [← InnerProductSpace.norm_sq_eq_inner] + apply sq_nonneg } + /-- Define `InnerProductSpace.Core` from `InnerProductSpace`. Defined to reuse lemmas about `InnerProductSpace.Core` for `InnerProductSpace`s. Note that the `Norm` instance provided by `InnerProductSpace.Core.norm` is propositionally but not definitionally equal to the original @@ -167,7 +193,9 @@ def InnerProductSpace.toCore [NormedAddCommGroup E] [c : InnerProductSpace 𝕜 namespace InnerProductSpace.Core -variable [AddCommGroup F] [Module 𝕜 F] [c : InnerProductSpace.Core 𝕜 F] +section PreInnerProductSpace.Core + +variable [AddCommGroup F] [Module 𝕜 F] [c : PreInnerProductSpace.Core 𝕜 F] local notation "⟪" x ", " y "⟫" => @inner 𝕜 F _ x y @@ -179,15 +207,15 @@ local notation "ext_iff" => @RCLike.ext_iff 𝕜 _ local postfix:90 "†" => starRingEnd _ -/-- Inner product defined by the `InnerProductSpace.Core` structure. We can't reuse -`InnerProductSpace.Core.toInner` because it takes `InnerProductSpace.Core` as an explicit +/-- Inner product defined by the `PreInnerProductSpace.Core` structure. We can't reuse +`PreInnerProductSpace.Core.toInner` because it takes `PreInnerProductSpace.Core` as an explicit argument. -/ -def toInner' : Inner 𝕜 F := +def toPreInner' : Inner 𝕜 F := c.toInner -attribute [local instance] toInner' +attribute [local instance] toPreInner' -/-- The norm squared function for `InnerProductSpace.Core` structure. -/ +/-- The norm squared function for `PreInnerProductSpace.Core` structure. -/ def normSq (x : F) := reK ⟪x, x⟫ @@ -231,19 +259,16 @@ theorem inner_zero_left (x : F) : ⟪0, x⟫ = 0 := by theorem inner_zero_right (x : F) : ⟪x, 0⟫ = 0 := by rw [← inner_conj_symm, inner_zero_left]; simp only [RingHom.map_zero] -theorem inner_self_eq_zero {x : F} : ⟪x, x⟫ = 0 ↔ x = 0 := - ⟨c.definite _, by - rintro rfl - exact inner_zero_left _⟩ - -theorem normSq_eq_zero {x : F} : normSqF x = 0 ↔ x = 0 := - Iff.trans - (by simp only [normSq, ext_iff, map_zero, inner_self_im, eq_self_iff_true, and_true_iff]) - (@inner_self_eq_zero 𝕜 _ _ _ _ _ x) +theorem inner_self_of_eq_zero {x : F} : x = 0 → ⟪x, x⟫ = 0 := by + rintro rfl + exact inner_zero_left _ -theorem inner_self_ne_zero {x : F} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := - inner_self_eq_zero.not +theorem normSq_eq_zero_of_eq_zero {x : F} : x = 0 → normSqF x = 0 := by + rintro rfl + simp [normSq, inner_self_of_eq_zero] +theorem ne_zero_of_inner_self_ne_zero {x : F} : ⟪x, x⟫ ≠ 0 → x ≠ 0 := + mt inner_self_of_eq_zero theorem inner_self_ofReal_re (x : F) : (re ⟪x, x⟫ : 𝕜) = ⟪x, x⟫ := by norm_num [ext_iff, inner_self_im] @@ -274,12 +299,41 @@ theorem inner_add_add_self (x y : F) : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y theorem inner_sub_sub_self (x y : F) : ⟪x - y, x - y⟫ = ⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫ := by simp only [inner_sub_left, inner_sub_right]; ring -/-- An auxiliary equality useful to prove the **Cauchy–Schwarz inequality**: the square of the norm -of `⟪x, y⟫ • x - ⟪x, x⟫ • y` is equal to `‖x‖ ^ 2 * (‖x‖ ^ 2 * ‖y‖ ^ 2 - ‖⟪x, y⟫‖ ^ 2)`. We use -`InnerProductSpace.ofCore.normSq x` etc (defeq to `is_R_or_C.re ⟪x, x⟫`) instead of `‖x‖ ^ 2` +theorem inner_smul_ofReal_left (x y : F) {t : ℝ} : ⟪(t : 𝕜) • x, y⟫ = ⟪x, y⟫ * t := by + rw [inner_smul_left, conj_ofReal, mul_comm] + +theorem inner_smul_ofReal_right (x y : F) {t : ℝ} : ⟪x, (t : 𝕜) • y⟫ = ⟪x, y⟫ * t := by + rw [inner_smul_right, mul_comm] + +theorem re_inner_smul_ofReal_smul_self (x : F) {t : ℝ} : + re ⟪(t : 𝕜) • x, (t : 𝕜) • x⟫ = normSqF x * t * t := by + apply ofReal_injective (K := 𝕜) + simp [inner_self_ofReal_re, inner_smul_ofReal_left, inner_smul_ofReal_right, normSq] + +/-- An auxiliary equality useful to prove the **Cauchy–Schwarz inequality**. Here we use the +standard argument involving the discriminant of quadratic form. -/ +lemma cauchy_schwarz_aux' (x y : F) (t : ℝ) : 0 ≤ normSqF x * t * t + 2 * re ⟪x, y⟫ * t + + normSqF y := by + calc 0 ≤ re ⟪(ofReal t : 𝕜) • x + y, (ofReal t : 𝕜) • x + y⟫ := inner_self_nonneg + _ = re (⟪(ofReal t : 𝕜) • x, (ofReal t : 𝕜) • x⟫ + ⟪(ofReal t : 𝕜) • x, y⟫ + + ⟪y, (ofReal t : 𝕜) • x⟫ + ⟪y, y⟫) := by rw [inner_add_add_self ((ofReal t : 𝕜) • x) y] + _ = re ⟪(ofReal t : 𝕜) • x, (ofReal t : 𝕜) • x⟫ + + re ⟪(ofReal t : 𝕜) • x, y⟫ + re ⟪y, (ofReal t : 𝕜) • x⟫ + re ⟪y, y⟫ := by + simp only [map_add] + _ = normSq x * t * t + re (⟪x, y⟫ * t) + re (⟪y, x⟫ * t) + re ⟪y, y⟫ := by rw + [re_inner_smul_ofReal_smul_self, inner_smul_ofReal_left, inner_smul_ofReal_right] + _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪y, x⟫ * t + re ⟪y, y⟫ := by rw [mul_comm ⟪x,y⟫ _, + RCLike.re_ofReal_mul, mul_comm t _, mul_comm ⟪y,x⟫ _, RCLike.re_ofReal_mul, mul_comm t _] + _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪y, x⟫ * t + normSq y := by exact rfl + _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪x, y⟫ * t + normSq y := by rw [inner_re_symm] + _ = normSq x * t * t + 2 * re ⟪x, y⟫ * t + normSq y := by ring + +/-- Another auxiliary equality related with the **Cauchy–Schwarz inequality**: the square of the +seminorm of `⟪x, y⟫ • x - ⟪x, x⟫ • y` is equal to `‖x‖ ^ 2 * (‖x‖ ^ 2 * ‖y‖ ^ 2 - ‖⟪x, y⟫‖ ^ 2)`. +We use `InnerProductSpace.ofCore.normSq x` etc (defeq to `is_R_or_C.re ⟪x, x⟫`) instead of `‖x‖ ^ 2` etc to avoid extra rewrites when applying it to an `InnerProductSpace`. -/ -theorem cauchy_schwarz_aux (x y : F) : - normSqF (⟪x, y⟫ • x - ⟪x, x⟫ • y) = normSqF x * (normSqF x * normSqF y - ‖⟪x, y⟫‖ ^ 2) := by +theorem cauchy_schwarz_aux (x y : F) : normSqF (⟪x, y⟫ • x - ⟪x, x⟫ • y) + = normSqF x * (normSqF x * normSqF y - ‖⟪x, y⟫‖ ^ 2) := by rw [← @ofReal_inj 𝕜, ofReal_normSq_eq_inner_self] simp only [inner_sub_sub_self, inner_smul_left, inner_smul_right, conj_ofReal, mul_sub, ← ofReal_normSq_eq_inner_self x, ← ofReal_normSq_eq_inner_self y] @@ -288,19 +342,33 @@ theorem cauchy_schwarz_aux (x y : F) : ring /-- **Cauchy–Schwarz inequality**. -We need this for the `Core` structure to prove the triangle inequality below when -showing the core is a normed group. +We need this for the `PreInnerProductSpace.Core` structure to prove the triangle inequality below +when showing the core is a normed group and to take the quotient. -/ theorem inner_mul_inner_self_le (x y : F) : ‖⟪x, y⟫‖ * ‖⟪y, x⟫‖ ≤ re ⟪x, x⟫ * re ⟪y, y⟫ := by - rcases eq_or_ne x 0 with (rfl | hx) - · simpa only [inner_zero_left, map_zero, zero_mul, norm_zero] using le_rfl - · have hx' : 0 < normSqF x := inner_self_nonneg.lt_of_ne' (mt normSq_eq_zero.1 hx) - rw [← sub_nonneg, ← mul_nonneg_iff_right_nonneg_of_pos hx', ← normSq, ← normSq, - norm_inner_symm y, ← sq, ← cauchy_schwarz_aux] - exact inner_self_nonneg - -/-- Norm constructed from an `InnerProductSpace.Core` structure, defined to be the square root -of the scalar product. -/ + suffices discrim (normSqF x) (2 * ‖⟪x, y⟫_𝕜‖) (normSqF y) ≤ 0 by + rw [norm_inner_symm y x] + rw [discrim, normSq, normSq, sq] at this + linarith + refine discrim_le_zero fun t ↦ ?_ + by_cases hzero : ⟪x, y⟫ = 0 + · simp only [mul_assoc, ← sq, hzero, norm_zero, mul_zero, zero_mul, add_zero, ge_iff_le] + obtain ⟨hx, hy⟩ : (0 ≤ normSqF x ∧ 0 ≤ normSqF y) := ⟨inner_self_nonneg, inner_self_nonneg⟩ + positivity + · have hzero' : ‖⟪x, y⟫‖ ≠ 0 := norm_ne_zero_iff.2 hzero + convert cauchy_schwarz_aux' (𝕜 := 𝕜) (⟪x, y⟫ • x) y (t / ‖⟪x, y⟫‖) using 3 + · field_simp + rw [← sq, normSq, normSq, inner_smul_right, inner_smul_left, ← mul_assoc _ _ ⟪x, x⟫, + mul_conj] + nth_rw 2 [sq] + rw [← ofReal_mul, re_ofReal_mul] + ring + · field_simp + rw [inner_smul_left, mul_comm _ ⟪x, y⟫_𝕜, mul_conj, ← ofReal_pow, ofReal_re] + ring + +/-- (Semi)norm constructed from an `PreInnerProductSpace.Core` structure, defined to be the square +root of the scalar product. -/ def toNorm : Norm F where norm x := √(re ⟪x, x⟫) attribute [local instance] toNorm @@ -320,6 +388,73 @@ theorem norm_inner_le_norm (x y : F) : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := _ ≤ re ⟪x, x⟫ * re ⟪y, y⟫ := inner_mul_inner_self_le x y _ = ‖x‖ * ‖y‖ * (‖x‖ * ‖y‖) := by simp only [inner_self_eq_norm_mul_norm]; ring +/-- Seminormed group structure constructed from an `PreInnerProductSpace.Core` structure -/ +def toSeminormedAddCommGroup : SeminormedAddCommGroup F := + AddGroupSeminorm.toSeminormedAddCommGroup + { toFun := fun x => √(re ⟪x, x⟫) + map_zero' := by simp only [sqrt_zero, inner_zero_right, map_zero] + neg' := fun x => by simp only [inner_neg_left, neg_neg, inner_neg_right] + add_le' := fun x y => by + have h₁ : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := norm_inner_le_norm _ _ + have h₂ : re ⟪x, y⟫ ≤ ‖⟪x, y⟫‖ := re_le_norm _ + have h₃ : re ⟪x, y⟫ ≤ ‖x‖ * ‖y‖ := h₂.trans h₁ + have h₄ : re ⟪y, x⟫ ≤ ‖x‖ * ‖y‖ := by rwa [← inner_conj_symm, conj_re] + have : ‖x + y‖ * ‖x + y‖ ≤ (‖x‖ + ‖y‖) * (‖x‖ + ‖y‖) := by + simp only [← inner_self_eq_norm_mul_norm, inner_add_add_self, mul_add, mul_comm, map_add] + linarith + exact nonneg_le_nonneg_of_sq_le_sq (add_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) this } + +attribute [local instance] toSeminormedAddCommGroup + +/-- Normed space (which is actually a seminorm) structure constructed from an +`PreInnerProductSpace.Core` structure -/ +def toSeminormedSpace : NormedSpace 𝕜 F where + norm_smul_le r x := by + rw [norm_eq_sqrt_inner, inner_smul_left, inner_smul_right, ← mul_assoc] + rw [RCLike.conj_mul, ← ofReal_pow, re_ofReal_mul, sqrt_mul, ← ofReal_normSq_eq_inner_self, + ofReal_re] + · simp [sqrt_normSq_eq_norm, RCLike.sqrt_normSq_eq_norm] + · positivity + +end PreInnerProductSpace.Core + +section InnerProductSpace.Core + +variable [AddCommGroup F] [Module 𝕜 F] [cd : InnerProductSpace.Core 𝕜 F] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 F _ x y + +local notation "normSqK" => @RCLike.normSq 𝕜 _ + +local notation "reK" => @RCLike.re 𝕜 _ + +local notation "ext_iff" => @RCLike.ext_iff 𝕜 _ + +local postfix:90 "†" => starRingEnd _ + +/-- Inner product defined by the `InnerProductSpace.Core` structure. We can't reuse +`InnerProductSpace.Core.toInner` because it takes `InnerProductSpace.Core` as an explicit +argument. -/ +def toInner' : Inner 𝕜 F := + cd.toInner + +attribute [local instance] toInner' + +local notation "normSqF" => @normSq 𝕜 F _ _ _ _ + +theorem inner_self_eq_zero {x : F} : ⟪x, x⟫ = 0 ↔ x = 0 := + ⟨cd.definite _, inner_self_of_eq_zero⟩ + +theorem normSq_eq_zero {x : F} : normSqF x = 0 ↔ x = 0 := + Iff.trans + (by simp only [normSq, ext_iff, map_zero, inner_self_im, eq_self_iff_true, and_true_iff]) + (@inner_self_eq_zero 𝕜 _ _ _ _ _ x) + +theorem inner_self_ne_zero {x : F} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := + inner_self_eq_zero.not + +attribute [local instance] toNorm + /-- Normed group structure constructed from an `InnerProductSpace.Core` structure -/ def toNormedAddCommGroup : NormedAddCommGroup F := AddGroupNorm.toNormedAddCommGroup @@ -351,6 +486,8 @@ def toNormedSpace : NormedSpace 𝕜 F where end InnerProductSpace.Core +end InnerProductSpace.Core + section attribute [local instance] InnerProductSpace.Core.toNormedAddCommGroup @@ -358,13 +495,13 @@ attribute [local instance] InnerProductSpace.Core.toNormedAddCommGroup /-- Given an `InnerProductSpace.Core` structure on a space, one can use it to turn the space into an inner product space. The `NormedAddCommGroup` structure is expected to already be defined with `InnerProductSpace.ofCore.toNormedAddCommGroup`. -/ -def InnerProductSpace.ofCore [AddCommGroup F] [Module 𝕜 F] (c : InnerProductSpace.Core 𝕜 F) : +def InnerProductSpace.ofCore [AddCommGroup F] [Module 𝕜 F] (cd : InnerProductSpace.Core 𝕜 F) : InnerProductSpace 𝕜 F := - letI : NormedSpace 𝕜 F := @InnerProductSpace.Core.toNormedSpace 𝕜 F _ _ _ c - { c with + letI : NormedSpace 𝕜 F := @InnerProductSpace.Core.toNormedSpace 𝕜 F _ _ _ cd + { cd with norm_sq_eq_inner := fun x => by - have h₁ : ‖x‖ ^ 2 = √(re (c.inner x x)) ^ 2 := rfl - have h₂ : 0 ≤ re (c.inner x x) := InnerProductSpace.Core.inner_self_nonneg + have h₁ : ‖x‖ ^ 2 = √(re (cd.inner x x)) ^ 2 := rfl + have h₂ : 0 ≤ re (cd.inner x x) := InnerProductSpace.Core.inner_self_nonneg simp [h₁, sq_sqrt, h₂] } end @@ -372,8 +509,10 @@ end /-! ### Properties of inner product spaces -/ -variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] -variable [NormedAddCommGroup F] [InnerProductSpace ℝ F] +section BasicProperties_Seminormed + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] +variable [SeminormedAddCommGroup F] [InnerProductSpace ℝ F] local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y @@ -383,8 +522,6 @@ local postfix:90 "†" => starRingEnd _ export InnerProductSpace (norm_sq_eq_inner) -section BasicProperties - @[simp] theorem inner_conj_symm (x y : E) : ⟪y, x⟫† = ⟪x, y⟫ := InnerProductSpace.conj_symm _ _ @@ -490,7 +627,7 @@ theorem inner_re_zero_right (x : E) : re ⟪x, 0⟫ = 0 := by simp only [inner_zero_right, AddMonoidHom.map_zero] theorem inner_self_nonneg {x : E} : 0 ≤ re ⟪x, x⟫ := - InnerProductSpace.toCore.nonneg_re x + PreInnerProductSpace.toCore.nonneg_re x theorem real_inner_self_nonneg {x : F} : 0 ≤ ⟪x, x⟫_ℝ := @inner_self_nonneg ℝ F _ _ _ x @@ -514,20 +651,6 @@ theorem inner_self_ofReal_norm (x : E) : (‖⟪x, x⟫‖ : 𝕜) = ⟪x, x⟫ theorem real_inner_self_abs (x : F) : |⟪x, x⟫_ℝ| = ⟪x, x⟫_ℝ := @inner_self_ofReal_norm ℝ F _ _ _ x -@[simp] -theorem inner_self_eq_zero {x : E} : ⟪x, x⟫ = 0 ↔ x = 0 := by - rw [inner_self_eq_norm_sq_to_K, sq_eq_zero_iff, ofReal_eq_zero, norm_eq_zero] - -theorem inner_self_ne_zero {x : E} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := - inner_self_eq_zero.not - -@[simp] -theorem inner_self_nonpos {x : E} : re ⟪x, x⟫ ≤ 0 ↔ x = 0 := by - rw [← norm_sq_eq_inner, (sq_nonneg _).le_iff_eq, sq_eq_zero_iff, norm_eq_zero] - -theorem real_inner_self_nonpos {x : F} : ⟪x, x⟫_ℝ ≤ 0 ↔ x = 0 := - @inner_self_nonpos ℝ F _ _ _ x - theorem norm_inner_symm (x y : E) : ‖⟪x, y⟫‖ = ‖⟪y, x⟫‖ := by rw [← inner_conj_symm, norm_conj] @[simp] @@ -576,16 +699,6 @@ theorem real_inner_sub_sub_self (x y : F) : simp only [inner_sub_sub_self, this, add_left_inj] ring -variable (𝕜) - -theorem ext_inner_left {x y : E} (h : ∀ v, ⟪v, x⟫ = ⟪v, y⟫) : x = y := by - rw [← sub_eq_zero, ← @inner_self_eq_zero 𝕜, inner_sub_right, sub_eq_zero, h (x - y)] - -theorem ext_inner_right {x y : E} (h : ∀ v, ⟪x, v⟫ = ⟪y, v⟫) : x = y := by - rw [← sub_eq_zero, ← @inner_self_eq_zero 𝕜, inner_sub_left, sub_eq_zero, h (x - y)] - -variable {𝕜} - /-- Parallelogram law -/ theorem parallelogram_law {x y : E} : ⟪x + y, x + y⟫ + ⟪x - y, x - y⟫ = 2 * (⟪x, x⟫ + ⟪y, y⟫) := by simp only [inner_add_add_self, inner_sub_sub_self] @@ -593,7 +706,7 @@ theorem parallelogram_law {x y : E} : ⟪x + y, x + y⟫ + ⟪x - y, x - y⟫ = /-- **Cauchy–Schwarz inequality**. -/ theorem inner_mul_inner_self_le (x y : E) : ‖⟪x, y⟫‖ * ‖⟪y, x⟫‖ ≤ re ⟪x, x⟫ * re ⟪y, y⟫ := - letI c : InnerProductSpace.Core 𝕜 E := InnerProductSpace.toCore + letI cd : PreInnerProductSpace.Core 𝕜 E := PreInnerProductSpace.toCore InnerProductSpace.Core.inner_mul_inner_self_le x y /-- Cauchy–Schwarz inequality for real inner products. -/ @@ -604,6 +717,46 @@ theorem real_inner_mul_inner_self_le (x y : F) : ⟪x, y⟫_ℝ * ⟪x, y⟫_ℝ exact le_abs_self _ _ ≤ ⟪x, x⟫_ℝ * ⟪y, y⟫_ℝ := @inner_mul_inner_self_le ℝ _ _ _ _ x y + +end BasicProperties_Seminormed + +section BasicProperties + +variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] +variable [NormedAddCommGroup F] [InnerProductSpace ℝ F] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ + +export InnerProductSpace (norm_sq_eq_inner) + +@[simp] +theorem inner_self_eq_zero {x : E} : ⟪x, x⟫ = 0 ↔ x = 0 := by + rw [inner_self_eq_norm_sq_to_K, sq_eq_zero_iff, ofReal_eq_zero, norm_eq_zero] + +theorem inner_self_ne_zero {x : E} : ⟪x, x⟫ ≠ 0 ↔ x ≠ 0 := + inner_self_eq_zero.not + +variable (𝕜) + +theorem ext_inner_left {x y : E} (h : ∀ v, ⟪v, x⟫ = ⟪v, y⟫) : x = y := by + rw [← sub_eq_zero, ← @inner_self_eq_zero 𝕜, inner_sub_right, sub_eq_zero, h (x - y)] + +theorem ext_inner_right {x y : E} (h : ∀ v, ⟪x, v⟫ = ⟪y, v⟫) : x = y := by + rw [← sub_eq_zero, ← @inner_self_eq_zero 𝕜, inner_sub_left, sub_eq_zero, h (x - y)] + +variable {𝕜} + +@[simp] +theorem inner_self_nonpos {x : E} : re ⟪x, x⟫ ≤ 0 ↔ x = 0 := by + rw [← norm_sq_eq_inner, (sq_nonneg _).le_iff_eq, sq_eq_zero_iff, norm_eq_zero] + +theorem real_inner_self_nonpos {x : F} : ⟪x, x⟫_ℝ ≤ 0 ↔ x = 0 := + @inner_self_nonpos ℝ F _ _ _ x + /-- A family of vectors is linearly independent if they are nonzero and orthogonal. -/ theorem linearIndependent_of_ne_zero_of_inner_eq_zero {ι : Type*} {v : ι → E} (hz : ∀ i, v i ≠ 0) @@ -622,7 +775,16 @@ theorem linearIndependent_of_ne_zero_of_inner_eq_zero {ι : Type*} {v : ι → E end BasicProperties -section OrthonormalSets +section OrthonormalSets_Seminormed + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] +variable [SeminormedAddCommGroup F] [InnerProductSpace ℝ F] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ variable {ι : Type*} (𝕜) @@ -826,12 +988,6 @@ theorem exists_maximal_orthonormal {s : Set E} (hs : Orthonormal 𝕜 (Subtype.v · exact orthonormal_sUnion_of_directed cc.directedOn fun x xc => hc xc · exact fun _ => Set.subset_sUnion_of_mem -theorem Orthonormal.ne_zero {v : ι → E} (hv : Orthonormal 𝕜 v) (i : ι) : v i ≠ 0 := by - have : ‖v i‖ ≠ 0 := by - rw [hv.1 i] - norm_num - simpa using this - open FiniteDimensional /-- A family of orthonormal vectors with the correct cardinality forms a basis. -/ @@ -845,9 +1001,38 @@ theorem coe_basisOfOrthonormalOfCardEqFinrank [Fintype ι] [Nonempty ι] {v : ι (basisOfOrthonormalOfCardEqFinrank hv card_eq : ι → E) = v := coe_basisOfLinearIndependentOfCardEqFinrank _ _ +end OrthonormalSets_Seminormed + +section OrthonormalSets + +variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] +variable [NormedAddCommGroup F] [InnerProductSpace ℝ F] +variable {ι : Type*} + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ + +theorem Orthonormal.ne_zero {v : ι → E} (hv : Orthonormal 𝕜 v) (i : ι) : v i ≠ 0 := by + have : ‖v i‖ ≠ 0 := by + rw [hv.1 i] + norm_num + simpa using this + end OrthonormalSets -section Norm +section Norm_Seminormed + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] +variable [SeminormedAddCommGroup F] [InnerProductSpace ℝ F] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ theorem norm_eq_sqrt_inner (x : E) : ‖x‖ = √(re ⟪x, x⟫) := calc @@ -930,7 +1115,7 @@ theorem norm_sub_mul_self_real (x y : F) : /-- Cauchy–Schwarz inequality with norm -/ theorem norm_inner_le_norm (x y : E) : ‖⟪x, y⟫‖ ≤ ‖x‖ * ‖y‖ := by rw [norm_eq_sqrt_inner (𝕜 := 𝕜) x, norm_eq_sqrt_inner (𝕜 := 𝕜) y] - letI : InnerProductSpace.Core 𝕜 E := InnerProductSpace.toCore + letI : PreInnerProductSpace.Core 𝕜 E := PreInnerProductSpace.toCore exact InnerProductSpace.Core.norm_inner_le_norm x y theorem nnnorm_inner_le_nnnorm (x y : E) : ‖⟪x, y⟫‖₊ ≤ ‖x‖₊ * ‖y‖₊ := @@ -994,25 +1179,6 @@ theorem inner_eq_sum_norm_sq_div_four (x y : E) : push_cast simp only [sq, ← mul_div_right_comm, ← add_div] -/-- Formula for the distance between the images of two nonzero points under an inversion with center -zero. See also `EuclideanGeometry.dist_inversion_inversion` for inversions around a general -point. -/ -theorem dist_div_norm_sq_smul {x y : F} (hx : x ≠ 0) (hy : y ≠ 0) (R : ℝ) : - dist ((R / ‖x‖) ^ 2 • x) ((R / ‖y‖) ^ 2 • y) = R ^ 2 / (‖x‖ * ‖y‖) * dist x y := - have hx' : ‖x‖ ≠ 0 := norm_ne_zero_iff.2 hx - have hy' : ‖y‖ ≠ 0 := norm_ne_zero_iff.2 hy - calc - dist ((R / ‖x‖) ^ 2 • x) ((R / ‖y‖) ^ 2 • y) = - √(‖(R / ‖x‖) ^ 2 • x - (R / ‖y‖) ^ 2 • y‖ ^ 2) := by - rw [dist_eq_norm, sqrt_sq (norm_nonneg _)] - _ = √((R ^ 2 / (‖x‖ * ‖y‖)) ^ 2 * ‖x - y‖ ^ 2) := - congr_arg sqrt <| by - field_simp [sq, norm_sub_mul_self_real, norm_smul, real_inner_smul_left, inner_smul_right, - Real.norm_of_nonneg (mul_self_nonneg _)] - ring - _ = R ^ 2 / (‖x‖ * ‖y‖) * dist x y := by - rw [sqrt_mul, sqrt_sq, sqrt_sq, dist_eq_norm] <;> positivity - -- See note [lower instance priority] instance (priority := 100) InnerProductSpace.toUniformConvexSpace : UniformConvexSpace F := ⟨fun ε hε => by @@ -1080,8 +1246,8 @@ end Complex section variable {ι : Type*} {ι' : Type*} {ι'' : Type*} -variable {E' : Type*} [NormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] -variable {E'' : Type*} [NormedAddCommGroup E''] [InnerProductSpace 𝕜 E''] +variable {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] +variable {E'' : Type*} [SeminormedAddCommGroup E''] [InnerProductSpace 𝕜 E''] /-- A linear isometry preserves the inner product. -/ @[simp] @@ -1211,19 +1377,6 @@ theorem Orthonormal.equiv_apply {ι' : Type*} {v : Basis ι 𝕜 E} (hv : Orthon hv.equiv hv' e (v i) = v' (e i) := Basis.equiv_apply _ _ _ _ -@[simp] -theorem Orthonormal.equiv_refl {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) : - hv.equiv hv (Equiv.refl ι) = LinearIsometryEquiv.refl 𝕜 E := - v.ext_linearIsometryEquiv fun i => by - simp only [Orthonormal.equiv_apply, Equiv.coe_refl, id, LinearIsometryEquiv.coe_refl] - -@[simp] -theorem Orthonormal.equiv_symm {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) {v' : Basis ι' 𝕜 E'} - (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') : (hv.equiv hv' e).symm = hv'.equiv hv e.symm := - v'.ext_linearIsometryEquiv fun i => - (hv.equiv hv' e).injective <| by - simp only [LinearIsometryEquiv.apply_symm_apply, Orthonormal.equiv_apply, e.apply_symm_apply] - @[simp] theorem Orthonormal.equiv_trans {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) {v' : Basis ι' 𝕜 E'} (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') {v'' : Basis ι'' 𝕜 E''} (hv'' : Orthonormal 𝕜 v'') @@ -1330,29 +1483,168 @@ theorem real_inner_smul_self_left (x : F) (r : ℝ) : ⟪r • x, x⟫_ℝ = r * theorem real_inner_smul_self_right (x : F) (r : ℝ) : ⟪x, r • x⟫_ℝ = r * (‖x‖ * ‖x‖) := by rw [inner_smul_right, ← real_inner_self_eq_norm_mul_norm] -/-- The inner product of a nonzero vector with a nonzero multiple of -itself, divided by the product of their norms, has absolute value -1. -/ -theorem norm_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul {x : E} {r : 𝕜} (hx : x ≠ 0) - (hr : r ≠ 0) : ‖⟪x, r • x⟫‖ / (‖x‖ * ‖r • x‖) = 1 := by - have hx' : ‖x‖ ≠ 0 := by simp [hx] - have hr' : ‖r‖ ≠ 0 := by simp [hr] - rw [inner_smul_right, norm_mul, ← inner_self_re_eq_norm, inner_self_eq_norm_mul_norm, norm_smul] - rw [← mul_assoc, ← div_div, mul_div_cancel_right₀ _ hx', ← div_div, mul_comm, - mul_div_cancel_right₀ _ hr', div_self hx'] +variable (𝕜) -/-- The inner product of a nonzero vector with a nonzero multiple of -itself, divided by the product of their norms, has absolute value -1. -/ -theorem abs_real_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul {x : F} {r : ℝ} - (hx : x ≠ 0) (hr : r ≠ 0) : |⟪x, r • x⟫_ℝ| / (‖x‖ * ‖r • x‖) = 1 := - norm_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul hx hr +/-- The inner product as a sesquilinear map. -/ +def innerₛₗ : E →ₗ⋆[𝕜] E →ₗ[𝕜] 𝕜 := + LinearMap.mk₂'ₛₗ _ _ (fun v w => ⟪v, w⟫) inner_add_left (fun _ _ _ => inner_smul_left _ _ _) + inner_add_right fun _ _ _ => inner_smul_right _ _ _ -/-- The inner product of a nonzero vector with a positive multiple of -itself, divided by the product of their norms, has value 1. -/ -theorem real_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_pos_mul {x : F} {r : ℝ} (hx : x ≠ 0) - (hr : 0 < r) : ⟪x, r • x⟫_ℝ / (‖x‖ * ‖r • x‖) = 1 := by - rw [real_inner_smul_self_right, norm_smul, Real.norm_eq_abs, ← mul_assoc ‖x‖, mul_comm _ |r|, +@[simp] +theorem innerₛₗ_apply_coe (v : E) : ⇑(innerₛₗ 𝕜 v) = fun w => ⟪v, w⟫ := + rfl + +@[simp] +theorem innerₛₗ_apply (v w : E) : innerₛₗ 𝕜 v w = ⟪v, w⟫ := + rfl + +variable (F) +/-- The inner product as a bilinear map in the real case. -/ +def innerₗ : F →ₗ[ℝ] F →ₗ[ℝ] ℝ := innerₛₗ ℝ + +@[simp] lemma flip_innerₗ : (innerₗ F).flip = innerₗ F := by + ext v w + exact real_inner_comm v w + +variable {F} + +@[simp] lemma innerₗ_apply (v w : F) : innerₗ F v w = ⟪v, w⟫_ℝ := rfl + +/-- The inner product as a continuous sesquilinear map. Note that `toDualMap` (resp. `toDual`) +in `InnerProductSpace.Dual` is a version of this given as a linear isometry (resp. linear +isometric equivalence). -/ +def innerSL : E →L⋆[𝕜] E →L[𝕜] 𝕜 := + LinearMap.mkContinuous₂ (innerₛₗ 𝕜) 1 fun x y => by + simp only [norm_inner_le_norm, one_mul, innerₛₗ_apply] + +@[simp] +theorem innerSL_apply_coe (v : E) : ⇑(innerSL 𝕜 v) = fun w => ⟪v, w⟫ := + rfl + +@[simp] +theorem innerSL_apply (v w : E) : innerSL 𝕜 v w = ⟪v, w⟫ := + rfl + +/-- The inner product as a continuous sesquilinear map, with the two arguments flipped. -/ +def innerSLFlip : E →L[𝕜] E →L⋆[𝕜] 𝕜 := + @ContinuousLinearMap.flipₗᵢ' 𝕜 𝕜 𝕜 E E 𝕜 _ _ _ _ _ _ _ _ _ (RingHom.id 𝕜) (starRingEnd 𝕜) _ _ + (innerSL 𝕜) + +@[simp] +theorem innerSLFlip_apply (x y : E) : innerSLFlip 𝕜 x y = ⟪y, x⟫ := + rfl + +variable (F) in +@[simp] lemma innerSL_real_flip : (innerSL ℝ (E := F)).flip = innerSL ℝ := by + ext v w + exact real_inner_comm _ _ + +variable {𝕜} + +namespace ContinuousLinearMap + +variable {E' : Type*} [NormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] + +-- Note: odd and expensive build behavior is explicitly turned off using `noncomputable` +/-- Given `f : E →L[𝕜] E'`, construct the continuous sesquilinear form `fun x y ↦ ⟪x, A y⟫`, given +as a continuous linear map. -/ +noncomputable def toSesqForm : (E →L[𝕜] E') →L[𝕜] E' →L⋆[𝕜] E →L[𝕜] 𝕜 := + (ContinuousLinearMap.flipₗᵢ' E E' 𝕜 (starRingEnd 𝕜) (RingHom.id 𝕜)).toContinuousLinearEquiv ∘L + ContinuousLinearMap.compSL E E' (E' →L⋆[𝕜] 𝕜) (RingHom.id 𝕜) (RingHom.id 𝕜) (innerSLFlip 𝕜) + +@[simp] +theorem toSesqForm_apply_coe (f : E →L[𝕜] E') (x : E') : toSesqForm f x = (innerSL 𝕜 x).comp f := + rfl + +theorem toSesqForm_apply_norm_le {f : E →L[𝕜] E'} {v : E'} : ‖toSesqForm f v‖ ≤ ‖f‖ * ‖v‖ := by + refine opNorm_le_bound _ (by positivity) fun x ↦ ?_ + have h₁ : ‖f x‖ ≤ ‖f‖ * ‖x‖ := le_opNorm _ _ + have h₂ := @norm_inner_le_norm 𝕜 E' _ _ _ v (f x) + calc + ‖⟪v, f x⟫‖ ≤ ‖v‖ * ‖f x‖ := h₂ + _ ≤ ‖v‖ * (‖f‖ * ‖x‖) := mul_le_mul_of_nonneg_left h₁ (norm_nonneg v) + _ = ‖f‖ * ‖v‖ * ‖x‖ := by ring + + +end ContinuousLinearMap + +end Norm_Seminormed + +section Norm + +variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] +variable [NormedAddCommGroup F] [InnerProductSpace ℝ F] +variable {ι : Type*} {ι' : Type*} {ι'' : Type*} + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ + +/-- Formula for the distance between the images of two nonzero points under an inversion with center +zero. See also `EuclideanGeometry.dist_inversion_inversion` for inversions around a general +point. -/ +theorem dist_div_norm_sq_smul {x y : F} (hx : x ≠ 0) (hy : y ≠ 0) (R : ℝ) : + dist ((R / ‖x‖) ^ 2 • x) ((R / ‖y‖) ^ 2 • y) = R ^ 2 / (‖x‖ * ‖y‖) * dist x y := + have hx' : ‖x‖ ≠ 0 := norm_ne_zero_iff.2 hx + have hy' : ‖y‖ ≠ 0 := norm_ne_zero_iff.2 hy + calc + dist ((R / ‖x‖) ^ 2 • x) ((R / ‖y‖) ^ 2 • y) = + √(‖(R / ‖x‖) ^ 2 • x - (R / ‖y‖) ^ 2 • y‖ ^ 2) := by + rw [dist_eq_norm, sqrt_sq (norm_nonneg _)] + _ = √((R ^ 2 / (‖x‖ * ‖y‖)) ^ 2 * ‖x - y‖ ^ 2) := + congr_arg sqrt <| by + field_simp [sq, norm_sub_mul_self_real, norm_smul, real_inner_smul_left, inner_smul_right, + Real.norm_of_nonneg (mul_self_nonneg _)] + ring + _ = R ^ 2 / (‖x‖ * ‖y‖) * dist x y := by + rw [sqrt_mul, sqrt_sq, sqrt_sq, dist_eq_norm] <;> positivity + +section + +variable {ι : Type*} {ι' : Type*} {ι'' : Type*} +variable {E' : Type*} [NormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] +variable {E'' : Type*} [NormedAddCommGroup E''] [InnerProductSpace 𝕜 E''] + +@[simp] +theorem Orthonormal.equiv_refl {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) : + hv.equiv hv (Equiv.refl ι) = LinearIsometryEquiv.refl 𝕜 E := + v.ext_linearIsometryEquiv fun i => by + simp only [Orthonormal.equiv_apply, Equiv.coe_refl, id, LinearIsometryEquiv.coe_refl] + +@[simp] +theorem Orthonormal.equiv_symm {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) {v' : Basis ι' 𝕜 E'} + (hv' : Orthonormal 𝕜 v') (e : ι ≃ ι') : (hv.equiv hv' e).symm = hv'.equiv hv e.symm := + v'.ext_linearIsometryEquiv fun i => + (hv.equiv hv' e).injective <| by + simp only [LinearIsometryEquiv.apply_symm_apply, Orthonormal.equiv_apply, e.apply_symm_apply] + +end + +/-- The inner product of a nonzero vector with a nonzero multiple of +itself, divided by the product of their norms, has absolute value +1. -/ +theorem norm_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul {x : E} {r : 𝕜} (hx : x ≠ 0) + (hr : r ≠ 0) : ‖⟪x, r • x⟫‖ / (‖x‖ * ‖r • x‖) = 1 := by + have hx' : ‖x‖ ≠ 0 := by simp [hx] + have hr' : ‖r‖ ≠ 0 := by simp [hr] + rw [inner_smul_right, norm_mul, ← inner_self_re_eq_norm, inner_self_eq_norm_mul_norm, norm_smul] + rw [← mul_assoc, ← div_div, mul_div_cancel_right₀ _ hx', ← div_div, mul_comm, + mul_div_cancel_right₀ _ hr', div_self hx'] + +/-- The inner product of a nonzero vector with a nonzero multiple of +itself, divided by the product of their norms, has absolute value +1. -/ +theorem abs_real_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul {x : F} {r : ℝ} + (hx : x ≠ 0) (hr : r ≠ 0) : |⟪x, r • x⟫_ℝ| / (‖x‖ * ‖r • x‖) = 1 := + norm_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_ne_zero_mul hx hr + +/-- The inner product of a nonzero vector with a positive multiple of +itself, divided by the product of their norms, has value 1. -/ +theorem real_inner_div_norm_mul_norm_eq_one_of_ne_zero_of_pos_mul {x : F} {r : ℝ} (hx : x ≠ 0) + (hr : 0 < r) : ⟪x, r • x⟫_ℝ / (‖x‖ * ‖r • x‖) = 1 := by + rw [real_inner_smul_self_right, norm_smul, Real.norm_eq_abs, ← mul_assoc ‖x‖, mul_comm _ |r|, mul_assoc, abs_of_nonneg hr.le, div_self] exact mul_ne_zero hr.ne' (mul_self_ne_zero.2 (norm_ne_zero_iff.2 hx)) @@ -1519,46 +1811,6 @@ theorem inner_sum_smul_sum_smul_of_sum_eq_zero {ι₁ : Type*} {s₁ : Finset ι variable (𝕜) -/-- The inner product as a sesquilinear map. -/ -def innerₛₗ : E →ₗ⋆[𝕜] E →ₗ[𝕜] 𝕜 := - LinearMap.mk₂'ₛₗ _ _ (fun v w => ⟪v, w⟫) inner_add_left (fun _ _ _ => inner_smul_left _ _ _) - inner_add_right fun _ _ _ => inner_smul_right _ _ _ - -@[simp] -theorem innerₛₗ_apply_coe (v : E) : ⇑(innerₛₗ 𝕜 v) = fun w => ⟪v, w⟫ := - rfl - -@[simp] -theorem innerₛₗ_apply (v w : E) : innerₛₗ 𝕜 v w = ⟪v, w⟫ := - rfl - -variable (F) -/-- The inner product as a bilinear map in the real case. -/ -def innerₗ : F →ₗ[ℝ] F →ₗ[ℝ] ℝ := innerₛₗ ℝ - -@[simp] lemma flip_innerₗ : (innerₗ F).flip = innerₗ F := by - ext v w - exact real_inner_comm v w - -variable {F} - -@[simp] lemma innerₗ_apply (v w : F) : innerₗ F v w = ⟪v, w⟫_ℝ := rfl - -/-- The inner product as a continuous sesquilinear map. Note that `toDualMap` (resp. `toDual`) -in `InnerProductSpace.Dual` is a version of this given as a linear isometry (resp. linear -isometric equivalence). -/ -def innerSL : E →L⋆[𝕜] E →L[𝕜] 𝕜 := - LinearMap.mkContinuous₂ (innerₛₗ 𝕜) 1 fun x y => by - simp only [norm_inner_le_norm, one_mul, innerₛₗ_apply] - -@[simp] -theorem innerSL_apply_coe (v : E) : ⇑(innerSL 𝕜 v) = fun w => ⟪v, w⟫ := - rfl - -@[simp] -theorem innerSL_apply (v w : E) : innerSL 𝕜 v w = ⟪v, w⟫ := - rfl - /-- `innerSL` is an isometry. Note that the associated `LinearIsometry` is defined in `InnerProductSpace.Dual` as `toDualMap`. -/ @[simp] @@ -1576,48 +1828,8 @@ theorem innerSL_apply_norm (x : E) : ‖innerSL 𝕜 x‖ = ‖x‖ := by lemma norm_innerSL_le : ‖innerSL 𝕜 (E := E)‖ ≤ 1 := ContinuousLinearMap.opNorm_le_bound _ zero_le_one (by simp) -/-- The inner product as a continuous sesquilinear map, with the two arguments flipped. -/ -def innerSLFlip : E →L[𝕜] E →L⋆[𝕜] 𝕜 := - @ContinuousLinearMap.flipₗᵢ' 𝕜 𝕜 𝕜 E E 𝕜 _ _ _ _ _ _ _ _ _ (RingHom.id 𝕜) (starRingEnd 𝕜) _ _ - (innerSL 𝕜) - -@[simp] -theorem innerSLFlip_apply (x y : E) : innerSLFlip 𝕜 x y = ⟪y, x⟫ := - rfl - -variable (F) in -@[simp] lemma innerSL_real_flip : (innerSL ℝ (E := F)).flip = innerSL ℝ := by - ext v w - exact real_inner_comm _ _ - variable {𝕜} -namespace ContinuousLinearMap - -variable {E' : Type*} [NormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] - --- Note: odd and expensive build behavior is explicitly turned off using `noncomputable` -/-- Given `f : E →L[𝕜] E'`, construct the continuous sesquilinear form `fun x y ↦ ⟪x, A y⟫`, given -as a continuous linear map. -/ -noncomputable def toSesqForm : (E →L[𝕜] E') →L[𝕜] E' →L⋆[𝕜] E →L[𝕜] 𝕜 := - (ContinuousLinearMap.flipₗᵢ' E E' 𝕜 (starRingEnd 𝕜) (RingHom.id 𝕜)).toContinuousLinearEquiv ∘L - ContinuousLinearMap.compSL E E' (E' →L⋆[𝕜] 𝕜) (RingHom.id 𝕜) (RingHom.id 𝕜) (innerSLFlip 𝕜) - -@[simp] -theorem toSesqForm_apply_coe (f : E →L[𝕜] E') (x : E') : toSesqForm f x = (innerSL 𝕜 x).comp f := - rfl - -theorem toSesqForm_apply_norm_le {f : E →L[𝕜] E'} {v : E'} : ‖toSesqForm f v‖ ≤ ‖f‖ * ‖v‖ := by - refine opNorm_le_bound _ (by positivity) fun x ↦ ?_ - have h₁ : ‖f x‖ ≤ ‖f‖ * ‖x‖ := le_opNorm _ _ - have h₂ := @norm_inner_le_norm 𝕜 E' _ _ _ v (f x) - calc - ‖⟪v, f x⟫‖ ≤ ‖v‖ * ‖f x‖ := h₂ - _ ≤ ‖v‖ * (‖f‖ * ‖x‖) := mul_le_mul_of_nonneg_left h₁ (norm_nonneg v) - _ = ‖f‖ * ‖v‖ * ‖x‖ := by ring - -end ContinuousLinearMap - /-- When an inner product space `E` over `𝕜` is considered as a real normed space, its inner product satisfies `IsBoundedBilinearMap`. @@ -1643,8 +1855,16 @@ end Norm section BesselsInequality +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + variable {ι : Type*} (x : E) {v : ι → E} +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ + /-- Bessel's inequality for finite sums. -/ theorem Orthonormal.sum_inner_products_le {s : Finset ι} (hv : Orthonormal 𝕜 v) : ∑ i ∈ s, ‖⟪v i, x⟫‖ ^ 2 ≤ ‖x‖ ^ 2 := by @@ -1684,6 +1904,10 @@ theorem Orthonormal.inner_products_summable (hv : Orthonormal 𝕜 v) : end BesselsInequality +section RCLike + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + /-- A field `𝕜` satisfying `RCLike` is itself a `𝕜`-inner product space. -/ instance RCLike.innerProductSpace : InnerProductSpace 𝕜 𝕜 where inner x y := conj x * y @@ -1696,8 +1920,19 @@ instance RCLike.innerProductSpace : InnerProductSpace 𝕜 𝕜 where theorem RCLike.inner_apply (x y : 𝕜) : ⟪x, y⟫ = conj x * y := rfl -/-! ### Inner product space structure on subspaces -/ +end RCLike + +section Submodule + +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ +/-! ### Inner product space structure on subspaces -/ /-- Induced inner product on a submodule. -/ instance Submodule.innerProductSpace (W : Submodule 𝕜 E) : InnerProductSpace 𝕜 W := @@ -1723,10 +1958,19 @@ theorem orthonormal_span {ι : Type*} {v : ι → E} (hv : Orthonormal 𝕜 v) : hv.codRestrict (Submodule.span 𝕜 (Set.range v)) fun i => Submodule.subset_span (Set.mem_range_self i) +end Submodule + /-! ### Families of mutually-orthogonal subspaces of an inner product space -/ +section OrthogonalFamily_Seminormed -section OrthogonalFamily +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ variable {ι : Type*} (𝕜) @@ -1744,7 +1988,7 @@ product space structure on each of the submodules is important -- for example, w their Hilbert sum (`PiLp V 2`). For example, given an orthonormal set of vectors `v : ι → E`, we have an associated orthogonal family of one-dimensional subspaces of `E`, which it is convenient to be able to discuss using `ι → 𝕜` rather than `Π i : ι, span 𝕜 (v i)`. -/ -def OrthogonalFamily (G : ι → Type*) [∀ i, NormedAddCommGroup (G i)] +def OrthogonalFamily (G : ι → Type*) [∀ i, SeminormedAddCommGroup (G i)] [∀ i, InnerProductSpace 𝕜 (G i)] (V : ∀ i, G i →ₗᵢ[𝕜] E) : Prop := Pairwise fun i j => ∀ v : G i, ∀ w : G j, ⟪V i v, V j w⟫ = 0 @@ -1894,6 +2138,22 @@ theorem OrthogonalFamily.summable_iff_norm_sq_summable [CompleteSpace E] (f : · exact fun i => sq_nonneg _ linarith +end OrthogonalFamily_Seminormed + +section OrthogonalFamily + +variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +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)] + /-- 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 pairwise intersections of elements of the family are 0. -/ @@ -1927,6 +2187,13 @@ section RCLikeToReal variable {G : Type*} variable (𝕜 E) +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ /-- A general inner product implies a real inner product. This is not registered as an instance since it creates problems with the case `𝕜 = ℝ`. -/ @@ -1962,7 +2229,7 @@ theorem real_inner_I_smul_self (x : E) : /-- A complex inner product implies a real inner product. This cannot be an instance since it creates a diamond with `PiLp.innerProductSpace` because `re (sum i, inner (x i) (y i))` and `sum i, re (inner (x i) (y i))` are not defeq. -/ -def InnerProductSpace.complexToReal [NormedAddCommGroup G] [InnerProductSpace ℂ G] : +def InnerProductSpace.complexToReal [SeminormedAddCommGroup G] [InnerProductSpace ℂ G] : InnerProductSpace ℝ G := InnerProductSpace.rclikeToReal ℂ G @@ -1974,13 +2241,23 @@ protected theorem Complex.inner (w z : ℂ) : ⟪w, z⟫_ℝ = (conj w * z).re : /-- The inner product on an inner product space of dimension 2 can be evaluated in terms of a complex-number representation of the space. -/ -theorem inner_map_complex [NormedAddCommGroup G] [InnerProductSpace ℝ G] (f : G ≃ₗᵢ[ℝ] ℂ) +theorem inner_map_complex [SeminormedAddCommGroup G] [InnerProductSpace ℝ G] (f : G ≃ₗᵢ[ℝ] ℂ) (x y : G) : ⟪x, y⟫_ℝ = (conj (f x) * f y).re := by rw [← Complex.inner, f.inner_map_map] +instance : InnerProductSpace ℝ ℂ := InnerProductSpace.complexToReal + end RCLikeToReal section Continuous +variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ + /-! ### Continuity of the inner product -/ @@ -2017,6 +2294,14 @@ end Continuous section ReApplyInnerSelf +variable [SeminormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ + /-- Extract a real bilinear form from an operator `T`, by taking the pairing `fun x ↦ re ⟪T x, x⟫`. -/ def ContinuousLinearMap.reApplyInnerSelf (T : E →L[𝕜] E) (x : E) : ℝ := @@ -2026,6 +2311,18 @@ theorem ContinuousLinearMap.reApplyInnerSelf_apply (T : E →L[𝕜] E) (x : E) T.reApplyInnerSelf x = re ⟪T x, x⟫ := rfl +end ReApplyInnerSelf + +section ReApplyInnerSelf_Seminormed + +variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ + theorem ContinuousLinearMap.reApplyInnerSelf_continuous (T : E →L[𝕜] E) : Continuous T.reApplyInnerSelf := reCLM.continuous.comp <| T.continuous.inner continuous_id @@ -2036,7 +2333,17 @@ theorem ContinuousLinearMap.reApplyInnerSelf_smul (T : E →L[𝕜] E) (x : E) { inner_smul_left, inner_smul_right, ← mul_assoc, mul_conj, ← ofReal_pow, ← smul_re, Algebra.smul_def (‖c‖ ^ 2) ⟪T x, x⟫, algebraMap_eq_ofReal] -end ReApplyInnerSelf +end ReApplyInnerSelf_Seminormed + +section UniformSpace.Completion + +variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] + +local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y + +local notation "IK" => @RCLike.I 𝕜 _ + +local postfix:90 "†" => starRingEnd _ namespace UniformSpace.Completion @@ -2096,3 +2403,5 @@ instance innerProductSpace : InnerProductSpace 𝕜 (Completion E) where fun a b => by simp only [← coe_smul c a, inner_coe, inner_smul_left] end UniformSpace.Completion + +end UniformSpace.Completion diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index f53cb12c2f7c0..2d222bd312cb6 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -405,7 +405,7 @@ noncomputable abbrev NormedAddCommGroup.ofMatrix {M : Matrix n n 𝕜} (hM : M.P /-- A positive definite matrix `M` induces an inner product `⟪x, y⟫ = xᴴMy`. -/ def InnerProductSpace.ofMatrix {M : Matrix n n 𝕜} (hM : M.PosDef) : - @InnerProductSpace 𝕜 (n → 𝕜) _ (NormedAddCommGroup.ofMatrix hM) := + @InnerProductSpace 𝕜 (n → 𝕜) _ (NormedAddCommGroup.ofMatrix hM).toSeminormedAddCommGroup := InnerProductSpace.ofCore _ end Matrix diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index 3cdc5fe127909..84c7d158bcbab 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -7,7 +7,7 @@ Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean : line 1 : ERR_NUM_LIN : 1 Mathlib/Analysis/Asymptotics/Asymptotics.lean : line 1 : ERR_NUM_LIN : 2100 file contains 1955 lines, try to split it up Mathlib/Analysis/Calculus/ContDiff/Basic.lean : line 1 : ERR_NUM_LIN : 2100 file contains 1952 lines, try to split it up Mathlib/Analysis/Calculus/ContDiff/Defs.lean : line 1 : ERR_NUM_LIN : 1700 file contains 1595 lines, try to split it up -Mathlib/Analysis/InnerProductSpace/Basic.lean : line 1 : ERR_NUM_LIN : 2200 file contains 2098 lines, try to split it up +Mathlib/Analysis/InnerProductSpace/Basic.lean : line 1 : ERR_NUM_LIN : 2500 file contains 2433 lines, try to split it up Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean : line 1 : ERR_NUM_LIN : 2100 file contains 1998 lines, try to split it up Mathlib/Computability/TMToPartrec.lean : line 1 : ERR_NUM_LIN : 2200 file contains 2071 lines, try to split it up Mathlib/Computability/TuringMachine.lean : line 1 : ERR_NUM_LIN : 2700 file contains 2561 lines, try to split it up From fdbc62955ed6c5f13b81c9b58c8839932610158d Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 7 Aug 2024 01:50:26 +0000 Subject: [PATCH 071/359] feat: modify `AEEqFun.cast` so that it chooses constant functions (#12949) Co-authored-by: Yury G. Kudryashov --- Mathlib/MeasureTheory/Function/AEEqFun.lean | 45 +++++++++++++++------ 1 file changed, 33 insertions(+), 12 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/AEEqFun.lean b/Mathlib/MeasureTheory/Function/AEEqFun.lean index 78474106e00ff..a95d9cfc9f057 100644 --- a/Mathlib/MeasureTheory/Function/AEEqFun.lean +++ b/Mathlib/MeasureTheory/Function/AEEqFun.lean @@ -110,24 +110,31 @@ variable [TopologicalSpace β] [TopologicalSpace γ] [TopologicalSpace δ] def mk {β : Type*} [TopologicalSpace β] (f : α → β) (hf : AEStronglyMeasurable f μ) : α →ₘ[μ] β := Quotient.mk'' ⟨f, hf⟩ +open scoped Classical in /-- Coercion from a space of equivalence classes of almost everywhere strongly measurable -functions to functions. -/ +functions to functions. We ensure that if `f` has a constant representative, +then we choose that one. -/ @[coe] def cast (f : α →ₘ[μ] β) : α → β := - AEStronglyMeasurable.mk _ (Quotient.out' f : { f : α → β // AEStronglyMeasurable f μ }).2 + if h : ∃ (b : β), f = mk (const α b) aestronglyMeasurable_const then + const α <| Classical.choose h else + AEStronglyMeasurable.mk _ (Quotient.out' f : { f : α → β // AEStronglyMeasurable f μ }).2 /-- A measurable representative of an `AEEqFun` [f] -/ instance instCoeFun : CoeFun (α →ₘ[μ] β) fun _ => α → β := ⟨cast⟩ -protected theorem stronglyMeasurable (f : α →ₘ[μ] β) : StronglyMeasurable f := - AEStronglyMeasurable.stronglyMeasurable_mk _ +protected theorem stronglyMeasurable (f : α →ₘ[μ] β) : StronglyMeasurable f := by + simp only [cast] + split_ifs with h + · exact stronglyMeasurable_const + · apply AEStronglyMeasurable.stronglyMeasurable_mk protected theorem aestronglyMeasurable (f : α →ₘ[μ] β) : AEStronglyMeasurable f μ := f.stronglyMeasurable.aestronglyMeasurable protected theorem measurable [PseudoMetrizableSpace β] [MeasurableSpace β] [BorelSpace β] (f : α →ₘ[μ] β) : Measurable f := - AEStronglyMeasurable.measurable_mk _ + f.stronglyMeasurable.measurable protected theorem aemeasurable [PseudoMetrizableSpace β] [MeasurableSpace β] [BorelSpace β] (f : α →ₘ[μ] β) : AEMeasurable f μ := @@ -144,10 +151,11 @@ theorem mk_eq_mk {f g : α → β} {hf hg} : (mk f hf : α →ₘ[μ] β) = mk g @[simp] theorem mk_coeFn (f : α →ₘ[μ] β) : mk f f.aestronglyMeasurable = f := by + conv_lhs => simp only [cast] + split_ifs with h + · exact Classical.choose_spec h |>.symm conv_rhs => rw [← Quotient.out_eq' f] - set g : { f : α → β // AEStronglyMeasurable f μ } := Quotient.out' f - have : g = ⟨g.1, g.2⟩ := Subtype.eq rfl - rw [this, ← mk, mk_eq_mk] + rw [← mk, mk_eq_mk] exact (AEStronglyMeasurable.ae_eq_mk _).symm @[ext] @@ -155,8 +163,7 @@ theorem ext {f g : α →ₘ[μ] β} (h : f =ᵐ[μ] g) : f = g := by rwa [← f.mk_coeFn, ← g.mk_coeFn, mk_eq_mk] theorem coeFn_mk (f : α → β) (hf) : (mk f hf : α →ₘ[μ] β) =ᵐ[μ] f := by - apply (AEStronglyMeasurable.ae_eq_mk _).symm.trans - exact @Quotient.mk_out' _ (μ.aeEqSetoid β) (⟨f, hf⟩ : { f // AEStronglyMeasurable f μ }) + rw [← mk_eq_mk, mk_coeFn] @[elab_as_elim] theorem induction_on (f : α →ₘ[μ] β) {p : (α →ₘ[μ] β) → Prop} (H : ∀ f hf, p (mk f hf)) : p f := @@ -548,11 +555,21 @@ variable (α) /-- The equivalence class of a constant function: `[fun _ : α => b]`, based on the equivalence relation of being almost everywhere equal -/ def const (b : β) : α →ₘ[μ] β := - mk (fun _ : α => b) aestronglyMeasurable_const + mk (fun _ : α ↦ b) aestronglyMeasurable_const theorem coeFn_const (b : β) : (const α b : α →ₘ[μ] β) =ᵐ[μ] Function.const α b := coeFn_mk _ _ +/-- If the measure is nonzero, we can strengthen `coeFn_const` to get an equality. -/ +@[simp] +theorem coeFn_const_eq [NeZero μ] (b : β) (x : α) : (const α b : α →ₘ[μ] β) x = b := by + simp only [cast] + split_ifs with h; swap; exact h.elim ⟨b, rfl⟩ + have := Classical.choose_spec h + set b' := Classical.choose h + simp_rw [const, mk_eq_mk, EventuallyEq, ← const_def, eventually_const] at this + rw [Function.const, this] + variable {α} instance instInhabited [Inhabited β] : Inhabited (α →ₘ[μ] β) := @@ -568,7 +585,11 @@ theorem one_def [One β] : (1 : α →ₘ[μ] β) = mk (fun _ : α => 1) aestron @[to_additive] theorem coeFn_one [One β] : ⇑(1 : α →ₘ[μ] β) =ᵐ[μ] 1 := - coeFn_const _ _ + coeFn_const .. + +@[to_additive (attr := simp)] +theorem coeFn_one_eq [NeZero μ] [One β] {x : α} : (1 : α →ₘ[μ] β) x = 1 := + coeFn_const_eq .. @[to_additive (attr := simp)] theorem one_toGerm [One β] : (1 : α →ₘ[μ] β).toGerm = 1 := From e2aa79d02456af2f925b7637259efc247a14e166 Mon Sep 17 00:00:00 2001 From: james18lpc Date: Wed, 7 Aug 2024 02:12:23 +0000 Subject: [PATCH 072/359] feat: Measurability of functions w.r.t supremums of sigma-algebras (#15326) From GibbsMeasure --- Mathlib/MeasureTheory/MeasurableSpace/Basic.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index f79c457d0f780..7d3dad830959e 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -190,6 +190,21 @@ theorem Measurable.mono {ma ma' : MeasurableSpace α} {mb mb' : MeasurableSpace (hf : @Measurable α β ma mb f) (ha : ma ≤ ma') (hb : mb' ≤ mb) : @Measurable α β ma' mb' f := fun _t ht => ha _ <| hf <| hb _ ht +lemma Measurable.iSup' {mα : ι → MeasurableSpace α} {_ : MeasurableSpace β} {f : α → β} (i₀ : ι) + (h : Measurable[mα i₀] f) : + Measurable[⨆ i, mα i] f := + h.mono (le_iSup mα i₀) le_rfl + +lemma Measurable.sup_of_left {mα mα' : MeasurableSpace α} {_ : MeasurableSpace β} {f : α → β} + (h : Measurable[mα] f) : + Measurable[mα ⊔ mα'] f := + h.mono le_sup_left le_rfl + +lemma Measurable.sup_of_right {mα mα' : MeasurableSpace α} {_ : MeasurableSpace β} {f : α → β} + (h : Measurable[mα'] f) : + Measurable[mα ⊔ mα'] f := + h.mono le_sup_right le_rfl + theorem measurable_id'' {m mα : MeasurableSpace α} (hm : m ≤ mα) : @Measurable α α mα m id := measurable_id.mono le_rfl hm From c7dfc55407acc6697e2f8d8ece42887d9634a000 Mon Sep 17 00:00:00 2001 From: grunweg Date: Wed, 7 Aug 2024 02:12:24 +0000 Subject: [PATCH 073/359] chore(Geometry/Manifold): split variable statements (#15439) Make sure to perform binder annotation updates and declaring new variables in separate variable commands. This is currently recommended because of leanprover/lean4#2789: mixing these can yield confusing errors. Found by a linter in #15400. --- Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean | 4 ++-- Mathlib/Geometry/Manifold/BumpFunction.lean | 10 +++++----- Mathlib/Geometry/Manifold/ChartedSpace.lean | 5 ++--- Mathlib/Geometry/Manifold/ContMDiff/Basic.lean | 2 +- Mathlib/Geometry/Manifold/PartitionOfUnity.lean | 3 ++- Mathlib/Geometry/Manifold/VectorBundle/Basic.lean | 9 +++++---- Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean | 6 ++++-- Mathlib/Topology/FiberBundle/Basic.lean | 12 ++++++++---- Mathlib/Topology/FiberBundle/Trivialization.lean | 4 ++-- 9 files changed, 31 insertions(+), 24 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean index 9e4a3803e6535..51e531175bf03 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Prod.lean @@ -71,8 +71,8 @@ end OpNorm section Prod -variable (M₁ M₂ M₃ M₄ : Type*) (𝕜) -variable +variable (𝕜) +variable (M₁ M₂ M₃ M₄ : Type*) [SeminormedAddCommGroup M₁] [NormedSpace 𝕜 M₁] [SeminormedAddCommGroup M₂] [NormedSpace 𝕜 M₂] [SeminormedAddCommGroup M₃] [NormedSpace 𝕜 M₃] diff --git a/Mathlib/Geometry/Manifold/BumpFunction.lean b/Mathlib/Geometry/Manifold/BumpFunction.lean index 753de631d7746..0cc7636a1ddc0 100644 --- a/Mathlib/Geometry/Manifold/BumpFunction.lean +++ b/Mathlib/Geometry/Manifold/BumpFunction.lean @@ -31,7 +31,7 @@ manifold, smooth bump function universe uE uF uH uM variable {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] - {H : Type uH} [TopologicalSpace H] (I : ModelWithCorners ℝ E H) {M : Type uM} [TopologicalSpace M] + {H : Type uH} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} {M : Type uM} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] open Function Filter FiniteDimensional Set Metric @@ -46,6 +46,7 @@ noncomputable section In this section we define a structure for a bundled smooth bump function and prove its properties. -/ +variable (I) in /-- Given a smooth manifold modelled on a finite dimensional space `E`, `f : SmoothBumpFunction I M` is a smooth function on `M` such that in the extended chart `e` at `f.c`: @@ -63,7 +64,7 @@ structure SmoothBumpFunction (c : M) extends ContDiffBump (extChartAt I c c) whe namespace SmoothBumpFunction -variable {c : M} (f : SmoothBumpFunction I c) {x : M} {I} +variable {c : M} (f : SmoothBumpFunction I c) {x : M} /-- The function defined by `f : SmoothBumpFunction c`. Use automatic coercion to function instead. -/ @@ -243,8 +244,9 @@ protected theorem hasCompactSupport : HasCompactSupport f := f.isCompact_symm_image_closedBall.of_isClosed_subset isClosed_closure f.tsupport_subset_symm_image_closedBall -variable (I c) +variable (I) +variable (c) in /-- The closures of supports of smooth bump functions centered at `c` form a basis of `𝓝 c`. In other words, each of these closures is a neighborhood of `c` and each neighborhood of `c` includes `tsupport f` for some `f : SmoothBumpFunction I c`. -/ @@ -258,8 +260,6 @@ theorem nhds_basis_tsupport : exact this.to_hasBasis' (fun f _ => ⟨f, trivial, f.tsupport_subset_symm_image_closedBall⟩) fun f _ => f.tsupport_mem_nhds -variable {c} - /-- Given `s ∈ 𝓝 c`, the supports of smooth bump functions `f : SmoothBumpFunction I c` such that `tsupport f ⊆ s` form a basis of `𝓝 c`. In other words, each of these supports is a neighborhood of `c` and each neighborhood of `c` includes `support f` for some diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index 82e72c61ddd29..c1ce4ba8838db 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -968,16 +968,15 @@ theorem StructureGroupoid.trans_restricted {e e' : PartialHomeomorph M H} {G : S section MaximalAtlas -variable (M) (G : StructureGroupoid H) +variable (G : StructureGroupoid H) +variable (M) in /-- Given a charted space admitting a structure groupoid, the maximal atlas associated to this structure groupoid is the set of all charts that are compatible with the atlas, i.e., such that changing coordinates with an atlas member gives an element of the groupoid. -/ def StructureGroupoid.maximalAtlas : Set (PartialHomeomorph M H) := { e | ∀ e' ∈ atlas H M, e.symm ≫ₕ e' ∈ G ∧ e'.symm ≫ₕ e ∈ G } -variable {M} - /-- The elements of the atlas belong to the maximal atlas for any structure groupoid. -/ theorem StructureGroupoid.subset_maximalAtlas [HasGroupoid M G] : atlas H M ⊆ G.maximalAtlas M := fun _ he _ he' ↦ ⟨G.compatible he he', G.compatible he' he⟩ diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index db57f53db85c1..d6a88bd9a46f7 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -354,7 +354,7 @@ end Inclusion section variable (I) - [Nonempty M] {e : M → H} (h : OpenEmbedding e) +variable [Nonempty M] {e : M → H} (h : OpenEmbedding e) [Nonempty M'] {e' : M' → H'} (h' : OpenEmbedding e') {n : WithTop ℕ} diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index 16809926abe3c..3e1374a8c5049 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -278,7 +278,8 @@ type if for each `i` the closure of the support of `f i` is a subset of `U i`. - def IsSubordinate (f : SmoothPartitionOfUnity ι I M s) (U : ι → Set M) := ∀ i, tsupport (f i) ⊆ U i -variable {f} {U : ι → Set M} +variable {f} +variable {U : ι → Set M} @[simp] theorem isSubordinate_toPartitionOfUnity : diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean index 0053a9e9c80ca..3be66afa23f67 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean @@ -403,8 +403,9 @@ protected theorem Smooth.coordChange (hf : Smooth IM IB f) Smooth IM 𝓘(𝕜, F) (fun y ↦ e.coordChange e' (f y) (g y)) := fun x ↦ (hf x).coordChange (hg x) (he x) (he' x) -variable (IB e e') +variable (e e') +variable (IB) in theorem Trivialization.contMDiffOn_symm_trans : ContMDiffOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) n (e.toPartialHomeomorph.symm ≫ₕ e'.toPartialHomeomorph) (e.target ∩ e'.target) := by @@ -420,7 +421,7 @@ theorem Trivialization.contMDiffOn_symm_trans : simp_all only [Trivialization.mem_target, mfld_simps] exact (e'.coe_fst' this).trans (e.proj_symm_apply hb.1) -variable {IB e e'} +variable {e e'} theorem ContMDiffWithinAt.change_section_trivialization {f : M → TotalSpace F E} (hp : ContMDiffWithinAt IM IB n (π F E ∘ f) s x) @@ -551,8 +552,8 @@ end namespace VectorBundleCore -variable {ι : Type*} {F} -variable (Z : VectorBundleCore 𝕜 B F ι) +variable {F} +variable {ι : Type*} (Z : VectorBundleCore 𝕜 B F ι) /-- Mixin for a `VectorBundleCore` stating smoothness (of transition functions). -/ class IsSmooth (IB : ModelWithCorners 𝕜 EB HB) : Prop where diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index 792a16535a08a..fed71bc4299df 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -196,7 +196,8 @@ does not pick wrong instances. In this section, we record the right instances fo them, noting in particular that the tangent bundle is a smooth manifold. -/ section -variable {M} (x : M) +variable {M} +variable (x : M) instance : Module 𝕜 (TangentSpace I x) := inferInstanceAs (Module 𝕜 E) @@ -408,7 +409,8 @@ theorem tangentBundleModelSpaceHomeomorph_coe_symm : section inTangentCoordinates -variable (I') {M H} {N : Type*} +variable (I') {M H} +variable {N : Type*} /-- The map `in_coordinates` for the tangent bundle is trivial on the model spaces -/ theorem inCoordinates_tangent_bundle_core_model_space (x₀ x : H) (y₀ y : H') (ϕ : E →L[𝕜] E') : diff --git a/Mathlib/Topology/FiberBundle/Basic.lean b/Mathlib/Topology/FiberBundle/Basic.lean index 5a646a2cf74a9..4606e9604e100 100644 --- a/Mathlib/Topology/FiberBundle/Basic.lean +++ b/Mathlib/Topology/FiberBundle/Basic.lean @@ -171,7 +171,8 @@ open TopologicalSpace Filter Set Bundle Topology section FiberBundle -variable (F) [TopologicalSpace B] [TopologicalSpace F] (E : B → Type*) +variable (F) +variable [TopologicalSpace B] [TopologicalSpace F] (E : B → Type*) [TopologicalSpace (TotalSpace F E)] [∀ b, TopologicalSpace (E b)] /-- A (topological) fiber bundle with fiber `F` over a base `B` is a space projecting on `B` @@ -207,7 +208,8 @@ end FiberBundle export FiberBundle (totalSpaceMk_inducing trivializationAtlas trivializationAt mem_baseSet_trivializationAt trivialization_mem_atlas) -variable {F E} +variable {F} +variable {E} /-- Given a type `E` equipped with a fiber bundle structure, this is a `Prop` typeclass for trivializations of `E`, expressing that a trivialization is in the designated atlas for the @@ -297,7 +299,8 @@ theorem continuousAt_totalSpace (f : X → TotalSpace F E) {x₀ : X} : end FiberBundle -variable (F E) +variable (F) +variable (E) /-- If `E` is a fiber bundle over a conditionally complete linear order, then it is trivial over any closed interval. -/ @@ -700,7 +703,8 @@ end FiberBundleCore /-! ### Prebundle construction for constructing fiber bundles -/ -variable (F) (E : B → Type*) [TopologicalSpace B] [TopologicalSpace F] +variable (F) +variable (E : B → Type*) [TopologicalSpace B] [TopologicalSpace F] [∀ x, TopologicalSpace (E x)] /-- This structure permits to define a fiber bundle when trivializations are given as local diff --git a/Mathlib/Topology/FiberBundle/Trivialization.lean b/Mathlib/Topology/FiberBundle/Trivialization.lean index cd79bc62c044f..d5a5d504ff3af 100644 --- a/Mathlib/Topology/FiberBundle/Trivialization.lean +++ b/Mathlib/Topology/FiberBundle/Trivialization.lean @@ -50,8 +50,8 @@ type of linear trivializations is not even particularly well-behaved. open TopologicalSpace Filter Set Bundle Function open scoped Topology -variable {ι : Type*} {B : Type*} {F : Type*} {E : B → Type*} -variable (F) {Z : Type*} [TopologicalSpace B] [TopologicalSpace F] {proj : Z → B} +variable {ι : Type*} {B : Type*} (F : Type*) {E : B → Type*} +variable {Z : Type*} [TopologicalSpace B] [TopologicalSpace F] {proj : Z → B} /-- This structure contains the information left for a local trivialization (which is implemented below as `Trivialization F proj`) if the total space has not been given a topology, but we From 4bb053f6b63c15e0a514f58f383baeefabe21135 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 7 Aug 2024 03:16:57 +0000 Subject: [PATCH 074/359] chore: backports for leanprover/lean4#4814 (part 30) (#15574) --- Mathlib/Order/Filter/AtTopBot.lean | 65 ++++++++++++++++++------------ 1 file changed, 39 insertions(+), 26 deletions(-) diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index b0c8240c7c620..0f53605218201 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -216,18 +216,32 @@ theorem tendsto_atBot_pure [PartialOrder α] [OrderBot α] (f : α → β) : @tendsto_atTop_pure αᵒᵈ _ _ _ _ section IsDirected -variable [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] {p : α → Prop} +variable [Preorder α] [IsDirected α (· ≤ ·)] {p : α → Prop} -theorem hasAntitoneBasis_atTop : (@atTop α _).HasAntitoneBasis Ici := +theorem hasAntitoneBasis_atTop [Nonempty α] : (@atTop α _).HasAntitoneBasis Ici := .iInf_principal fun _ _ ↦ Ici_subset_Ici.2 -theorem atTop_basis : (@atTop α _).HasBasis (fun _ => True) Ici := hasAntitoneBasis_atTop.1 +theorem atTop_basis [Nonempty α] : (@atTop α _).HasBasis (fun _ => True) Ici := + hasAntitoneBasis_atTop.1 theorem atTop_eq_generate_Ici : atTop = generate (range (Ici (α := α))) := by rcases isEmpty_or_nonempty α with hα|hα · simp only [eq_iff_true_of_subsingleton] · simp [(atTop_basis (α := α)).eq_generate, range] +lemma atTop_basis_Ioi [Nonempty α] [NoMaxOrder α] : (@atTop α _).HasBasis (fun _ => True) Ioi := + atTop_basis.to_hasBasis (fun a ha => ⟨a, ha, Ioi_subset_Ici_self⟩) fun a ha => + (exists_gt a).imp fun _b hb => ⟨ha, Ici_subset_Ioi.2 hb⟩ + +lemma atTop_basis_Ioi' [NoMaxOrder α] (a : α) : atTop.HasBasis (a < ·) Ioi := by + have : Nonempty α := ⟨a⟩ + refine atTop_basis_Ioi.to_hasBasis (fun b _ ↦ ?_) fun b _ ↦ ⟨b, trivial, Subset.rfl⟩ + obtain ⟨c, hac, hbc⟩ := exists_ge_ge a b + obtain ⟨d, hcd⟩ := exists_gt c + exact ⟨d, hac.trans_lt hcd, Ioi_subset_Ioi (hbc.trans hcd.le)⟩ + +variable [Nonempty α] + theorem atTop_basis' (a : α) : (@atTop α _).HasBasis (fun x => a ≤ x) Ici := by refine atTop_basis.to_hasBasis (fun b _ ↦ ?_) fun b _ ↦ ⟨b, trivial, Subset.rfl⟩ obtain ⟨c, hac, hbc⟩ := exists_ge_ge a b @@ -256,17 +270,6 @@ lemma exists_eventually_atTop {r : α → β → Prop} : theorem map_atTop_eq {f : α → β} : atTop.map f = ⨅ a, 𝓟 (f '' { a' | a ≤ a' }) := (atTop_basis.map f).eq_iInf -lemma atTop_basis_Ioi [NoMaxOrder α] : (@atTop α _).HasBasis (fun _ => True) Ioi := - atTop_basis.to_hasBasis (fun a ha => ⟨a, ha, Ioi_subset_Ici_self⟩) fun a ha => - (exists_gt a).imp fun _b hb => ⟨ha, Ici_subset_Ioi.2 hb⟩ - -lemma atTop_basis_Ioi' [NoMaxOrder α] (a : α) : atTop.HasBasis (a < ·) Ioi := by - have : Nonempty α := ⟨a⟩ - refine atTop_basis_Ioi.to_hasBasis (fun b _ ↦ ?_) fun b _ ↦ ⟨b, trivial, Subset.rfl⟩ - obtain ⟨c, hac, hbc⟩ := exists_ge_ge a b - obtain ⟨d, hcd⟩ := exists_gt c - exact ⟨d, hac.trans_lt hcd, Ioi_subset_Ioi (hbc.trans hcd.le)⟩ - theorem frequently_atTop' [NoMaxOrder α] : (∃ᶠ x in atTop, p x) ↔ ∀ a, ∃ b > a, p b := atTop_basis_Ioi.frequently_iff.trans <| by simp @@ -277,7 +280,15 @@ lemma atTop_countable_basis [Countable α] : end IsDirected section IsCodirected -variable [Nonempty α] [Preorder α] [IsDirected α (· ≥ ·)] {p : α → Prop} +variable [Preorder α] [IsDirected α (· ≥ ·)] {p : α → Prop} + +lemma atBot_basis_Iio [Nonempty α] [NoMinOrder α] : (@atBot α _).HasBasis (fun _ => True) Iio := + atTop_basis_Ioi (α := αᵒᵈ) + +lemma atBot_basis_Iio' [NoMinOrder α] (a : α) : atBot.HasBasis (· < a) Iio := + atTop_basis_Ioi' (α := αᵒᵈ) a + +variable [Nonempty α] lemma atBot_basis : (@atBot α _).HasBasis (fun _ => True) Iic := atTop_basis (α := αᵒᵈ) @@ -304,12 +315,6 @@ lemma exists_eventually_atBot {r : α → β → Prop} : theorem map_atBot_eq {f : α → β} : atBot.map f = ⨅ a, 𝓟 (f '' { a' | a' ≤ a }) := map_atTop_eq (α := αᵒᵈ) -lemma atBot_basis_Iio [NoMinOrder α] : (@atBot α _).HasBasis (fun _ => True) Iio := - atTop_basis_Ioi (α := αᵒᵈ) - -lemma atBot_basis_Iio' [NoMinOrder α] (a : α) : atBot.HasBasis (· < a) Iio := - atTop_basis_Ioi' (α := αᵒᵈ) a - theorem frequently_atBot' [NoMinOrder α] : (∃ᶠ x in atBot, p x) ↔ ∀ a, ∃ b < a, p b := frequently_atTop' (α := αᵒᵈ) @@ -459,11 +464,14 @@ theorem Eventually.atTop_of_arithmetic {p : ℕ → Prop} {n : ℕ} (hn : n ≠ exact (Finset.le_sup (f := (n * N ·)) (Finset.mem_range.2 hlt)).trans hb section IsDirected -variable [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] [Preorder β] {F : Filter β} {u : α → β} +variable [Preorder α] [IsDirected α (· ≤ ·)] {F : Filter β} {u : α → β} -theorem inf_map_atTop_neBot_iff : NeBot (F ⊓ map u atTop) ↔ ∀ U ∈ F, ∀ N, ∃ n ≥ N, u n ∈ U := by +theorem inf_map_atTop_neBot_iff [Nonempty α] : + NeBot (F ⊓ map u atTop) ↔ ∀ U ∈ F, ∀ N, ∃ n ≥ N, u n ∈ U := by simp_rw [inf_neBot_iff_frequently_left, frequently_map, frequently_atTop]; rfl +variable [Preorder β] + lemma exists_le_of_tendsto_atTop (h : Tendsto u atTop atTop) (a : α) (b : β) : ∃ a' ≥ a, b ≤ u a' := by have : Nonempty α := ⟨a⟩ @@ -1185,7 +1193,7 @@ theorem tendsto_atBot_atTop_of_antitone [Preorder α] [Preorder β] {f : α → @tendsto_atBot_atBot_of_monotone _ βᵒᵈ _ _ _ hf h section IsDirected -variable [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] [Preorder β] {f : α → β} {l : Filter β} +variable [Nonempty α] [Preorder α] [IsDirected α (· ≤ ·)] {f : α → β} {l : Filter β} theorem tendsto_atTop' : Tendsto f atTop l ↔ ∀ s ∈ l, ∃ a, ∀ b ≥ a, f b ∈ s := by simp only [tendsto_def, mem_atTop_sets, mem_preimage] @@ -1193,6 +1201,8 @@ theorem tendsto_atTop' : Tendsto f atTop l ↔ ∀ s ∈ l, ∃ a, ∀ b ≥ a, theorem tendsto_atTop_principal {s : Set β} : Tendsto f atTop (𝓟 s) ↔ ∃ N, ∀ n ≥ N, f n ∈ s := by simp_rw [tendsto_iff_comap, comap_principal, le_principal_iff, mem_atTop_sets, mem_preimage] +variable [Preorder β] + /-- A function `f` grows to `+∞` independent of an order-preserving embedding `e`. -/ theorem tendsto_atTop_atTop : Tendsto f atTop atTop ↔ ∀ b : β, ∃ i : α, ∀ a : α, i ≤ a → b ≤ f a := tendsto_iInf.trans <| forall_congr' fun _ => tendsto_atTop_principal @@ -1212,7 +1222,7 @@ theorem tendsto_atTop_atBot_iff_of_antitone (hf : Antitone f) : end IsDirected section IsCodirected -variable [Nonempty α] [Preorder α] [IsDirected α (· ≥ ·)] [Preorder β] {f : α → β} {l : Filter β} +variable [Nonempty α] [Preorder α] [IsDirected α (· ≥ ·)] {f : α → β} {l : Filter β} theorem tendsto_atBot' : Tendsto f atBot l ↔ ∀ s ∈ l, ∃ a, ∀ b ≤ a, f b ∈ s := tendsto_atTop' (α := αᵒᵈ) @@ -1220,6 +1230,8 @@ theorem tendsto_atBot' : Tendsto f atBot l ↔ ∀ s ∈ l, ∃ a, ∀ b ≤ a, theorem tendsto_atBot_principal {s : Set β} : Tendsto f atBot (𝓟 s) ↔ ∃ N, ∀ n ≤ N, f n ∈ s := tendsto_atTop_principal (α := αᵒᵈ) (β := βᵒᵈ) +variable [Preorder β] + theorem tendsto_atBot_atTop : Tendsto f atBot atTop ↔ ∀ b : β, ∃ i : α, ∀ a : α, a ≤ i → b ≤ f a := tendsto_atTop_atTop (α := αᵒᵈ) @@ -1392,7 +1404,8 @@ theorem map_atTop_eq_of_gc [SemilatticeSup α] [SemilatticeSup β] {f : α → refine le_antisymm (hf.tendsto_atTop_atTop fun b => ⟨g (b ⊔ b'), le_sup_left.trans <| hgi _ le_sup_right⟩) ?_ - rw [@map_atTop_eq _ _ ⟨g b'⟩] + have : Nonempty α := ⟨g b'⟩ + rw [map_atTop_eq] refine le_iInf fun a => iInf_le_of_le (f a ⊔ b') <| principal_mono.2 fun b hb => ?_ rw [mem_Ici, sup_le_iff] at hb exact ⟨g b, (gc _ _ hb.2).1 hb.1, le_antisymm ((gc _ _ hb.2).2 le_rfl) (hgi _ hb.2)⟩ From 53f5571e327fb37c1da77b0ddd80289f1bde8890 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 7 Aug 2024 05:28:34 +0000 Subject: [PATCH 075/359] chore: backports for leanprover/lean4#4814 (part 17) (#15429) see https://github.com/leanprover-community/mathlib4/pull/15245 Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib/Algebra/CharP/ExpChar.lean | 32 +++---- Mathlib/Algebra/Polynomial/AlgebraMap.lean | 2 +- Mathlib/Algebra/Polynomial/Degree/Lemmas.lean | 28 +++--- Mathlib/Algebra/Polynomial/Module/Basic.lean | 4 +- Mathlib/Algebra/Polynomial/Smeval.lean | 72 ++++++++------ Mathlib/Algebra/Quaternion.lean | 65 ++++++++++--- Mathlib/GroupTheory/HNNExtension.lean | 2 +- Mathlib/GroupTheory/PushoutI.lean | 94 +++++++++---------- Mathlib/RingTheory/Adjoin/Basic.lean | 2 +- .../RingTheory/Derivation/ToSquareZero.lean | 12 +-- Mathlib/RingTheory/Idempotents.lean | 12 +-- .../MvPolynomial/NewtonIdentities.lean | 50 +++++----- Mathlib/RingTheory/PowerSeries/WellKnown.lean | 4 +- Mathlib/RingTheory/PrincipalIdealDomain.lean | 50 +++++----- .../RingTheory/UniqueFactorizationDomain.lean | 52 +++++----- .../RingTheory/WittVector/WittPolynomial.lean | 2 +- Mathlib/Tactic/ComputeDegree.lean | 9 +- 17 files changed, 277 insertions(+), 215 deletions(-) diff --git a/Mathlib/Algebra/CharP/ExpChar.lean b/Mathlib/Algebra/CharP/ExpChar.lean index f25938834d0bd..ef45c8099756c 100644 --- a/Mathlib/Algebra/CharP/ExpChar.lean +++ b/Mathlib/Algebra/CharP/ExpChar.lean @@ -91,6 +91,22 @@ theorem char_eq_expChar_iff (p q : ℕ) [hp : CharP R p] [hq : ExpChar R q] : p decide · exact ⟨fun hpq => hpq.symm ▸ hq_prime, fun _ => CharP.eq R hp hq_hchar⟩ +/-- The exponential characteristic is a prime number or one. +See also `CharP.char_is_prime_or_zero`. -/ +theorem expChar_is_prime_or_one (q : ℕ) [hq : ExpChar R q] : Nat.Prime q ∨ q = 1 := by + cases hq with + | zero => exact .inr rfl + | prime hp => exact .inl hp + +/-- The exponential characteristic is positive. -/ +theorem expChar_pos (q : ℕ) [ExpChar R q] : 0 < q := by + rcases expChar_is_prime_or_one R q with h | rfl + exacts [Nat.Prime.pos h, Nat.one_pos] + +/-- Any power of the exponential characteristic is positive. -/ +theorem expChar_pow_pos (q : ℕ) [ExpChar R q] (n : ℕ) : 0 < q ^ n := + Nat.pos_pow_of_pos n (expChar_pos R q) + section Nontrivial variable [Nontrivial R] @@ -126,22 +142,6 @@ theorem char_prime_of_ne_zero {p : ℕ} [hp : CharP R p] (p_ne_zero : p ≠ 0) : · exact h · contradiction -/-- The exponential characteristic is a prime number or one. -See also `CharP.char_is_prime_or_zero`. -/ -theorem expChar_is_prime_or_one (q : ℕ) [hq : ExpChar R q] : Nat.Prime q ∨ q = 1 := by - cases hq with - | zero => exact .inr rfl - | prime hp => exact .inl hp - -/-- The exponential characteristic is positive. -/ -theorem expChar_pos (q : ℕ) [ExpChar R q] : 0 < q := by - rcases expChar_is_prime_or_one R q with h | rfl - exacts [Nat.Prime.pos h, Nat.one_pos] - -/-- Any power of the exponential characteristic is positive. -/ -theorem expChar_pow_pos (q : ℕ) [ExpChar R q] (n : ℕ) : 0 < q ^ n := - Nat.pos_pow_of_pos n (expChar_pos R q) - end NoZeroDivisors end Nontrivial diff --git a/Mathlib/Algebra/Polynomial/AlgebraMap.lean b/Mathlib/Algebra/Polynomial/AlgebraMap.lean index 2c532eb9256f4..488fbdbe25fad 100644 --- a/Mathlib/Algebra/Polynomial/AlgebraMap.lean +++ b/Mathlib/Algebra/Polynomial/AlgebraMap.lean @@ -217,7 +217,7 @@ end CommSemiring section aeval variable [CommSemiring R] [Semiring A] [CommSemiring A'] [Semiring B] -variable [Algebra R A] [Algebra R A'] [Algebra R B] +variable [Algebra R A] [Algebra R B] variable {p q : R[X]} (x : A) /-- Given a valuation `x` of the variable in an `R`-algebra `A`, `aeval R A x` is diff --git a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean index bdbf349cc8a11..f49898f66a5bd 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean @@ -305,19 +305,7 @@ end Ring section NoZeroDivisors -variable [Semiring R] [NoZeroDivisors R] {p q : R[X]} {a : R} - -theorem degree_mul_C (a0 : a ≠ 0) : (p * C a).degree = p.degree := by - rw [degree_mul, degree_C a0, add_zero] - -theorem degree_C_mul (a0 : a ≠ 0) : (C a * p).degree = p.degree := by - rw [degree_mul, degree_C a0, zero_add] - -theorem natDegree_mul_C (a0 : a ≠ 0) : (p * C a).natDegree = p.natDegree := by - simp only [natDegree, degree_mul_C a0] - -theorem natDegree_C_mul (a0 : a ≠ 0) : (C a * p).natDegree = p.natDegree := by - simp only [natDegree, degree_C_mul a0] +variable [Semiring R] {p q : R[X]} {a : R} @[simp] lemma nextCoeff_C_mul_X_add_C (ha : a ≠ 0) (c : R) : nextCoeff (C a * X + C c) = c := by @@ -337,6 +325,20 @@ lemma natDegree_eq_one : p.natDegree = 1 ↔ ∃ a ≠ 0, ∃ b, C a * X + C b = · rintro ⟨a, ha, b, rfl⟩ simp [ha] +variable [NoZeroDivisors R] + +theorem degree_mul_C (a0 : a ≠ 0) : (p * C a).degree = p.degree := by + rw [degree_mul, degree_C a0, add_zero] + +theorem degree_C_mul (a0 : a ≠ 0) : (C a * p).degree = p.degree := by + rw [degree_mul, degree_C a0, zero_add] + +theorem natDegree_mul_C (a0 : a ≠ 0) : (p * C a).natDegree = p.natDegree := by + simp only [natDegree, degree_mul_C a0] + +theorem natDegree_C_mul (a0 : a ≠ 0) : (C a * p).natDegree = p.natDegree := by + simp only [natDegree, degree_C_mul a0] + theorem natDegree_comp : natDegree (p.comp q) = natDegree p * natDegree q := by by_cases q0 : q.natDegree = 0 · rw [degree_le_zero_iff.mp (natDegree_eq_zero_iff_degree_le_zero.mp q0), comp_C, natDegree_C, diff --git a/Mathlib/Algebra/Polynomial/Module/Basic.lean b/Mathlib/Algebra/Polynomial/Module/Basic.lean index a549dcebffb49..4beded12152cd 100644 --- a/Mathlib/Algebra/Polynomial/Module/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Module/Basic.lean @@ -221,7 +221,7 @@ lemma equivPolynomial_single {S : Type*} [CommRing S] [Algebra R S] (n : ℕ) (x equivPolynomial (single R n x) = monomial n x := rfl variable (R' : Type*) {M' : Type*} [CommRing R'] [AddCommGroup M'] [Module R' M'] -variable [Algebra R R'] [Module R M'] [IsScalarTower R R' M'] +variable [Module R M'] /-- The image of a polynomial under a linear map. -/ noncomputable def map (f : M →ₗ[R] M') : PolynomialModule R M →ₗ[R] PolynomialModule R' M' := @@ -231,6 +231,8 @@ noncomputable def map (f : M →ₗ[R] M') : PolynomialModule R M →ₗ[R] Poly theorem map_single (f : M →ₗ[R] M') (i : ℕ) (m : M) : map R' f (single R i m) = single R' i (f m) := Finsupp.mapRange_single (hf := f.map_zero) +variable [Algebra R R'] [IsScalarTower R R' M'] + theorem map_smul (f : M →ₗ[R] M') (p : R[X]) (q : PolynomialModule R M) : map R' f (p • q) = p.map (algebraMap R R') • map R' f q := by apply induction_linear q diff --git a/Mathlib/Algebra/Polynomial/Smeval.lean b/Mathlib/Algebra/Polynomial/Smeval.lean index f37be4cc5c4c5..4f3e4012ca414 100644 --- a/Mathlib/Algebra/Polynomial/Smeval.lean +++ b/Mathlib/Algebra/Polynomial/Smeval.lean @@ -181,8 +181,16 @@ octonions), we replace the `[Semiring S]` with `[NonAssocSemiring S] [Pow S ℕ] -/ variable (R : Type*) [Semiring R] {p : R[X]} (r : R) (p q : R[X]) {S : Type*} - [NonAssocSemiring S] [Module R S] [IsScalarTower R S S] [SMulCommClass R S S] [Pow S ℕ] - [NatPowAssoc S] (x : S) + [NonAssocSemiring S] [Module R S] [Pow S ℕ] (x : S) + +theorem smeval_C_mul : (C r * p).smeval x = r • p.smeval x := by + induction p using Polynomial.induction_on' with + | h_add p q ph qh => + simp only [mul_add, smeval_add, ph, qh, smul_add] + | h_monomial n b => + simp only [C_mul_monomial, smeval_monomial, mul_smul] + +variable [NatPowAssoc S] theorem smeval_at_natCast (q : ℕ[X]) : ∀(n : ℕ), q.smeval (n : S) = q.smeval n := by induction q using Polynomial.induction_on' with @@ -206,13 +214,8 @@ theorem smeval_at_zero : p.smeval (0 : S) = (p.coeff 0) • (1 : S) := by | succ n => rw [coeff_monomial_succ, smeval_monomial, npow_add, npow_one, mul_zero, zero_smul, smul_zero] -theorem smeval_mul_X : (p * X).smeval x = p.smeval x * x := by - induction p using Polynomial.induction_on' with - | h_add p q ph qh => - simp only [add_mul, smeval_add, ph, qh] - | h_monomial n a => - simp only [← monomial_one_one_eq_X, monomial_mul_monomial, smeval_monomial, mul_one, pow_succ', - mul_assoc, npow_add, smul_mul_assoc, npow_one] +section +variable [SMulCommClass R S S] theorem smeval_X_mul : (X * p).smeval x = x * p.smeval x := by induction p using Polynomial.induction_on' with @@ -222,21 +225,6 @@ theorem smeval_X_mul : (X * p).smeval x = x * p.smeval x := by rw [← monomial_one_one_eq_X, monomial_mul_monomial, smeval_monomial, one_mul, npow_add, npow_one, ← mul_smul_comm, smeval_monomial] -theorem smeval_assoc_X_pow (m n : ℕ) : - p.smeval x * x ^ m * x ^ n = p.smeval x * (x ^ m * x ^ n) := by - induction p using Polynomial.induction_on' with - | h_add p q ph qh => - simp only [smeval_add, ph, qh, add_mul] - | h_monomial n a => - rw [smeval_monomial, smul_mul_assoc, smul_mul_assoc, npow_mul_assoc, ← smul_mul_assoc] - -theorem smeval_mul_X_pow : ∀ (n : ℕ), (p * X^n).smeval x = p.smeval x * x^n - | 0 => by - simp only [npow_zero, mul_one] - | n + 1 => by - rw [npow_add, ← mul_assoc, npow_one, smeval_mul_X, smeval_mul_X_pow n, npow_add, - ← smeval_assoc_X_pow, npow_one] - theorem smeval_X_pow_assoc (m n : ℕ) : x ^ m * x ^ n * p.smeval x = x ^ m * (x ^ n * p.smeval x) := by induction p using Polynomial.induction_on' with @@ -252,13 +240,6 @@ theorem smeval_X_pow_mul : ∀ (n : ℕ), (X^n * p).smeval x = x^n * p.smeval x rw [add_comm, npow_add, mul_assoc, npow_one, smeval_X_mul, smeval_X_pow_mul n, npow_add, smeval_X_pow_assoc, npow_one] -theorem smeval_C_mul : (C r * p).smeval x = r • p.smeval x := by - induction p using Polynomial.induction_on' with - | h_add p q ph qh => - simp only [mul_add, smeval_add, ph, qh, smul_add] - | h_monomial n b => - simp only [C_mul_monomial, smeval_monomial, mul_smul] - theorem smeval_monomial_mul (n : ℕ) : (monomial n r * p).smeval x = r • (x ^ n * p.smeval x) := by induction p using Polynomial.induction_on' with @@ -268,6 +249,35 @@ theorem smeval_monomial_mul (n : ℕ) : | h_monomial n a => rw [smeval_monomial, monomial_mul_monomial, smeval_monomial, npow_add, mul_smul, mul_smul_comm] +end + +variable [IsScalarTower R S S] + +theorem smeval_mul_X : (p * X).smeval x = p.smeval x * x := by + induction p using Polynomial.induction_on' with + | h_add p q ph qh => + simp only [add_mul, smeval_add, ph, qh] + | h_monomial n a => + simp only [← monomial_one_one_eq_X, monomial_mul_monomial, smeval_monomial, mul_one, pow_succ', + mul_assoc, npow_add, smul_mul_assoc, npow_one] + +theorem smeval_assoc_X_pow (m n : ℕ) : + p.smeval x * x ^ m * x ^ n = p.smeval x * (x ^ m * x ^ n) := by + induction p using Polynomial.induction_on' with + | h_add p q ph qh => + simp only [smeval_add, ph, qh, add_mul] + | h_monomial n a => + rw [smeval_monomial, smul_mul_assoc, smul_mul_assoc, npow_mul_assoc, ← smul_mul_assoc] + +theorem smeval_mul_X_pow : ∀ (n : ℕ), (p * X^n).smeval x = p.smeval x * x^n + | 0 => by + simp only [npow_zero, mul_one] + | n + 1 => by + rw [npow_add, ← mul_assoc, npow_one, smeval_mul_X, smeval_mul_X_pow n, npow_add, + ← smeval_assoc_X_pow, npow_one] + +variable [SMulCommClass R S S] + theorem smeval_mul : (p * q).smeval x = p.smeval x * q.smeval x := by induction p using Polynomial.induction_on' with | h_add r s hr hs => diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index df41bba69c9e6..177560a1d9ab8 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -89,11 +89,14 @@ 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*} [CommRing R] {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 z : R) (a b c : ℍ[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 +section Zero +variable [Zero R] + /-- The imaginary part of a quaternion. -/ def im (x : ℍ[R,c₁,c₂]) : ℍ[R,c₁,c₂] := ⟨0, x.imI, x.imJ, x.imK⟩ @@ -159,6 +162,9 @@ theorem coe_zero : ((0 : R) : ℍ[R,c₁,c₂]) = 0 := rfl instance : Inhabited ℍ[R,c₁,c₂] := ⟨0⟩ +section One +variable [One R] + -- Porting note: removed `simps`, added simp lemmas manually instance : One ℍ[R,c₁,c₂] := ⟨⟨1, 0, 0, 0⟩⟩ @@ -175,6 +181,11 @@ instance : One ℍ[R,c₁,c₂] := ⟨⟨1, 0, 0, 0⟩⟩ @[simp, norm_cast] theorem coe_one : ((1 : R) : ℍ[R,c₁,c₂]) = 1 := rfl +end One +end Zero +section Add +variable [Add R] + -- Porting note: removed `simps`, added simp lemmas manually instance : Add ℍ[R,c₁,c₂] := ⟨fun a b => ⟨a.1 + b.1, a.2 + b.2, a.3 + b.3, a.4 + b.4⟩⟩ @@ -187,17 +198,27 @@ instance : Add ℍ[R,c₁,c₂] := @[simp] theorem add_imK : (a + b).imK = a.imK + b.imK := rfl -@[simp] theorem add_im : (a + b).im = a.im + b.im := - QuaternionAlgebra.ext (zero_add _).symm rfl rfl rfl - @[simp] theorem mk_add_mk (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) : (mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) + mk b₁ b₂ b₃ b₄ = mk (a₁ + b₁) (a₂ + b₂) (a₃ + b₃) (a₄ + b₄) := rfl +end Add + +section AddZeroClass +variable [AddZeroClass R] + +@[simp] theorem add_im : (a + b).im = a.im + b.im := + QuaternionAlgebra.ext (zero_add _).symm rfl rfl rfl + @[simp, norm_cast] theorem coe_add : ((x + y : R) : ℍ[R,c₁,c₂]) = x + y := by ext <;> simp +end AddZeroClass + +section Neg +variable [Neg R] + -- Porting note: removed `simps`, added simp lemmas manually instance : Neg ℍ[R,c₁,c₂] := ⟨fun a => ⟨-a.1, -a.2, -a.3, -a.4⟩⟩ @@ -209,13 +230,18 @@ instance : Neg ℍ[R,c₁,c₂] := ⟨fun a => ⟨-a.1, -a.2, -a.3, -a.4⟩⟩ @[simp] theorem neg_imK : (-a).imK = -a.imK := rfl -@[simp] theorem neg_im : (-a).im = -a.im := - QuaternionAlgebra.ext neg_zero.symm rfl rfl rfl - @[simp] theorem neg_mk (a₁ a₂ a₃ a₄ : R) : -(mk a₁ a₂ a₃ a₄ : ℍ[R,c₁,c₂]) = ⟨-a₁, -a₂, -a₃, -a₄⟩ := rfl +end Neg + +section AddGroup +variable [AddGroup R] + +@[simp] theorem neg_im : (-a).im = -a.im := + QuaternionAlgebra.ext neg_zero.symm rfl rfl rfl + @[simp, norm_cast] theorem coe_neg : ((-x : R) : ℍ[R,c₁,c₂]) = -x := by ext <;> simp @@ -254,6 +280,11 @@ theorem sub_self_im : a - a.im = a.re := theorem sub_self_re : a - a.re = a.im := QuaternionAlgebra.ext (sub_self _) (sub_zero _) (sub_zero _) (sub_zero _) +end AddGroup + +section Ring +variable [Ring R] + /-- Multiplication is given by * `1 * x = x * 1 = x`; @@ -290,11 +321,11 @@ theorem mk_mul_mk (a₁ a₂ a₃ a₄ b₁ b₂ b₃ b₄ : R) : a₁ * b₃ + c₁ * a₂ * b₄ + a₃ * b₁ - c₁ * a₄ * b₂, a₁ * b₄ + a₂ * b₃ - a₃ * b₂ + a₄ * b₁⟩ := rfl -section +end Ring +section SMul variable [SMul S R] [SMul T R] (s : S) --- Porting note: Lean 4 auto drops the unused `[Ring R]` argument instance : SMul S ℍ[R,c₁,c₂] where smul s a := ⟨s • a.1, s • a.2, s • a.3, s • a.4⟩ instance [SMul S T] [IsScalarTower S T R] : IsScalarTower S T ℍ[R,c₁,c₂] where @@ -311,7 +342,7 @@ instance [SMulCommClass S T R] : SMulCommClass S T ℍ[R,c₁,c₂] where @[simp] theorem smul_imK : (s • a).imK = s • a.imK := rfl -@[simp] theorem smul_im {S} [SMulZeroClass S R] (s : S) : (s • a).im = s • a.im := +@[simp] theorem smul_im {S} [CommRing R] [SMulZeroClass S R] (s : S) : (s • a).im = s • a.im := QuaternionAlgebra.ext (smul_zero s).symm rfl rfl rfl @[simp] @@ -319,17 +350,20 @@ theorem smul_mk (re im_i im_j im_k : R) : s • (⟨re, im_i, im_j, im_k⟩ : ℍ[R,c₁,c₂]) = ⟨s • re, s • im_i, s • im_j, s • im_k⟩ := rfl -end +end SMul @[simp, norm_cast] -theorem coe_smul [SMulZeroClass S R] (s : S) (r : R) : +theorem coe_smul [Zero R] [SMulZeroClass S R] (s : S) (r : R) : (↑(s • r) : ℍ[R,c₁,c₂]) = s • (r : ℍ[R,c₁,c₂]) := QuaternionAlgebra.ext rfl (smul_zero s).symm (smul_zero s).symm (smul_zero s).symm -instance : AddCommGroup ℍ[R,c₁,c₂] := +instance [AddCommGroup R] : AddCommGroup ℍ[R,c₁,c₂] := (equivProd c₁ c₂).injective.addCommGroup _ rfl (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) +section AddCommGroupWithOne +variable [AddCommGroupWithOne R] + instance : AddCommGroupWithOne ℍ[R,c₁,c₂] where natCast n := ((n : R) : ℍ[R,c₁,c₂]) natCast_zero := by simp @@ -424,6 +458,11 @@ theorem coe_intCast (z : ℤ) : ↑(z : R) = (z : ℍ[R,c₁,c₂]) := @[deprecated (since := "2024-04-17")] alias coe_int_cast := coe_intCast +end AddCommGroupWithOne + +-- For the remainder of the file we assume `CommRing R`. +variable [CommRing R] + instance instRing : Ring ℍ[R,c₁,c₂] where __ := inferInstanceAs (AddCommGroupWithOne ℍ[R,c₁,c₂]) left_distrib _ _ _ := by ext <;> simp <;> ring diff --git a/Mathlib/GroupTheory/HNNExtension.lean b/Mathlib/GroupTheory/HNNExtension.lean index eaf0e222c3698..f6df864683e3c 100644 --- a/Mathlib/GroupTheory/HNNExtension.lean +++ b/Mathlib/GroupTheory/HNNExtension.lean @@ -345,7 +345,7 @@ theorem unitsSMulGroup_snd (u : ℤˣ) (g : G) : (unitsSMulGroup φ d u g).2 = ((d.compl u).equiv g).2 := by rcases Int.units_eq_one_or u with rfl | rfl <;> rfl -variable {d} [DecidableEq G] +variable {d} /-- `Cancels u w` is a predicate expressing whether `t^u` cancels with some occurence of `t^-u` when when we multiply `t^u` by `w`. -/ diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean index 851351d1966d0..2d9497effe2a7 100644 --- a/Mathlib/GroupTheory/PushoutI.lean +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -270,8 +270,6 @@ instance (i : ι) : Inhabited (Pair d i) := head := 1, fstIdx_ne := fun h => by cases h }⟩ -variable [DecidableEq ι] [∀ i, DecidableEq (G i)] - @[ext] theorem ext {w₁ w₂ : NormalWord d} (hhead : w₁.head = w₂.head) (hlist : w₁.toList = w₂.toList) : w₁ = w₂ := by @@ -281,6 +279,53 @@ theorem ext {w₁ w₂ : NormalWord d} (hhead : w₁.head = w₂.head) open Subgroup.IsComplement +instance baseAction : MulAction H (NormalWord d) := + { smul := fun h w => { w with head := h * w.head }, + one_smul := by simp [instHSMul] + mul_smul := by simp [instHSMul, mul_assoc] } + +theorem base_smul_def' (h : H) (w : NormalWord d) : + h • w = { w with head := h * w.head } := rfl +/-- Take the product of a normal word as an element of the `PushoutI`. We show that this is +bijective, in `NormalWord.equiv`. -/ +def prod (w : NormalWord d) : PushoutI φ := + base φ w.head * ofCoprodI (w.toWord).prod + +@[simp] +theorem prod_base_smul (h : H) (w : NormalWord d) : + (h • w).prod = base φ h * w.prod := by + simp only [base_smul_def', prod, map_mul, mul_assoc] + +@[simp] +theorem prod_empty : (empty : NormalWord d).prod = 1 := by + simp [prod, empty] + +/-- A constructor that multiplies a `NormalWord` by an element, with condition to make +sure the underlying list does get longer. -/ +@[simps!] +noncomputable def cons {i} (g : G i) (w : NormalWord d) (hmw : w.fstIdx ≠ some i) + (hgr : g ∉ (φ i).range) : NormalWord d := + letI n := (d.compl i).equiv (g * (φ i w.head)) + letI w' := Word.cons (n.2 : G i) w.toWord hmw + (mt (coe_equiv_snd_eq_one_iff_mem _ (d.one_mem _)).1 + (mt (mul_mem_cancel_right (by simp)).1 hgr)) + { toWord := w' + head := (MonoidHom.ofInjective (d.injective i)).symm n.1 + normalized := fun i g hg => by + simp only [w', Word.cons, mem_cons, Sigma.mk.inj_iff] at hg + rcases hg with ⟨rfl, hg | hg⟩ + · simp + · exact w.normalized _ _ (by assumption) } + +@[simp] +theorem prod_cons {i} (g : G i) (w : NormalWord d) (hmw : w.fstIdx ≠ some i) + (hgr : g ∉ (φ i).range) : (cons g w hmw hgr).prod = of i g * w.prod := by + simp only [prod, cons, Word.prod, List.map, ← of_apply_eq_base φ i, equiv_fst_eq_mul_inv, + mul_assoc, MonoidHom.apply_ofInjective_symm, List.prod_cons, map_mul, map_inv, + ofCoprodI_of, inv_mul_cancel_left] + +variable [DecidableEq ι] [∀ i, DecidableEq (G i)] + /-- Given a word in `CoprodI`, if every letter is in the transversal and when we multiply by an element of the base group it still has this property, then the element of the base group we multiplied by was one. -/ @@ -338,23 +383,6 @@ theorem ext_smul {w₁ w₂ : NormalWord d} (i : ι) rw [inv_mul_eq_one] at this; subst this simp -/-- A constructor that multiplies a `NormalWord` by an element, with condition to make -sure the underlying list does get longer. -/ -@[simps!] -noncomputable def cons {i} (g : G i) (w : NormalWord d) (hmw : w.fstIdx ≠ some i) - (hgr : g ∉ (φ i).range) : NormalWord d := - letI n := (d.compl i).equiv (g * (φ i w.head)) - letI w' := Word.cons (n.2 : G i) w.toWord hmw - (mt (coe_equiv_snd_eq_one_iff_mem _ (d.one_mem _)).1 - (mt (mul_mem_cancel_right (by simp)).1 hgr)) - { toWord := w' - head := (MonoidHom.ofInjective (d.injective i)).symm n.1 - normalized := fun i g hg => by - simp only [w', Word.cons, mem_cons, Sigma.mk.inj_iff] at hg - rcases hg with ⟨rfl, hg | hg⟩ - · simp - · exact w.normalized _ _ (by assumption) } - /-- Given a pair `(head, tail)`, we can form a word by prepending `head` to `tail`, but putting head into normal form first, by making sure it is expressed as an element of the base group multiplied by an element of the transversal. -/ @@ -417,14 +445,6 @@ noncomputable instance summandAction (i : ι) : MulAction (G i) (NormalWord d) : dsimp [instHSMul] simp [mul_assoc, Equiv.apply_symm_apply, Function.End.mul_def] } -instance baseAction : MulAction H (NormalWord d) := - { smul := fun h w => { w with head := h * w.head }, - one_smul := by simp [instHSMul] - mul_smul := by simp [instHSMul, mul_assoc] } - -theorem base_smul_def' (h : H) (w : NormalWord d) : - h • w = { w with head := h * w.head } := rfl - theorem summand_smul_def' {i : ι} (g : G i) (w : NormalWord d) : g • w = (equivPair i).symm { equivPair i w with @@ -500,10 +520,6 @@ noncomputable def consRecOn {motive : NormalWord d → Sort _} (w : NormalWord d (equiv_snd_eq_self_iff_mem (d.compl i) (one_mem _)).2 (h3 _ _ (List.mem_cons_self _ _))] -/-- Take the product of a normal word as an element of the `PushoutI`. We show that this is -bijective, in `NormalWord.equiv`. -/ -def prod (w : NormalWord d) : PushoutI φ := - base φ w.head * ofCoprodI (w.toWord).prod theorem cons_eq_smul {i : ι} (g : G i) (w : NormalWord d) (hmw : w.fstIdx ≠ some i) @@ -523,11 +539,6 @@ theorem prod_summand_smul {i : ι} (g : G i) (w : NormalWord d) : MonoidHom.apply_ofInjective_symm, equiv_fst_eq_mul_inv, mul_assoc, map_mul, map_inv, Word.prod_smul, ofCoprodI_of, inv_mul_cancel_left, mul_inv_cancel_left] -@[simp] -theorem prod_base_smul (h : H) (w : NormalWord d) : - (h • w).prod = base φ h * w.prod := by - simp only [base_smul_def', prod, map_mul, mul_assoc] - @[simp] theorem prod_smul (g : PushoutI φ) (w : NormalWord d) : (g • w).prod = g * w.prod := by @@ -536,17 +547,6 @@ theorem prod_smul (g : PushoutI φ) (w : NormalWord d) : | base h => rw [base_smul_eq_smul, prod_base_smul] | mul x y ihx ihy => rw [mul_smul, ihx, ihy, mul_assoc] -@[simp] -theorem prod_empty : (empty : NormalWord d).prod = 1 := by - simp [prod, empty] - -@[simp] -theorem prod_cons {i} (g : G i) (w : NormalWord d) (hmw : w.fstIdx ≠ some i) - (hgr : g ∉ (φ i).range) : (cons g w hmw hgr).prod = of i g * w.prod := by - simp only [prod, cons, Word.prod, List.map, ← of_apply_eq_base φ i, equiv_fst_eq_mul_inv, - mul_assoc, MonoidHom.apply_ofInjective_symm, List.prod_cons, map_mul, map_inv, - ofCoprodI_of, inv_mul_cancel_left] - theorem prod_smul_empty (w : NormalWord d) : w.prod • empty = w := by induction w using consRecOn with | h_empty => simp diff --git a/Mathlib/RingTheory/Adjoin/Basic.lean b/Mathlib/RingTheory/Adjoin/Basic.lean index 12d762fa8add8..74d1dcbf2f39d 100644 --- a/Mathlib/RingTheory/Adjoin/Basic.lean +++ b/Mathlib/RingTheory/Adjoin/Basic.lean @@ -319,7 +319,7 @@ theorem adjoin_adjoin_of_tower (s : Set A) : adjoin S (adjoin R s : Set A) = adj exact subset_adjoin @[simp] -theorem adjoin_top : +theorem adjoin_top {A} [Semiring A] [Algebra S A] (t : Set A) : adjoin (⊤ : Subalgebra R S) t = (adjoin S t).restrictScalars (⊤ : Subalgebra R S) := let equivTop : Subalgebra (⊤ : Subalgebra R S) A ≃o Subalgebra S A := { toFun := fun s => { s with algebraMap_mem' := fun r => s.algebraMap_mem ⟨r, trivial⟩ } diff --git a/Mathlib/RingTheory/Derivation/ToSquareZero.lean b/Mathlib/RingTheory/Derivation/ToSquareZero.lean index 1fc6b65f92b46..eba48921f14f4 100644 --- a/Mathlib/RingTheory/Derivation/ToSquareZero.lean +++ b/Mathlib/RingTheory/Derivation/ToSquareZero.lean @@ -38,11 +38,11 @@ theorem diffToIdealOfQuotientCompEq_apply (f₁ f₂ : A →ₐ[R] B) ((diffToIdealOfQuotientCompEq I f₁ f₂ e) x : B) = f₁ x - f₂ x := rfl -variable [Algebra A B] [IsScalarTower R A B] +variable [Algebra A B] /-- Given a tower of algebras `R → A → B`, and a square-zero `I : Ideal B`, each lift `A →ₐ[R] B` of the canonical map `A →ₐ[R] B ⧸ I` corresponds to an `R`-derivation from `A` to `I`. -/ -def derivationToSquareZeroOfLift (f : A →ₐ[R] B) +def derivationToSquareZeroOfLift [IsScalarTower R A B] (f : A →ₐ[R] B) (e : (Ideal.Quotient.mkₐ R I).comp f = IsScalarTower.toAlgHom R A (B ⧸ I)) : Derivation R A I := by refine @@ -66,7 +66,7 @@ def derivationToSquareZeroOfLift (f : A →ₐ[R] B) IsScalarTower.coe_toAlgHom'] ring -theorem derivationToSquareZeroOfLift_apply (f : A →ₐ[R] B) +theorem derivationToSquareZeroOfLift_apply [IsScalarTower R A B] (f : A →ₐ[R] B) (e : (Ideal.Quotient.mkₐ R I).comp f = IsScalarTower.toAlgHom R A (B ⧸ I)) (x : A) : (derivationToSquareZeroOfLift I hI f e x : B) = f x - algebraMap A B x := rfl @@ -74,7 +74,7 @@ theorem derivationToSquareZeroOfLift_apply (f : A →ₐ[R] B) /-- Given a tower of algebras `R → A → B`, and a square-zero `I : Ideal B`, each `R`-derivation from `A` to `I` corresponds to a lift `A →ₐ[R] B` of the canonical map `A →ₐ[R] B ⧸ I`. -/ @[simps (config := .lemmasOnly)] -def liftOfDerivationToSquareZero (f : Derivation R A I) : A →ₐ[R] B := +def liftOfDerivationToSquareZero [IsScalarTower R A B] (f : Derivation R A I) : A →ₐ[R] B := { ((I.restrictScalars R).subtype.comp f.toLinearMap + (IsScalarTower.toAlgHom R A B).toLinearMap : A →ₗ[R] B) with toFun := fun x => f x + algebraMap A B x @@ -96,7 +96,7 @@ def liftOfDerivationToSquareZero (f : Derivation R A I) : A →ₐ[R] B := (IsScalarTower.toAlgHom R A B).toLinearMap).map_zero } -- @[simp] -- Porting note: simp normal form is `liftOfDerivationToSquareZero_mk_apply'` -theorem liftOfDerivationToSquareZero_mk_apply (d : Derivation R A I) (x : A) : +theorem liftOfDerivationToSquareZero_mk_apply [IsScalarTower R A B] (d : Derivation R A I) (x : A) : Ideal.Quotient.mk I (liftOfDerivationToSquareZero I hI d x) = algebraMap A (B ⧸ I) x := by rw [liftOfDerivationToSquareZero_apply, map_add, Ideal.Quotient.eq_zero_iff_mem.mpr (d x).prop, zero_add] @@ -111,7 +111,7 @@ theorem liftOfDerivationToSquareZero_mk_apply' (d : Derivation R A I) (x : A) : there is a 1-1 correspondence between `R`-derivations from `A` to `I` and lifts `A →ₐ[R] B` of the canonical map `A →ₐ[R] B ⧸ I`. -/ @[simps!] -def derivationToSquareZeroEquivLift : Derivation R A I ≃ +def derivationToSquareZeroEquivLift [IsScalarTower R A B] : Derivation R A I ≃ { f : A →ₐ[R] B // (Ideal.Quotient.mkₐ R I).comp f = IsScalarTower.toAlgHom R A (B ⧸ I) } := by refine ⟨fun d => ⟨liftOfDerivationToSquareZero I hI d, ?_⟩, fun f => (derivationToSquareZeroOfLift I hI f.1 f.2 : _), ?_, ?_⟩ diff --git a/Mathlib/RingTheory/Idempotents.lean b/Mathlib/RingTheory/Idempotents.lean index 0278341a7ee7a..84922b3c3220e 100644 --- a/Mathlib/RingTheory/Idempotents.lean +++ b/Mathlib/RingTheory/Idempotents.lean @@ -112,7 +112,6 @@ structure OrthogonalIdempotents : Prop where ortho : ∀ i j, i ≠ j → e i * e j = 0 variable {e} -variable (he : OrthogonalIdempotents e) lemma OrthogonalIdempotents.mul_eq [DecidableEq I] (he : OrthogonalIdempotents e) (i j) : e i * e j = if i = j then e i else 0 := by @@ -124,11 +123,12 @@ lemma OrthogonalIdempotents.iff_mul_eq [DecidableEq I] : OrthogonalIdempotents e ↔ ∀ i j, e i * e j = if i = j then e i else 0 := ⟨mul_eq, fun H ↦ ⟨fun i ↦ by simpa using H i i, fun i j e ↦ by simpa [e] using H i j⟩⟩ -lemma OrthogonalIdempotents.isIdempotentElem_sum [Fintype I] : IsIdempotentElem (∑ i, e i) := by +lemma OrthogonalIdempotents.isIdempotentElem_sum [Fintype I] (he : OrthogonalIdempotents e) : + IsIdempotentElem (∑ i, e i) := by classical simp [IsIdempotentElem, Finset.sum_mul, Finset.mul_sum, he.mul_eq] -lemma OrthogonalIdempotents.map : +lemma OrthogonalIdempotents.map (he : OrthogonalIdempotents e) : OrthogonalIdempotents (f ∘ e) := by classical simp [iff_mul_eq, he.mul_eq, ← map_mul f, apply_ite f] @@ -138,7 +138,7 @@ lemma OrthogonalIdempotents.map_injective_iff (hf : Function.Injective f) : classical simp [iff_mul_eq, ← hf.eq_iff, apply_ite] -lemma OrthogonalIdempotents.embedding {J} (i : J ↪ I) : +lemma OrthogonalIdempotents.embedding (he : OrthogonalIdempotents e) {J} (i : J ↪ I) : OrthogonalIdempotents (e ∘ i) := by classical simp [iff_mul_eq, he.mul_eq] @@ -152,7 +152,7 @@ lemma OrthogonalIdempotents.unique [Unique I] : OrthogonalIdempotents e ↔ IsIdempotentElem (e default) := by simp [orthogonalIdempotents_iff, Unique.forall_iff] -lemma OrthogonalIdempotents.option [Fintype I] (x) +lemma OrthogonalIdempotents.option (he : OrthogonalIdempotents e) [Fintype I] (x) (hx : IsIdempotentElem x) (hx₁ : x * ∑ i, e i = 0) (hx₂ : (∑ i, e i) * x = 0) : OrthogonalIdempotents (Option.elim · x e) where idem i := i.rec hx he.idem @@ -240,7 +240,7 @@ lemma CompleteOrthogonalIdempotents.single {I : Type*} [Fintype I] [DecidableEq · subst hi; simp [hij] · simp [hi] -lemma CompleteOrthogonalIdempotents.map : +lemma CompleteOrthogonalIdempotents.map (he : CompleteOrthogonalIdempotents e) : CompleteOrthogonalIdempotents (f ∘ e) where __ := he.toOrthogonalIdempotents.map f complete := by simp only [Function.comp_apply, ← map_sum, he.complete, map_one] diff --git a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean index d9f1a7bdaee22..e9b925a2e5ed6 100644 --- a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean @@ -52,18 +52,7 @@ open Finset Nat namespace NewtonIdentities -variable (σ : Type*) [Fintype σ] [DecidableEq σ] (R : Type*) [CommRing R] - -private def pairs (k : ℕ) : Finset (Finset σ × σ) := - univ.filter (fun t ↦ card t.fst ≤ k ∧ (card t.fst = k → t.snd ∈ t.fst)) - -@[simp] -private lemma mem_pairs (k : ℕ) (t : Finset σ × σ) : - t ∈ pairs σ k ↔ card t.fst ≤ k ∧ (card t.fst = k → t.snd ∈ t.fst) := by - simp [pairs] - -private def weight (k : ℕ) (t : Finset σ × σ) : MvPolynomial σ R := - (-1) ^ card t.fst * ((∏ a ∈ t.fst, X a) * X t.snd ^ (k - card t.fst)) +variable (σ : Type*) [DecidableEq σ] (R : Type*) [CommRing R] private def pairMap (t : Finset σ × σ) : Finset σ × σ := if h : t.snd ∈ t.fst then (t.fst.erase t.snd, t.snd) else (t.fst.cons t.snd h, t.snd) @@ -81,6 +70,29 @@ private lemma pairMap_of_snd_nmem_fst {t : Finset σ × σ} (h : t.snd ∉ t.fst pairMap σ t = (t.fst.cons t.snd h, t.snd) := by simp [pairMap, h] +@[simp] +private theorem pairMap_involutive : (pairMap σ).Involutive := by + intro t + rw [pairMap, pairMap] + split_ifs with h1 h2 h3 + · simp at h2 + · simp [insert_erase h1] + · simp_all + · simp at h3 + +variable [Fintype σ] + +private def pairs (k : ℕ) : Finset (Finset σ × σ) := + univ.filter (fun t ↦ card t.fst ≤ k ∧ (card t.fst = k → t.snd ∈ t.fst)) + +@[simp] +private lemma mem_pairs (k : ℕ) (t : Finset σ × σ) : + t ∈ pairs σ k ↔ card t.fst ≤ k ∧ (card t.fst = k → t.snd ∈ t.fst) := by + simp [pairs] + +private def weight (k : ℕ) (t : Finset σ × σ) : MvPolynomial σ R := + (-1) ^ card t.fst * ((∏ a ∈ t.fst, X a) * X t.snd ^ (k - card t.fst)) + private theorem pairMap_mem_pairs {k : ℕ} (t : Finset σ × σ) (h : t ∈ pairs σ k) : pairMap σ t ∈ pairs σ k := by rw [mem_pairs] at h ⊢ @@ -98,16 +110,6 @@ private theorem pairMap_mem_pairs {k : ℕ} (t : Finset σ × σ) (h : t ∈ pai simp only [card_cons, mem_cons, true_or, implies_true, and_true] exact (le_iff_eq_or_lt.mp h.left).resolve_left h.right -@[simp] -private theorem pairMap_involutive : (pairMap σ).Involutive := by - intro t - rw [pairMap, pairMap] - split_ifs with h1 h2 h3 - · simp at h2 - · simp [insert_erase h1] - · simp_all - · simp at h3 - private theorem weight_add_weight_pairMap {k : ℕ} (t : Finset σ × σ) (h : t ∈ pairs σ k) : weight σ R k t + weight σ R k (pairMap σ t) = 0 := by rw [weight, weight] @@ -189,7 +191,7 @@ private theorem esymm_summand_to_weight (k : ℕ) (A : Finset σ) (h : A ∈ pow ∑ j ∈ A, weight σ R k (A, j) = k * (-1) ^ k * (∏ i ∈ A, X i : MvPolynomial σ R) := by simp [weight, mem_powersetCard_univ.mp h, mul_assoc] -private theorem esymm_to_weight (k : ℕ) : k * esymm σ R k = +private theorem esymm_to_weight [DecidableEq σ] (k : ℕ) : k * esymm σ R k = (-1) ^ k * ∑ t ∈ filter (fun t ↦ card t.fst = k) (pairs σ k), weight σ R k t := by rw [esymm, sum_filter_pairs_eq_sum_powersetCard_sum σ R k (fun t ↦ weight σ R k t), sum_congr rfl (esymm_summand_to_weight σ R k), mul_comm (k : MvPolynomial σ R) ((-1) ^ k), @@ -205,7 +207,7 @@ private theorem esymm_mul_psum_summand_to_weight (k : ℕ) (a : ℕ × ℕ) (ha refine sum_congr rfl fun s hs ↦ ?_ rw [mem_powersetCard_univ.mp hs, ← mem_antidiagonal.mp ha, add_sub_self_left] -private theorem esymm_mul_psum_to_weight (k : ℕ) : +private theorem esymm_mul_psum_to_weight [DecidableEq σ] (k : ℕ) : ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst < k), (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd = ∑ t ∈ filter (fun t ↦ card t.fst < k) (pairs σ k), weight σ R k t := by diff --git a/Mathlib/RingTheory/PowerSeries/WellKnown.lean b/Mathlib/RingTheory/PowerSeries/WellKnown.lean index 4fb2f9ee739a8..6ef5b6366a493 100644 --- a/Mathlib/RingTheory/PowerSeries/WellKnown.lean +++ b/Mathlib/RingTheory/PowerSeries/WellKnown.lean @@ -151,7 +151,7 @@ def sin : PowerSeries A := def cos : PowerSeries A := mk fun n => if Even n then algebraMap ℚ A ((-1) ^ (n / 2) / n !) else 0 -variable {A A'} [Ring A] [Ring A'] [Algebra ℚ A] [Algebra ℚ A'] (n : ℕ) (f : A →+* A') +variable {A A'} (n : ℕ) @[simp] theorem coeff_exp : coeff A n (exp A) = algebraMap ℚ A (1 / n !) := @@ -162,6 +162,8 @@ theorem constantCoeff_exp : constantCoeff A (exp A) = 1 := by rw [← coeff_zero_eq_constantCoeff_apply, coeff_exp] simp +variable (f : A →+* A') + @[simp] theorem map_exp : map (f : A →+* A') (exp A) = exp A' := by ext diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index b8d6dd6c7c672..e2acb507eec86 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -373,11 +373,35 @@ section open Ideal -variable [CommRing R] [IsDomain R] +variable [CommRing R] section Bezout variable [IsBezout R] +theorem isCoprime_of_dvd (x y : R) (nonzero : ¬(x = 0 ∧ y = 0)) + (H : ∀ z ∈ nonunits R, z ≠ 0 → z ∣ x → ¬z ∣ y) : IsCoprime x y := + (isRelPrime_of_no_nonunits_factors nonzero H).isCoprime + +theorem dvd_or_coprime (x y : R) (h : Irreducible x) : x ∣ y ∨ IsCoprime x y := + h.dvd_or_isRelPrime.imp_right IsRelPrime.isCoprime + +/-- See also `Irreducible.isRelPrime_iff_not_dvd`. -/ +theorem Irreducible.coprime_iff_not_dvd {p n : R} (hp : Irreducible p) : + IsCoprime p n ↔ ¬p ∣ n := by rw [← isRelPrime_iff_isCoprime, hp.isRelPrime_iff_not_dvd] + +/-- See also `Irreducible.coprime_iff_not_dvd'`. -/ +theorem Irreducible.dvd_iff_not_coprime {p n : R} (hp : Irreducible p) : p ∣ n ↔ ¬IsCoprime p n := + iff_not_comm.2 hp.coprime_iff_not_dvd + +theorem Irreducible.coprime_pow_of_not_dvd {p a : R} (m : ℕ) (hp : Irreducible p) (h : ¬p ∣ a) : + IsCoprime a (p ^ m) := + (hp.coprime_iff_not_dvd.2 h).symm.pow_right + +theorem Irreducible.coprime_or_dvd {p : R} (hp : Irreducible p) (i : R) : IsCoprime p i ∨ p ∣ i := + (_root_.em _).imp_right hp.dvd_iff_not_coprime.2 + +variable [IsDomain R] + section GCD variable [GCDMonoid R] @@ -404,31 +428,9 @@ theorem gcd_isUnit_iff (x y : R) : IsUnit (gcd x y) ↔ IsCoprime x y := by end GCD -theorem isCoprime_of_dvd (x y : R) (nonzero : ¬(x = 0 ∧ y = 0)) - (H : ∀ z ∈ nonunits R, z ≠ 0 → z ∣ x → ¬z ∣ y) : IsCoprime x y := - (isRelPrime_of_no_nonunits_factors nonzero H).isCoprime - -theorem dvd_or_coprime (x y : R) (h : Irreducible x) : x ∣ y ∨ IsCoprime x y := - h.dvd_or_isRelPrime.imp_right IsRelPrime.isCoprime - -/-- See also `Irreducible.isRelPrime_iff_not_dvd`. -/ -theorem Irreducible.coprime_iff_not_dvd {p n : R} (hp : Irreducible p) : - IsCoprime p n ↔ ¬p ∣ n := by rw [← isRelPrime_iff_isCoprime, hp.isRelPrime_iff_not_dvd] - theorem Prime.coprime_iff_not_dvd {p n : R} (hp : Prime p) : IsCoprime p n ↔ ¬p ∣ n := hp.irreducible.coprime_iff_not_dvd -/-- See also `Irreducible.coprime_iff_not_dvd'`. -/ -theorem Irreducible.dvd_iff_not_coprime {p n : R} (hp : Irreducible p) : p ∣ n ↔ ¬IsCoprime p n := - iff_not_comm.2 hp.coprime_iff_not_dvd - -theorem Irreducible.coprime_pow_of_not_dvd {p a : R} (m : ℕ) (hp : Irreducible p) (h : ¬p ∣ a) : - IsCoprime a (p ^ m) := - (hp.coprime_iff_not_dvd.2 h).symm.pow_right - -theorem Irreducible.coprime_or_dvd {p : R} (hp : Irreducible p) (i : R) : IsCoprime p i ∨ p ∣ i := - (_root_.em _).imp_right hp.dvd_iff_not_coprime.2 - theorem exists_associated_pow_of_mul_eq_pow' {a b c : R} (hab : IsCoprime a b) {k : ℕ} (h : a * b = c ^ k) : ∃ d : R, Associated (d ^ k) a := by classical @@ -443,7 +445,7 @@ theorem exists_associated_pow_of_associated_pow_mul {a b c : R} (hab : IsCoprime end Bezout -variable [IsPrincipalIdealRing R] +variable [IsDomain R] [IsPrincipalIdealRing R] theorem isCoprime_of_irreducible_dvd {x y : R} (nonzero : ¬(x = 0 ∧ y = 0)) (H : ∀ z : R, Irreducible z → z ∣ x → ¬z ∣ y) : IsCoprime x y := diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index 9ec78d1c04325..ca7bf9b421ff0 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -1584,6 +1584,20 @@ theorem prime_pow_le_iff_le_bcount [DecidableEq (Associates α)] {m p : Associat rw [bcount, factors_mk, Multiset.le_count_iff_replicate_le, ← factors_le, factors_prime_pow, factors_mk, WithTop.coe_le_coe] <;> assumption +@[simp] +theorem factors_one [Nontrivial α] : factors (1 : Associates α) = 0 := by + apply eq_of_prod_eq_prod + rw [Associates.factors_prod] + exact Multiset.prod_zero + +@[simp] +theorem pow_factors [Nontrivial α] {a : Associates α} {k : ℕ} : + (a ^ k).factors = k • a.factors := by + induction' k with n h + · rw [zero_nsmul, pow_zero] + exact factors_one + · rw [pow_succ, succ_nsmul, factors_mul, h] + section count variable [DecidableEq (Associates α)] [∀ p : Associates α, Decidable (Irreducible p)] @@ -1615,16 +1629,16 @@ theorem count_ne_zero_iff_dvd {a p : α} (ha0 : a ≠ 0) (hp : Irreducible p) : (Associates.irreducible_mk.mpr hp)] at h exact (zero_lt_one.trans_le h).ne' -theorem count_self [Nontrivial α] [DecidableEq (Associates α)] {p : Associates α} +theorem count_self [Nontrivial α] {p : Associates α} (hp : Irreducible p) : p.count p.factors = 1 := by simp [factors_self hp, Associates.count_some hp] -theorem count_eq_zero_of_ne [DecidableEq (Associates α)] {p q : Associates α} (hp : Irreducible p) +theorem count_eq_zero_of_ne {p q : Associates α} (hp : Irreducible p) (hq : Irreducible q) (h : p ≠ q) : p.count q.factors = 0 := not_ne_iff.mp fun h' ↦ h <| associated_iff_eq.mp <| hp.associated_of_dvd hq <| le_of_count_ne_zero hq.ne_zero hp h' -theorem count_mul [DecidableEq (Associates α)] {a : Associates α} (ha : a ≠ 0) {b : Associates α} +theorem count_mul {a : Associates α} (ha : a ≠ 0) {b : Associates α} (hb : b ≠ 0) {p : Associates α} (hp : Irreducible p) : count p (factors (a * b)) = count p a.factors + count p b.factors := by obtain ⟨a0, nza, rfl⟩ := exists_non_zero_rep ha @@ -1632,7 +1646,7 @@ theorem count_mul [DecidableEq (Associates α)] {a : Associates α} (ha : a ≠ rw [factors_mul, factors_mk a0 nza, factors_mk b0 nzb, ← FactorSet.coe_add, count_some hp, Multiset.count_add, count_some hp, count_some hp] -theorem count_of_coprime [DecidableEq (Associates α)] {a : Associates α} (ha : a ≠ 0) +theorem count_of_coprime {a : Associates α} (ha : a ≠ 0) {b : Associates α} (hb : b ≠ 0) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {p : Associates α} (hp : Irreducible p) : count p a.factors = 0 ∨ count p b.factors = 0 := by rw [or_iff_not_imp_left, ← Ne] @@ -1641,7 +1655,7 @@ theorem count_of_coprime [DecidableEq (Associates α)] {a : Associates α} (ha : exact ⟨p, le_of_count_ne_zero ha hp hca, le_of_count_ne_zero hb hp hcb, UniqueFactorizationMonoid.irreducible_iff_prime.mp hp⟩ -theorem count_mul_of_coprime [DecidableEq (Associates α)] {a : Associates α} {b : Associates α} +theorem count_mul_of_coprime {a : Associates α} {b : Associates α} (hb : b ≠ 0) {p : Associates α} (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) : count p a.factors = 0 ∨ count p a.factors = count p (a * b).factors := by by_cases ha : a = 0 @@ -1650,7 +1664,7 @@ theorem count_mul_of_coprime [DecidableEq (Associates α)] {a : Associates α} { apply Or.intro_right rw [count_mul ha hb hp, hb0, add_zero] -theorem count_mul_of_coprime' [DecidableEq (Associates α)] {a b : Associates α} {p : Associates α} +theorem count_mul_of_coprime' {a b : Associates α} {p : Associates α} (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) : count p (a * b).factors = count p a.factors ∨ count p (a * b).factors = count p b.factors := by by_cases ha : a = 0 @@ -1664,7 +1678,7 @@ theorem count_mul_of_coprime' [DecidableEq (Associates α)] {a b : Associates α · apply Or.intro_left rw [hb0, add_zero] -theorem dvd_count_of_dvd_count_mul [DecidableEq (Associates α)] {a b : Associates α} (hb : b ≠ 0) +theorem dvd_count_of_dvd_count_mul {a b : Associates α} (hb : b ≠ 0) {p : Associates α} (hp : Irreducible p) (hab : ∀ d, d ∣ a → d ∣ b → ¬Prime d) {k : ℕ} (habk : k ∣ count p (a * b).factors) : k ∣ count p a.factors := by by_cases ha : a = 0 @@ -1675,21 +1689,7 @@ theorem dvd_count_of_dvd_count_mul [DecidableEq (Associates α)] {a b : Associat · rw [count_mul ha hb hp, h] at habk exact habk -@[simp] -theorem factors_one [Nontrivial α] : factors (1 : Associates α) = 0 := by - apply eq_of_prod_eq_prod - rw [Associates.factors_prod] - exact Multiset.prod_zero - -@[simp] -theorem pow_factors [Nontrivial α] {a : Associates α} {k : ℕ} : - (a ^ k).factors = k • a.factors := by - induction' k with n h - · rw [zero_nsmul, pow_zero] - exact factors_one - · rw [pow_succ, succ_nsmul, factors_mul, h] - -theorem count_pow [Nontrivial α] [DecidableEq (Associates α)] {a : Associates α} (ha : a ≠ 0) +theorem count_pow [Nontrivial α] {a : Associates α} (ha : a ≠ 0) {p : Associates α} (hp : Irreducible p) (k : ℕ) : count p (a ^ k).factors = k * count p a.factors := by induction' k with n h @@ -1697,12 +1697,12 @@ theorem count_pow [Nontrivial α] [DecidableEq (Associates α)] {a : Associates · rw [pow_succ', count_mul ha (pow_ne_zero _ ha) hp, h] ring -theorem dvd_count_pow [Nontrivial α] [DecidableEq (Associates α)] {a : Associates α} (ha : a ≠ 0) +theorem dvd_count_pow [Nontrivial α] {a : Associates α} (ha : a ≠ 0) {p : Associates α} (hp : Irreducible p) (k : ℕ) : k ∣ count p (a ^ k).factors := by rw [count_pow ha hp] apply dvd_mul_right -theorem is_pow_of_dvd_count [DecidableEq (Associates α)] {a : Associates α} +theorem is_pow_of_dvd_count {a : Associates α} (ha : a ≠ 0) {k : ℕ} (hk : ∀ p : Associates α, Irreducible p → k ∣ count p a.factors) : ∃ b : Associates α, a = b ^ k := by nontriviality α @@ -1721,7 +1721,7 @@ theorem is_pow_of_dvd_count [DecidableEq (Associates α)] {a : Associates α} /-- The only divisors of prime powers are prime powers. See `eq_pow_find_of_dvd_irreducible_pow` for an explicit expression as a p-power (without using `count`). -/ -theorem eq_pow_count_factors_of_dvd_pow [DecidableEq (Associates α)] {p a : Associates α} +theorem eq_pow_count_factors_of_dvd_pow {p a : Associates α} (hp : Irreducible p) {n : ℕ} (h : a ∣ p ^ n) : a = p ^ p.count a.factors := by nontriviality α have hph := pow_ne_zero n hp.ne_zero @@ -1738,7 +1738,7 @@ theorem eq_pow_count_factors_of_dvd_pow [DecidableEq (Associates α)] {p a : Ass · rw [h, count_self hp, mul_one] · rw [count_eq_zero_of_ne hq hp h, mul_zero, eq_zero_of_ne q hq h] -theorem count_factors_eq_find_of_dvd_pow [DecidableEq (Associates α)] {a p : Associates α} +theorem count_factors_eq_find_of_dvd_pow {a p : Associates α} (hp : Irreducible p) [∀ n : ℕ, Decidable (a ∣ p ^ n)] {n : ℕ} (h : a ∣ p ^ n) : @Nat.find (fun n => a ∣ p ^ n) _ ⟨n, h⟩ = p.count a.factors := by apply le_antisymm diff --git a/Mathlib/RingTheory/WittVector/WittPolynomial.lean b/Mathlib/RingTheory/WittVector/WittPolynomial.lean index 0ef21ce915858..ff8d6143ceae0 100644 --- a/Mathlib/RingTheory/WittVector/WittPolynomial.lean +++ b/Mathlib/RingTheory/WittVector/WittPolynomial.lean @@ -65,7 +65,7 @@ open Finsupp (single) --attribute [-simp] coe_eval₂_hom variable (p : ℕ) -variable (R : Type*) [CommRing R] [DecidableEq R] +variable (R : Type*) [CommRing R] /-- `wittPolynomial p R n` is the `n`-th Witt polynomial with respect to a prime `p` with coefficients in a commutative ring `R`. diff --git a/Mathlib/Tactic/ComputeDegree.lean b/Mathlib/Tactic/ComputeDegree.lean index 34514dba89c8a..3dc94baf457c7 100644 --- a/Mathlib/Tactic/ComputeDegree.lean +++ b/Mathlib/Tactic/ComputeDegree.lean @@ -164,10 +164,13 @@ theorem degree_eq_of_le_of_coeff_ne_zero' {deg m o : WithBot ℕ} {c : R} {p : R · obtain ⟨m, rfl⟩ := WithBot.ne_bot_iff_exists.mp hh exact degree_eq_of_le_of_coeff_ne_zero ‹_› ‹_› -variable {m n : ℕ} {f : R[X]} {r : R} (h : coeff f m = r) (natDeg_eq_coeff : m = n) +variable {m n : ℕ} {f : R[X]} {r : R} -theorem coeff_congr_lhs : coeff f n = r := natDeg_eq_coeff ▸ h -theorem coeff_congr {s : R} (rs : r = s) : coeff f n = s := natDeg_eq_coeff ▸ rs ▸ h +theorem coeff_congr_lhs (h : coeff f m = r) (natDeg_eq_coeff : m = n) : coeff f n = r := + natDeg_eq_coeff ▸ h +theorem coeff_congr (h : coeff f m = r) (natDeg_eq_coeff : m = n) {s : R} (rs : r = s) : + coeff f n = s := + natDeg_eq_coeff ▸ rs ▸ h end congr_lemmas From c0d0e84f369ac7f4457b5f0383a6b82295526603 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 7 Aug 2024 06:03:40 +0000 Subject: [PATCH 076/359] chore: backports for leanprover/lean4#4814 (part 18) (#15431) see https://github.com/leanprover-community/mathlib4/pull/15245 Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- .../ShortComplex/ConcreteCategory.lean | 11 ++-- .../Galois/Prorepresentability.lean | 51 ++++++++++--------- .../Triangulated/HomologicalFunctor.lean | 2 - .../Manifold/LocalInvariantProperties.lean | 3 ++ Mathlib/RingTheory/FractionalIdeal/Basic.lean | 10 ++-- Mathlib/RingTheory/Ideal/Maps.lean | 16 +++--- .../Category/TopCat/Limits/Cofiltered.lean | 3 +- Mathlib/Topology/Covering.lean | 12 ++--- Mathlib/Topology/Sets/Closeds.lean | 2 - Mathlib/Topology/StoneCech.lean | 33 +++++++----- Mathlib/Topology/UniformSpace/Cauchy.lean | 8 +-- 11 files changed, 84 insertions(+), 67 deletions(-) diff --git a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean index 282e51089144b..2c1cd6db62c31 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean @@ -38,7 +38,10 @@ lemma ShortComplex.zero_apply section preadditive variable [Preadditive C] [(forget₂ C Ab).Additive] [(forget₂ C Ab).PreservesHomology] - [HasZeroObject C] (S : ShortComplex C) + (S : ShortComplex C) + +section +variable [HasZeroObject C] lemma Preadditive.mono_iff_injective {X Y : C} (f : X ⟶ Y) : Mono f ↔ Function.Injective ((forget₂ C Ab).map f) := by @@ -70,6 +73,8 @@ lemma Preadditive.epi_iff_surjective' {X Y : C} (f : X ⟶ Y) : have e : forget₂ C Ab ⋙ forget Ab ≅ forget C := eqToIso (HasForget₂.forget_comp) exact Arrow.isoOfNatIso e (Arrow.mk f) +end + namespace ShortComplex lemma exact_iff_exact_map_forget₂ [S.HasHomology] : @@ -84,12 +89,12 @@ lemma exact_iff_of_concreteCategory [S.HasHomology] : variable {S} -lemma ShortExact.injective_f (hS : S.ShortExact) : +lemma ShortExact.injective_f [HasZeroObject C] (hS : S.ShortExact) : Function.Injective ((forget₂ C Ab).map S.f) := by rw [← Preadditive.mono_iff_injective] exact hS.mono_f -lemma ShortExact.surjective_g (hS : S.ShortExact) : +lemma ShortExact.surjective_g [HasZeroObject C] (hS : S.ShortExact) : Function.Surjective ((forget₂ C Ab).map S.g) := by rw [← Preadditive.epi_iff_surjective] exact hS.epi_g diff --git a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean index cee99b90c06b8..e02cc45fc92e2 100644 --- a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean +++ b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean @@ -49,8 +49,7 @@ namespace PreGaloisCategory open Limits Functor variable {C : Type u₁} [Category.{u₂} C] [GaloisCategory C] -variable (F : C ⥤ FintypeCat.{u₂}) [FiberFunctor F] - +variable (F : C ⥤ FintypeCat.{u₂}) /-- A pointed Galois object is a Galois object with a fixed point of its fiber. -/ structure PointedGaloisObject : Type (max u₁ u₂) where /-- The underlying object of `C`. -/ @@ -105,20 +104,6 @@ lemma comp_val {A B C : PointedGaloisObject F} (f : A ⟶ B) (g : B ⟶ C) : variable (F) -/-- The category of pointed Galois objects is cofiltered. -/ -instance : IsCofilteredOrEmpty (PointedGaloisObject F) where - cone_objs := fun ⟨A, a, _⟩ ⟨B, b, _⟩ ↦ by - obtain ⟨Z, f, z, hgal, hfz⟩ := exists_hom_from_galois_of_fiber F (A ⨯ B) - <| (fiberBinaryProductEquiv F A B).symm (a, b) - refine ⟨⟨Z, z, hgal⟩, ⟨f ≫ prod.fst, ?_⟩, ⟨f ≫ prod.snd, ?_⟩, trivial⟩ - · simp only [F.map_comp, hfz, FintypeCat.comp_apply, fiberBinaryProductEquiv_symm_fst_apply] - · simp only [F.map_comp, hfz, FintypeCat.comp_apply, fiberBinaryProductEquiv_symm_snd_apply] - cone_maps := fun ⟨A, a, _⟩ ⟨B, b, _⟩ ⟨f, hf⟩ ⟨g, hg⟩ ↦ by - obtain ⟨Z, h, z, hgal, hhz⟩ := exists_hom_from_galois_of_fiber F A a - refine ⟨⟨Z, z, hgal⟩, ⟨h, hhz⟩, hom_ext ?_⟩ - apply evaluation_injective_of_isConnected F Z B z - simp [hhz, hf, hg] - /-- The canonical functor from pointed Galois objects to `C`. -/ def incl : PointedGaloisObject F ⥤ C where obj := fun A ↦ A @@ -149,6 +134,22 @@ lemma cocone_app (A : PointedGaloisObject F) (B : C) (f : (A : C) ⟶ B) : ((cocone F).ι.app ⟨A⟩).app B f = F.map f A.pt := rfl +variable [FiberFunctor F] + +/-- The category of pointed Galois objects is cofiltered. -/ +instance : IsCofilteredOrEmpty (PointedGaloisObject F) where + cone_objs := fun ⟨A, a, _⟩ ⟨B, b, _⟩ ↦ by + obtain ⟨Z, f, z, hgal, hfz⟩ := exists_hom_from_galois_of_fiber F (A ⨯ B) + <| (fiberBinaryProductEquiv F A B).symm (a, b) + refine ⟨⟨Z, z, hgal⟩, ⟨f ≫ prod.fst, ?_⟩, ⟨f ≫ prod.snd, ?_⟩, trivial⟩ + · simp only [F.map_comp, hfz, FintypeCat.comp_apply, fiberBinaryProductEquiv_symm_fst_apply] + · simp only [F.map_comp, hfz, FintypeCat.comp_apply, fiberBinaryProductEquiv_symm_snd_apply] + cone_maps := fun ⟨A, a, _⟩ ⟨B, b, _⟩ ⟨f, hf⟩ ⟨g, hg⟩ ↦ by + obtain ⟨Z, h, z, hgal, hhz⟩ := exists_hom_from_galois_of_fiber F A a + refine ⟨⟨Z, z, hgal⟩, ⟨h, hhz⟩, hom_ext ?_⟩ + apply evaluation_injective_of_isConnected F Z B z + simp [hhz, hf, hg] + /-- `cocone F` is a colimit cocone, i.e. `F` is pro-represented by `incl F`. -/ noncomputable def isColimit : IsColimit (cocone F) := by refine evaluationJointlyReflectsColimits _ (fun X ↦ ?_) @@ -213,14 +214,6 @@ lemma autGaloisSystem_map_surjective ⦃A B : PointedGaloisObject F⦄ (f : A simp only [autGaloisSystem_map] exact hψ -/-- `autGalois.π` is surjective for every pointed Galois object. -/ -theorem AutGalois.π_surjective (A : PointedGaloisObject F) : - Function.Surjective (AutGalois.π F A) := fun (σ : Aut A.obj) ↦ by - have (i : PointedGaloisObject F) : Finite ((autGaloisSystem F ⋙ forget _).obj i) := - inferInstanceAs <| Finite (Aut (i.obj)) - exact eval_section_surjective_of_surjective - (autGaloisSystem F ⋙ forget _) (autGaloisSystem_map_surjective F) A σ - /-- Equality of elements of `AutGalois F` can be checked on the projections on each pointed Galois object. -/ lemma AutGalois.ext {f g : AutGalois F} @@ -229,6 +222,16 @@ lemma AutGalois.ext {f g : AutGalois F} ext A exact h A +/-- `autGalois.π` is surjective for every pointed Galois object. -/ +theorem AutGalois.π_surjective [FiberFunctor F] (A : PointedGaloisObject F) : + Function.Surjective (AutGalois.π F A) := fun (σ : Aut A.obj) ↦ by + have (i : PointedGaloisObject F) : Finite ((autGaloisSystem F ⋙ forget _).obj i) := + inferInstanceAs <| Finite (Aut (i.obj)) + exact eval_section_surjective_of_surjective + (autGaloisSystem F ⋙ forget _) (autGaloisSystem_map_surjective F) A σ + +variable [FiberFunctor F] + section EndAutGaloisIsomorphism /-! diff --git a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean index 6e411eb3a4025..ee62de01bca66 100644 --- a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean +++ b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean @@ -86,8 +86,6 @@ lemma IsHomological.mk' [F.PreservesZeroMorphisms] exact (ShortComplex.exact_iff_of_iso (F.mapShortComplex.mapIso ((shortComplexOfDistTriangleIsoOfIso e hT)))).2 h' -variable [F.IsHomological] - lemma IsHomological.of_iso {F₁ F₂ : C ⥤ A} [F₁.IsHomological] (e : F₁ ≅ F₂) : F₂.IsHomological := have := preservesZeroMorphisms_of_iso e diff --git a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean index 9f702ef2c1989..6f17294b2d254 100644 --- a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean +++ b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean @@ -212,6 +212,7 @@ theorem liftPropWithinAt_self_target {f : M → H'} : namespace LocalInvariantProp +section variable (hG : G.LocalInvariantProp G' P) /-- `LiftPropWithinAt P f s x` is equivalent to a definition where we restrict the set we are @@ -397,6 +398,8 @@ theorem liftPropOn_congr (h : LiftPropOn P g s) (h₁ : ∀ y ∈ s, g' y = g y) theorem liftPropOn_congr_iff (h₁ : ∀ y ∈ s, g' y = g y) : LiftPropOn P g' s ↔ LiftPropOn P g s := ⟨fun h ↦ hG.liftPropOn_congr h fun y hy ↦ (h₁ y hy).symm, fun h ↦ hG.liftPropOn_congr h h₁⟩ +end + theorem liftPropWithinAt_mono_of_mem (mono_of_mem : ∀ ⦃s x t⦄ ⦃f : H → H'⦄, s ∈ 𝓝[t] x → P f s x → P f t x) (h : LiftPropWithinAt P g s x) (hst : s ∈ 𝓝[t] x) : LiftPropWithinAt P g t x := by diff --git a/Mathlib/RingTheory/FractionalIdeal/Basic.lean b/Mathlib/RingTheory/FractionalIdeal/Basic.lean index c8993fe0443f5..155b25b1b2ad7 100644 --- a/Mathlib/RingTheory/FractionalIdeal/Basic.lean +++ b/Mathlib/RingTheory/FractionalIdeal/Basic.lean @@ -84,7 +84,7 @@ namespace FractionalIdeal open Set Submodule variable {R : Type*} [CommRing R] {S : Submonoid R} {P : Type*} [CommRing P] -variable [Algebra R P] [loc : IsLocalization S P] +variable [Algebra R P] /-- Map a fractional ideal `I` to a submodule by forgetting that `∃ a, a I ⊆ R`. @@ -289,15 +289,15 @@ theorem coe_zero : ↑(0 : FractionalIdeal S P) = (⊥ : Submodule R P) := theorem coeIdeal_bot : ((⊥ : Ideal R) : FractionalIdeal S P) = 0 := rfl -variable (P) +section +variable [loc : IsLocalization S P] +variable (P) in @[simp] theorem exists_mem_algebraMap_eq {x : R} {I : Ideal R} (h : S ≤ nonZeroDivisors R) : (∃ x', x' ∈ I ∧ algebraMap R P x' = algebraMap R P x) ↔ x ∈ I := ⟨fun ⟨_, hx', Eq⟩ => IsLocalization.injective _ h Eq ▸ hx', fun h => ⟨x, h, rfl⟩⟩ -variable {P} - theorem coeIdeal_injective' (h : S ≤ nonZeroDivisors R) : Function.Injective (fun (I : Ideal R) ↦ (I : FractionalIdeal S P)) := fun _ _ h' => ((coeIdeal_le_coeIdeal' S h).mp h'.le).antisymm ((coeIdeal_le_coeIdeal' S h).mp @@ -316,6 +316,8 @@ theorem coeIdeal_ne_zero' {I : Ideal R} (h : S ≤ nonZeroDivisors R) : (I : FractionalIdeal S P) ≠ 0 ↔ I ≠ (⊥ : Ideal R) := not_iff_not.mpr <| coeIdeal_eq_zero' h +end + theorem coeToSubmodule_eq_bot {I : FractionalIdeal S P} : (I : Submodule R P) = ⊥ ↔ I = 0 := ⟨fun h => coeToSubmodule_injective (by simp [h]), fun h => by simp [h]⟩ diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 071fae5bc69ae..1d74105198490 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -25,7 +25,7 @@ variable {R : Type u} {S : Type v} section Semiring variable {F : Type*} [Semiring R] [Semiring S] -variable [FunLike F R S] [rc : RingHomClass F R S] +variable [FunLike F R S] [RingHomClass F R S] variable (f : F) variable {I J : Ideal R} {K L : Ideal S} @@ -74,20 +74,24 @@ variable (f) theorem comap_ne_top (hK : K ≠ ⊤) : comap f K ≠ ⊤ := (ne_top_iff_one _).2 <| by rw [mem_comap, map_one]; exact (ne_top_iff_one _).1 hK -variable {G : Type*} [FunLike G S R] [rcg : RingHomClass G S R] +variable {G : Type*} [FunLike G S R] -theorem map_le_comap_of_inv_on (g : G) (I : Ideal R) (hf : Set.LeftInvOn g f I) : +theorem map_le_comap_of_inv_on [RingHomClass G S R] (g : G) (I : Ideal R) + (hf : Set.LeftInvOn g f I) : I.map f ≤ I.comap g := by refine Ideal.span_le.2 ?_ rintro x ⟨x, hx, rfl⟩ rw [SetLike.mem_coe, mem_comap, hf hx] exact hx -theorem comap_le_map_of_inv_on (g : G) (I : Ideal S) (hf : Set.LeftInvOn g f (f ⁻¹' I)) : - I.comap f ≤ I.map g := fun x (hx : f x ∈ I) => hf hx ▸ Ideal.mem_map_of_mem g hx +theorem comap_le_map_of_inv_on (g : G) (I : Ideal S) + (hf : Set.LeftInvOn g f (f ⁻¹' I)) : + I.comap f ≤ I.map g := + fun x (hx : f x ∈ I) => hf hx ▸ Ideal.mem_map_of_mem g hx /-- The `Ideal` version of `Set.image_subset_preimage_of_inverse`. -/ -theorem map_le_comap_of_inverse (g : G) (I : Ideal R) (h : Function.LeftInverse g f) : +theorem map_le_comap_of_inverse [RingHomClass G S R] (g : G) (I : Ideal R) + (h : Function.LeftInverse g f) : I.map f ≤ I.comap g := map_le_comap_of_inv_on _ _ _ <| h.leftInvOn _ diff --git a/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean b/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean index f37763b4d22dd..a6ef38a679ede 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Cofiltered.lean @@ -30,13 +30,12 @@ namespace TopCat section CofilteredLimit variable {J : Type v} [SmallCategory J] [IsCofiltered J] (F : J ⥤ TopCat.{max v u}) (C : Cone F) - (hC : IsLimit C) /-- Given a *compatible* collection of topological bases for the factors in a cofiltered limit which contain `Set.univ` and are closed under intersections, the induced *naive* collection of sets in the limit is, in fact, a topological basis. -/ -theorem isTopologicalBasis_cofiltered_limit (T : ∀ j, Set (Set (F.obj j))) +theorem isTopologicalBasis_cofiltered_limit (hC : IsLimit C) (T : ∀ j, Set (Set (F.obj j))) (hT : ∀ j, IsTopologicalBasis (T j)) (univ : ∀ i : J, Set.univ ∈ T i) (inter : ∀ (i) (U1 U2 : Set (F.obj i)), U1 ∈ T i → U2 ∈ T i → U1 ∩ U2 ∈ T i) (compat : ∀ (i j : J) (f : i ⟶ j) (V : Set (F.obj j)) (_hV : V ∈ T j), F.map f ⁻¹' V ∈ T i) : diff --git a/Mathlib/Topology/Covering.lean b/Mathlib/Topology/Covering.lean index 45396bca7dc2d..7a1b2b120f702 100644 --- a/Mathlib/Topology/Covering.lean +++ b/Mathlib/Topology/Covering.lean @@ -166,21 +166,21 @@ protected theorem isSeparatedMap : IsSeparatedMap f := refine Prod.ext ?_ (h₁.2.symm.trans h₂.2) rwa [t.proj_toFun e₁ he₁, t.proj_toFun e₂ he₂] -variable {A} [TopologicalSpace A] {s : Set A} (hs : IsPreconnected s) {g g₁ g₂ : A → E} +variable {A} [TopologicalSpace A] {s : Set A} {g g₁ g₂ : A → E} theorem eq_of_comp_eq [PreconnectedSpace A] (h₁ : Continuous g₁) (h₂ : Continuous g₂) (he : f ∘ g₁ = f ∘ g₂) (a : A) (ha : g₁ a = g₂ a) : g₁ = g₂ := hf.isSeparatedMap.eq_of_comp_eq hf.isLocalHomeomorph.isLocallyInjective h₁ h₂ he a ha -theorem eqOn_of_comp_eqOn (h₁ : ContinuousOn g₁ s) (h₂ : ContinuousOn g₂ s) - (he : s.EqOn (f ∘ g₁) (f ∘ g₂)) {a : A} (has : a ∈ s) (ha : g₁ a = g₂ a) : s.EqOn g₁ g₂ := - hf.isSeparatedMap.eqOn_of_comp_eqOn hf.isLocalHomeomorph.isLocallyInjective hs h₁ h₂ he has ha - theorem const_of_comp [PreconnectedSpace A] (cont : Continuous g) (he : ∀ a a', f (g a) = f (g a')) (a a') : g a = g a' := hf.isSeparatedMap.const_of_comp hf.isLocalHomeomorph.isLocallyInjective cont he a a' -theorem constOn_of_comp (cont : ContinuousOn g s) +theorem eqOn_of_comp_eqOn (hs : IsPreconnected s) (h₁ : ContinuousOn g₁ s) (h₂ : ContinuousOn g₂ s) + (he : s.EqOn (f ∘ g₁) (f ∘ g₂)) {a : A} (has : a ∈ s) (ha : g₁ a = g₂ a) : s.EqOn g₁ g₂ := + hf.isSeparatedMap.eqOn_of_comp_eqOn hf.isLocalHomeomorph.isLocallyInjective hs h₁ h₂ he has ha + +theorem constOn_of_comp (hs : IsPreconnected s) (cont : ContinuousOn g s) (he : ∀ a ∈ s, ∀ a' ∈ s, f (g a) = f (g a')) {a a'} (ha : a ∈ s) (ha' : a' ∈ s) : g a = g a' := hf.isSeparatedMap.constOn_of_comp hf.isLocalHomeomorph.isLocallyInjective hs cont he ha ha' diff --git a/Mathlib/Topology/Sets/Closeds.lean b/Mathlib/Topology/Sets/Closeds.lean index 6c7fd6bd1046a..5953aae8fe8e1 100644 --- a/Mathlib/Topology/Sets/Closeds.lean +++ b/Mathlib/Topology/Sets/Closeds.lean @@ -310,8 +310,6 @@ instance : BooleanAlgebra (Clopens α) := instance : Inhabited (Clopens α) := ⟨⊥⟩ -variable [TopologicalSpace β] - instance : SProd (Clopens α) (Clopens β) (Clopens (α × β)) where sprod s t := ⟨s ×ˢ t, s.2.prod t.2⟩ diff --git a/Mathlib/Topology/StoneCech.lean b/Mathlib/Topology/StoneCech.lean index b109e1f527f20..5d91050675249 100644 --- a/Mathlib/Topology/StoneCech.lean +++ b/Mathlib/Topology/StoneCech.lean @@ -254,8 +254,17 @@ theorem continuous_preStoneCechUnit : Continuous (preStoneCechUnit : α → PreS theorem denseRange_preStoneCechUnit : DenseRange (preStoneCechUnit : α → PreStoneCech α) := (surjective_quot_mk _).denseRange.comp denseRange_pure continuous_coinduced_rng + section Extension -variable {β : Type v} [TopologicalSpace β] [T2Space β] [CompactSpace β] +variable {β : Type v} [TopologicalSpace β] [T2Space β] + +theorem preStoneCech_hom_ext {g₁ g₂ : PreStoneCech α → β} (h₁ : Continuous g₁) (h₂ : Continuous g₂) + (h : g₁ ∘ preStoneCechUnit = g₂ ∘ preStoneCechUnit) : g₁ = g₂ := by + apply Continuous.ext_on denseRange_preStoneCechUnit h₁ h₂ + rintro x ⟨x, rfl⟩ + apply congr_fun h x + +variable [CompactSpace β] variable {g : α → β} (hg : Continuous g) lemma preStoneCechCompat {F G : Ultrafilter α} {x : α} (hF : ↑F ≤ 𝓝 x) (hG : ↑G ≤ 𝓝 x) : @@ -289,12 +298,6 @@ lemma eq_if_preStoneCechUnit_eq {a b : α} (h : preStoneCechUnit a = preStoneCec theorem continuous_preStoneCechExtend : Continuous (preStoneCechExtend hg) := continuous_quot_lift _ (continuous_ultrafilter_extend g) -theorem preStoneCech_hom_ext {g₁ g₂ : PreStoneCech α → β} (h₁ : Continuous g₁) (h₂ : Continuous g₂) - (h : g₁ ∘ preStoneCechUnit = g₂ ∘ preStoneCechUnit) : g₁ = g₂ := by - apply Continuous.ext_on denseRange_preStoneCechUnit h₁ h₂ - rintro x ⟨x, rfl⟩ - apply congr_fun h x - end Extension end PreStoneCech @@ -338,9 +341,17 @@ theorem denseRange_stoneCechUnit : DenseRange (stoneCechUnit : α → StoneCech section Extension -variable {β : Type v} [TopologicalSpace β] [T2Space β] [CompactSpace β] +variable {β : Type v} [TopologicalSpace β] [T2Space β] variable {g : α → β} (hg : Continuous g) +theorem stoneCech_hom_ext {g₁ g₂ : StoneCech α → β} (h₁ : Continuous g₁) (h₂ : Continuous g₂) + (h : g₁ ∘ stoneCechUnit = g₂ ∘ stoneCechUnit) : g₁ = g₂ := by + apply h₁.ext_on denseRange_stoneCechUnit h₂ + rintro _ ⟨x, rfl⟩ + exact congr_fun h x + +variable [CompactSpace β] + /-- The extension of a continuous function from `α` to a compact Hausdorff space `β` to the Stone-Čech compactification of `α`. This extension implements the universal property of this compactification. -/ @@ -360,12 +371,6 @@ lemma eq_if_stoneCechUnit_eq {a b : α} {f : α → β} (hcf : Continuous f) rw [← congrFun (stoneCechExtend_extends hcf), ← congrFun (stoneCechExtend_extends hcf)] exact congrArg (stoneCechExtend hcf) h -theorem stoneCech_hom_ext {g₁ g₂ : StoneCech α → β} (h₁ : Continuous g₁) (h₂ : Continuous g₂) - (h : g₁ ∘ stoneCechUnit = g₂ ∘ stoneCechUnit) : g₁ = g₂ := by - apply h₁.ext_on denseRange_stoneCechUnit h₂ - rintro _ ⟨x, rfl⟩ - exact congr_fun h x - end Extension end StoneCech diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index eadf9a7d228a9..36a5bea483aed 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -99,7 +99,7 @@ lemma cauchy_iInf_uniformSpace' {ι : Sort*} {u : ι → UniformSpace β} Cauchy (uniformSpace := ⨅ i, u i) l ↔ ∀ i, Cauchy (uniformSpace := u i) l := by simp_rw [cauchy_iff_le (uniformSpace := _), iInf_uniformity, le_iInf_iff] -lemma cauchy_comap_uniformSpace {u : UniformSpace β} {f : α → β} {l : Filter α} : +lemma cauchy_comap_uniformSpace {u : UniformSpace β} {α} {f : α → β} {l : Filter α} : Cauchy (uniformSpace := comap f u) l ↔ Cauchy (map f l) := by simp only [Cauchy, map_neBot_iff, prod_map_map_eq, map_le_iff_le_comap] rfl @@ -661,7 +661,6 @@ that this is a Cauchy sequence. If this sequence converges to some `a`, then `f namespace SequentiallyComplete variable {f : Filter α} (hf : Cauchy f) {U : ℕ → Set (α × α)} (U_mem : ∀ n, U n ∈ 𝓤 α) - (U_le : ∀ s ∈ 𝓤 α, ∃ n, U n ⊆ s) open Set Finset @@ -705,11 +704,12 @@ theorem seq_pair_mem ⦃N m n : ℕ⦄ (hm : N ≤ m) (hn : N ≤ n) : (seq hf U_mem m, seq hf U_mem n) ∈ U N := setSeq_prod_subset hf U_mem hm hn ⟨seq_mem hf U_mem m, seq_mem hf U_mem n⟩ -theorem seq_is_cauchySeq : CauchySeq <| seq hf U_mem := +theorem seq_is_cauchySeq (U_le : ∀ s ∈ 𝓤 α, ∃ n, U n ⊆ s) : CauchySeq <| seq hf U_mem := cauchySeq_of_controlled U U_le <| seq_pair_mem hf U_mem /-- If the sequence `SequentiallyComplete.seq` converges to `a`, then `f ≤ 𝓝 a`. -/ -theorem le_nhds_of_seq_tendsto_nhds ⦃a : α⦄ (ha : Tendsto (seq hf U_mem) atTop (𝓝 a)) : f ≤ 𝓝 a := +theorem le_nhds_of_seq_tendsto_nhds (U_le : ∀ s ∈ 𝓤 α, ∃ n, U n ⊆ s) + ⦃a : α⦄ (ha : Tendsto (seq hf U_mem) atTop (𝓝 a)) : f ≤ 𝓝 a := le_nhds_of_cauchy_adhp_aux (fun s hs => by rcases U_le s hs with ⟨m, hm⟩ From 97719f577ba17bd1443fd5ca5df568e9662fa435 Mon Sep 17 00:00:00 2001 From: Peter Nelson <71660771+apnelson1@users.noreply.github.com> Date: Wed, 7 Aug 2024 06:03:41 +0000 Subject: [PATCH 077/359] chore(Data/Set/Basic): predicate is equal to membership in setOf. (#15554) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We add a lemma stating that for `p : α -> Prop` we have `p = (· ∈ {a | p a})`. This lemma fills a gap that came up a couple of times when adapting `minimals` to the new `Minimal`, the alternative being an awkward `show`/`change` or `funext`. (See [this example from mathlib](https://github.com/leanprover-community/mathlib4/pull/14721#pullrequestreview-2208354111) (@j-loreaux 's first comment) and [this example](https://github.com/fpvandoorn/carleson/pull/109/files#r1705631209) from the Carleson project) --- Mathlib/Data/Set/Basic.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 740784235067c..c2505bdf4d84b 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -199,6 +199,11 @@ theorem forall_in_swap {p : α → β → Prop} : (∀ a ∈ s, ∀ (b), p a b) theorem mem_setOf {a : α} {p : α → Prop} : a ∈ { x | p x } ↔ p a := Iff.rfl +/-- This lemma is intended for use with `rw` where a membership predicate is needed, +hence the explicit argument and the equality in the reverse direction from normal. +See also `Set.mem_setOf_eq` for the reverse direction applied to an argument. -/ +theorem eq_mem_setOf (p : α → Prop) : p = (· ∈ {a | p a}) := rfl + /-- If `h : a ∈ {x | p x}` then `h.out : p x`. These are definitionally equal, but this can nevertheless be useful for various reasons, e.g. to apply further projection notation or in an argument to `simp`. -/ From 20f7e68b59c460e68e823c44d74390f992d92e28 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 7 Aug 2024 07:08:02 +0000 Subject: [PATCH 078/359] feat(RelSeries): {head,last}_map (#15387) from the Carleson project --- Mathlib/Order/RelSeries.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index 45f5e2f12ec4c..802ccaa4fc078 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -305,6 +305,10 @@ def map (p : RelSeries r) (f : r →r s) : RelSeries s where @[simp] lemma map_apply (p : RelSeries r) (f : r →r s) (i : Fin (p.length + 1)) : p.map f i = f (p i) := rfl +@[simp] lemma head_map (p : RelSeries r) (f : r →r s) : (p.map f).head = f p.head := rfl + +@[simp] lemma last_map (p : RelSeries r) (f : r →r s) : (p.map f).last = f p.last := rfl + /-- If `a₀ -r→ a₁ -r→ ... -r→ aₙ` is an `r`-series and `a` is such that `aᵢ -r→ a -r→ a_ᵢ₊₁`, then @@ -638,6 +642,12 @@ can be pushed out to a strict chain of `β` by def map (p : LTSeries α) (f : α → β) (hf : StrictMono f) : LTSeries β := LTSeries.mk p.length (f.comp p) (hf.comp p.strictMono) +@[simp] lemma head_map (p : LTSeries α) (f : α → β) (hf : StrictMono f) : + (p.map f hf).head = f p.head := rfl + +@[simp] lemma last_map (p : LTSeries α) (f : α → β) (hf : StrictMono f) : + (p.map f hf).last = f p.last := rfl + /-- For two preorders `α, β`, if `f : α → β` is surjective and strictly comonotonic, then a strict series of `β` can be pulled back to a strict chain of `α` by From cfba24713d4f27da026ce1c78d56f6237177b24b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 7 Aug 2024 07:08:03 +0000 Subject: [PATCH 079/359] chore: backports for leanprover/lean4#4814 (part 21) (#15510) --- Mathlib/Algebra/Algebra/Subalgebra/Rank.lean | 28 +++++++---- Mathlib/Algebra/Category/Grp/Injective.lean | 1 - Mathlib/Algebra/Module/CharacterModule.lean | 1 - Mathlib/Algebra/Star/NonUnitalSubalgebra.lean | 2 +- .../ContinuousFunctionalCalculus/Unital.lean | 49 +++++++++++-------- .../Analysis/CStarAlgebra/Unitization.lean | 3 +- .../Analysis/Normed/Algebra/Unitization.lean | 3 +- .../NormedSpace/Multilinear/Basic.lean | 15 ++++-- .../NormedSpace/OperatorNorm/Bilinear.lean | 4 +- .../PiTensorProduct/ProjectiveSeminorm.lean | 4 +- .../Sites/Coherent/Equivalence.lean | 17 +++---- .../AffineSpace/ContinuousAffineEquiv.lean | 4 +- .../LinearAlgebra/FiniteDimensional/Defs.lean | 8 +-- .../AdicCompletion/Functoriality.lean | 2 +- .../FractionalIdeal/Operations.lean | 31 +++++++----- Mathlib/Topology/Algebra/Algebra.lean | 12 +++-- .../Topology/Algebra/InfiniteSum/NatInt.lean | 4 +- .../Topology/Algebra/InfiniteSum/Real.lean | 4 +- Mathlib/Topology/Algebra/Module/Basic.lean | 5 +- .../Topology/ContinuousFunction/Compact.lean | 2 +- .../CompactlySupported.lean | 2 - .../ContinuousFunction/ContinuousMapZero.lean | 2 +- Mathlib/Topology/Instances/AddCircle.lean | 45 +++++++++-------- Mathlib/Topology/Instances/ENNReal.lean | 2 +- Mathlib/Topology/MetricSpace/Gluing.lean | 5 +- 25 files changed, 146 insertions(+), 109 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Rank.lean b/Mathlib/Algebra/Algebra/Subalgebra/Rank.lean index 560a5447624aa..1dac252bbf591 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Rank.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Rank.lean @@ -23,9 +23,10 @@ open FiniteDimensional namespace Subalgebra variable {R S : Type*} [CommRing R] [CommRing S] [Algebra R S] - (A B : Subalgebra R S) [Module.Free R A] [Module.Free R B] - [Module.Free A (Algebra.adjoin A (B : Set S))] - [Module.Free B (Algebra.adjoin B (A : Set S))] + (A B : Subalgebra R S) + +section +variable [Module.Free R A] [Module.Free A (Algebra.adjoin A (B : Set S))] theorem rank_sup_eq_rank_left_mul_rank_of_free : Module.rank R ↥(A ⊔ B) = Module.rank R A * Module.rank A (Algebra.adjoin A (B : Set S)) := by @@ -40,22 +41,29 @@ theorem rank_sup_eq_rank_left_mul_rank_of_free : change _ = Module.rank R ((Algebra.adjoin A (B : Set S)).restrictScalars R) rw [Algebra.restrictScalars_adjoin]; rfl -theorem rank_sup_eq_rank_right_mul_rank_of_free : - Module.rank R ↥(A ⊔ B) = Module.rank R B * Module.rank B (Algebra.adjoin B (A : Set S)) := by - rw [sup_comm, rank_sup_eq_rank_left_mul_rank_of_free] - theorem finrank_sup_eq_finrank_left_mul_finrank_of_free : finrank R ↥(A ⊔ B) = finrank R A * finrank A (Algebra.adjoin A (B : Set S)) := by simpa only [map_mul] using congr(Cardinal.toNat $(rank_sup_eq_rank_left_mul_rank_of_free A B)) +theorem finrank_left_dvd_finrank_sup_of_free : + finrank R A ∣ finrank R ↥(A ⊔ B) := ⟨_, finrank_sup_eq_finrank_left_mul_finrank_of_free A B⟩ + +end + +section +variable [Module.Free R B] [Module.Free B (Algebra.adjoin B (A : Set S))] + +theorem rank_sup_eq_rank_right_mul_rank_of_free : + Module.rank R ↥(A ⊔ B) = Module.rank R B * Module.rank B (Algebra.adjoin B (A : Set S)) := by + rw [sup_comm, rank_sup_eq_rank_left_mul_rank_of_free] + theorem finrank_sup_eq_finrank_right_mul_finrank_of_free : finrank R ↥(A ⊔ B) = finrank R B * finrank B (Algebra.adjoin B (A : Set S)) := by rw [sup_comm, finrank_sup_eq_finrank_left_mul_finrank_of_free] -theorem finrank_left_dvd_finrank_sup_of_free : - finrank R A ∣ finrank R ↥(A ⊔ B) := ⟨_, finrank_sup_eq_finrank_left_mul_finrank_of_free A B⟩ - theorem finrank_right_dvd_finrank_sup_of_free : finrank R B ∣ finrank R ↥(A ⊔ B) := ⟨_, finrank_sup_eq_finrank_right_mul_finrank_of_free A B⟩ +end + end Subalgebra diff --git a/Mathlib/Algebra/Category/Grp/Injective.lean b/Mathlib/Algebra/Category/Grp/Injective.lean index 1cc9915df9083..04d0f6d61c4a0 100644 --- a/Mathlib/Algebra/Category/Grp/Injective.lean +++ b/Mathlib/Algebra/Category/Grp/Injective.lean @@ -8,7 +8,6 @@ import Mathlib.Algebra.EuclideanDomain.Int import Mathlib.Algebra.Module.Injective import Mathlib.RingTheory.PrincipalIdealDomain import Mathlib.Topology.Instances.AddCircle -import Mathlib.Topology.Instances.Rat /-! # Injective objects in the category of abelian groups diff --git a/Mathlib/Algebra/Module/CharacterModule.lean b/Mathlib/Algebra/Module/CharacterModule.lean index 951be6418e88c..bd0b99a07a4fe 100644 --- a/Mathlib/Algebra/Module/CharacterModule.lean +++ b/Mathlib/Algebra/Module/CharacterModule.lean @@ -7,7 +7,6 @@ Authors: Jujian Zhang, Junyan Xu import Mathlib.Algebra.Category.ModuleCat.Basic import Mathlib.Algebra.Category.Grp.Injective import Mathlib.Topology.Instances.AddCircle -import Mathlib.Topology.Instances.Rat import Mathlib.LinearAlgebra.Isomorphisms /-! diff --git a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean index 83a147ec797e3..ab0f694338b18 100644 --- a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean @@ -685,7 +685,7 @@ lemma adjoin_eq_span (s : Set A) : rw [adjoin_toNonUnitalSubalgebra, NonUnitalAlgebra.adjoin_eq_span] @[simp] -lemma span_eq_toSubmodule (s : NonUnitalStarSubalgebra R A) : +lemma span_eq_toSubmodule {R} [CommSemiring R] [Module R A] (s : NonUnitalStarSubalgebra R A) : Submodule.span R (s : Set A) = s.toSubmodule := by simp [SetLike.ext'_iff, Submodule.coe_span_eq_self] diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index a824595d91b6c..7defe194cdd22 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -322,14 +322,16 @@ lemma cfc_cases (P : A → Prop) (a : A) (f : R → R) (h₀ : P 0) · rwa [cfc_apply_of_not_continuousOn _ h] variable (R) in -lemma cfc_id : cfc (id : R → R) a = a := +lemma cfc_id (ha : p a := by cfc_tac) : cfc (id : R → R) a = a := cfc_apply (id : R → R) a ▸ cfcHom_id (p := p) ha variable (R) in -lemma cfc_id' : cfc (fun x : R ↦ x) a = a := cfc_id R a +lemma cfc_id' (ha : p a := by cfc_tac) : cfc (fun x : R ↦ x) a = a := cfc_id R a /-- The **spectral mapping theorem** for the continuous functional calculus. -/ -lemma cfc_map_spectrum : spectrum R (cfc f a) = f '' spectrum R a := by +lemma cfc_map_spectrum (ha : p a := by cfc_tac) + (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) : + spectrum R (cfc f a) = f '' spectrum R a := by simp [cfc_apply f a, cfcHom_map_spectrum (p := p)] lemma cfc_const (r : R) (a : A) (ha : p a := by cfc_tac) : @@ -372,15 +374,18 @@ lemma eqOn_of_cfc_eq_cfc {f g : R → R} {a : A} (h : cfc f a = cfc g a) congrm($(this) ⟨x, hx⟩) variable {a f g} in -lemma cfc_eq_cfc_iff_eqOn : cfc f a = cfc g a ↔ (spectrum R a).EqOn f g := +lemma cfc_eq_cfc_iff_eqOn (ha : p a := by cfc_tac) + (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) + (hg : ContinuousOn g (spectrum R a) := by cfc_cont_tac) : + cfc f a = cfc g a ↔ (spectrum R a).EqOn f g := ⟨eqOn_of_cfc_eq_cfc, cfc_congr⟩ variable (R) -lemma cfc_one : cfc (1 : R → R) a = 1 := +lemma cfc_one (ha : p a := by cfc_tac) : cfc (1 : R → R) a = 1 := cfc_apply (1 : R → R) a ▸ map_one (cfcHom (show p a from ha)) -lemma cfc_const_one : cfc (fun _ : R ↦ 1) a = 1 := cfc_one R a +lemma cfc_const_one (ha : p a := by cfc_tac) : cfc (fun _ : R ↦ 1) a = 1 := cfc_one R a @[simp] lemma cfc_zero : cfc (0 : R → R) a = 0 := by @@ -488,7 +493,7 @@ lemma cfc_smul_id {S : Type*} [SMul S R] [ContinuousConstSMul S R] lemma cfc_const_mul_id (r : R) (a : A) (ha : p a := by cfc_tac) : cfc (r * ·) a = r • a := cfc_smul_id r a -lemma cfc_star_id : cfc (star · : R → R) a = star a := by +lemma cfc_star_id (ha : p a := by cfc_tac) : cfc (star · : R → R) a = star a := by rw [cfc_star .., cfc_id' ..] section Polynomial @@ -518,6 +523,19 @@ lemma cfc_polynomial (q : R[X]) (a : A) (ha : p a := by cfc_tac) : end Polynomial +lemma CFC.eq_algebraMap_of_spectrum_subset_singleton (r : R) (h_spec : spectrum R a ⊆ {r}) + (ha : p a := by cfc_tac) : a = algebraMap R A r := by + simpa [cfc_id R a, cfc_const r a] using + cfc_congr (f := id) (g := fun _ : R ↦ r) (a := a) fun x hx ↦ by simpa using h_spec hx + +lemma CFC.eq_zero_of_spectrum_subset_zero (h_spec : spectrum R a ⊆ {0}) (ha : p a := by cfc_tac) : + a = 0 := by + simpa using eq_algebraMap_of_spectrum_subset_singleton a 0 h_spec + +lemma CFC.eq_one_of_spectrum_subset_one (h_spec : spectrum R a ⊆ {1}) (ha : p a := by cfc_tac) : + a = 1 := by + simpa using eq_algebraMap_of_spectrum_subset_singleton a 1 h_spec + variable [UniqueContinuousFunctionalCalculus R A] lemma cfc_comp (g f : R → R) (a : A) (ha : p a := by cfc_tac) @@ -569,19 +587,6 @@ lemma cfc_comp_polynomial (q : R[X]) (f : R → R) (a : A) cfc (f <| q.eval ·) a = cfc f (aeval a q) := by rw [cfc_comp' .., cfc_polynomial ..] -lemma CFC.eq_algebraMap_of_spectrum_subset_singleton (r : R) (h_spec : spectrum R a ⊆ {r}) - (ha : p a := by cfc_tac) : a = algebraMap R A r := by - simpa [cfc_id R a, cfc_const r a] using - cfc_congr (f := id) (g := fun _ : R ↦ r) (a := a) fun x hx ↦ by simpa using h_spec hx - -lemma CFC.eq_zero_of_spectrum_subset_zero (h_spec : spectrum R a ⊆ {0}) (ha : p a := by cfc_tac) : - a = 0 := by - simpa using eq_algebraMap_of_spectrum_subset_singleton a 0 h_spec - -lemma CFC.eq_one_of_spectrum_subset_one (h_spec : spectrum R a ⊆ {1}) (ha : p a := by cfc_tac) : - a = 1 := by - simpa using eq_algebraMap_of_spectrum_subset_singleton a 1 h_spec - lemma CFC.spectrum_algebraMap_subset (r : R) : spectrum R (algebraMap R A r) ⊆ {r} := by rw [← cfc_const r 0 (cfc_predicate_zero R), cfc_map_spectrum (fun _ ↦ r) 0 (cfc_predicate_zero R)] @@ -637,7 +642,7 @@ end Basic section Inv variable {R A : Type*} {p : A → Prop} [Semifield R] [StarRing R] [MetricSpace R] -variable [TopologicalSemiring R] [ContinuousStar R] [HasContinuousInv₀ R] [TopologicalSpace A] +variable [TopologicalSemiring R] [ContinuousStar R] [TopologicalSpace A] variable [Ring A] [StarRing A] [Algebra R A] [ContinuousFunctionalCalculus R p] variable (f : R → R) (a : A) @@ -648,6 +653,8 @@ lemma isUnit_cfc_iff (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) alias ⟨_, isUnit_cfc⟩ := isUnit_cfc_iff +variable [HasContinuousInv₀ R] + /-- Bundle `cfc f a` into a unit given a proof that `f` is nonzero on the spectrum of `a`. -/ @[simps] noncomputable def cfcUnits (hf' : ∀ x ∈ spectrum R a, f x ≠ 0) diff --git a/Mathlib/Analysis/CStarAlgebra/Unitization.lean b/Mathlib/Analysis/CStarAlgebra/Unitization.lean index 85446f1873905..0ccc38b7262f0 100644 --- a/Mathlib/Analysis/CStarAlgebra/Unitization.lean +++ b/Mathlib/Analysis/CStarAlgebra/Unitization.lean @@ -76,7 +76,7 @@ instance CStarRing.instRegularNormedAlgebra : RegularNormedAlgebra 𝕜 E where section CStarProperty -variable [StarRing 𝕜] [CStarRing 𝕜] [StarModule 𝕜 E] +variable [StarRing 𝕜] [StarModule 𝕜 E] variable {E} /-- This is the key lemma used to establish the instance `Unitization.instCStarRing` @@ -122,6 +122,7 @@ theorem Unitization.norm_splitMul_snd_sq (x : Unitization 𝕜 E) : simp only [smul_smul, smul_mul_assoc, ← add_assoc, ← mul_assoc, mul_smul_comm] variable {𝕜} +variable [CStarRing 𝕜] /-- The norm on `Unitization 𝕜 E` satisfies the C⋆-property -/ instance Unitization.instCStarRing : CStarRing (Unitization 𝕜 E) where diff --git a/Mathlib/Analysis/Normed/Algebra/Unitization.lean b/Mathlib/Analysis/Normed/Algebra/Unitization.lean index 68b0fbc2e2b4a..789e764e8c4e0 100644 --- a/Mathlib/Analysis/Normed/Algebra/Unitization.lean +++ b/Mathlib/Analysis/Normed/Algebra/Unitization.lean @@ -208,7 +208,8 @@ def uniformEquivProd : (Unitization 𝕜 A) ≃ᵤ (𝕜 × A) := instance instBornology : Bornology (Unitization 𝕜 A) := Bornology.induced <| addEquiv 𝕜 A -theorem uniformEmbedding_addEquiv : UniformEmbedding (addEquiv 𝕜 A) where +theorem uniformEmbedding_addEquiv {𝕜} [NontriviallyNormedField 𝕜] : + UniformEmbedding (addEquiv 𝕜 A) where comap_uniformity := rfl inj := (addEquiv 𝕜 A).injective diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index e1c5202a1d788..c3efa1e0e4436 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -90,10 +90,9 @@ theorem ContinuousMultilinearMap.continuous_eval {𝕜 ι : Type*} {E : ι → T section Seminorm variable {𝕜 : Type u} {ι : Type v} {ι' : Type v'} {E : ι → Type wE} {E₁ : ι → Type wE₁} - {E' : ι' → Type wE'} {G : Type wG} {G' : Type wG'} [Fintype ι] + {E' : ι' → Type wE'} {G : Type wG} {G' : Type wG'} [Fintype ι'] [NontriviallyNormedField 𝕜] [∀ i, SeminormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] [∀ i, SeminormedAddCommGroup (E₁ i)] [∀ i, NormedSpace 𝕜 (E₁ i)] - [∀ i, SeminormedAddCommGroup (E' i)] [∀ i, NormedSpace 𝕜 (E' i)] [SeminormedAddCommGroup G] [NormedSpace 𝕜 G] [SeminormedAddCommGroup G'] [NormedSpace 𝕜 G'] /-! @@ -121,6 +120,8 @@ lemma norm_map_coord_zero (hf : Continuous f) {m : ∀ i, E i} {i : ι} (hi : (forall_update_iff m fun i a ↦ Inseparable a (m i)).2 ⟨hi.symm, fun _ _ ↦ rfl⟩ simpa only [map_update_zero] using this.symm.map hf +variable [Fintype ι] + theorem bound_of_shell_of_norm_map_coord_zero (hf₀ : ∀ {m i}, ‖m i‖ = 0 → ‖f m‖ = 0) {ε : ι → ℝ} {C : ℝ} (hε : ∀ i, 0 < ε i) {c : ι → 𝕜} (hc : ∀ i, 1 < ‖c i‖) (hf : ∀ m : ∀ i, E i, (∀ i, ε i / ‖c i‖ ≤ ‖m i‖) → (∀ i, ‖m i‖ < ε i) → ‖f m‖ ≤ C * ∏ i, ‖m i‖) @@ -297,7 +298,7 @@ defines a normed space structure on `ContinuousMultilinearMap 𝕜 E G`. namespace ContinuousMultilinearMap -variable (c : 𝕜) (f g : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) +variable [Fintype ι] (c : 𝕜) (f g : ContinuousMultilinearMap 𝕜 E G) (m : ∀ i, E i) theorem bound : ∃ C : ℝ, 0 < C ∧ ∀ m, ‖f m‖ ≤ C * ∏ i, ‖m i‖ := f.toMultilinearMap.exists_bound_of_continuous f.2 @@ -687,6 +688,8 @@ theorem norm_image_sub_le (m₁ m₂ : ∀ i, E i) : end ContinuousMultilinearMap +variable [Fintype ι] + /-- If a continuous multilinear map is constructed from a multilinear map via the constructor `mkContinuous`, then its norm is bounded by the bound given to the constructor if it is nonnegative. -/ @@ -990,6 +993,8 @@ theorem mkContinuousLinear_norm_le (f : G →ₗ[𝕜] MultilinearMap 𝕜 E G') (H : ∀ x m, ‖f x m‖ ≤ C * ‖x‖ * ∏ i, ‖m i‖) : ‖mkContinuousLinear f C H‖ ≤ C := (mkContinuousLinear_norm_le' f C H).trans_eq (max_eq_left hC) +variable [∀ i, SeminormedAddCommGroup (E' i)] [∀ i, NormedSpace 𝕜 (E' i)] + /-- Given a map `f : MultilinearMap 𝕜 E (MultilinearMap 𝕜 E' G)` and an estimate `H : ∀ m m', ‖f m m'‖ ≤ C * ∏ i, ‖m i‖ * ∏ i, ‖m' i‖`, upgrade all `MultilinearMap`s in the type to `ContinuousMultilinearMap`s. -/ @@ -1157,10 +1162,10 @@ noncomputable def compContinuousLinearMapContinuousMultilinear : ContinuousMultilinearMap 𝕜 (fun i ↦ E i →L[𝕜] E₁ i) ((ContinuousMultilinearMap 𝕜 E₁ G) →L[𝕜] ContinuousMultilinearMap 𝕜 E G) := @MultilinearMap.mkContinuous 𝕜 ι (fun i ↦ E i →L[𝕜] E₁ i) - ((ContinuousMultilinearMap 𝕜 E₁ G) →L[𝕜] ContinuousMultilinearMap 𝕜 E G) _ _ + ((ContinuousMultilinearMap 𝕜 E₁ G) →L[𝕜] ContinuousMultilinearMap 𝕜 E G) _ (fun _ ↦ ContinuousLinearMap.toSeminormedAddCommGroup) (fun _ ↦ ContinuousLinearMap.toNormedSpace) _ _ - (compContinuousLinearMapMultilinear 𝕜 E E₁ G) 1 + (compContinuousLinearMapMultilinear 𝕜 E E₁ G) _ 1 fun f ↦ by simpa using norm_compContinuousLinearMapL_le G f variable {𝕜 E E₁} diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean index 8be1705a830fc..07c05d0a62e73 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Bilinear.lean @@ -84,14 +84,14 @@ end ContinuousLinearMap namespace LinearMap -variable [RingHomIsometric σ₂₃] - lemma norm_mkContinuous₂_aux (f : E →ₛₗ[σ₁₃] F →ₛₗ[σ₂₃] G) (C : ℝ) (h : ∀ x y, ‖f x y‖ ≤ C * ‖x‖ * ‖y‖) (x : E) : ‖(f x).mkContinuous (C * ‖x‖) (h x)‖ ≤ max C 0 * ‖x‖ := (mkContinuous_norm_le' (f x) (h x)).trans_eq <| by rw [max_mul_of_nonneg _ _ (norm_nonneg x), zero_mul] +variable [RingHomIsometric σ₂₃] + /-- Create a bilinear map (represented as a map `E →L[𝕜] F →L[𝕜] G`) from the corresponding linear map and existence of a bound on the norm of the image. The linear map can be constructed using `LinearMap.mk₂`. diff --git a/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean b/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean index 74ef28683a937..ec8cbeca4457f 100644 --- a/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean +++ b/Mathlib/Analysis/NormedSpace/PiTensorProduct/ProjectiveSeminorm.lean @@ -39,7 +39,7 @@ universe uι u𝕜 uE uF variable {ι : Type uι} [Fintype ι] variable {𝕜 : Type u𝕜} [NontriviallyNormedField 𝕜] -variable {E : ι → Type uE} [∀ i, SeminormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] +variable {E : ι → Type uE} [∀ i, SeminormedAddCommGroup (E i)] variable {F : Type uF} [SeminormedAddCommGroup F] [NormedSpace 𝕜 F] open scoped TensorProduct @@ -81,6 +81,8 @@ theorem projectiveSeminormAux_smul (p : FreeAddMonoid (𝕜 × Π i, E i)) (a : simp only [Function.comp_apply, norm_mul, smul_eq_mul] rw [mul_assoc] +variable [∀ i, NormedSpace 𝕜 (E i)] + theorem bddBelow_projectiveSemiNormAux (x : ⨂[𝕜] i, E i) : BddBelow (Set.range (fun (p : lifts x) ↦ projectiveSeminormAux p.1)) := by existsi 0 diff --git a/Mathlib/CategoryTheory/Sites/Coherent/Equivalence.lean b/Mathlib/CategoryTheory/Sites/Coherent/Equivalence.lean index f5e91a85cbb1e..e66a91cfd4c63 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/Equivalence.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/Equivalence.lean @@ -22,19 +22,18 @@ open GrothendieckTopology namespace Equivalence variable {D : Type*} [Category D] -variable (e : C ≌ D) section Coherent variable [Precoherent C] /-- `Precoherent` is preserved by equivalence of categories. -/ -theorem precoherent : Precoherent D := e.inverse.reflects_precoherent +theorem precoherent (e : C ≌ D) : Precoherent D := e.inverse.reflects_precoherent instance [EssentiallySmall C] : Precoherent (SmallModel C) := (equivSmallModel C).precoherent -instance : haveI := precoherent e +instance (e : C ≌ D) : haveI := precoherent e e.inverse.IsDenseSubsite (coherentTopology D) (coherentTopology C) where functorPushforward_mem_iff := by rw [coherentTopology.eq_induced e.inverse] @@ -46,7 +45,7 @@ variable (A : Type*) [Category A] Equivalent precoherent categories give equivalent coherent toposes. -/ @[simps!] -def sheafCongrPrecoherent : haveI := e.precoherent +def sheafCongrPrecoherent (e : C ≌ D) : haveI := e.precoherent Sheaf (coherentTopology C) A ≌ Sheaf (coherentTopology D) A := e.sheafCongr _ _ _ open Presheaf @@ -54,7 +53,7 @@ open Presheaf /-- The coherent sheaf condition can be checked after precomposing with the equivalence. -/ -theorem precoherent_isSheaf_iff (F : Cᵒᵖ ⥤ A) : haveI := e.precoherent +theorem precoherent_isSheaf_iff (e : C ≌ D) (F : Cᵒᵖ ⥤ A) : haveI := e.precoherent IsSheaf (coherentTopology C) F ↔ IsSheaf (coherentTopology D) (e.inverse.op ⋙ F) := by refine ⟨fun hF ↦ ((e.sheafCongrPrecoherent A).functor.obj ⟨F, hF⟩).cond, fun hF ↦ ?_⟩ rw [isSheaf_of_iso_iff (P' := e.functor.op ⋙ e.inverse.op ⋙ F)] @@ -77,12 +76,12 @@ section Regular variable [Preregular C] /-- `Preregular` is preserved by equivalence of categories. -/ -theorem preregular : Preregular D := e.inverse.reflects_preregular +theorem preregular (e : C ≌ D) : Preregular D := e.inverse.reflects_preregular instance [EssentiallySmall C] : Preregular (SmallModel C) := (equivSmallModel C).preregular -instance : haveI := preregular e +instance (e : C ≌ D) : haveI := preregular e e.inverse.IsDenseSubsite (regularTopology D) (regularTopology C) where functorPushforward_mem_iff := by rw [regularTopology.eq_induced e.inverse] @@ -94,7 +93,7 @@ variable (A : Type*) [Category A] Equivalent preregular categories give equivalent regular toposes. -/ @[simps!] -def sheafCongrPreregular : haveI := e.preregular +def sheafCongrPreregular (e : C ≌ D) : haveI := e.preregular Sheaf (regularTopology C) A ≌ Sheaf (regularTopology D) A := e.sheafCongr _ _ _ open Presheaf @@ -102,7 +101,7 @@ open Presheaf /-- The regular sheaf condition can be checked after precomposing with the equivalence. -/ -theorem preregular_isSheaf_iff (F : Cᵒᵖ ⥤ A) : haveI := e.preregular +theorem preregular_isSheaf_iff (e : C ≌ D) (F : Cᵒᵖ ⥤ A) : haveI := e.preregular IsSheaf (regularTopology C) F ↔ IsSheaf (regularTopology D) (e.inverse.op ⋙ F) := by refine ⟨fun hF ↦ ((e.sheafCongrPreregular A).functor.obj ⟨F, hF⟩).cond, fun hF ↦ ?_⟩ rw [isSheaf_of_iso_iff (P' := e.functor.op ⋙ e.inverse.op ⋙ F)] diff --git a/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean index ab0bbf619ae27..0360e2114a8d9 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/ContinuousAffineEquiv.lean @@ -49,8 +49,8 @@ variable {k P₁ P₂ P₃ P₄ V₁ V₂ V₃ V₄ : Type*} [Ring k] [AddCommGroup V₂] [Module k V₂] [AddTorsor V₂ P₂] [AddCommGroup V₃] [Module k V₃] [AddTorsor V₃ P₃] [AddCommGroup V₄] [Module k V₄] [AddTorsor V₄ P₄] - [TopologicalSpace P₁] [AddCommMonoid P₁] [Module k P₁] - [TopologicalSpace P₂] [AddCommMonoid P₂] [Module k P₂] + [TopologicalSpace P₁] + [TopologicalSpace P₂] [TopologicalSpace P₃] [TopologicalSpace P₄] namespace ContinuousAffineEquiv diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index f1fb7c9f708a9..ed9305efcd20a 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -473,17 +473,19 @@ theorem eq_of_le_of_finrank_eq {S₁ S₂ : Submodule K V} [FiniteDimensional K section Subalgebra variable {K L : Type*} [Field K] [Ring L] [Algebra K L] {F E : Subalgebra K L} - [hfin : FiniteDimensional K E] (h_le : F ≤ E) + [hfin : FiniteDimensional K E] /-- If a subalgebra is contained in a finite-dimensional subalgebra with the same or smaller dimension, they are equal. -/ -theorem _root_.Subalgebra.eq_of_le_of_finrank_le (h_finrank : finrank K E ≤ finrank K F) : F = E := +theorem _root_.Subalgebra.eq_of_le_of_finrank_le (h_le : F ≤ E) + (h_finrank : finrank K E ≤ finrank K F) : F = E := haveI : Module.Finite K (Subalgebra.toSubmodule E) := hfin Subalgebra.toSubmodule_injective <| FiniteDimensional.eq_of_le_of_finrank_le h_le h_finrank /-- If a subalgebra is contained in a finite-dimensional subalgebra with the same dimension, they are equal. -/ -theorem _root_.Subalgebra.eq_of_le_of_finrank_eq (h_finrank : finrank K F = finrank K E) : F = E := +theorem _root_.Subalgebra.eq_of_le_of_finrank_eq (h_le : F ≤ E) + (h_finrank : finrank K F = finrank K E) : F = E := Subalgebra.eq_of_le_of_finrank_le h_le h_finrank.ge end Subalgebra diff --git a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean index cbe6c2575aaa0..8455083320dad 100644 --- a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean +++ b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean @@ -140,7 +140,7 @@ theorem map_val_apply (f : M →ₗ[R] N) {n : ℕ} (x : AdicCompletion I M) : rfl /-- Equality of maps out of an adic completion can be checked on Cauchy sequences. -/ -theorem map_ext {f g : AdicCompletion I M → N} +theorem map_ext {N} {f g : AdicCompletion I M → N} (h : ∀ (a : AdicCauchySequence I M), f (AdicCompletion.mk I M a) = g (AdicCompletion.mk I M a)) : f = g := by diff --git a/Mathlib/RingTheory/FractionalIdeal/Operations.lean b/Mathlib/RingTheory/FractionalIdeal/Operations.lean index 1676d72c3620b..6836737387416 100644 --- a/Mathlib/RingTheory/FractionalIdeal/Operations.lean +++ b/Mathlib/RingTheory/FractionalIdeal/Operations.lean @@ -38,12 +38,12 @@ namespace FractionalIdeal open Set Submodule variable {R : Type*} [CommRing R] {S : Submonoid R} {P : Type*} [CommRing P] -variable [Algebra R P] [loc : IsLocalization S P] +variable [Algebra R P] section -variable {P' : Type*} [CommRing P'] [Algebra R P'] [loc' : IsLocalization S P'] -variable {P'' : Type*} [CommRing P''] [Algebra R P''] [loc'' : IsLocalization S P''] +variable {P' : Type*} [CommRing P'] [Algebra R P'] +variable {P'' : Type*} [CommRing P''] [Algebra R P''] theorem _root_.IsFractional.map (g : P →ₐ[R] P') {I : Submodule R P} : IsFractional S I → IsFractional S (Submodule.map g.toLinearMap I) @@ -165,7 +165,8 @@ theorem isFractional_span_iff {s : Set P} : rw [smul_comm] exact isInteger_smul hx⟩⟩ -theorem isFractional_of_fg {I : Submodule R P} (hI : I.FG) : IsFractional S I := by +theorem isFractional_of_fg [IsLocalization S P] {I : Submodule R P} (hI : I.FG) : + IsFractional S I := by rcases hI with ⟨I, rfl⟩ rcases exist_integer_multiples_of_finset S I with ⟨⟨s, hs1⟩, hs⟩ rw [isFractional_span_iff] @@ -196,6 +197,8 @@ theorem _root_.Ideal.fg_of_isUnit (inj : Function.Injective (algebraMap R P)) (I variable (S P P') +variable [IsLocalization S P] [IsLocalization S P'] + /-- `canonicalEquiv f f'` is the canonical equivalence between the fractional ideals in `P` and in `P'`, which are both localizations of `R` at `S`. -/ noncomputable irreducible_def canonicalEquiv : FractionalIdeal S P ≃+* FractionalIdeal S P' := @@ -329,7 +332,7 @@ is a field because `R` is a domain. -/ variable {R₁ : Type*} [CommRing R₁] {K : Type*} [Field K] -variable [Algebra R₁ K] [frac : IsFractionRing R₁ K] +variable [Algebra R₁ K] instance : Nontrivial (FractionalIdeal R₁⁰ K) := ⟨⟨0, 1, fun h => @@ -344,7 +347,7 @@ theorem ne_zero_of_mul_eq_one (I J : FractionalIdeal R₁⁰ K) (h : I * J = 1) convert h simp [hI]) -variable [IsDomain R₁] +variable [IsFractionRing R₁ K] [IsDomain R₁] theorem _root_.IsFractional.div_of_nonzero {I J : Submodule R₁ K} : IsFractional R₁⁰ I → IsFractional R₁⁰ J → J ≠ 0 → IsFractional R₁⁰ (I / J) @@ -541,6 +544,8 @@ theorem spanFinset_ne_zero {ι : Type*} {s : Finset ι} {f : ι → K} : open Submodule.IsPrincipal +variable [IsLocalization S P] + theorem isFractional_span_singleton (x : P) : IsFractional S (span R {x} : Submodule R P) := let ⟨a, ha⟩ := exists_integer_multiple S x isFractional_span_iff.mpr ⟨a, a.2, fun _ hx' => (Set.mem_singleton_iff.mp hx').symm ▸ ha⟩ @@ -814,7 +819,7 @@ theorem num_le (I : FractionalIdeal S P) : end PrincipalIdeal variable {R₁ : Type*} [CommRing R₁] -variable {K : Type*} [Field K] [Algebra R₁ K] [frac : IsFractionRing R₁ K] +variable {K : Type*} [Field K] [Algebra R₁ K] attribute [local instance] Classical.propDecidable @@ -835,7 +840,7 @@ theorem isNoetherian_coeIdeal [IsNoetherianRing R₁] (I : Ideal R₁) : obtain ⟨J, rfl⟩ := le_one_iff_exists_coeIdeal.mp (le_trans hJ coeIdeal_le_one) exact (IsNoetherian.noetherian J).map _ -variable [IsDomain R₁] +variable [IsFractionRing R₁ K] [IsDomain R₁] theorem isNoetherian_spanSingleton_inv_to_map_mul (x : R₁) {I : FractionalIdeal R₁⁰ K} (hI : IsNoetherian R₁ I) : @@ -864,10 +869,10 @@ theorem isNoetherian [IsNoetherianRing R₁] (I : FractionalIdeal R₁⁰ K) : I section Adjoin variable (S) -variable (x : P) (hx : IsIntegral R x) +variable [IsLocalization S P] (x : P) /-- `A[x]` is a fractional ideal for every integral `x`. -/ -theorem isFractional_adjoin_integral : +theorem isFractional_adjoin_integral (hx : IsIntegral R x) : IsFractional S (Subalgebra.toSubmodule (Algebra.adjoin R ({x} : Set P))) := isFractional_of_fg hx.fg_adjoin_singleton @@ -875,16 +880,16 @@ theorem isFractional_adjoin_integral : where `hx` is a proof that `x : P` is integral over `R`. -/ -- Porting note: `@[simps]` generated a `Subtype.val` coercion instead of a -- `FractionalIdeal.coeToSubmodule` coercion -def adjoinIntegral : FractionalIdeal S P := +def adjoinIntegral (hx : IsIntegral R x) : FractionalIdeal S P := ⟨_, isFractional_adjoin_integral S x hx⟩ @[simp] -theorem adjoinIntegral_coe : +theorem adjoinIntegral_coe (hx : IsIntegral R x) : (adjoinIntegral S x hx : Submodule R P) = (Subalgebra.toSubmodule (Algebra.adjoin R ({x} : Set P))) := rfl -theorem mem_adjoinIntegral_self : x ∈ adjoinIntegral S x hx := +theorem mem_adjoinIntegral_self (hx : IsIntegral R x) : x ∈ adjoinIntegral S x hx := Algebra.subset_adjoin (Set.mem_singleton x) end Adjoin diff --git a/Mathlib/Topology/Algebra/Algebra.lean b/Mathlib/Topology/Algebra/Algebra.lean index 5870483dbea15..b4367b4042dc1 100644 --- a/Mathlib/Topology/Algebra/Algebra.lean +++ b/Mathlib/Topology/Algebra/Algebra.lean @@ -78,8 +78,8 @@ section TopologicalAlgebra section -variable (R : Type*) [CommSemiring R] [TopologicalSpace R] [TopologicalSemiring R] - (A : Type*) [Semiring A] [TopologicalSpace A] [TopologicalSemiring A] +variable (R : Type*) [CommSemiring R] + (A : Type*) [Semiring A] /-- Continuous algebra homomorphisms between algebras. We only put the type classes that are necessary for the definition, although in applications `M` and `B` will be topological algebras @@ -98,7 +98,7 @@ namespace ContinuousAlgHom section Semiring variable {R} {A} -variable [TopologicalSpace R] [TopologicalSemiring R] +variable [TopologicalSpace A] variable {B : Type*} [Semiring B] [TopologicalSpace B] [Algebra R A] [Algebra R B] @@ -207,6 +207,8 @@ theorem ext_on [T2Space B] {s : Set A} (hs : Dense (Algebra.adjoin R s : Set A)) {f g : A →A[R] B} (h : Set.EqOn f g s) : f = g := ext fun x => eqOn_closure_adjoin h (hs x) +variable [TopologicalSemiring A] + /-- The topological closure of a subalgebra -/ def _root_.Subalgebra.topologicalClosure (s : Subalgebra R A) : Subalgebra R A where toSubsemiring := s.toSubsemiring.topologicalClosure @@ -244,6 +246,7 @@ end Semiring section id +variable [TopologicalSpace A] variable [Algebra R A] /-- The identity map as a continuous algebra homomorphism. -/ @@ -274,6 +277,7 @@ end id section comp variable {R} {A} +variable [TopologicalSpace A] variable {B : Type*} [Semiring B] [TopologicalSpace B] [Algebra R A] [Algebra R B] {C : Type*} [Semiring C] [Algebra R C] [TopologicalSpace C] @@ -331,6 +335,7 @@ end comp section prod variable {R} {A} +variable [TopologicalSpace A] variable {B : Type*} [Semiring B] [TopologicalSpace B] [Algebra R A] [Algebra R B] {C : Type*} [Semiring C] [Algebra R C] [TopologicalSpace C] @@ -431,6 +436,7 @@ end prod section subalgebra variable {R A} +variable [TopologicalSpace A] variable {B : Type*} [Semiring B] [TopologicalSpace B] [Algebra R A] [Algebra R B] /-- Restrict codomain of a continuous algebra morphism. -/ diff --git a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean index be2896bac4115..09d1d25020184 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/NatInt.lean @@ -108,7 +108,7 @@ end Multipliable section tprod -variable [T2Space M] {α β γ : Type*} +variable {α β γ : Type*} section Encodable @@ -175,7 +175,7 @@ end Countable section ContinuousMul -variable [ContinuousMul M] +variable [T2Space M] [ContinuousMul M] @[to_additive] theorem prod_mul_tprod_nat_mul' diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Real.lean b/Mathlib/Topology/Algebra/InfiniteSum/Real.lean index 12800c8af15fb..1ba9e35f1b1fc 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Real.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Real.lean @@ -64,7 +64,7 @@ theorem summable_iff_not_tendsto_nat_atTop_of_nonneg {f : ℕ → ℝ} (hf : ∀ Summable f ↔ ¬Tendsto (fun n : ℕ => ∑ i ∈ Finset.range n, f i) atTop atTop := by rw [← not_iff_not, Classical.not_not, not_summable_iff_tendsto_nat_atTop_of_nonneg hf] -theorem summable_sigma_of_nonneg {β : α → Type*} {f : (Σ x, β x) → ℝ} (hf : ∀ x, 0 ≤ f x) : +theorem summable_sigma_of_nonneg {α} {β : α → Type*} {f : (Σ x, β x) → ℝ} (hf : ∀ x, 0 ≤ f x) : Summable f ↔ (∀ x, Summable fun y => f ⟨x, y⟩) ∧ Summable fun x => ∑' y, f ⟨x, y⟩ := by lift f to (Σx, β x) → ℝ≥0 using hf exact mod_cast NNReal.summable_sigma @@ -74,7 +74,7 @@ lemma summable_partition {α β : Type*} {f : β → ℝ} (hf : 0 ≤ f) {s : α (∀ j, Summable fun i : s j ↦ f i) ∧ Summable fun j ↦ ∑' i : s j, f i := by simpa only [← (Set.sigmaEquiv s hs).summable_iff] using summable_sigma_of_nonneg (fun _ ↦ hf _) -theorem summable_prod_of_nonneg {f : (α × β) → ℝ} (hf : 0 ≤ f) : +theorem summable_prod_of_nonneg {α β} {f : (α × β) → ℝ} (hf : 0 ≤ f) : Summable f ↔ (∀ x, Summable fun y ↦ f (x, y)) ∧ Summable fun x ↦ ∑' y, f (x, y) := (Equiv.sigmaEquivProd _ _).summable_iff.symm.trans <| summable_sigma_of_nonneg fun _ ↦ hf _ diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index e50b1de578992..f4ca295850c70 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -2076,9 +2076,6 @@ section /-! The next theorems cover the identification between `M ≃L[𝕜] M`and the group of units of the ring `M →L[R] M`. -/ - -variable [TopologicalAddGroup M] - /-- An invertible continuous linear map `f` determines a continuous equivalence from `M` to itself. -/ def ofUnit (f : (M →L[R] M)ˣ) : M ≃L[R] M where @@ -2273,7 +2270,7 @@ end section variable [Ring R] -variable [AddCommGroup M] [TopologicalAddGroup M] [Module R M] +variable [AddCommGroup M] [Module R M] variable [AddCommGroup M₂] [Module R M₂] @[simp] diff --git a/Mathlib/Topology/ContinuousFunction/Compact.lean b/Mathlib/Topology/ContinuousFunction/Compact.lean index 24fddb06c48a9..5257cc87a452a 100644 --- a/Mathlib/Topology/ContinuousFunction/Compact.lean +++ b/Mathlib/Topology/ContinuousFunction/Compact.lean @@ -429,7 +429,7 @@ of `C(X, E)` (i.e. locally uniform convergence). -/ open TopologicalSpace -variable {X : Type*} [TopologicalSpace X] [T2Space X] [LocallyCompactSpace X] +variable {X : Type*} [TopologicalSpace X] [LocallyCompactSpace X] variable {E : Type*} [NormedAddCommGroup E] [CompleteSpace E] theorem summable_of_locally_summable_norm {ι : Type*} {F : ι → C(X, E)} diff --git a/Mathlib/Topology/ContinuousFunction/CompactlySupported.lean b/Mathlib/Topology/ContinuousFunction/CompactlySupported.lean index a5e36fed1b6b0..d8d72123235af 100644 --- a/Mathlib/Topology/ContinuousFunction/CompactlySupported.lean +++ b/Mathlib/Topology/ContinuousFunction/CompactlySupported.lean @@ -437,8 +437,6 @@ theorem coe_comp_to_continuous_fun (f : C_c(γ, δ)) (g : β →co γ) : ((f.com theorem comp_id (f : C_c(γ, δ)) : f.comp (CocompactMap.id γ) = f := ext fun _ => rfl -variable [T2Space β] - @[simp] theorem comp_assoc (f : C_c(γ, δ)) (g : β →co γ) (h : α →co β) : (f.comp g).comp h = f.comp (g.comp h) := diff --git a/Mathlib/Topology/ContinuousFunction/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousFunction/ContinuousMapZero.lean index c34d6f692c3d4..f02fe2b8cc5f7 100644 --- a/Mathlib/Topology/ContinuousFunction/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousFunction/ContinuousMapZero.lean @@ -57,7 +57,7 @@ lemma ext {f g : C(X, R)₀} (h : ∀ x, f x = g x) : f = g := DFunLike.ext f g @[simp] lemma coe_mk {f : C(X, R)} {h0 : f 0 = 0} : ⇑(mk f h0) = f := rfl -lemma toContinuousMap_injective [Zero R] : Injective ((↑) : C(X, R)₀ → C(X, R)) := +lemma toContinuousMap_injective : Injective ((↑) : C(X, R)₀ → C(X, R)) := fun _ _ h ↦ congr(.mk $(h) _) lemma range_toContinuousMap : range ((↑) : C(X, R)₀ → C(X, R)) = {f : C(X, R) | f 0 = 0} := diff --git a/Mathlib/Topology/Instances/AddCircle.lean b/Mathlib/Topology/Instances/AddCircle.lean index 9f33cbe774d18..21303dc60bf02 100644 --- a/Mathlib/Topology/Instances/AddCircle.lean +++ b/Mathlib/Topology/Instances/AddCircle.lean @@ -85,23 +85,24 @@ theorem continuous_left_toIocMod : ContinuousWithinAt (toIocMod hp a) (Iic x) x (continuous_sub_left _).continuousAt.comp_continuousWithinAt <| (continuous_right_toIcoMod _ _ _).comp continuous_neg.continuousWithinAt fun y => neg_le_neg -variable {x} (hx : (x : 𝕜 ⧸ zmultiples p) ≠ a) +variable {x} -theorem toIcoMod_eventuallyEq_toIocMod : toIcoMod hp a =ᶠ[𝓝 x] toIocMod hp a := +theorem toIcoMod_eventuallyEq_toIocMod (hx : (x : 𝕜 ⧸ zmultiples p) ≠ a) : + toIcoMod hp a =ᶠ[𝓝 x] toIocMod hp a := IsOpen.mem_nhds (by rw [Ico_eq_locus_Ioc_eq_iUnion_Ioo] exact isOpen_iUnion fun i => isOpen_Ioo) <| (not_modEq_iff_toIcoMod_eq_toIocMod hp).1 <| not_modEq_iff_ne_mod_zmultiples.2 hx -theorem continuousAt_toIcoMod : ContinuousAt (toIcoMod hp a) x := +theorem continuousAt_toIcoMod (hx : (x : 𝕜 ⧸ zmultiples p) ≠ a) : ContinuousAt (toIcoMod hp a) x := let h := toIcoMod_eventuallyEq_toIocMod hp a hx continuousAt_iff_continuous_left_right.2 <| ⟨(continuous_left_toIocMod hp a x).congr_of_eventuallyEq (h.filter_mono nhdsWithin_le_nhds) h.eq_of_nhds, continuous_right_toIcoMod hp a x⟩ -theorem continuousAt_toIocMod : ContinuousAt (toIocMod hp a) x := +theorem continuousAt_toIocMod (hx : (x : 𝕜 ⧸ zmultiples p) ≠ a) : ContinuousAt (toIocMod hp a) x := let h := toIcoMod_eventuallyEq_toIocMod hp a hx continuousAt_iff_continuous_left_right.2 <| ⟨continuous_left_toIocMod hp a x, @@ -112,14 +113,14 @@ end Continuity /-- The "additive circle": `𝕜 ⧸ (ℤ ∙ p)`. See also `Circle` and `Real.angle`. -/ @[nolint unusedArguments] -abbrev AddCircle [LinearOrderedAddCommGroup 𝕜] [TopologicalSpace 𝕜] [OrderTopology 𝕜] (p : 𝕜) := +abbrev AddCircle [LinearOrderedAddCommGroup 𝕜] (p : 𝕜) := 𝕜 ⧸ zmultiples p namespace AddCircle section LinearOrderedAddCommGroup -variable [LinearOrderedAddCommGroup 𝕜] [TopologicalSpace 𝕜] [OrderTopology 𝕜] (p : 𝕜) +variable [LinearOrderedAddCommGroup 𝕜] (p : 𝕜) theorem coe_nsmul {n : ℕ} {x : 𝕜} : (↑(n • x) : AddCircle p) = n • (x : AddCircle p) := rfl @@ -160,7 +161,7 @@ theorem coe_add_period (x : 𝕜) : ((x + p : 𝕜) : AddCircle p) = x := by rw [coe_add, ← eq_sub_iff_add_eq', sub_self, coe_period] @[continuity, nolint unusedArguments] -protected theorem continuous_mk' : +protected theorem continuous_mk' [TopologicalSpace 𝕜] : Continuous (QuotientAddGroup.mk' (zmultiples p) : 𝕜 → AddCircle p) := continuous_coinduced_rng @@ -227,6 +228,8 @@ variable (p a) section Continuity +variable [TopologicalSpace 𝕜] + @[continuity] theorem continuous_equivIco_symm : Continuous (equivIco p a).symm := continuous_quotient_mk'.comp continuous_subtype_val @@ -235,14 +238,14 @@ theorem continuous_equivIco_symm : Continuous (equivIco p a).symm := theorem continuous_equivIoc_symm : Continuous (equivIoc p a).symm := continuous_quotient_mk'.comp continuous_subtype_val -variable {x : AddCircle p} (hx : x ≠ a) +variable [OrderTopology 𝕜] {x : AddCircle p} -theorem continuousAt_equivIco : ContinuousAt (equivIco p a) x := by +theorem continuousAt_equivIco (hx : x ≠ a) : ContinuousAt (equivIco p a) x := by induction x using QuotientAddGroup.induction_on rw [ContinuousAt, Filter.Tendsto, QuotientAddGroup.nhds_eq, Filter.map_map] exact (continuousAt_toIcoMod hp.out a hx).codRestrict _ -theorem continuousAt_equivIoc : ContinuousAt (equivIoc p a) x := by +theorem continuousAt_equivIoc (hx : x ≠ a) : ContinuousAt (equivIoc p a) x := by induction x using QuotientAddGroup.induction_on rw [ContinuousAt, Filter.Tendsto, QuotientAddGroup.nhds_eq, Filter.map_map] exact (continuousAt_toIocMod hp.out a hx).codRestrict _ @@ -304,7 +307,7 @@ end LinearOrderedAddCommGroup section LinearOrderedField -variable [LinearOrderedField 𝕜] [TopologicalSpace 𝕜] [OrderTopology 𝕜] (p q : 𝕜) +variable [LinearOrderedField 𝕜] (p q : 𝕜) /-- The rescaling equivalence between additive circles with different periods. -/ def equivAddCircle (hp : p ≠ 0) (hq : q ≠ 0) : AddCircle p ≃+ AddCircle q := @@ -322,6 +325,9 @@ theorem equivAddCircle_symm_apply_mk (hp : p ≠ 0) (hq : q ≠ 0) (x : 𝕜) : (equivAddCircle p q hp hq).symm (x : 𝕜) = (x * (q⁻¹ * p) : 𝕜) := rfl +section +variable [TopologicalSpace 𝕜] [OrderTopology 𝕜] + /-- The rescaling homeomorphism between additive circles with different periods. -/ def homeomorphAddCircle (hp : p ≠ 0) (hq : q ≠ 0) : AddCircle p ≃ₜ AddCircle q := ⟨equivAddCircle p q hp hq, @@ -337,6 +343,7 @@ theorem homeomorphAddCircle_apply_mk (hp : p ≠ 0) (hq : q ≠ 0) (x : 𝕜) : theorem homeomorphAddCircle_symm_apply_mk (hp : p ≠ 0) (hq : q ≠ 0) (x : 𝕜) : (homeomorphAddCircle p q hp hq).symm (x : 𝕜) = (x * (q⁻¹ * p) : 𝕜) := rfl +end variable [hp : Fact (0 < p)] @@ -526,7 +533,7 @@ by the equivalence relation identifying the endpoints. -/ namespace AddCircle -variable [LinearOrderedAddCommGroup 𝕜] [TopologicalSpace 𝕜] [OrderTopology 𝕜] (p a : 𝕜) +variable [LinearOrderedAddCommGroup 𝕜] (p a : 𝕜) [hp : Fact (0 < p)] local notation "𝕋" => AddCircle p @@ -582,7 +589,7 @@ theorem equivIccQuot_comp_mk_eq_toIocMod : /-- The natural map from `[a, a + p] ⊂ 𝕜` with endpoints identified to `𝕜 / ℤ • p`, as a homeomorphism of topological spaces. -/ -def homeoIccQuot : 𝕋 ≃ₜ Quot (EndpointIdent p a) where +def homeoIccQuot [TopologicalSpace 𝕜] [OrderTopology 𝕜] : 𝕋 ≃ₜ Quot (EndpointIdent p a) where toEquiv := equivIccQuot p a continuous_toFun := by -- Porting note: was `simp_rw` @@ -615,23 +622,21 @@ theorem liftIco_eq_lift_Icc {f : 𝕜 → B} (h : f a = f (a + p)) : equivIccQuot p a := rfl +theorem liftIco_zero_coe_apply {f : 𝕜 → B} {x : 𝕜} (hx : x ∈ Ico 0 p) : liftIco p 0 f ↑x = f x := + liftIco_coe_apply (by rwa [zero_add]) + +variable [TopologicalSpace 𝕜] [OrderTopology 𝕜] + theorem liftIco_continuous [TopologicalSpace B] {f : 𝕜 → B} (hf : f a = f (a + p)) (hc : ContinuousOn f <| Icc a (a + p)) : Continuous (liftIco p a f) := by rw [liftIco_eq_lift_Icc hf] refine Continuous.comp ?_ (homeoIccQuot p a).continuous_toFun exact continuous_coinduced_dom.mpr (continuousOn_iff_continuous_restrict.mp hc) -section ZeroBased - -theorem liftIco_zero_coe_apply {f : 𝕜 → B} {x : 𝕜} (hx : x ∈ Ico 0 p) : liftIco p 0 f ↑x = f x := - liftIco_coe_apply (by rwa [zero_add]) - theorem liftIco_zero_continuous [TopologicalSpace B] {f : 𝕜 → B} (hf : f 0 = f p) (hc : ContinuousOn f <| Icc 0 p) : Continuous (liftIco p 0 f) := liftIco_continuous (by rwa [zero_add] : f 0 = f (0 + p)) (by rwa [zero_add]) -end ZeroBased - end AddCircle end IdentifyIccEnds diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 338cc5c4ba0ef..7d4f1adcd31f9 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -1290,7 +1290,7 @@ theorem Filter.Tendsto.edist {f g : β → α} {x : Filter β} {a b : α} (hf : /-- If the extended distance between consecutive points of a sequence is estimated by a summable series of `NNReal`s, then the original sequence is a Cauchy sequence. -/ -theorem cauchySeq_of_edist_le_of_summable [PseudoEMetricSpace α] {f : ℕ → α} (d : ℕ → ℝ≥0) +theorem cauchySeq_of_edist_le_of_summable {f : ℕ → α} (d : ℕ → ℝ≥0) (hf : ∀ n, edist (f n) (f n.succ) ≤ d n) (hd : Summable d) : CauchySeq f := by refine EMetric.cauchySeq_iff_NNReal.2 fun ε εpos ↦ ?_ -- Actually we need partial sums of `d` to be a Cauchy sequence. diff --git a/Mathlib/Topology/MetricSpace/Gluing.lean b/Mathlib/Topology/MetricSpace/Gluing.lean index 3c2c725c49738..41444fef866b6 100644 --- a/Mathlib/Topology/MetricSpace/Gluing.lean +++ b/Mathlib/Topology/MetricSpace/Gluing.lean @@ -103,6 +103,7 @@ theorem le_glueDist_inr_inl (Φ : Z → X) (Ψ : Z → Y) (ε : ℝ) (x y) : ε ≤ glueDist Φ Ψ ε (.inr x) (.inl y) := by rw [glueDist_comm]; apply le_glueDist_inl_inr +section variable [Nonempty Z] private theorem glueDist_triangle_inl_inr_inr (Φ : Z → X) (Ψ : Z → Y) (ε : ℝ) (x : X) (y z : Y) : @@ -145,6 +146,8 @@ private theorem glueDist_triangle (Φ : Z → X) (Ψ : Z → Y) (ε : ℝ) apply glueDist_triangle_inl_inr_inl simpa only [abs_sub_comm] +end + private theorem eq_of_glueDist_eq_zero (Φ : Z → X) (Ψ : Z → Y) (ε : ℝ) (ε0 : 0 < ε) : ∀ p q : X ⊕ Y, glueDist Φ Ψ ε p q = 0 → p = q | .inl x, .inl y, h => by rw [eq_of_dist_eq_zero h] @@ -170,7 +173,7 @@ theorem Sum.mem_uniformity_iff_glueDist (hε : 0 < ε) (s : Set ((X ⊕ Y) × (X `Φ p` and `Φ q`, and between `Ψ p` and `Ψ q`, coincide up to `2 ε` where `ε > 0`, one can almost glue the two spaces `X` and `Y` along the images of `Φ` and `Ψ`, so that `Φ p` and `Ψ p` are at distance `ε`. -/ -def glueMetricApprox (Φ : Z → X) (Ψ : Z → Y) (ε : ℝ) (ε0 : 0 < ε) +def glueMetricApprox [Nonempty Z] (Φ : Z → X) (Ψ : Z → Y) (ε : ℝ) (ε0 : 0 < ε) (H : ∀ p q, |dist (Φ p) (Φ q) - dist (Ψ p) (Ψ q)| ≤ 2 * ε) : MetricSpace (X ⊕ Y) where dist := glueDist Φ Ψ ε dist_self := glueDist_self Φ Ψ ε From 064231772afc930f87bca2c80644399ed1995dfd Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 7 Aug 2024 07:08:04 +0000 Subject: [PATCH 080/359] chore: backports for leanprover/lean4#4814 (part 22) (#15511) --- .../Category/ModuleCat/Sheaf/Limits.lean | 6 +- .../BoxIntegral/Partition/Filter.lean | 7 ++- Mathlib/Analysis/Convex/Radon.lean | 4 +- Mathlib/Data/Matrix/ColumnRowPartitioned.lean | 57 +++++++++---------- .../Matrix/NonsingularInverse.lean | 13 ++--- .../Matrix/SesquilinearForm.lean | 2 +- .../LinearAlgebra/QuadraticForm/Basic.lean | 15 +++-- Mathlib/RepresentationTheory/Maschke.lean | 8 ++- Mathlib/RingTheory/SimpleModule.lean | 18 +++--- Mathlib/RingTheory/TensorProduct/Basic.lean | 13 +++-- Mathlib/Topology/MetricSpace/Perfect.lean | 11 ++-- Mathlib/Topology/TietzeExtension.lean | 18 +++--- 12 files changed, 94 insertions(+), 78 deletions(-) diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean index bace928c718ea..d82c0027f62ba 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean @@ -30,11 +30,11 @@ namespace PresheafOfModules variable {R : Cᵒᵖ ⥤ RingCat.{u}} {F : D ⥤ PresheafOfModules.{v} R} [∀ X, Small.{v} ((F ⋙ evaluation R X) ⋙ forget _).sections] - {c : Cone F} (hc : IsLimit c) - (hF : ∀ j, Presheaf.IsSheaf J (F.obj j).presheaf) + {c : Cone F} [HasLimitsOfShape D AddCommGrp.{v}] -lemma isSheaf_of_isLimit : Presheaf.IsSheaf J (c.pt.presheaf) := by +lemma isSheaf_of_isLimit (hc : IsLimit c) (hF : ∀ j, Presheaf.IsSheaf J (F.obj j).presheaf) : + Presheaf.IsSheaf J (c.pt.presheaf) := by let G : D ⥤ Sheaf J AddCommGrp.{v} := { obj := fun j => ⟨(F.obj j).presheaf, hF j⟩ map := fun φ => ⟨(PresheafOfModules.toPresheaf R).map (F.map φ)⟩ } diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean index 1489ce4929c4c..03f25141416a6 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean @@ -173,8 +173,7 @@ noncomputable section namespace BoxIntegral -variable {ι : Type*} [Fintype ι] {I J : Box ι} {c c₁ c₂ : ℝ≥0} {r r₁ r₂ : (ι → ℝ) → Ioi (0 : ℝ)} - {π π₁ π₂ : TaggedPrepartition I} +variable {ι : Type*} [Fintype ι] {I J : Box ι} {c c₁ c₂ : ℝ≥0} open TaggedPrepartition @@ -327,6 +326,8 @@ theorem toFilter_inf_iUnion_eq (l : IntegrationParams) (I : Box ι) (π₀ : Pre l.toFilter I ⊓ 𝓟 { π | π.iUnion = π₀.iUnion } = l.toFilteriUnion I π₀ := (iSup_inf_principal _ _).symm +variable {r r₁ r₂ : (ι → ℝ) → Ioi (0 : ℝ)} {π π₁ π₂ : TaggedPrepartition I} + theorem MemBaseSet.mono' (I : Box ι) (h : l₁ ≤ l₂) (hc : c₁ ≤ c₂) {π : TaggedPrepartition I} (hr : ∀ J ∈ π, r₁ (π.tag J) ≤ r₂ (π.tag J)) (hπ : l₁.MemBaseSet I c₁ r₁ π) : l₂.MemBaseSet I c₂ r₂ π := @@ -345,7 +346,7 @@ theorem MemBaseSet.exists_common_compl (h₁ : l.MemBaseSet I c₁ r₁ π₁) ( (l.bDistortion → π.distortion ≤ c₁) ∧ (l.bDistortion → π.distortion ≤ c₂) := by wlog hc : c₁ ≤ c₂ with H · simpa [hU, _root_.and_comm] using - @H _ _ I J c c₂ c₁ r r₂ r₁ π π₂ π₁ _ l₂ l₁ h₂ h₁ hU.symm (le_of_not_le hc) + @H _ _ I J c c₂ c₁ _ l₂ l₁ r r₂ r₁ π π₂ π₁ h₂ h₁ hU.symm (le_of_not_le hc) by_cases hD : (l.bDistortion : Prop) · rcases h₁.4 hD with ⟨π, hπU, hπc⟩ exact ⟨π, hπU, fun _ => hπc, fun _ => hπc.trans hc⟩ diff --git a/Mathlib/Analysis/Convex/Radon.lean b/Mathlib/Analysis/Convex/Radon.lean index e11bb4b933256..ee505ba41b251 100644 --- a/Mathlib/Analysis/Convex/Radon.lean +++ b/Mathlib/Analysis/Convex/Radon.lean @@ -64,14 +64,14 @@ theorem radon_partition {f : ι → E} (h : ¬ AffineIndependent 𝕜 f) : open FiniteDimensional -variable [FiniteDimensional 𝕜 E] - /-- Corner case for `helly_theorem'`. -/ private lemma helly_theorem_corner {F : ι → Set E} {s : Finset ι} (h_card_small : s.card ≤ finrank 𝕜 E + 1) (h_inter : ∀ I ⊆ s, I.card ≤ finrank 𝕜 E + 1 → (⋂ i ∈ I, F i).Nonempty) : (⋂ i ∈ s, F i).Nonempty := h_inter s (by simp) h_card_small +variable [FiniteDimensional 𝕜 E] + /-- **Helly's theorem** for finite families of convex sets. If `F` is a finite family of convex sets in a vector space of finite dimension `d`, and any diff --git a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean index c0382722e49e6..d8f1117d5f756 100644 --- a/Mathlib/Data/Matrix/ColumnRowPartitioned.lean +++ b/Mathlib/Data/Matrix/ColumnRowPartitioned.lean @@ -26,10 +26,6 @@ namespace Matrix variable {R : Type*} variable {m m₁ m₂ n n₁ n₂ : Type*} -variable [Fintype m] [Fintype m₁] [Fintype m₂] -variable [Fintype n] [Fintype n₁] [Fintype n₂] -variable [DecidableEq m] [DecidableEq m₁] [DecidableEq m₂] -variable [DecidableEq n] [DecidableEq n₁] [DecidableEq n₂] /-- Concatenate together two matrices A₁[m₁ × N] and A₂[m₂ × N] with the same columns (N) to get a bigger matrix indexed by [(m₁ ⊕ m₂) × N] -/ @@ -158,41 +154,53 @@ lemma fromColumns_neg (A₁ : Matrix n m₁ R) (A₂ : Matrix n m₂ R) : end Neg +@[simp] +lemma fromColumns_fromRows_eq_fromBlocks (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) + (B₂₁ : Matrix m₂ n₁ R) (B₂₂ : Matrix m₂ n₂ R) : + fromColumns (fromRows B₁₁ B₂₁) (fromRows B₁₂ B₂₂) = fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ := by + ext (_ | _) (_ | _) <;> simp + +@[simp] +lemma fromRows_fromColumn_eq_fromBlocks (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) + (B₂₁ : Matrix m₂ n₁ R) (B₂₂ : Matrix m₂ n₂ R) : + fromRows (fromColumns B₁₁ B₁₂) (fromColumns B₂₁ B₂₂) = fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ := by + ext (_ | _) (_ | _) <;> simp + section Semiring variable [Semiring R] @[simp] -lemma fromRows_mulVec (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) (v : n → R) : +lemma fromRows_mulVec [Fintype n] (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) (v : n → R) : fromRows A₁ A₂ *ᵥ v = Sum.elim (A₁ *ᵥ v) (A₂ *ᵥ v) := by ext (_ | _) <;> rfl @[simp] -lemma vecMul_fromColumns (B₁ : Matrix m n₁ R) (B₂ : Matrix m n₂ R) (v : m → R) : +lemma vecMul_fromColumns [Fintype m] (B₁ : Matrix m n₁ R) (B₂ : Matrix m n₂ R) (v : m → R) : v ᵥ* fromColumns B₁ B₂ = Sum.elim (v ᵥ* B₁) (v ᵥ* B₂) := by ext (_ | _) <;> rfl @[simp] -lemma sum_elim_vecMul_fromRows (B₁ : Matrix m₁ n R) (B₂ : Matrix m₂ n R) +lemma sum_elim_vecMul_fromRows [Fintype m₁] [Fintype m₂] (B₁ : Matrix m₁ n R) (B₂ : Matrix m₂ n R) (v₁ : m₁ → R) (v₂ : m₂ → R) : Sum.elim v₁ v₂ ᵥ* fromRows B₁ B₂ = v₁ ᵥ* B₁ + v₂ ᵥ* B₂ := by ext simp [Matrix.vecMul, fromRows, dotProduct] @[simp] -lemma fromColumns_mulVec_sum_elim (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) - (v₁ : n₁ → R) (v₂ : n₂ → R) : +lemma fromColumns_mulVec_sum_elim [Fintype n₁] [Fintype n₂] + (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) (v₁ : n₁ → R) (v₂ : n₂ → R) : fromColumns A₁ A₂ *ᵥ Sum.elim v₁ v₂ = A₁ *ᵥ v₁ + A₂ *ᵥ v₂ := by ext simp [Matrix.mulVec, fromColumns] @[simp] -lemma fromRows_mul (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) (B : Matrix n m R) : +lemma fromRows_mul [Fintype n] (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) (B : Matrix n m R) : fromRows A₁ A₂ * B = fromRows (A₁ * B) (A₂ * B) := by ext (_ | _) _ <;> simp [mul_apply] @[simp] -lemma mul_fromColumns (A : Matrix m n R) (B₁ : Matrix n n₁ R) (B₂ : Matrix n n₂ R) : +lemma mul_fromColumns [Fintype n] (A : Matrix m n R) (B₁ : Matrix n n₁ R) (B₂ : Matrix n n₂ R) : A * fromColumns B₁ B₂ = fromColumns (A * B₁) (A * B₂) := by ext _ (_ | _) <;> simp [mul_apply] @@ -204,20 +212,8 @@ lemma fromRows_zero : fromRows (0 : Matrix m₁ n R) (0 : Matrix m₂ n R) = 0 : lemma fromColumns_zero : fromColumns (0 : Matrix m n₁ R) (0 : Matrix m n₂ R) = 0 := by ext _ (_ | _) <;> simp -@[simp] -lemma fromColumns_fromRows_eq_fromBlocks (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) - (B₂₁ : Matrix m₂ n₁ R) (B₂₂ : Matrix m₂ n₂ R) : - fromColumns (fromRows B₁₁ B₂₁) (fromRows B₁₂ B₂₂) = fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ := by - ext (_ | _) (_ | _) <;> simp - -@[simp] -lemma fromRows_fromColumn_eq_fromBlocks (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) - (B₂₁ : Matrix m₂ n₁ R) (B₂₂ : Matrix m₂ n₂ R) : - fromRows (fromColumns B₁₁ B₁₂) (fromColumns B₂₁ B₂₂) = fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ := by - ext (_ | _) (_ | _) <;> simp - /-- A row partitioned matrix multiplied by a column partioned matrix gives a 2 by 2 block matrix -/ -lemma fromRows_mul_fromColumns (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) +lemma fromRows_mul_fromColumns [Fintype n] (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) (B₁ : Matrix n n₁ R) (B₂ : Matrix n n₂ R) : (fromRows A₁ A₂) * (fromColumns B₁ B₂) = fromBlocks (A₁ * B₁) (A₁ * B₂) (A₂ * B₁) (A₂ * B₂) := by @@ -225,21 +221,21 @@ lemma fromRows_mul_fromColumns (A₁ : Matrix m₁ n R) (A₂ : Matrix m₂ n R) /-- A column partitioned matrix mulitplied by a row partitioned matrix gives the sum of the "outer" products of the block matrices -/ -lemma fromColumns_mul_fromRows (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) +lemma fromColumns_mul_fromRows [Fintype n₁] [Fintype n₂] (A₁ : Matrix m n₁ R) (A₂ : Matrix m n₂ R) (B₁ : Matrix n₁ n R) (B₂ : Matrix n₂ n R) : fromColumns A₁ A₂ * fromRows B₁ B₂ = (A₁ * B₁ + A₂ * B₂) := by ext simp [mul_apply] /-- A column partitioned matrix multipiled by a block matrix results in a column partioned matrix -/ -lemma fromColumns_mul_fromBlocks (A₁ : Matrix m m₁ R) (A₂ : Matrix m m₂ R) +lemma fromColumns_mul_fromBlocks [Fintype m₁] [Fintype m₂] (A₁ : Matrix m m₁ R) (A₂ : Matrix m m₂ R) (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) (B₂₁ : Matrix m₂ n₁ R) (B₂₂ : Matrix m₂ n₂ R) : (fromColumns A₁ A₂) * fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ = fromColumns (A₁ * B₁₁ + A₂ * B₂₁) (A₁ * B₁₂ + A₂ * B₂₂) := by ext _ (_ | _) <;> simp [mul_apply] /-- A block matrix mulitplied by a row partitioned matrix gives a row partitioned matrix -/ -lemma fromBlocks_mul_fromRows (A₁ : Matrix n₁ n R) (A₂ : Matrix n₂ n R) +lemma fromBlocks_mul_fromRows [Fintype n₁] [Fintype n₂] (A₁ : Matrix n₁ n R) (A₂ : Matrix n₂ n R) (B₁₁ : Matrix m₁ n₁ R) (B₁₂ : Matrix m₁ n₂ R) (B₂₁ : Matrix m₂ n₁ R) (B₂₂ : Matrix m₂ n₂ R) : fromBlocks B₁₁ B₁₂ B₂₁ B₂₂ * (fromRows A₁ A₂) = fromRows (B₁₁ * A₁ + B₁₂ * A₂) (B₂₁ * A₁ + B₂₂ * A₂) := by @@ -256,7 +252,9 @@ This is the column and row partitioned matrix form of `Matrix.mul_eq_one_comm`. The condition `e : n ≃ n₁ ⊕ n₂` states that `fromColumns A₁ A₂` and `fromRows B₁ B₂` are "square". -/ -lemma fromColumns_mul_fromRows_eq_one_comm (e : n ≃ n₁ ⊕ n₂) +lemma fromColumns_mul_fromRows_eq_one_comm + [Fintype n₁] [Fintype n₂] [Fintype n] [DecidableEq n] [DecidableEq n₁] [DecidableEq n₂] + (e : n ≃ n₁ ⊕ n₂) (A₁ : Matrix n n₁ R) (A₂ : Matrix n n₂ R) (B₁ : Matrix n₁ n R) (B₂ : Matrix n₂ n R) : fromColumns A₁ A₂ * fromRows B₁ B₂ = 1 ↔ fromRows B₁ B₂ * fromColumns A₁ A₂ = 1 := by calc fromColumns A₁ A₂ * fromRows B₁ B₂ = 1 @@ -272,7 +270,8 @@ lemma fromColumns_mul_fromRows_eq_one_comm (e : n ≃ n₁ ⊕ n₂) /-- The lemma `fromColumns_mul_fromRows_eq_one_comm` specialized to the case where the index sets n₁ and n₂, are the result of subtyping by a predicate and its complement. -/ -lemma equiv_compl_fromColumns_mul_fromRows_eq_one_comm (p : n → Prop)[DecidablePred p] +lemma equiv_compl_fromColumns_mul_fromRows_eq_one_comm + [Fintype n] [DecidableEq n] (p : n → Prop) [DecidablePred p] (A₁ : Matrix n {i // p i} R) (A₂ : Matrix n {i // ¬p i} R) (B₁ : Matrix {i // p i} n R) (B₂ : Matrix {i // ¬p i} n R) : fromColumns A₁ A₂ * fromRows B₁ B₂ = 1 ↔ fromRows B₁ B₂ * fromColumns A₁ A₂ = 1 := diff --git a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean index 745b0e49846d7..e91010d1bc343 100644 --- a/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean +++ b/Mathlib/LinearAlgebra/Matrix/NonsingularInverse.lean @@ -320,7 +320,6 @@ end Inv section InjectiveMul variable [Fintype n] [Fintype m] [DecidableEq m] [CommRing α] -variable [Fintype l] [DecidableEq l] lemma mul_left_injective_of_inv (A : Matrix m n α) (B : Matrix n m α) (h : A * B = 1) : Function.Injective (fun x : Matrix l m α => x * A) := fun _ _ g => by @@ -334,13 +333,12 @@ end InjectiveMul section vecMul -variable [DecidableEq m] [DecidableEq n] - section Semiring variable {R : Type*} [Semiring R] -theorem vecMul_surjective_iff_exists_left_inverse [Fintype m] [Finite n] {A : Matrix m n R} : +theorem vecMul_surjective_iff_exists_left_inverse + [DecidableEq n] [Fintype m] [Finite n] {A : Matrix m n R} : Function.Surjective A.vecMul ↔ ∃ B : Matrix n m R, B * A = 1 := by cases nonempty_fintype n refine ⟨fun h ↦ ?_, fun ⟨B, hBA⟩ y ↦ ⟨y ᵥ* B, by simp [hBA]⟩⟩ @@ -349,7 +347,8 @@ theorem vecMul_surjective_iff_exists_left_inverse [Fintype m] [Finite n] {A : Ma rw [mul_apply_eq_vecMul, one_eq_pi_single, ← hrows] rfl -theorem mulVec_surjective_iff_exists_right_inverse [Finite m] [Fintype n] {A : Matrix m n R} : +theorem mulVec_surjective_iff_exists_right_inverse + [DecidableEq m] [Finite m] [Fintype n] {A : Matrix m n R} : Function.Surjective A.mulVec ↔ ∃ B : Matrix n m R, A * B = 1 := by cases nonempty_fintype m refine ⟨fun h ↦ ?_, fun ⟨B, hBA⟩ y ↦ ⟨B *ᵥ y, by simp [hBA]⟩⟩ @@ -360,7 +359,7 @@ theorem mulVec_surjective_iff_exists_right_inverse [Finite m] [Fintype n] {A : M end Semiring -variable {R K : Type*} [CommRing R] [Field K] [Fintype m] +variable [DecidableEq m] {R K : Type*} [CommRing R] [Field K] [Fintype m] theorem vecMul_surjective_iff_isUnit {A : Matrix m m R} : Function.Surjective A.vecMul ↔ IsUnit A := by @@ -711,7 +710,7 @@ end Submatrix section Det -variable [Fintype m] [DecidableEq m] [CommRing α] +variable [Fintype m] [DecidableEq m] /-- A variant of `Matrix.det_units_conj`. -/ theorem det_conj {M : Matrix m m α} (h : IsUnit M) (N : Matrix m m α) : diff --git a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean index e14a9a860b1f3..53360f1ca3856 100644 --- a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean @@ -137,7 +137,7 @@ This section deals with the conversion between matrices and sesquilinear maps on -/ variable [CommSemiring R] [AddCommMonoid N₂] [Module R N₂] [Semiring R₁] [Semiring R₂] - [SMulCommClass R R N₂] [Semiring S₁] [Semiring S₂] [Module S₁ N₂] [Module S₂ N₂] + [Semiring S₁] [Semiring S₂] [Module S₁ N₂] [Module S₂ N₂] [SMulCommClass S₁ R N₂] [SMulCommClass S₂ R N₂] [SMulCommClass S₂ S₁ N₂] variable {σ₁ : R₁ →+* S₁} {σ₂ : R₂ →+* S₂} variable [Fintype n] [Fintype m] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 8686fa1fe4bb1..bddba63da1c4f 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -780,13 +780,15 @@ theorem _root_.QuadraticMap.polarBilin_comp (Q : QuadraticMap R N' N) (f : M → end -variable {N' : Type*} [AddCommGroup N'] [Module R N'] +variable {N' : Type*} [AddCommGroup N'] theorem _root_.LinearMap.compQuadraticMap_polar [CommSemiring S] [Algebra S R] [Module S N] [Module S N'] [IsScalarTower S R N] [Module S M] [IsScalarTower S R M] (f : N →ₗ[S] N') (Q : QuadraticMap R M N) (x y : M) : polar (f.compQuadraticMap' Q) x y = f (polar Q x y) := by simp [polar] +variable [Module R N'] + theorem _root_.LinearMap.compQuadraticMap_polarBilin (f : N →ₗ[R] N') (Q : QuadraticMap R M N) : (f.compQuadraticMap' Q).polarBilin = Q.polarBilin.compr₂ f := by ext @@ -1032,7 +1034,7 @@ end Anisotropic section PosDef variable {R₂ : Type u} [CommSemiring R₂] [AddCommMonoid M] [Module R₂ M] -variable [PartialOrder N] [AddCommMonoid N] [Module R₂ N] [CovariantClass N N (· + ·) (· < ·)] +variable [PartialOrder N] [AddCommMonoid N] [Module R₂ N] variable {Q₂ : QuadraticMap R₂ M N} /-- A positive definite quadratic form is positive on nonzero vectors. -/ @@ -1057,13 +1059,16 @@ theorem PosDef.anisotropic {Q : QuadraticMap R₂ M N} (hQ : Q.PosDef) : Q.Aniso exact this theorem posDef_of_nonneg {Q : QuadraticMap R₂ M N} (h : ∀ x, 0 ≤ Q x) (h0 : Q.Anisotropic) : - PosDef Q := fun x hx => lt_of_le_of_ne (h x) (Ne.symm fun hQx => hx <| h0 _ hQx) + PosDef Q := + fun x hx => lt_of_le_of_ne (h x) (Ne.symm fun hQx => hx <| h0 _ hQx) theorem posDef_iff_nonneg {Q : QuadraticMap R₂ M N} : PosDef Q ↔ (∀ x, 0 ≤ Q x) ∧ Q.Anisotropic := ⟨fun h => ⟨h.nonneg, h.anisotropic⟩, fun ⟨n, a⟩ => posDef_of_nonneg n a⟩ -theorem PosDef.add (Q Q' : QuadraticMap R₂ M N) (hQ : PosDef Q) (hQ' : PosDef Q') : - PosDef (Q + Q') := fun x hx => add_pos (hQ x hx) (hQ' x hx) +theorem PosDef.add [CovariantClass N N (· + ·) (· < ·)] + (Q Q' : QuadraticMap R₂ M N) (hQ : PosDef Q) (hQ' : PosDef Q') : + PosDef (Q + Q') := + fun x hx => add_pos (hQ x hx) (hQ' x hx) theorem linMulLinSelfPosDef {R} [LinearOrderedCommRing R] [Module R M] [LinearOrderedSemiring A] [ExistsAddOfLE A] [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] (f : M →ₗ[R] A) diff --git a/Mathlib/RepresentationTheory/Maschke.lean b/Mathlib/RepresentationTheory/Maschke.lean index caeae09e6064a..1c9739d487652 100644 --- a/Mathlib/RepresentationTheory/Maschke.lean +++ b/Mathlib/RepresentationTheory/Maschke.lean @@ -71,11 +71,12 @@ theorem conjugate_apply (g : G) (v : W) : π.conjugate g v = MonoidAlgebra.single g⁻¹ (1 : k) • π (MonoidAlgebra.single g (1 : k) • v) := rfl -variable (i : V →ₗ[MonoidAlgebra k G] W) (h : ∀ v : V, (π : W → V) (i v) = v) +variable (i : V →ₗ[MonoidAlgebra k G] W) section -theorem conjugate_i (g : G) (v : V) : (conjugate π g : W → V) (i v) = v := by +theorem conjugate_i (h : ∀ v : V, (π : W → V) (i v) = v) (g : G) (v : V) : + (conjugate π g : W → V) (i v) = v := by rw [conjugate_apply, ← i.map_smul, h, ← mul_smul, single_mul_single, mul_one, mul_left_inv, ← one_def, one_smul] @@ -119,7 +120,8 @@ theorem equivariantProjection_apply (v : W) : π.equivariantProjection G v = ⅟(Fintype.card G : k) • ∑ g : G, π.conjugate g v := by simp only [equivariantProjection, smul_apply, sumOfConjugatesEquivariant_apply] -theorem equivariantProjection_condition (v : V) : (π.equivariantProjection G) (i v) = v := by +theorem equivariantProjection_condition (h : ∀ v : V, (π : W → V) (i v) = v) (v : V) : + (π.equivariantProjection G) (i v) = v := by rw [equivariantProjection_apply] simp only [conjugate_i π i h] rw [Finset.sum_const, Finset.card_univ, ← Nat.cast_smul_eq_nsmul k, smul_smul, diff --git a/Mathlib/RingTheory/SimpleModule.lean b/Mathlib/RingTheory/SimpleModule.lean index 8e0ca5c07d30b..6bd08e9df362c 100644 --- a/Mathlib/RingTheory/SimpleModule.lean +++ b/Mathlib/RingTheory/SimpleModule.lean @@ -68,12 +68,14 @@ theorem IsSimpleModule.nontrivial [IsSimpleModule R M] : Nontrivial M := ext x simp [Submodule.mem_bot, Submodule.mem_top, h x]⟩⟩ -variable {m : Submodule R M} {N : Type*} [AddCommGroup N] [Module R N] {R S M} +variable {m : Submodule R M} {N : Type*} [AddCommGroup N] {R S M} theorem LinearMap.isSimpleModule_iff_of_bijective [Module S N] {σ : R →+* S} [RingHomSurjective σ] (l : M →ₛₗ[σ] N) (hl : Function.Bijective l) : IsSimpleModule R M ↔ IsSimpleModule S N := (Submodule.orderIsoMapComapOfBijective l hl).isSimpleOrder_iff +variable [Module R N] + theorem IsSimpleModule.congr (l : M ≃ₗ[R] N) [IsSimpleModule R N] : IsSimpleModule R M := (Submodule.orderIsoMapComap l).isSimpleOrder @@ -204,7 +206,7 @@ instance submodule {m : Submodule R M} : IsSemisimpleModule R m := variable {R M} open LinearMap -theorem congr [IsSemisimpleModule R N] (e : M ≃ₗ[R] N) : IsSemisimpleModule R M := +theorem congr (e : N ≃ₗ[R] M) : IsSemisimpleModule R N := (Submodule.orderIsoMapComap e.symm).complementedLattice instance quotient : IsSemisimpleModule R (M ⧸ m) := @@ -217,19 +219,21 @@ protected theorem range (f : M →ₗ[R] N) : IsSemisimpleModule R (range f) := section -variable [Module S N] {σ : R →+* S} [RingHomSurjective σ] (l : M →ₛₗ[σ] N) +variable {M' : Type*} [AddCommGroup M'] [Module R M'] {N'} [AddCommGroup N'] [Module S N'] + {σ : R →+* S} (l : M' →ₛₗ[σ] N') -theorem _root_.LinearMap.isSemisimpleModule_iff_of_bijective (hl : Function.Bijective l) : - IsSemisimpleModule R M ↔ IsSemisimpleModule S N := +theorem _root_.LinearMap.isSemisimpleModule_iff_of_bijective + [RingHomSurjective σ] (hl : Function.Bijective l) : + IsSemisimpleModule R M' ↔ IsSemisimpleModule S N' := (Submodule.orderIsoMapComapOfBijective l hl).complementedLattice_iff -- TODO: generalize Submodule.equivMapOfInjective from InvPair to RingHomSurjective proof_wanted _root_.LinearMap.isSemisimpleModule_of_injective (_ : Function.Injective l) - [IsSemisimpleModule S N] : IsSemisimpleModule R M + [IsSemisimpleModule S N'] : IsSemisimpleModule R M' --TODO: generalize LinearMap.quotKerEquivOfSurjective to SemilinearMaps + RingHomSurjective proof_wanted _root_.LinearMap.isSemisimpleModule_of_surjective (_ : Function.Surjective l) - [IsSemisimpleModule R M] : IsSemisimpleModule S N + [IsSemisimpleModule R M'] : IsSemisimpleModule S N' end diff --git a/Mathlib/RingTheory/TensorProduct/Basic.lean b/Mathlib/RingTheory/TensorProduct/Basic.lean index c061e1411a588..e215d9a539db5 100644 --- a/Mathlib/RingTheory/TensorProduct/Basic.lean +++ b/Mathlib/RingTheory/TensorProduct/Basic.lean @@ -610,7 +610,7 @@ section variable [CommSemiring R] [CommSemiring S] [Algebra R S] variable [Semiring A] [Algebra R A] [Algebra S A] [IsScalarTower R S A] variable [Semiring B] [Algebra R B] -variable [Semiring C] [Algebra R C] [Algebra S C] +variable [Semiring C] [Algebra S C] variable [Semiring D] [Algebra R D] /-- Build an algebra morphism from a linear map out of a tensor product, and evidence that on pure @@ -656,6 +656,7 @@ theorem algEquivOfLinearEquivTensorProduct_apply (f h_mul h_one x) : (algEquivOfLinearEquivTensorProduct f h_mul h_one : A ⊗[R] B ≃ₐ[S] C) x = f x := rfl +variable [Algebra R C] /-- Build an algebra equivalence from a linear equivalence out of a triple tensor product, and evidence of multiplicativity on pure tensors. -/ @@ -741,7 +742,7 @@ end variable [CommSemiring R] [CommSemiring S] [Algebra R S] variable [Semiring A] [Algebra R A] [Algebra S A] [IsScalarTower R S A] variable [Semiring B] [Algebra R B] [Algebra S B] [IsScalarTower R S B] -variable [Semiring C] [Algebra R C] [Algebra S C] [IsScalarTower R S C] +variable [Semiring C] [Algebra R C] variable [Semiring D] [Algebra R D] variable [Semiring E] [Algebra R E] variable [Semiring F] [Algebra R F] @@ -882,7 +883,8 @@ theorem map_tmul (f : A →ₐ[S] B) (g : C →ₐ[R] D) (a : A) (c : C) : map f theorem map_id : map (.id S A) (.id R C) = .id S _ := ext (AlgHom.ext fun _ => rfl) (AlgHom.ext fun _ => rfl) -theorem map_comp (f₂ : B →ₐ[S] C) (f₁ : A →ₐ[S] B) (g₂ : E →ₐ[R] F) (g₁ : D →ₐ[R] E) : +theorem map_comp [Algebra S C] [IsScalarTower R S C] + (f₂ : B →ₐ[S] C) (f₁ : A →ₐ[S] B) (g₂ : E →ₐ[R] F) (g₁ : D →ₐ[R] E) : map (f₂.comp f₁) (g₂.comp g₁) = (map f₂ g₂).comp (map f₁ g₁) := ext (AlgHom.ext fun _ => rfl) (AlgHom.ext fun _ => rfl) @@ -936,7 +938,8 @@ theorem congr_symm_apply (f : A ≃ₐ[S] B) (g : C ≃ₐ[R] D) (x) : theorem congr_refl : congr (.refl : A ≃ₐ[S] A) (.refl : C ≃ₐ[R] C) = .refl := AlgEquiv.coe_algHom_injective <| map_id -theorem congr_trans (f₁ : A ≃ₐ[S] B) (f₂ : B ≃ₐ[S] C) (g₁ : D ≃ₐ[R] E) (g₂ : E ≃ₐ[R] F) : +theorem congr_trans [Algebra S C] [IsScalarTower R S C] + (f₁ : A ≃ₐ[S] B) (f₂ : B ≃ₐ[S] C) (g₁ : D ≃ₐ[R] E) (g₂ : E ≃ₐ[R] F) : congr (f₁.trans f₂) (g₁.trans g₂) = (congr f₁ g₁).trans (congr f₂ g₂) := AlgEquiv.coe_algHom_injective <| map_comp f₂.toAlgHom f₁.toAlgHom g₂.toAlgHom g₁.toAlgHom @@ -1121,7 +1124,7 @@ namespace LinearMap open Algebra.TensorProduct variable {R M₁ M₂ ι ι₂ : Type*} (A : Type*) - [Fintype ι] [Finite ι₂] [DecidableEq ι] [DecidableEq ι₂] + [Fintype ι] [Finite ι₂] [DecidableEq ι] [CommSemiring R] [CommSemiring A] [Algebra R A] [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] diff --git a/Mathlib/Topology/MetricSpace/Perfect.lean b/Mathlib/Topology/MetricSpace/Perfect.lean index 9cad323c4bca8..c90e6e4ce4bd6 100644 --- a/Mathlib/Topology/MetricSpace/Perfect.lean +++ b/Mathlib/Topology/MetricSpace/Perfect.lean @@ -35,9 +35,9 @@ section CantorInjMetric open Function ENNReal -variable {α : Type*} [MetricSpace α] {C : Set α} (hC : Perfect C) {ε : ℝ≥0∞} +variable {α : Type*} [MetricSpace α] {C : Set α} {ε : ℝ≥0∞} -private theorem Perfect.small_diam_aux (ε_pos : 0 < ε) {x : α} (xC : x ∈ C) : +private theorem Perfect.small_diam_aux (hC : Perfect C) (ε_pos : 0 < ε) {x : α} (xC : x ∈ C) : let D := closure (EMetric.ball x (ε / 2) ∩ C) Perfect D ∧ D.Nonempty ∧ D ⊆ C ∧ EMetric.diam D ≤ ε := by have : x ∈ EMetric.ball x (ε / 2) := by @@ -53,11 +53,9 @@ private theorem Perfect.small_diam_aux (ε_pos : 0 < ε) {x : α} (xC : x ∈ C) convert EMetric.diam_ball (x := x) rw [mul_comm, ENNReal.div_mul_cancel] <;> norm_num -variable (hnonempty : C.Nonempty) - /-- A refinement of `Perfect.splitting` for metric spaces, where we also control the diameter of the new perfect sets. -/ -theorem Perfect.small_diam_splitting (ε_pos : 0 < ε) : +theorem Perfect.small_diam_splitting (hC : Perfect C) (hnonempty : C.Nonempty) (ε_pos : 0 < ε) : ∃ C₀ C₁ : Set α, (Perfect C₀ ∧ C₀.Nonempty ∧ C₀ ⊆ C ∧ EMetric.diam C₀ ≤ ε) ∧ (Perfect C₁ ∧ C₁.Nonempty ∧ C₁ ⊆ C ∧ EMetric.diam C₁ ≤ ε) ∧ Disjoint C₀ C₁ := by rcases hC.splitting hnonempty with ⟨D₀, D₁, ⟨perf0, non0, sub0⟩, ⟨perf1, non1, sub1⟩, hdisj⟩ @@ -74,7 +72,8 @@ open CantorScheme /-- Any nonempty perfect set in a complete metric space admits a continuous injection from the Cantor space, `ℕ → Bool`. -/ -theorem Perfect.exists_nat_bool_injection [CompleteSpace α] : +theorem Perfect.exists_nat_bool_injection + (hC : Perfect C) (hnonempty : C.Nonempty) [CompleteSpace α] : ∃ f : (ℕ → Bool) → α, range f ⊆ C ∧ Continuous f ∧ Injective f := by obtain ⟨u, -, upos', hu⟩ := exists_seq_strictAnti_tendsto' (zero_lt_one' ℝ≥0∞) have upos := fun n => (upos' n).1 diff --git a/Mathlib/Topology/TietzeExtension.lean b/Mathlib/Topology/TietzeExtension.lean index 9552e73b4aece..ebbe95a8f4538 100644 --- a/Mathlib/Topology/TietzeExtension.lean +++ b/Mathlib/Topology/TietzeExtension.lean @@ -52,22 +52,23 @@ class TietzeExtension (Y : Type v) [TopologicalSpace Y] : Prop where (hs : IsClosed s) (f : C(s, Y)) : ∃ (g : C(X, Y)), g.restrict s = f variable {X₁ : Type u₁} [TopologicalSpace X₁] -variable {X : Type u} [TopologicalSpace X] [NormalSpace X] {s : Set X} (hs : IsClosed s) -variable {e : X₁ → X} (he : ClosedEmbedding e) +variable {X : Type u} [TopologicalSpace X] [NormalSpace X] {s : Set X} +variable {e : X₁ → X} variable {Y : Type v} [TopologicalSpace Y] [TietzeExtension.{u, v} Y] /-- **Tietze extension theorem** for `TietzeExtension` spaces, a version for a closed set. Let `s` be a closed set in a normal topological space `X`. Let `f` be a continuous function on `s` with values in a `TietzeExtension` space `Y`. Then there exists a continuous function `g : C(X, Y)` such that `g.restrict s = f`. -/ -theorem ContinuousMap.exists_restrict_eq (f : C(s, Y)) : ∃ (g : C(X, Y)), g.restrict s = f := +theorem ContinuousMap.exists_restrict_eq (hs : IsClosed s) (f : C(s, Y)) : + ∃ (g : C(X, Y)), g.restrict s = f := TietzeExtension.exists_restrict_eq' s hs f /-- **Tietze extension theorem** for `TietzeExtension` spaces. Let `e` be a closed embedding of a nonempty topological space `X₁` into a normal topological space `X`. Let `f` be a continuous function on `X₁` with values in a `TietzeExtension` space `Y`. Then there exists a continuous function `g : C(X, Y)` such that `g ∘ e = f`. -/ -theorem ContinuousMap.exists_extension (f : C(X₁, Y)) : +theorem ContinuousMap.exists_extension (he : ClosedEmbedding e) (f : C(X₁, Y)) : ∃ (g : C(X, Y)), g.comp ⟨e, he.continuous⟩ = f := by let e' : X₁ ≃ₜ Set.range e := Homeomorph.ofEmbedding _ he.toEmbedding obtain ⟨g, hg⟩ := (f.comp e'.symm).exists_restrict_eq he.isClosed_range @@ -80,7 +81,8 @@ continuous function `g : C(X, Y)` such that `g ∘ e = f`. This version is provided for convenience and backwards compatibility. Here the composition is phrased in terms of bare functions. -/ -theorem ContinuousMap.exists_extension' (f : C(X₁, Y)) : ∃ (g : C(X, Y)), g ∘ e = f := +theorem ContinuousMap.exists_extension' (he : ClosedEmbedding e) (f : C(X₁, Y)) : + ∃ (g : C(X, Y)), g ∘ e = f := f.exists_extension he |>.imp fun g hg ↦ by ext x; congrm($(hg) x) /-- This theorem is not intended to be used directly because it is rare for a set alone to @@ -89,7 +91,8 @@ the radius is strictly positive, so finding this as an instance will fail. Instead, it is intended to be used as a constructor for theorems about sets which *do* satisfy `[TietzeExtension t]` under some hypotheses. -/ -theorem ContinuousMap.exists_forall_mem_restrict_eq {Y : Type v} [TopologicalSpace Y] (f : C(s, Y)) +theorem ContinuousMap.exists_forall_mem_restrict_eq (hs : IsClosed s) + {Y : Type v} [TopologicalSpace Y] (f : C(s, Y)) {t : Set Y} (hf : ∀ x, f x ∈ t) [ht : TietzeExtension.{u, v} t] : ∃ (g : C(X, Y)), (∀ x, g x ∈ t) ∧ g.restrict s = f := by obtain ⟨g, hg⟩ := mk _ (map_continuous f |>.codRestrict hf) |>.exists_restrict_eq hs @@ -101,7 +104,8 @@ the radius is strictly positive, so finding this as an instance will fail. Instead, it is intended to be used as a constructor for theorems about sets which *do* satisfy `[TietzeExtension t]` under some hypotheses. -/ -theorem ContinuousMap.exists_extension_forall_mem {Y : Type v} [TopologicalSpace Y] (f : C(X₁, Y)) +theorem ContinuousMap.exists_extension_forall_mem (he : ClosedEmbedding e) + {Y : Type v} [TopologicalSpace Y] (f : C(X₁, Y)) {t : Set Y} (hf : ∀ x, f x ∈ t) [ht : TietzeExtension.{u, v} t] : ∃ (g : C(X, Y)), (∀ x, g x ∈ t) ∧ g.comp ⟨e, he.continuous⟩ = f := by obtain ⟨g, hg⟩ := mk _ (map_continuous f |>.codRestrict hf) |>.exists_extension he From 45121e17219b729a8d2bf890391885cbb6b1210a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 7 Aug 2024 07:08:06 +0000 Subject: [PATCH 081/359] chore: backports for leanprover/lean4#4814 (part 23) (#15514) --- Mathlib/Algebra/DirectSum/LinearMap.lean | 11 ++++---- .../ApproximatesLinearOn.lean | 9 ++++--- Mathlib/Analysis/Calculus/LHopital.lean | 16 ++++++------ Mathlib/Analysis/Matrix.lean | 3 +-- Mathlib/Condensed/Epi.lean | 18 +++++++------ Mathlib/FieldTheory/Cardinality.lean | 2 +- .../Manifold/SmoothManifoldWithCorners.lean | 8 +++--- .../BilinearForm/Properties.lean | 25 ++++++++++--------- .../IsIntegralClosure/Basic.lean | 5 ++-- Mathlib/RingTheory/LocalRing/Module.lean | 11 +++++--- .../RingTheory/Localization/Cardinality.lean | 7 +++--- 11 files changed, 62 insertions(+), 53 deletions(-) diff --git a/Mathlib/Algebra/DirectSum/LinearMap.lean b/Mathlib/Algebra/DirectSum/LinearMap.lean index 56f1f82a84bac..24dc43c284d75 100644 --- a/Mathlib/Algebra/DirectSum/LinearMap.lean +++ b/Mathlib/Algebra/DirectSum/LinearMap.lean @@ -24,7 +24,7 @@ variable {ι R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] {N : ι → section IsInternal -variable [DecidableEq ι] (h : IsInternal N) +variable [DecidableEq ι] /-- If a linear map `f : M₁ → M₂` respects direct sum decompositions of `M₁` and `M₂`, then it has a block diagonal matrix with respect to bases compatible with the direct sum decompositions. -/ @@ -62,7 +62,7 @@ variable [∀ i, Module.Finite R (N i)] [∀ i, Module.Free R (N i)] /-- The trace of an endomorphism of a direct sum is the sum of the traces on each component. See also `LinearMap.trace_restrict_eq_sum_trace_restrict`. -/ -lemma trace_eq_sum_trace_restrict [Fintype ι] +lemma trace_eq_sum_trace_restrict (h : IsInternal N) [Fintype ι] {f : M →ₗ[R] M} (hf : ∀ i, MapsTo f (N i) (N i)) : trace R M f = ∑ i, trace R (N i) (f.restrict (hf i)) := by let b : (i : ι) → Basis _ R (N i) := fun i ↦ Module.Free.chooseBasis R (N i) @@ -70,7 +70,7 @@ lemma trace_eq_sum_trace_restrict [Fintype ι] toMatrix_directSum_collectedBasis_eq_blockDiagonal' h h b b hf, Matrix.trace_blockDiagonal', ← trace_eq_matrix_trace] -lemma trace_eq_sum_trace_restrict' (hN : {i | N i ≠ ⊥}.Finite) +lemma trace_eq_sum_trace_restrict' (h : IsInternal N) (hN : {i | N i ≠ ⊥}.Finite) {f : M →ₗ[R] M} (hf : ∀ i, MapsTo f (N i) (N i)) : trace R M f = ∑ i ∈ hN.toFinset, trace R (N i) (f.restrict (hf i)) := by let _ : Fintype {i // N i ≠ ⊥} := hN.fintype @@ -78,7 +78,7 @@ lemma trace_eq_sum_trace_restrict' (hN : {i | N i ≠ ⊥}.Finite) rw [← Finset.sum_coe_sort, trace_eq_sum_trace_restrict (isInternal_ne_bot_iff.mpr h) _] exact Fintype.sum_equiv hN.subtypeEquivToFinset _ _ (fun i ↦ rfl) -lemma trace_eq_zero_of_mapsTo_ne [IsNoetherian R M] +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 @@ -122,7 +122,8 @@ lemma trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero rw [restrict_comp, trace_comp_eq_mul_of_commute_of_isNilpotent μ h_comm (f.isNilpotent_restrict_iSup_sub_algebraMap μ), hg, mul_zero] -lemma mapsTo_biSup_of_mapsTo (s : Set ι) {f : Module.End R M} (hf : ∀ i, MapsTo f (N i) (N i)) : +lemma mapsTo_biSup_of_mapsTo {ι : Type*} {N : ι → Submodule R M} + (s : Set ι) {f : Module.End R M} (hf : ∀ i, MapsTo f (N i) (N i)) : MapsTo f ↑(⨆ i ∈ s, N i) ↑(⨆ i ∈ s, N i) := by replace hf : ∀ i, (N i).map f ≤ N i := fun i ↦ Submodule.map_le_iff_le_comap.mpr (hf i) suffices (⨆ i ∈ s, N i).map f ≤ ⨆ i ∈ s, N i from Submodule.map_le_iff_le_comap.mp this diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean index 378ee5517382e..7994a8712c5d2 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean @@ -75,7 +75,7 @@ theorem approximatesLinearOn_empty (f : E → F) (f' : E →L[𝕜] F) (c : ℝ namespace ApproximatesLinearOn -variable [CompleteSpace E] {f : E → F} +variable {f : E → F} /-! First we prove some properties of a function that `ApproximatesLinearOn` a (not necessarily invertible) continuous linear map. -/ @@ -127,12 +127,13 @@ equivalence, for the local inverse theorem, but also whenever the approximating by Banach's open mapping theorem. -/ -variable {s : Set E} {c : ℝ≥0} {f' : E →L[𝕜] F} +variable [CompleteSpace E] {s : Set E} {c : ℝ≥0} {f' : E →L[𝕜] F} /-- If a function is linearly approximated by a continuous linear map with a (possibly nonlinear) right inverse, then it is locally onto: a ball of an explicit radius is included in the image of the map. -/ -theorem surjOn_closedBall_of_nonlinearRightInverse (hf : ApproximatesLinearOn f f' s c) +theorem surjOn_closedBall_of_nonlinearRightInverse + (hf : ApproximatesLinearOn f f' s c) (f'symm : f'.NonlinearRightInverse) {ε : ℝ} {b : E} (ε0 : 0 ≤ ε) (hε : closedBall b ε ⊆ s) : SurjOn f (closedBall b ε) (closedBall (f b) (((f'symm.nnnorm : ℝ)⁻¹ - c) * ε)) := by intro y hy @@ -374,6 +375,8 @@ theorem to_inv (hf : ApproximatesLinearOn f (f' : E →L[𝕜] F) s c) (hc : Sub _ = (N * (N⁻¹ - c)⁻¹ * c : ℝ≥0) * ‖A x' - A y'‖ := by simp only [norm_sub_rev, NNReal.coe_mul]; ring +variable [CompleteSpace E] + section variable (f s) diff --git a/Mathlib/Analysis/Calculus/LHopital.lean b/Mathlib/Analysis/Calculus/LHopital.lean index 57597e1ff7984..0e18fd06df1c3 100644 --- a/Mathlib/Analysis/Calculus/LHopital.lean +++ b/Mathlib/Analysis/Calculus/LHopital.lean @@ -34,7 +34,7 @@ open Filter Set open scoped Filter Topology Pointwise -variable {a b : ℝ} (hab : a < b) {l : Filter ℝ} {f f' g g' : ℝ → ℝ} +variable {a b : ℝ} {l : Filter ℝ} {f f' g g' : ℝ → ℝ} /-! ## Interval-based versions @@ -46,7 +46,7 @@ to be satisfied on an explicitly-provided interval. namespace HasDerivAt -theorem lhopital_zero_right_on_Ioo (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x) x) +theorem lhopital_zero_right_on_Ioo (hab : a < b) (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x) x) (hgg' : ∀ x ∈ Ioo a b, HasDerivAt g (g' x) x) (hg' : ∀ x ∈ Ioo a b, g' x ≠ 0) (hfa : Tendsto f (𝓝[>] a) (𝓝 0)) (hga : Tendsto g (𝓝[>] a) (𝓝 0)) (hdiv : Tendsto (fun x => f' x / g' x) (𝓝[>] a) l) : @@ -89,7 +89,7 @@ theorem lhopital_zero_right_on_Ioo (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x try simp linarith [this] -theorem lhopital_zero_right_on_Ico (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x) x) +theorem lhopital_zero_right_on_Ico (hab : a < b) (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x) x) (hgg' : ∀ x ∈ Ioo a b, HasDerivAt g (g' x) x) (hcf : ContinuousOn f (Ico a b)) (hcg : ContinuousOn g (Ico a b)) (hg' : ∀ x ∈ Ioo a b, g' x ≠ 0) (hfa : f a = 0) (hga : g a = 0) (hdiv : Tendsto (fun x => f' x / g' x) (𝓝[>] a) l) : @@ -100,7 +100,7 @@ theorem lhopital_zero_right_on_Ico (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x · rw [← hga, ← nhdsWithin_Ioo_eq_nhdsWithin_Ioi hab] exact ((hcg a <| left_mem_Ico.mpr hab).mono Ioo_subset_Ico_self).tendsto -theorem lhopital_zero_left_on_Ioo (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x) x) +theorem lhopital_zero_left_on_Ioo (hab : a < b) (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x) x) (hgg' : ∀ x ∈ Ioo a b, HasDerivAt g (g' x) x) (hg' : ∀ x ∈ Ioo a b, g' x ≠ 0) (hfb : Tendsto f (𝓝[<] b) (𝓝 0)) (hgb : Tendsto g (𝓝[<] b) (𝓝 0)) (hdiv : Tendsto (fun x => f' x / g' x) (𝓝[<] b) l) : @@ -124,7 +124,7 @@ theorem lhopital_zero_left_on_Ioo (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x) unfold Function.comp at this simpa only [neg_neg] -theorem lhopital_zero_left_on_Ioc (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x) x) +theorem lhopital_zero_left_on_Ioc (hab : a < b) (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x) x) (hgg' : ∀ x ∈ Ioo a b, HasDerivAt g (g' x) x) (hcf : ContinuousOn f (Ioc a b)) (hcg : ContinuousOn g (Ioc a b)) (hg' : ∀ x ∈ Ioo a b, g' x ≠ 0) (hfb : f b = 0) (hgb : g b = 0) (hdiv : Tendsto (fun x => f' x / g' x) (𝓝[<] b) l) : @@ -195,7 +195,7 @@ end HasDerivAt namespace deriv -theorem lhopital_zero_right_on_Ioo (hdf : DifferentiableOn ℝ f (Ioo a b)) +theorem lhopital_zero_right_on_Ioo (hab : a < b) (hdf : DifferentiableOn ℝ f (Ioo a b)) (hg' : ∀ x ∈ Ioo a b, deriv g x ≠ 0) (hfa : Tendsto f (𝓝[>] a) (𝓝 0)) (hga : Tendsto g (𝓝[>] a) (𝓝 0)) (hdiv : Tendsto (fun x => (deriv f) x / (deriv g) x) (𝓝[>] a) l) : @@ -207,7 +207,7 @@ theorem lhopital_zero_right_on_Ioo (hdf : DifferentiableOn ℝ f (Ioo a b)) exact HasDerivAt.lhopital_zero_right_on_Ioo hab (fun x hx => (hdf x hx).hasDerivAt) (fun x hx => (hdg x hx).hasDerivAt) hg' hfa hga hdiv -theorem lhopital_zero_right_on_Ico (hdf : DifferentiableOn ℝ f (Ioo a b)) +theorem lhopital_zero_right_on_Ico (hab : a < b) (hdf : DifferentiableOn ℝ f (Ioo a b)) (hcf : ContinuousOn f (Ico a b)) (hcg : ContinuousOn g (Ico a b)) (hg' : ∀ x ∈ Ioo a b, (deriv g) x ≠ 0) (hfa : f a = 0) (hga : g a = 0) (hdiv : Tendsto (fun x => (deriv f) x / (deriv g) x) (𝓝[>] a) l) : @@ -218,7 +218,7 @@ theorem lhopital_zero_right_on_Ico (hdf : DifferentiableOn ℝ f (Ioo a b)) · rw [← hga, ← nhdsWithin_Ioo_eq_nhdsWithin_Ioi hab] exact ((hcg a <| left_mem_Ico.mpr hab).mono Ioo_subset_Ico_self).tendsto -theorem lhopital_zero_left_on_Ioo (hdf : DifferentiableOn ℝ f (Ioo a b)) +theorem lhopital_zero_left_on_Ioo (hab : a < b) (hdf : DifferentiableOn ℝ f (Ioo a b)) (hg' : ∀ x ∈ Ioo a b, (deriv g) x ≠ 0) (hfb : Tendsto f (𝓝[<] b) (𝓝 0)) (hgb : Tendsto g (𝓝[<] b) (𝓝 0)) (hdiv : Tendsto (fun x => (deriv f) x / (deriv g) x) (𝓝[<] b) l) : diff --git a/Mathlib/Analysis/Matrix.lean b/Mathlib/Analysis/Matrix.lean index af2323fa39d72..d692efa7722f7 100644 --- a/Mathlib/Analysis/Matrix.lean +++ b/Mathlib/Analysis/Matrix.lean @@ -389,7 +389,7 @@ protected def linftyOpNormedAlgebra [NormedField R] [SeminormedRing α] [NormedA section -variable [NormedDivisionRing α] [NormedAlgebra ℝ α] [CompleteSpace α] +variable [NormedDivisionRing α] [NormedAlgebra ℝ α] /-- Auxiliary construction; an element of norm 1 such that `a * unitOf a = ‖a‖`. -/ private def unitOf (a : α) : α := by classical exact if a = 0 then 1 else ‖a‖ • a⁻¹ @@ -401,7 +401,6 @@ private theorem norm_unitOf (a : α) : ‖unitOf a‖₊ = 1 := by · rw [← nnnorm_eq_zero] at h rw [nnnorm_smul, nnnorm_inv, nnnorm_norm, mul_inv_cancel h] -set_option tactic.skipAssignedInstances false in private theorem mul_unitOf (a : α) : a * unitOf a = algebraMap _ _ (‖a‖₊ : ℝ) := by simp only [unitOf, coe_nnnorm] split_ifs with h diff --git a/Mathlib/Condensed/Epi.lean b/Mathlib/Condensed/Epi.lean index 31b139daa9668..2b558cdb9c975 100644 --- a/Mathlib/Condensed/Epi.lean +++ b/Mathlib/Condensed/Epi.lean @@ -27,15 +27,15 @@ namespace Condensed variable (A : Type u') [Category.{v'} A] [ConcreteCategory.{v'} A] [HasFunctorialSurjectiveInjectiveFactorization A] + +variable {X Y : Condensed.{u} A} (f : X ⟶ Y) + +variable [(coherentTopology CompHaus).WEqualsLocallyBijective A] [HasSheafify (coherentTopology CompHaus) A] [(coherentTopology CompHaus.{u}).HasSheafCompose (CategoryTheory.forget A)] [Balanced (Sheaf (coherentTopology CompHaus) A)] - [PreservesFiniteProducts (CategoryTheory.forget A)] - [∀ (X : CompHausᵒᵖ), HasLimitsOfShape (StructuredArrow X Stonean.toCompHaus.op) A] - -variable {X Y : Condensed.{u} A} (f : X ⟶ Y) - + [PreservesFiniteProducts (CategoryTheory.forget A)] in lemma epi_iff_locallySurjective_on_compHaus : Epi f ↔ ∀ (S : CompHaus) (y : Y.val.obj ⟨S⟩), (∃ (S' : CompHaus) (φ : S' ⟶ S) (_ : Function.Surjective φ) (x : X.val.obj ⟨S'⟩), @@ -44,11 +44,13 @@ lemma epi_iff_locallySurjective_on_compHaus : Epi f ↔ regularTopology.isLocallySurjective_iff] simp_rw [((CompHaus.effectiveEpi_tfae _).out 0 2 :)] -variable [(extensiveTopology Stonean).WEqualsLocallyBijective A] +variable + [PreservesFiniteProducts (CategoryTheory.forget A)] + [∀ (X : CompHausᵒᵖ), HasLimitsOfShape (StructuredArrow X Stonean.toCompHaus.op) A] + [(extensiveTopology Stonean).WEqualsLocallyBijective A] [HasSheafify (extensiveTopology Stonean) A] [(extensiveTopology Stonean.{u}).HasSheafCompose (CategoryTheory.forget A)] - [Balanced (Sheaf (extensiveTopology Stonean) A)] - + [Balanced (Sheaf (extensiveTopology Stonean) A)] in lemma epi_iff_surjective_on_stonean : Epi f ↔ ∀ (S : Stonean), Function.Surjective (f.val.app (op S.compHaus)) := by rw [← (StoneanCompHaus.equivalence A).inverse.epi_map_iff_epi, diff --git a/Mathlib/FieldTheory/Cardinality.lean b/Mathlib/FieldTheory/Cardinality.lean index 9b73b9a872fb2..498e3b8d326c7 100644 --- a/Mathlib/FieldTheory/Cardinality.lean +++ b/Mathlib/FieldTheory/Cardinality.lean @@ -63,7 +63,7 @@ theorem Infinite.nonempty_field {α : Type u} [Infinite α] : Nonempty (Field α suffices #α = #K by obtain ⟨e⟩ := Cardinal.eq.1 this exact ⟨e.field⟩ - rw [← IsLocalization.card (MvPolynomial α <| ULift.{u} ℚ)⁰ K le_rfl] + rw [← IsLocalization.card K (MvPolynomial α <| ULift.{u} ℚ)⁰ le_rfl] apply le_antisymm · refine ⟨⟨fun a => MvPolynomial.monomial (Finsupp.single a 1) (1 : ULift.{u} ℚ), fun x y h => ?_⟩⟩ diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index f9d20aedccf3b..8d56a55352b79 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -1339,13 +1339,13 @@ section Topology variable {E : Type*} {𝕜 : Type*} [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] - [HasGroupoid M (contDiffGroupoid 0 I)] + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] /-- A finite-dimensional manifold modelled on a locally compact field (such as ℝ, ℂ or the `p`-adic numbers) is locally compact. -/ -lemma Manifold.locallyCompact_of_finiteDimensional [LocallyCompactSpace 𝕜] - [FiniteDimensional 𝕜 E] : LocallyCompactSpace M := by +lemma Manifold.locallyCompact_of_finiteDimensional + (I : ModelWithCorners 𝕜 E H) [LocallyCompactSpace 𝕜] [FiniteDimensional 𝕜 E] : + LocallyCompactSpace M := by have : ProperSpace E := FiniteDimensional.proper 𝕜 E have : LocallyCompactSpace H := I.locallyCompactSpace exact ChartedSpace.locallyCompactSpace H M diff --git a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean index edaadb55e2867..3b04b67dd0bca 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean @@ -56,9 +56,7 @@ def IsRefl (B : BilinForm R M) : Prop := LinearMap.IsRefl B namespace IsRefl -variable (H : B.IsRefl) - -theorem eq_zero : ∀ {x y : M}, B x y = 0 → B y x = 0 := fun {x y} => H x y +theorem eq_zero (H : B.IsRefl) : ∀ {x y : M}, B x y = 0 → B y x = 0 := fun {x y} => H x y protected theorem neg {B : BilinForm R₁ M₁} (hB : B.IsRefl) : (-B).IsRefl := fun x y => neg_eq_zero.mpr ∘ hB x y ∘ neg_eq_zero.mp @@ -87,12 +85,10 @@ def IsSymm (B : BilinForm R M) : Prop := LinearMap.IsSymm B namespace IsSymm -variable (H : B.IsSymm) - -protected theorem eq (x y : M) : B x y = B y x := +protected theorem eq (H : B.IsSymm) (x y : M) : B x y = B y x := H x y -theorem isRefl : B.IsRefl := fun x y H1 => H x y ▸ H1 +theorem isRefl (H : B.IsSymm) : B.IsRefl := fun x y H1 => H x y ▸ H1 protected theorem add {B₁ B₂ : BilinForm R M} (hB₁ : B₁.IsSymm) (hB₂ : B₂.IsSymm) : (B₁ + B₂).IsSymm := fun x y => (congr_arg₂ (· + ·) (hB₁ x y) (hB₂ x y) : _) @@ -397,6 +393,8 @@ lemma Nondegenerate.flip {B : BilinForm K V} (hB : B.Nondegenerate) : lemma nonDegenerateFlip_iff {B : BilinForm K V} : B.flip.Nondegenerate ↔ B.Nondegenerate := ⟨Nondegenerate.flip, Nondegenerate.flip⟩ +end FiniteDimensional + section DualBasis variable {ι : Type*} [DecidableEq ι] [Finite ι] @@ -410,10 +408,12 @@ noncomputable def dualBasis (B : BilinForm K V) (hB : B.Nondegenerate) (b : Basi b.dualBasis.map (B.toDual hB).symm @[simp] -theorem dualBasis_repr_apply (B : BilinForm K V) (hB : B.Nondegenerate) (b : Basis ι K V) (x i) : +theorem dualBasis_repr_apply + (B : BilinForm K V) (hB : B.Nondegenerate) (b : Basis ι K V) (x i) : (B.dualBasis hB b).repr x i = B x (b i) := by rw [dualBasis, Basis.map_repr, LinearEquiv.symm_symm, LinearEquiv.trans_apply, - Basis.dualBasis_repr, toDual_def] + Basis.dualBasis_repr] + rfl theorem apply_dualBasis_left (B : BilinForm K V) (hB : B.Nondegenerate) (b : Basis ι K V) (i j) : B (B.dualBasis hB b i) (b j) = if j = i then 1 else 0 := by @@ -426,7 +426,8 @@ theorem apply_dualBasis_right (B : BilinForm K V) (hB : B.Nondegenerate) (sym : rw [sym.eq, apply_dualBasis_left] @[simp] -lemma dualBasis_dualBasis_flip (B : BilinForm K V) (hB : B.Nondegenerate) {ι} +lemma dualBasis_dualBasis_flip [FiniteDimensional K V] + (B : BilinForm K V) (hB : B.Nondegenerate) {ι : Type*} [Finite ι] [DecidableEq ι] (b : Basis ι K V) : B.dualBasis hB (B.flip.dualBasis hB.flip b) = b := by ext i @@ -450,6 +451,8 @@ end DualBasis section LinearAdjoints +variable [FiniteDimensional K V] + /-- Given bilinear forms `B₁, B₂` where `B₂` is nondegenerate, `symmCompOfNondegenerate` is the linear map `B₂ ∘ B₁`. -/ noncomputable def symmCompOfNondegenerate (B₁ B₂ : BilinForm K V) (b₂ : B₂.Nondegenerate) : @@ -490,8 +493,6 @@ theorem isAdjointPair_iff_eq_of_nondegenerate (B : BilinForm K V) (b : B.Nondege end LinearAdjoints -end FiniteDimensional - end BilinForm end LinearMap diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean index e103e5fba4071..b1fb223b3f0ba 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean @@ -35,9 +35,10 @@ theorem IsIntegral.map_of_comp_eq {R S T U : Type*} [CommRing R] [Ring S] section variable {A B : Type*} [Ring A] [Ring B] [Algebra R A] [Algebra R B] -variable (f : A →ₐ[R] B) (hf : Function.Injective f) +variable (f : A →ₐ[R] B) -theorem Algebra.IsIntegral.of_injective [Algebra.IsIntegral R B] : Algebra.IsIntegral R A := +theorem Algebra.IsIntegral.of_injective (hf : Function.Injective f) [Algebra.IsIntegral R B] : + Algebra.IsIntegral R A := ⟨fun _ ↦ (isIntegral_algHom_iff f hf).mp (isIntegral _)⟩ end diff --git a/Mathlib/RingTheory/LocalRing/Module.lean b/Mathlib/RingTheory/LocalRing/Module.lean index 40342f8b2c8d6..f850665a24192 100644 --- a/Mathlib/RingTheory/LocalRing/Module.lean +++ b/Mathlib/RingTheory/LocalRing/Module.lean @@ -28,7 +28,7 @@ This file gathers various results about finite modules over a local ring `(R, `l` is a split injection if and only if `k ⊗ l` is a (split) injection. -/ -variable {R S} [CommRing R] [CommRing S] [Algebra R S] [LocalRing R] +variable {R S} [CommRing R] [CommRing S] [Algebra R S] section @@ -41,10 +41,11 @@ local notation "k" => ResidueField R local notation "𝔪" => maximalIdeal R variable {P} [AddCommGroup P] [Module R P] (f : M →ₗ[R] N) (g : N →ₗ[R] P) -variable (hg : Surjective g) (h : Exact f g) namespace LocalRing +variable [LocalRing R] + theorem map_mkQ_eq {N₁ N₂ : Submodule R M} (h : N₁ ≤ N₂) (h' : N₂.FG) : N₁.map (Submodule.mkQ (𝔪 • N₂)) = N₂.map (Submodule.mkQ (𝔪 • N₂)) ↔ N₁ = N₂ := by constructor @@ -135,6 +136,8 @@ theorem lTensor_injective_of_exact_of_exact_of_rTensor_injective namespace Module +variable [LocalRing R] + /-- If `M` is a finitely presented module over a local ring `(R, 𝔪)` such that `m ⊗ M → M` is injective, then `M` is free. @@ -199,7 +202,7 @@ theorem free_of_flat_of_localRing [Module.FinitePresentation R P] [Module.Flat R If `M → N → P → 0` is a presentation of `P` over a local ring `(R, 𝔪, k)` with `M` finite and `N` finite free, then injectivity of `k ⊗ M → k ⊗ N` implies that `P` is free. -/ -theorem free_of_lTensor_residueField_injective +theorem free_of_lTensor_residueField_injective (hg : Surjective g) (h : Exact f g) [Module.Finite R M] [Module.Finite R N] [Module.Free R N] (hf : Function.Injective (f.lTensor k)) : Module.Free R P := by @@ -219,7 +222,7 @@ Given a linear map `l : M → N` over a local ring `(R, 𝔪, k)` with `M` finite and `N` finite free, `l` is a split injection if and only if `k ⊗ l` is a (split) injection. -/ -theorem LocalRing.split_injective_iff_lTensor_residueField_injective +theorem LocalRing.split_injective_iff_lTensor_residueField_injective [LocalRing R] [Module.Finite R M] [Module.Finite R N] [Module.Free R N] (l : M →ₗ[R] N) : (∃ l', l' ∘ₗ l = LinearMap.id) ↔ Function.Injective (l.lTensor (ResidueField R)) := by constructor diff --git a/Mathlib/RingTheory/Localization/Cardinality.lean b/Mathlib/RingTheory/Localization/Cardinality.lean index 3c4e0d9d2850d..bc4e2828f5f1b 100644 --- a/Mathlib/RingTheory/Localization/Cardinality.lean +++ b/Mathlib/RingTheory/Localization/Cardinality.lean @@ -29,11 +29,10 @@ universe u v namespace IsLocalization -variable {R : Type u} [CommRing R] (S : Submonoid R) {L : Type u} [CommRing L] [Algebra R L] - [IsLocalization S L] +variable {R : Type u} [CommRing R] {L : Type u} [CommRing L] [Algebra R L] /-- A localization always has cardinality less than or equal to the base ring. -/ -theorem card_le : #L ≤ #R := by +theorem card_le (S : Submonoid R) [IsLocalization S L] : #L ≤ #R := by classical cases fintypeOrInfinite R · exact Cardinal.mk_le_of_surjective (IsArtinianRing.localization_surjective S _) @@ -48,7 +47,7 @@ theorem card_le : #L ≤ #R := by variable (L) /-- If you do not localize at any zero-divisors, localization preserves cardinality. -/ -theorem card (hS : S ≤ R⁰) : #R = #L := +theorem card (S : Submonoid R) [IsLocalization S L] (hS : S ≤ R⁰) : #R = #L := (Cardinal.mk_le_of_injective (IsLocalization.injective L hS)).antisymm (card_le S) end IsLocalization From a2aaae87227590693a1d9bb9a4696956c6ebd8d2 Mon Sep 17 00:00:00 2001 From: Shuhao Song Date: Wed, 7 Aug 2024 07:08:07 +0000 Subject: [PATCH 082/359] fix: `withMainContext` in `type_check` tactic (#15533) Fixing a problem about `type_check` tactic: If you don't use `withMainContext`, you can't use local hypotheses and definitions produced by tactics before the `type_check` tactic, only can use hypotheses written on `theorem` header. --- Mathlib/Tactic/TypeCheck.lean | 12 +++++++----- test/TypeCheck.lean | 5 ++++- 2 files changed, 11 insertions(+), 6 deletions(-) diff --git a/Mathlib/Tactic/TypeCheck.lean b/Mathlib/Tactic/TypeCheck.lean index 63fbf103f9658..b72e3f2079b0f 100644 --- a/Mathlib/Tactic/TypeCheck.lean +++ b/Mathlib/Tactic/TypeCheck.lean @@ -4,17 +4,19 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Lean.Elab.Tactic.Basic +import Lean.Elab.SyntheticMVars /-! # The `type_check` tactic Define the `type_check` tactic: it type checks a given expression, and traces its type. -/ -open Lean Meta +open Lean Elab Meta /-- Type check the given expression, and trace its type. -/ elab tk:"type_check " e:term : tactic => do - let e ← Lean.Elab.Term.elabTerm e none - check e - let type ← inferType e - Lean.logInfoAt tk f!"{← Lean.instantiateMVars type}" + Tactic.withMainContext do + let e ← Term.elabTermAndSynthesize e none + check e + let type ← inferType e + Lean.logInfoAt tk m!"{← Lean.instantiateMVars type}" diff --git a/test/TypeCheck.lean b/test/TypeCheck.lean index a5498d874644f..c707c493a6bd8 100644 --- a/test/TypeCheck.lean +++ b/test/TypeCheck.lean @@ -17,7 +17,9 @@ info: Prop --- info: Prop --- -info: Nat -> Nat +info: Nat → Nat +--- +info: List Nat -/ #guard_msgs in example : True := by @@ -28,5 +30,6 @@ example : True := by type_check (True : _) -- Prop type_check ∀ x y : Nat, x = y -- Prop type_check fun x : Nat => 2 * x + 1 -- Nat -> Nat + type_check [1] fail_if_success type_check wrong trivial From 99de650cf0efc373f69882333aacc7df2db18dca Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 7 Aug 2024 07:08:08 +0000 Subject: [PATCH 083/359] chore: backports for leanprover/lean4#4814 (part 27) (#15543) --- Mathlib/AlgebraicGeometry/Morphisms/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean index e62765a7496f5..e92624a6197d0 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean @@ -542,7 +542,7 @@ instance (priority := 900) : IsLocalAtTarget P := by exact of_isPullback (.of_hasPullback _ _) this open AffineTargetMorphismProperty in -protected theorem iff : +protected theorem iff {P : MorphismProperty Scheme} {Q : AffineTargetMorphismProperty} : HasAffineProperty P Q ↔ IsLocalAtTarget P ∧ Q = of P := ⟨fun _ ↦ ⟨inferInstance, ext fun _ _ _ ↦ iff_of_isAffine.symm⟩, fun ⟨_, e⟩ ↦ e ▸ of_isLocalAtTarget P⟩ From dd6114f627e2881ce72229be138f9f0de9b76a28 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 7 Aug 2024 07:08:09 +0000 Subject: [PATCH 084/359] feat(Topology/Maps): preimage of a dense set (#15552) ... under an open map is dense. --- Mathlib/Topology/Maps/Basic.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index f069635fbd1f8..568d2b9c373a5 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -381,6 +381,11 @@ theorem isOpenMap_iff_interior : IsOpenMap f ↔ ∀ s, f '' interior s ⊆ inte protected theorem Inducing.isOpenMap (hi : Inducing f) (ho : IsOpen (range f)) : IsOpenMap f := IsOpenMap.of_nhds_le fun _ => (hi.map_nhds_of_mem _ <| IsOpen.mem_nhds ho <| mem_range_self _).ge +/-- Preimage of a dense set under an open map is dense. -/ +protected theorem Dense.preimage {s : Set Y} (hs : Dense s) (hf : IsOpenMap f) : + Dense (f ⁻¹' s) := fun x ↦ + hf.preimage_closure_subset_closure_preimage <| hs (f x) + end OpenMap section IsClosedMap From 9efc41778779cc91ab76269e1b5877bab5672382 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 7 Aug 2024 07:08:10 +0000 Subject: [PATCH 085/359] feat(RelSeries): {last,head}_reverse (#15556) from the Carleson project. --- Mathlib/Order/RelSeries.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index 802ccaa4fc078..2438ce2f05ea8 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -387,6 +387,12 @@ def reverse (p : RelSeries r) : RelSeries (fun (a b : α) ↦ r b a) where @[simp] lemma reverse_apply (p : RelSeries r) (i : Fin (p.length + 1)) : p.reverse i = p i.rev := rfl +@[simp] lemma last_reverse (p : RelSeries r) : p.reverse.last = p.head := by + simp [RelSeries.last, RelSeries.head] + +@[simp] lemma head_reverse (p : RelSeries r) : p.reverse.head = p.last := by + simp [RelSeries.last, RelSeries.head] + /-- Given a series `a₀ -r→ a₁ -r→ ... -r→ aₙ` and an `a` such that `a₀ -r→ a` holds, there is a series of length `n+1`: `a -r→ a₀ -r→ a₁ -r→ ... -r→ aₙ`. From ad4663c3145dd1f46a643edfc85e87713fc2ca1e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 7 Aug 2024 08:06:51 +0000 Subject: [PATCH 086/359] feat: norm_num ext for `Int.floor` (#13647) I tried this a while ago in https://github.com/leanprover-community/mathlib/pull/16502; this is now just one of the handlers from that PR. --- Mathlib/Data/Rat/Floor.lean | 44 ++++++++++++++++++++++++++++++++++++- test/norm_num_ext.lean | 12 ++++++++++ 2 files changed, 55 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Rat/Floor.lean b/Mathlib/Data/Rat/Floor.lean index 7e7f210c19ae8..9c7accdfc0acb 100644 --- a/Mathlib/Data/Rat/Floor.lean +++ b/Mathlib/Data/Rat/Floor.lean @@ -49,7 +49,7 @@ instance : FloorRing ℚ := protected theorem floor_def {q : ℚ} : ⌊q⌋ = q.num / q.den := Rat.floor_def' q -theorem floor_int_div_nat_eq_div {n : ℤ} {d : ℕ} : ⌊(↑n : ℚ) / (↑d : ℚ)⌋ = n / (↑d : ℤ) := by +theorem floor_int_div_nat_eq_div (n : ℤ) (d : ℕ) : ⌊(↑n : ℚ) / (↑d : ℚ)⌋ = n / (↑d : ℤ) := by rw [Rat.floor_def] obtain rfl | hd := @eq_zero_or_pos _ _ d · simp @@ -78,6 +78,48 @@ theorem round_cast (x : ℚ) : round (x : α) = round x := by theorem cast_fract (x : ℚ) : (↑(fract x) : α) = fract (x : α) := by simp only [fract, cast_sub, cast_intCast, floor_cast] +section NormNum + +open Mathlib.Meta.NormNum Qq + +theorem isNat_intFloor {R} [LinearOrderedRing R] [FloorRing R] (r : R) (m : ℕ) : + IsNat r m → IsNat ⌊r⌋ m := by rintro ⟨⟨⟩⟩; exact ⟨by simp⟩ + +theorem isInt_intFloor {R} [LinearOrderedRing R] [FloorRing R] (r : R) (m : ℤ) : + IsInt r m → IsInt ⌊r⌋ m := by rintro ⟨⟨⟩⟩; exact ⟨by simp⟩ + +theorem isInt_intFloor_ofIsRat (r : α) (n : ℤ) (d : ℕ) : + IsRat r n d → IsInt ⌊r⌋ (n / d) := by + rintro ⟨inv, rfl⟩ + constructor + simp only [invOf_eq_inv, ← div_eq_mul_inv, Int.cast_id] + rw [← floor_int_div_nat_eq_div n d, ← floor_cast (α := α), Rat.cast_div, + cast_intCast, cast_natCast] + +/-- `norm_num` extension for `Int.floor` -/ +@[norm_num ⌊_⌋] +def evalIntFloor : NormNumExt where eval {u αZ} e := do + match u, αZ, e with + | 0, ~q(ℤ), ~q(@Int.floor $α $instR $instF $x) => + match ← derive x with + | .isBool .. => failure + | .isNat sα nb pb => do + assertInstancesCommute + return .isNat q(inferInstance) _ q(isNat_intFloor $x _ $pb) + | .isNegNat sα nb pb => do + assertInstancesCommute + -- floor always keeps naturals negative, so we can shortcut `.isInt` + return .isNegNat q(inferInstance) _ q(isInt_intFloor _ _ $pb) + | .isRat dα q n d h => do + let _i ← synthInstanceQ q(LinearOrderedField $α) + assertInstancesCommute + have z : Q(ℤ) := mkRawIntLit ⌊q⌋ + letI : $z =Q ⌊$n / $d⌋ := ⟨⟩ + return .isInt q(inferInstance) z ⌊q⌋ q(isInt_intFloor_ofIsRat _ $n $d $h) + | _, _, _ => failure + +end NormNum + end Rat theorem Int.mod_nat_eq_sub_mul_floor_rat_div {n : ℤ} {d : ℕ} : n % d = n - d * ⌊(n : ℚ) / d⌋ := by diff --git a/test/norm_num_ext.lean b/test/norm_num_ext.lean index e5e9341121490..4db23b8b748f9 100644 --- a/test/norm_num_ext.lean +++ b/test/norm_num_ext.lean @@ -389,6 +389,18 @@ example : ∏ i ∈ Finset.range 9, Nat.sqrt (i + 1) = 96 := by norm_num1 end big_operators +section floor + +variable (R : Type*) [LinearOrderedRing R] [FloorRing R] +variable (K : Type*) [LinearOrderedField K] [FloorRing K] + +example : ⌊(-1 : R)⌋ = -1 := by norm_num +example : ⌊(2 : R)⌋ = 2 := by norm_num +example : ⌊(15 / 16 : K)⌋ + 1 = 1 := by norm_num +example : ⌊(-15 / 16 : K)⌋ + 1 = 0 := by norm_num + +end floor + section jacobi -- Jacobi and Legendre symbols From 9a5c301bfa32ef1e1a1d493d55b408c27dd5f75a Mon Sep 17 00:00:00 2001 From: Lorenzo Luccioli Date: Wed, 7 Aug 2024 08:06:52 +0000 Subject: [PATCH 087/359] feat (Probability.Kernel): Composition with a constant kernel. (#15461) Add `Kernel.const_comp` and `Kernel.const_comp'`. --- Mathlib/Probability/Kernel/Composition.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index a8650f58e53c3..412687e4436d6 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -1033,6 +1033,17 @@ theorem comp_deterministic_eq_comap (κ : Kernel α β) (hg : Measurable g) : simp_rw [comap_apply' _ _ _ s, comp_apply' _ _ _ hs, deterministic_apply hg a, lintegral_dirac' _ (Kernel.measurable_coe κ hs)] +lemma const_comp (μ : Measure γ) (κ : Kernel α β) : + const β μ ∘ₖ κ = fun a ↦ (κ a) Set.univ • μ := by + ext _ _ hs + simp_rw [comp_apply' _ _ _ hs, const_apply, MeasureTheory.lintegral_const, Measure.smul_apply, + smul_eq_mul, mul_comm] + +@[simp] +lemma const_comp' (μ : Measure γ) (κ : Kernel α β) [IsMarkovKernel κ] : + const β μ ∘ₖ κ = const α μ := by + ext; simp_rw [const_comp, measure_univ, one_smul, const_apply] + end Comp section Prod From a38c21c21b5a747fb2ae8b1ff6e8adba89cfd207 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Wed, 7 Aug 2024 09:19:24 +0000 Subject: [PATCH 088/359] chore: backports for leanprover/lean4#4814 (part 28) (#15549) --- Mathlib/FieldTheory/IsPerfectClosure.lean | 31 +++++++++++++---------- 1 file changed, 18 insertions(+), 13 deletions(-) diff --git a/Mathlib/FieldTheory/IsPerfectClosure.lean b/Mathlib/FieldTheory/IsPerfectClosure.lean index b9d456514b0a7..70d3708472eb5 100644 --- a/Mathlib/FieldTheory/IsPerfectClosure.lean +++ b/Mathlib/FieldTheory/IsPerfectClosure.lean @@ -130,8 +130,7 @@ variable {K L M N : Type*} section CommSemiring variable [CommSemiring K] [CommSemiring L] [CommSemiring M] - (i : K →+* L) (j : K →+* M) (f : L →+* M) - (p : ℕ) [ExpChar K p] [ExpChar L p] [ExpChar M p] + (i : K →+* L) (j : K →+* M) (f : L →+* M) (p : ℕ) /-- If `i : K →+* L` is a ring homomorphism of characteristic `p` rings, then it is called `p`-radical if the following conditions are satisfied: @@ -186,17 +185,18 @@ Our definition makes it synonymous to `IsPRadical` if `PerfectRing L p` is prese that you need to write `[PerfectRing L p] [IsPerfectClosure i p]`. This is similar to `PerfectRing` which has `ExpChar` as a prerequisite. -/ @[nolint unusedArguments] -abbrev IsPerfectClosure [PerfectRing L p] := IsPRadical i p +abbrev IsPerfectClosure [ExpChar L p] [PerfectRing L p] := IsPRadical i p /-- If `i : K →+* L` is a ring homomorphism of exponential characteristic `p` rings, such that `L` is perfect, then the `p`-nilradical of `K` is contained in the kernel of `i`. -/ -theorem RingHom.pNilradical_le_ker_of_perfectRing [PerfectRing L p] : +theorem RingHom.pNilradical_le_ker_of_perfectRing [ExpChar L p] [PerfectRing L p] : pNilradical K p ≤ RingHom.ker i := fun x h ↦ by obtain ⟨n, h⟩ := mem_pNilradical.1 h replace h := congr((iterateFrobeniusEquiv L p n).symm (i $h)) rwa [map_pow, ← iterateFrobenius_def, ← iterateFrobeniusEquiv_apply, RingEquiv.symm_apply_apply, map_zero, map_zero] at h +variable [ExpChar L p] [ExpChar M p] in theorem IsPerfectClosure.ker_eq [PerfectRing L p] [IsPerfectClosure i p] : RingHom.ker i = pNilradical K p := IsPRadical.ker_le'.antisymm (i.pNilradical_le_ker_of_perfectRing p) @@ -206,7 +206,7 @@ namespace PerfectRing /- NOTE: To define `PerfectRing.lift_aux`, only the `IsPRadical.pow_mem` is required, but not `IsPRadical.ker_le`. But in order to use typeclass, here we require the whole `IsPRadical`. -/ -variable [PerfectRing M p] [IsPRadical i p] +variable [ExpChar M p] [PerfectRing M p] [IsPRadical i p] theorem lift_aux (x : L) : ∃ y : ℕ × K, i y.2 = x ^ p ^ y.1 := by obtain ⟨n, y, h⟩ := IsPRadical.pow_mem i p x @@ -220,12 +220,13 @@ def liftAux (x : L) : M := (iterateFrobeniusEquiv M p (Classical.choose (lift_au (j (Classical.choose (lift_aux i p x)).2) @[simp] -theorem liftAux_self_apply [PerfectRing L p] (x : L) : liftAux i i p x = x := by +theorem liftAux_self_apply [ExpChar L p] [PerfectRing L p] (x : L) : liftAux i i p x = x := by rw [liftAux, Classical.choose_spec (lift_aux i p x), ← iterateFrobenius_def, ← iterateFrobeniusEquiv_apply, RingEquiv.symm_apply_apply] @[simp] -theorem liftAux_self [PerfectRing L p] : liftAux i i p = id := funext (liftAux_self_apply i p) +theorem liftAux_self [ExpChar L p] [PerfectRing L p] : liftAux i i p = id := + funext (liftAux_self_apply i p) @[simp] theorem liftAux_id_apply (x : K) : liftAux (RingHom.id K) j p x = j x := by @@ -244,7 +245,8 @@ section CommRing variable [CommRing K] [CommRing L] [CommRing M] [CommRing N] (i : K →+* L) (j : K →+* M) (k : K →+* N) (f : L →+* M) (g : L →+* N) - (p : ℕ) [ExpChar K p] [ExpChar L p] [ExpChar M p] [ExpChar N p] + (p : ℕ) [ExpChar M p] + namespace IsPRadical @@ -279,7 +281,7 @@ end IsPRadical namespace PerfectRing -variable [PerfectRing M p] [IsPRadical i p] +variable [ExpChar K p] [PerfectRing M p] [IsPRadical i p] /-- If `i : K →+* L` and `j : K →+* M` are ring homomorphisms of characteristic `p` rings, such that `i` is `p`-radical, and `M` is a perfect ring, then `PerfectRing.liftAux` is well-defined. -/ @@ -300,6 +302,8 @@ theorem liftAux_apply (x : L) (n : ℕ) (y : K) (h : i y = x ^ p ^ n) : ← sub_eq_zero, ← map_pow, ← map_pow, ← map_sub, add_comm m, add_comm m, pow_add, pow_mul, pow_add, pow_mul, ← sub_pow_expChar_pow, h, map_zero] +variable [ExpChar L p] + /-- If `i : K →+* L` and `j : K →+* M` are ring homomorphisms of characteristic `p` rings, such that `i` is `p`-radical, and `M` is a perfect ring, then `PerfectRing.liftAux` is a ring homomorphism. This is similar to `IsAlgClosed.lift` and `IsSepClosed.lift`. -/ @@ -382,7 +386,7 @@ theorem liftEquiv_id : liftEquiv M (RingHom.id K) p = Equiv.refl _ := section comp -variable [PerfectRing N p] [IsPRadical j p] +variable [ExpChar N p] [PerfectRing N p] [IsPRadical j p] @[simp] theorem lift_comp_lift : (lift j k p).comp (lift i j p) = lift i k p := @@ -404,7 +408,7 @@ end comp section liftEquiv_comp -variable [IsPRadical g p] [IsPRadical (g.comp i) p] +variable [ExpChar N p] [IsPRadical g p] [IsPRadical (g.comp i) p] @[simp] theorem lift_lift : lift g (lift i j p) p = lift (g.comp i) j p := by @@ -429,7 +433,8 @@ end PerfectRing namespace IsPerfectClosure -variable [PerfectRing L p] [IsPerfectClosure i p] [PerfectRing M p] [IsPerfectClosure j p] +variable [ExpChar K p] [ExpChar L p] [PerfectRing L p] [IsPerfectClosure i p] [PerfectRing M p] + [IsPerfectClosure j p] /-- If `L` and `M` are both perfect closures of `K`, then there is a ring isomorphism `L ≃+* M`. This is similar to `IsAlgClosure.equiv` and `IsSepClosure.equiv`. -/ @@ -472,7 +477,7 @@ theorem equiv_comp : RingHom.comp (equiv i j p) i = j := section comp -variable [PerfectRing N p] [IsPerfectClosure k p] +variable [ExpChar N p] [PerfectRing N p] [IsPerfectClosure k p] @[simp] theorem equiv_comp_equiv_apply (x : L) : From 30dd0d401819f33578e7fb75222bfcd3da814806 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 7 Aug 2024 10:57:09 +0000 Subject: [PATCH 089/359] chore: backports for leanprover/lean4#4814 (part 25) (#15537) --- Mathlib/Algebra/Lie/Submodule.lean | 88 ++++++----- .../EllipticCurve/Jacobian.lean | 6 +- .../InnerProductSpace/Orientation.lean | 4 +- .../Normed/Algebra/TrivSqZeroExt.lean | 8 +- .../Normed/Algebra/UnitizationL1.lean | 4 +- Mathlib/Analysis/Normed/Lp/PiLp.lean | 4 +- Mathlib/Analysis/NormedSpace/DualNumber.lean | 2 +- .../Combinatorics/Additive/Corner/Roth.lean | 6 +- Mathlib/FieldTheory/Perfect.lean | 7 +- .../Geometry/Manifold/InteriorBoundary.lean | 2 +- .../QuadraticForm/TensorProduct.lean | 8 +- Mathlib/LinearAlgebra/Trace.lean | 10 +- Mathlib/MeasureTheory/Constructions/Pi.lean | 11 +- .../Constructions/Prod/Basic.lean | 6 +- Mathlib/MeasureTheory/Group/Measure.lean | 3 + .../Measure/HasOuterApproxClosed.lean | 4 +- Mathlib/RingTheory/Kaehler/Basic.lean | 13 +- Mathlib/RingTheory/Unramified/Finite.lean | 53 +++---- Mathlib/RingTheory/WittVector/Basic.lean | 16 +- Mathlib/RingTheory/WittVector/InitTail.lean | 10 +- .../RingTheory/WittVector/Teichmuller.lean | 4 +- .../RingTheory/WittVector/Verschiebung.lean | 9 +- .../MetricSpace/GromovHausdorffRealized.lean | 143 ++++++++++-------- 23 files changed, 246 insertions(+), 175 deletions(-) diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index 3bc21e9da307d..08978bd95d8ba 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -36,8 +36,8 @@ universe u v w w₁ w₂ section LieSubmodule variable (R : Type u) (L : Type v) (M : Type w) -variable [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] -variable [LieRingModule L M] [LieModule R L M] +variable [CommRing R] [LieRing L] [AddCommGroup M] [Module R M] +variable [LieRingModule L M] /-- A Lie submodule of a Lie module is a submodule that is closed under the Lie bracket. This is a sufficient condition for the subset itself to form a Lie module. -/ @@ -169,10 +169,6 @@ instance {S : Type*} [Semiring S] [SMul S R] [SMul Sᵐᵒᵖ R] [Module S M] [M [IsScalarTower S R M] [IsScalarTower Sᵐᵒᵖ R M] [IsCentralScalar S M] : IsCentralScalar S N := N.toSubmodule.isCentralScalar -instance instLieModule : LieModule R L N where - lie_smul := by intro t x y; apply SetCoe.ext; apply lie_smul - smul_lie := by intro t x y; apply SetCoe.ext; apply smul_lie - @[simp, norm_cast] theorem coe_zero : ((0 : N) : M) = (0 : M) := rfl @@ -197,12 +193,19 @@ theorem coe_smul (t : R) (m : N) : (↑(t • m) : M) = t • (m : M) := theorem coe_bracket (x : L) (m : N) : (↑⁅x, m⁆ : M) = ⁅x, ↑m⁆ := rfl +variable [LieAlgebra R L] [LieModule R L M] + +instance instLieModule : LieModule R L N where + lie_smul := by intro t x y; apply SetCoe.ext; apply lie_smul + smul_lie := by intro t x y; apply SetCoe.ext; apply smul_lie + instance [Subsingleton M] : Unique (LieSubmodule R L M) := ⟨⟨0⟩, fun _ ↦ (coe_toSubmodule_eq_iff _ _).mp (Subsingleton.elim _ _)⟩ end LieSubmodule section LieIdeal +variable [LieAlgebra R L] [LieModule R L M] /-- An ideal of a Lie algebra is a Lie submodule of the Lie algebra as a Lie module over itself. -/ abbrev LieIdeal := @@ -265,6 +268,7 @@ theorem Submodule.exists_lieSubmodule_coe_eq_iff (p : Submodule R M) : namespace LieSubalgebra variable {L} +variable [LieAlgebra R L] variable (K : LieSubalgebra R L) /-- Given a Lie subalgebra `K ⊆ L`, if we view `L` as a `K`-module by restriction, it contains @@ -301,9 +305,9 @@ end LieSubmodule namespace LieSubmodule variable {R : Type u} {L : Type v} {M : Type w} -variable [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] -variable [LieRingModule L M] [LieModule R L M] -variable (N N' : LieSubmodule R L M) (I J : LieIdeal R L) +variable [CommRing R] [LieRing L] [AddCommGroup M] [Module R M] +variable [LieRingModule L M] +variable (N N' : LieSubmodule R L M) section LatticeStructure @@ -741,9 +745,9 @@ end LieSubmodule section LieSubmoduleMapAndComap variable {R : Type u} {L : Type v} {L' : Type w₂} {M : Type w} {M' : Type w₁} -variable [CommRing R] [LieRing L] [LieAlgebra R L] [LieRing L'] [LieAlgebra R L'] -variable [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] -variable [AddCommGroup M'] [Module R M'] [LieRingModule L M'] [LieModule R L M'] +variable [CommRing R] [LieRing L] [LieRing L'] [LieAlgebra R L'] +variable [AddCommGroup M] [Module R M] [LieRingModule L M] +variable [AddCommGroup M'] [Module R M'] [LieRingModule L M'] namespace LieSubmodule @@ -884,6 +888,7 @@ Submodules. -/ end LieSubmodule namespace LieIdeal +variable [LieAlgebra R L] [LieModule R L M] [LieModule R L M'] variable (f : L →ₗ⁅R⁆ L') (I I₂ : LieIdeal R L) (J : LieIdeal R L') @@ -988,7 +993,7 @@ instance subsingleton_of_bot : Subsingleton (LieIdeal R (⊥ : LieIdeal R L)) := end LieIdeal namespace LieHom - +variable [LieAlgebra R L] [LieModule R L M] [LieModule R L M'] variable (f : L →ₗ⁅R⁆ L') (I : LieIdeal R L) (J : LieIdeal R L') /-- The kernel of a morphism of Lie algebras, as an ideal in the domain. -/ @@ -1090,7 +1095,7 @@ theorem isIdealMorphism_of_surjective (h : Function.Surjective f) : f.IsIdealMor end LieHom namespace LieIdeal - +variable [LieAlgebra R L] [LieModule R L M] [LieModule R L M'] variable {f : L →ₗ⁅R⁆ L'} {I : LieIdeal R L} {J : LieIdeal R L'} @[simp] @@ -1223,9 +1228,9 @@ end LieSubmoduleMapAndComap namespace LieModuleHom variable {R : Type u} {L : Type v} {M : Type w} {N : Type w₁} -variable [CommRing R] [LieRing L] [LieAlgebra R L] -variable [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] -variable [AddCommGroup N] [Module R N] [LieRingModule L N] [LieModule R L N] +variable [CommRing R] [LieRing L] +variable [AddCommGroup M] [Module R M] [LieRingModule L M] +variable [AddCommGroup N] [Module R N] [LieRingModule L N] variable (f : M →ₗ⁅R,L⁆ N) /-- The kernel of a morphism of Lie algebras, as an ideal in the domain. -/ @@ -1299,8 +1304,8 @@ end LieModuleHom namespace LieSubmodule variable {R : Type u} {L : Type v} {M : Type w} -variable [CommRing R] [LieRing L] [LieAlgebra R L] -variable [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] +variable [CommRing R] [LieRing L] +variable [AddCommGroup M] [Module R M] [LieRingModule L M] variable (N : LieSubmodule R L M) @[simp] @@ -1339,8 +1344,30 @@ end LieSubmodule section TopEquiv -variable {R : Type u} {L : Type v} -variable [CommRing R] [LieRing L] [LieAlgebra R L] +variable (R : Type u) (L : Type v) +variable [CommRing R] [LieRing L] + +variable (M : Type*) [AddCommGroup M] [Module R M] [LieRingModule L M] + +/-- The natural equivalence between the 'top' Lie submodule and the enclosing Lie module. -/ +def LieModuleEquiv.ofTop : (⊤ : LieSubmodule R L M) ≃ₗ⁅R,L⁆ M := + { LinearEquiv.ofTop ⊤ rfl with + map_lie' := rfl } + +variable {R L} + +-- This lemma has always been bad, but leanprover/lean4#2644 made `simp` start noticing +@[simp, nolint simpNF] lemma LieModuleEquiv.ofTop_apply (x : (⊤ : LieSubmodule R L M)) : + LieModuleEquiv.ofTop R L M x = x := + rfl + +@[simp] lemma LieModuleEquiv.range_coe {M' : Type*} + [AddCommGroup M'] [Module R M'] [LieRingModule L M'] (e : M ≃ₗ⁅R,L⁆ M') : + LieModuleHom.range (e : M →ₗ⁅R,L⁆ M') = ⊤ := by + rw [LieModuleHom.range_eq_top] + exact e.surjective + +variable [LieAlgebra R L] [LieModule R L M] /-- The natural equivalence between the 'top' Lie subalgebra and the enclosing Lie algebra. @@ -1366,23 +1393,4 @@ def LieIdeal.topEquiv : (⊤ : LieIdeal R L) ≃ₗ⁅R⁆ L := theorem LieIdeal.topEquiv_apply (x : (⊤ : LieIdeal R L)) : LieIdeal.topEquiv x = x := rfl -variable (R L) -variable (M : Type*) [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] - -/-- The natural equivalence between the 'top' Lie submodule and the enclosing Lie module. -/ -def LieModuleEquiv.ofTop : (⊤ : LieSubmodule R L M) ≃ₗ⁅R,L⁆ M := - { LinearEquiv.ofTop ⊤ rfl with - map_lie' := rfl } - --- This lemma has always been bad, but leanprover/lean4#2644 made `simp` start noticing -@[simp, nolint simpNF] lemma LieModuleEquiv.ofTop_apply (x : (⊤ : LieSubmodule R L M)) : - LieModuleEquiv.ofTop R L M x = x := - rfl - -@[simp] lemma LieModuleEquiv.range_coe {M' : Type*} - [AddCommGroup M'] [Module R M'] [LieRingModule L M'] (e : M ≃ₗ⁅R,L⁆ M') : - LieModuleHom.range (e : M →ₗ⁅R,L⁆ M') = ⊤ := by - rw [LieModuleHom.range_eq_top] - exact e.surjective - end TopEquiv diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean index 2154fb228839b..38abc3bb0ed03 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -112,7 +112,7 @@ local macro "pderiv_simp" : tactic => pderiv_X_of_ne (by decide : z ≠ x), pderiv_X_of_ne (by decide : x ≠ z), pderiv_X_of_ne (by decide : z ≠ y), pderiv_X_of_ne (by decide : y ≠ z)]) -variable {R : Type u} [CommRing R] {W' : Jacobian R} {F : Type v} [Field F] {W : Jacobian F} +variable {R : Type u} {W' : Jacobian R} {F : Type v} [Field F] {W : Jacobian F} section Jacobian @@ -127,6 +127,8 @@ lemma fin3_def_ext (X Y Z : R) : ![X, Y, Z] x = X ∧ ![X, Y, Z] y = Y ∧ ![X, lemma comp_fin3 {S} (f : R → S) (X Y Z : R) : f ∘ ![X, Y, Z] = ![f X, f Y, f Z] := (FinVec.map_eq _ _).symm +variable [CommRing R] + /-- The scalar multiplication on a point representative. -/ scoped instance instSMulPoint : SMul R <| Fin 3 → R := ⟨fun u P => ![u ^ 2 * P x, u ^ 3 * P y, u * P z]⟩ @@ -224,6 +226,8 @@ lemma Y_eq_iff {P Q : Fin 3 → F} (hPz : P z ≠ 0) (hQz : Q z ≠ 0) : end Jacobian +variable [CommRing R] + section Equation /-! ### Weierstrass equations -/ diff --git a/Mathlib/Analysis/InnerProductSpace/Orientation.lean b/Mathlib/Analysis/InnerProductSpace/Orientation.lean index 4dcb286356c0b..60b27da658fe3 100644 --- a/Mathlib/Analysis/InnerProductSpace/Orientation.lean +++ b/Mathlib/Analysis/InnerProductSpace/Orientation.lean @@ -44,7 +44,7 @@ open scoped RealInnerProductSpace namespace OrthonormalBasis -variable {ι : Type*} [Fintype ι] [DecidableEq ι] [ne : Nonempty ι] (e f : OrthonormalBasis ι ℝ E) +variable {ι : Type*} [Fintype ι] [DecidableEq ι] (e f : OrthonormalBasis ι ℝ E) (x : Orientation ℝ E ι) /-- The change-of-basis matrix between two orthonormal bases with the same orientation has @@ -90,6 +90,8 @@ theorem det_eq_neg_det_of_opposite_orientation (h : e.toBasis.orientation ≠ f. simp [e.det_to_matrix_orthonormalBasis_of_opposite_orientation f h, neg_one_smul ℝ (M := E [⋀^ι]→ₗ[ℝ] ℝ)] +variable [Nonempty ι] + section AdjustToOrientation /-- `OrthonormalBasis.adjustToOrientation`, applied to an orthonormal basis, preserves the diff --git a/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean index 3f1aa98ec5d75..5a779b0580765 100644 --- a/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Analysis/Normed/Algebra/TrivSqZeroExt.lean @@ -192,9 +192,8 @@ noncomputable section Seminormed section Ring variable [SeminormedCommRing S] [SeminormedRing R] [SeminormedAddCommGroup M] -variable [Algebra S R] [Module S M] [Module R M] [Module Rᵐᵒᵖ M] -variable [BoundedSMul S R] [BoundedSMul S M] [BoundedSMul R M] [BoundedSMul Rᵐᵒᵖ M] -variable [SMulCommClass R Rᵐᵒᵖ M] [IsScalarTower S R M] [IsScalarTower S Rᵐᵒᵖ M] +variable [Algebra S R] [Module S M] +variable [BoundedSMul S R] [BoundedSMul S M] instance instL1SeminormedAddCommGroup : SeminormedAddCommGroup (tsze R M) := inferInstanceAs <| SeminormedAddCommGroup (WithLp 1 <| R × M) @@ -217,6 +216,9 @@ theorem nnnorm_def (x : tsze R M) : ‖x‖₊ = ‖fst x‖₊ + ‖snd x‖₊ @[simp] theorem nnnorm_inl (r : R) : ‖(inl r : tsze R M)‖₊ = ‖r‖₊ := by simp [nnnorm_def] @[simp] theorem nnnorm_inr (m : M) : ‖(inr m : tsze R M)‖₊ = ‖m‖₊ := by simp [nnnorm_def] +variable [Module R M] [BoundedSMul R M] [Module Rᵐᵒᵖ M] [BoundedSMul Rᵐᵒᵖ M] + [SMulCommClass R Rᵐᵒᵖ M] + instance instL1SeminormedRing : SeminormedRing (tsze R M) where norm_mul | ⟨r₁, m₁⟩, ⟨r₂, m₂⟩ => by diff --git a/Mathlib/Analysis/Normed/Algebra/UnitizationL1.lean b/Mathlib/Analysis/Normed/Algebra/UnitizationL1.lean index f92a28cdee905..4a0ab143ce2a1 100644 --- a/Mathlib/Analysis/Normed/Algebra/UnitizationL1.lean +++ b/Mathlib/Analysis/Normed/Algebra/UnitizationL1.lean @@ -24,7 +24,7 @@ non-unital Banach algebra is compact, which can be established by passing to the -/ variable (𝕜 A : Type*) [NormedField 𝕜] [NonUnitalNormedRing A] -variable [NormedSpace 𝕜 A] [IsScalarTower 𝕜 A A] [SMulCommClass 𝕜 A A] +variable [NormedSpace 𝕜 A] namespace WithLp @@ -79,6 +79,8 @@ lemma unitization_isometry_inr : ((WithLp.linearEquiv 1 𝕜 (Unitization 𝕜 A)).symm.comp <| Unitization.inrHom 𝕜 A) unitization_norm_inr +variable [IsScalarTower 𝕜 A A] [SMulCommClass 𝕜 A A] + instance instUnitizationRing : Ring (WithLp 1 (Unitization 𝕜 A)) := inferInstanceAs (Ring (Unitization 𝕜 A)) diff --git a/Mathlib/Analysis/Normed/Lp/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean index 5ceff479c7c2d..dda1bee40dc60 100644 --- a/Mathlib/Analysis/Normed/Lp/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -92,7 +92,7 @@ section for Pi types will not trigger. -/ variable {𝕜 p α} variable [SeminormedRing 𝕜] [∀ i, SeminormedAddCommGroup (β i)] -variable [∀ i, Module 𝕜 (β i)] [∀ i, BoundedSMul 𝕜 (β i)] (c : 𝕜) +variable [∀ i, Module 𝕜 (β i)] (c : 𝕜) variable (x y : PiLp p β) (i : ι) #adaptation_note @@ -744,7 +744,7 @@ variable {ι : Type*} {κ : ι → Type*} (p : ℝ≥0∞) [Fact (1 ≤ p)] variable (𝕜) in /-- `LinearEquiv.piCurry` for `PiLp`, as an isometry. -/ -def _root_.LinearIsometryEquiv.piLpCurry : +def _root_.LinearIsometryEquiv.piLpCurry : PiLp p (fun i : Sigma _ => α i.1 i.2) ≃ₗᵢ[𝕜] PiLp p (fun i => PiLp p (α i)) where toLinearEquiv := WithLp.linearEquiv _ _ _ diff --git a/Mathlib/Analysis/NormedSpace/DualNumber.lean b/Mathlib/Analysis/NormedSpace/DualNumber.lean index 4058d986db870..f81134aac4e9d 100644 --- a/Mathlib/Analysis/NormedSpace/DualNumber.lean +++ b/Mathlib/Analysis/NormedSpace/DualNumber.lean @@ -25,7 +25,7 @@ open TrivSqZeroExt variable (𝕜 : Type*) {R : Type*} variable [Field 𝕜] [CharZero 𝕜] [CommRing R] [Algebra 𝕜 R] -variable [UniformSpace R] [TopologicalRing R] [CompleteSpace R] [T2Space R] +variable [UniformSpace R] [TopologicalRing R] [T2Space R] @[simp] theorem exp_eps : exp 𝕜 (eps : DualNumber R) = 1 + eps := diff --git a/Mathlib/Combinatorics/Additive/Corner/Roth.lean b/Mathlib/Combinatorics/Additive/Corner/Roth.lean index 9abc01be4aa41..cc3a7145d5a4d 100644 --- a/Mathlib/Combinatorics/Additive/Corner/Roth.lean +++ b/Mathlib/Combinatorics/Additive/Corner/Roth.lean @@ -23,7 +23,7 @@ open Finset SimpleGraph TripartiteFromTriangles open Function hiding graph open Fintype (card) -variable {G : Type*} [AddCommGroup G] [Fintype G] {A B : Finset (G × G)} +variable {G : Type*} [AddCommGroup G] {A B : Finset (G × G)} {a b c d x y : G} {n : ℕ} {ε : ℝ} namespace Corners @@ -55,7 +55,7 @@ private lemma noAccidental (hs : IsCornerFree (A : Set (G × G))) : simp only [mk_mem_triangleIndices] at ha hb hc exact .inl $ hs ⟨hc.1, hb.1, ha.1, hb.2.symm.trans ha.2⟩ -private lemma farFromTriangleFree_graph [DecidableEq G] (hε : ε * card G ^ 2 ≤ A.card) : +private lemma farFromTriangleFree_graph [Fintype G] [DecidableEq G] (hε : ε * card G ^ 2 ≤ A.card) : (graph <| triangleIndices A).FarFromTriangleFree (ε / 9) := by refine farFromTriangleFree _ ?_ simp_rw [card_triangleIndices, mul_comm_div, Nat.cast_pow, Nat.cast_add] @@ -64,6 +64,8 @@ private lemma farFromTriangleFree_graph [DecidableEq G] (hε : ε * card G ^ 2 end Corners +variable [Fintype G] + open Corners diff --git a/Mathlib/FieldTheory/Perfect.lean b/Mathlib/FieldTheory/Perfect.lean index cf6c499537bf9..f0725115ada01 100644 --- a/Mathlib/FieldTheory/Perfect.lean +++ b/Mathlib/FieldTheory/Perfect.lean @@ -290,6 +290,7 @@ theorem roots_expand_image_frobenius_subset [DecidableEq R] : convert ← roots_expand_pow_image_iterateFrobenius_subset p 1 f apply pow_one +section PerfectRing variable {p n f} variable [PerfectRing R p] @@ -342,7 +343,9 @@ theorem roots_expand_image_frobenius [DecidableEq R] : rw [Finset.image_toFinset, roots_expand_map_frobenius, (roots f).toFinset_nsmul _ (expChar_pos R p).ne'] -variable (p n f) [DecidableEq R] +end PerfectRing + +variable [DecidableEq R] /-- If `f` is a polynomial over an integral domain `R` of characteristic `p`, then there is a map from the set of roots of `Polynomial.expand R p f` to the set of roots of `f`. @@ -367,6 +370,8 @@ noncomputable def rootsExpandPowToRoots : @[simp] theorem rootsExpandPowToRoots_apply (x) : (rootsExpandPowToRoots p n f x : R) = x ^ p ^ n := rfl +variable [PerfectRing R p] + /-- If `f` is a polynomial over a perfect integral domain `R` of characteristic `p`, then there is a bijection from the set of roots of `Polynomial.expand R p f` to the set of roots of `f`. It's given by `x ↦ x ^ p`, see `rootsExpandEquivRoots_apply`. -/ diff --git a/Mathlib/Geometry/Manifold/InteriorBoundary.lean b/Mathlib/Geometry/Manifold/InteriorBoundary.lean index aafb361508ac5..1bfcd8ec40979 100644 --- a/Mathlib/Geometry/Manifold/InteriorBoundary.lean +++ b/Mathlib/Geometry/Manifold/InteriorBoundary.lean @@ -159,7 +159,7 @@ variable {I} {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {N : Type*} [TopologicalSpace N] [ChartedSpace H' N] - (J : ModelWithCorners 𝕜 E' H') [SmoothManifoldWithCorners J N] {x : M} {y : N} + (J : ModelWithCorners 𝕜 E' H') {x : M} {y : N} /-- The interior of `M × N` is the product of the interiors of `M` and `N`. -/ lemma ModelWithCorners.interior_prod : diff --git a/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean b/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean index 54a0bb985a977..f4bca4eca0457 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/TensorProduct.lean @@ -31,9 +31,11 @@ section CommRing variable [CommRing R] [CommRing A] variable [AddCommGroup M₁] [AddCommGroup M₂] variable [Algebra R A] [Module R M₁] [Module A M₁] -variable [SMulCommClass R A M₁] [SMulCommClass A R M₁] [IsScalarTower R A M₁] -variable [Module R M₂] [Invertible (2 : R)] +variable [SMulCommClass R A M₁] [IsScalarTower R A M₁] +variable [Module R M₂] +section InvertibleTwo +variable [Invertible (2 : R)] variable (R A) in /-- The tensor product of two quadratic forms injects into quadratic forms on tensor products. @@ -106,6 +108,8 @@ theorem polarBilin_baseChange [Invertible (2 : A)] (Q : QuadraticForm R M₂) : ← LinearMap.map_smul, smul_tmul', ← two_nsmul_associated R, coe_associatedHom, associated_sq, smul_comm, ← smul_assoc, two_smul, invOf_two_add_invOf_two, one_smul] +end InvertibleTwo + /-- If two quadratic forms from `A ⊗[R] M₂` agree on elements of the form `1 ⊗ m`, they are equal. In other words, if a base change exists for a quadratic form, it is unique. diff --git a/Mathlib/LinearAlgebra/Trace.lean b/Mathlib/LinearAlgebra/Trace.lean index b9d42ad6884fe..9bbcd7e816372 100644 --- a/Mathlib/LinearAlgebra/Trace.lean +++ b/Mathlib/LinearAlgebra/Trace.lean @@ -145,9 +145,9 @@ theorem trace_eq_contract_of_basis' [Fintype ι] [DecidableEq ι] (b : Basis ι LinearMap.trace R M = contractLeft R M ∘ₗ (dualTensorHomEquivOfBasis b).symm.toLinearMap := by simp [LinearEquiv.eq_comp_toLinearMap_symm, trace_eq_contract_of_basis b] +section variable (R M) variable [Module.Free R M] [Module.Finite R M] [Module.Free R N] [Module.Finite R N] - [Module.Free R P] [Module.Finite R P] /-- When `M` is finite free, the trace of a linear map correspond to the contraction pairing under the isomorphism `End(M) ≃ M* ⊗ M`-/ @@ -257,10 +257,16 @@ theorem trace_comp_comm' (f : M →ₗ[R] N) (g : N →ₗ[R] M) : simp only [llcomp_apply', compr₂_apply, flip_apply] at h exact h +end + +variable {N P} + +variable [Module.Free R N] [Module.Finite R N] [Module.Free R P] [Module.Finite R P] in lemma trace_comp_cycle (f : M →ₗ[R] N) (g : N →ₗ[R] P) (h : P →ₗ[R] M) : trace R P (g ∘ₗ f ∘ₗ h) = trace R N (f ∘ₗ h ∘ₗ g) := by rw [trace_comp_comm', comp_assoc] +variable [Module.Free R M] [Module.Finite R M] [Module.Free R P] [Module.Finite R P] in lemma trace_comp_cycle' (f : M →ₗ[R] N) (g : N →ₗ[R] P) (h : P →ₗ[R] M) : trace R P ((g ∘ₗ f) ∘ₗ h) = trace R M ((h ∘ₗ g) ∘ₗ f) := by rw [trace_comp_comm', ← comp_assoc] @@ -286,6 +292,7 @@ theorem IsProj.trace {p : Submodule R M} {f : M →ₗ[R] M} (h : IsProj p f) [M trace R M f = (finrank R p : R) := by rw [h.eq_conj_prodMap, trace_conj', trace_prodMap', trace_id, map_zero, add_zero] +variable [Module.Free R M] [Module.Finite R M] in lemma isNilpotent_trace_of_isNilpotent {f : M →ₗ[R] M} (hf : IsNilpotent f) : IsNilpotent (trace R M f) := by let b : Basis _ R M := Module.Free.chooseBasis R M @@ -293,6 +300,7 @@ lemma isNilpotent_trace_of_isNilpotent {f : M →ₗ[R] M} (hf : IsNilpotent f) apply Matrix.isNilpotent_trace_of_isNilpotent simpa +variable [Module.Free R M] [Module.Finite R M] in lemma trace_comp_eq_mul_of_commute_of_isNilpotent [IsReduced R] {f g : Module.End R M} (μ : R) (h_comm : Commute f g) (hg : IsNilpotent (g - algebraMap R _ μ)) : trace R M (f ∘ₗ g) = μ * trace R M f := by diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 71b2a4b053c3d..3482a2535938d 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -867,7 +867,8 @@ theorem volume_preserving_pi_empty {ι : Type u} (α : ι → Type v) [Fintype MeasurePreserving (MeasurableEquiv.ofUniqueOfUnique (∀ i, α i) Unit) volume volume := measurePreserving_pi_empty fun _ => volume -theorem measurePreserving_piFinsetUnion [DecidableEq ι] {s t : Finset ι} (h : Disjoint s t) +theorem measurePreserving_piFinsetUnion {ι : Type*} {α : ι → Type*} + {_ : ∀ i, MeasurableSpace (α i)} [DecidableEq ι] {s t : Finset ι} (h : Disjoint s t) (μ : ∀ i, Measure (α i)) [∀ i, SigmaFinite (μ i)] : MeasurePreserving (MeasurableEquiv.piFinsetUnion α h) ((Measure.pi fun i : s ↦ μ i).prod (Measure.pi fun i : t ↦ μ i)) @@ -876,13 +877,15 @@ theorem measurePreserving_piFinsetUnion [DecidableEq ι] {s t : Finset ι} (h : measurePreserving_piCongrLeft (fun i : ↥(s ∪ t) ↦ μ i) e |>.comp <| measurePreserving_sumPiEquivProdPi_symm fun b ↦ μ (e b) -theorem volume_preserving_piFinsetUnion (α : ι → Type*) [DecidableEq ι] {s t : Finset ι} +theorem volume_preserving_piFinsetUnion {ι : Type*} [DecidableEq ι] (α : ι → Type*) {s t : Finset ι} (h : Disjoint s t) [∀ i, MeasureSpace (α i)] [∀ i, SigmaFinite (volume : Measure (α i))] : MeasurePreserving (MeasurableEquiv.piFinsetUnion α h) volume volume := measurePreserving_piFinsetUnion h (fun _ ↦ volume) -theorem measurePreserving_pi {β : ι → Type*} [∀ i, MeasurableSpace (β i)] - (ν : (i : ι) → Measure (β i)) {f : (i : ι) → (α i) → (β i)} [∀ i, SigmaFinite (ν i)] +theorem measurePreserving_pi {ι : Type*} [Fintype ι] {α : ι → Type v} {β : ι → Type*} + [∀ i, MeasureSpace (α i)] [∀ i, MeasurableSpace (β i)] + (μ : (i : ι) → Measure (α i)) (ν : (i : ι) → Measure (β i)) + {f : (i : ι) → (α i) → (β i)} [∀ i, SigmaFinite (ν i)] (hf : ∀ i, MeasurePreserving (f i) (μ i) (ν i)) : MeasurePreserving (fun a i ↦ f i (a i)) (Measure.pi μ) (Measure.pi ν) where measurable := diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean index dd948e06b1dfd..ebf5efd059fd2 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean @@ -466,12 +466,12 @@ theorem ae_ae_of_ae_prod {p : α × β → Prop} (h : ∀ᵐ z ∂μ.prod ν, p ∀ᵐ x ∂μ, ∀ᵐ y ∂ν, p (x, y) := measure_ae_null_of_prod_null h -theorem ae_ae_eq_curry_of_prod {f g : α × β → γ} (h : f =ᵐ[μ.prod ν] g) : +theorem ae_ae_eq_curry_of_prod {γ : Type*} {f g : α × β → γ} (h : f =ᵐ[μ.prod ν] g) : ∀ᵐ x ∂μ, curry f x =ᵐ[ν] curry g x := ae_ae_of_ae_prod h -theorem ae_ae_eq_of_ae_eq_uncurry {f g : α → β → γ} (h : uncurry f =ᵐ[μ.prod ν] uncurry g) : - ∀ᵐ x ∂μ, f x =ᵐ[ν] g x := +theorem ae_ae_eq_of_ae_eq_uncurry {γ : Type*} {f g : α → β → γ} + (h : uncurry f =ᵐ[μ.prod ν] uncurry g) : ∀ᵐ x ∂μ, f x =ᵐ[ν] g x := ae_ae_eq_curry_of_prod h theorem ae_prod_mem_iff_ae_ae_mem {s : Set (α × β)} (hs : MeasurableSet s) : diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index c45cc18acbfdf..8172325ce37ac 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -523,6 +523,7 @@ lemma tendsto_measure_smul_diff_isCompact_isClosed [LocallyCompactSpace G] ENNReal.nhds_zero_basis.tendsto_right_iff.mpr <| fun _ h ↦ eventually_nhds_one_measure_smul_diff_lt hk h'k h.ne' +section IsMulLeftInvariant variable [IsMulLeftInvariant μ] /-- If a left-invariant measure gives positive mass to a compact set, then it gives positive mass to @@ -682,6 +683,8 @@ lemma _root_.IsCompact.measure_closure_eq_of_group {k : Set G} (hk : IsCompact k μ (closure k) = μ k := hk.measure_closure μ +end IsMulLeftInvariant + @[to_additive] lemma innerRegularWRT_isCompact_isClosed_measure_ne_top_of_group [h : InnerRegularCompactLTTop μ] : InnerRegularWRT μ (fun s ↦ IsCompact s ∧ IsClosed s) (fun s ↦ MeasurableSet s ∧ μ s ≠ ∞) := by diff --git a/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean b/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean index 38a0aab6c0316..b3d825fc3b697 100644 --- a/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean +++ b/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean @@ -72,7 +72,7 @@ This formulation assumes: * boundedness holds almost everywhere. -/ theorem measure_of_cont_bdd_of_tendsto_filter_indicator {ι : Type*} {L : Filter ι} - [L.IsCountablyGenerated] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] (μ : Measure Ω) + [L.IsCountablyGenerated] (μ : Measure Ω) [IsFiniteMeasure μ] {c : ℝ≥0} {E : Set Ω} (E_mble : MeasurableSet E) (fs : ι → Ω →ᵇ ℝ≥0) (fs_bdd : ∀ᶠ i in L, ∀ᵐ ω : Ω ∂μ, fs i ω ≤ c) (fs_lim : ∀ᵐ ω ∂μ, Tendsto (fun i ↦ fs i ω) L (𝓝 (indicator E (fun _ ↦ (1 : ℝ≥0)) ω))) : @@ -90,7 +90,7 @@ measure of the set. A similar result with more general assumptions is `MeasureTheory.measure_of_cont_bdd_of_tendsto_filter_indicator`. -/ -theorem measure_of_cont_bdd_of_tendsto_indicator [OpensMeasurableSpace Ω] +theorem measure_of_cont_bdd_of_tendsto_indicator (μ : Measure Ω) [IsFiniteMeasure μ] {c : ℝ≥0} {E : Set Ω} (E_mble : MeasurableSet E) (fs : ℕ → Ω →ᵇ ℝ≥0) (fs_bdd : ∀ n ω, fs n ω ≤ c) (fs_lim : Tendsto (fun n ω ↦ fs n ω) atTop (𝓝 (indicator E fun _ ↦ (1 : ℝ≥0)))) : diff --git a/Mathlib/RingTheory/Kaehler/Basic.lean b/Mathlib/RingTheory/Kaehler/Basic.lean index f896a50f46ebc..d5a828b146a26 100644 --- a/Mathlib/RingTheory/Kaehler/Basic.lean +++ b/Mathlib/RingTheory/Kaehler/Basic.lean @@ -617,9 +617,8 @@ A --→ B R --→ S ``` -/ -variable (A B : Type*) [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] -variable [Algebra A B] [Algebra S B] [IsScalarTower R A B] [IsScalarTower R S B] -variable [SMulCommClass S A B] +variable (A B : Type*) [CommRing A] [CommRing B] [Algebra R A] +variable [Algebra A B] [Algebra S B] unsuppress_compilation in -- The map `(A →₀ A) →ₗ[A] (B →₀ B)` @@ -639,7 +638,8 @@ The kernel of the presentation `⊕ₓ B dx ↠ Ω_{B/S}` is spanned by the imag kernel of `⊕ₓ A dx ↠ Ω_{A/R}` and all `ds` with `s : S`. See `kerTotal_map'` for the special case where `R = S`. -/ -theorem KaehlerDifferential.kerTotal_map (h : Function.Surjective (algebraMap A B)) : +theorem KaehlerDifferential.kerTotal_map [Algebra R B] [IsScalarTower R A B] [IsScalarTower R S B] + (h : Function.Surjective (algebraMap A B)) : (KaehlerDifferential.kerTotal R A).map finsupp_map ⊔ Submodule.span A (Set.range fun x : S => .single (algebraMap S B x) (1 : B)) = (KaehlerDifferential.kerTotal S B).restrictScalars _ := by @@ -665,7 +665,8 @@ This is a special case of `kerTotal_map` where `R = S`. The kernel of the presentation `⊕ₓ B dx ↠ Ω_{B/R}` is spanned by the image of the kernel of `⊕ₓ A dx ↠ Ω_{A/R}` and all `da` with `a : A`. -/ -theorem KaehlerDifferential.kerTotal_map' (h : Function.Surjective (algebraMap A B)) : +theorem KaehlerDifferential.kerTotal_map' [Algebra R B] + [IsScalarTower R A B] (h : Function.Surjective (algebraMap A B)) : (KaehlerDifferential.kerTotal R A ⊔ Submodule.span A (Set.range fun x ↦ .single (algebraMap R A x) 1)).map finsupp_map = (KaehlerDifferential.kerTotal R B).restrictScalars _ := by @@ -674,6 +675,8 @@ theorem KaehlerDifferential.kerTotal_map' (h : Function.Surjective (algebraMap A refine congr_arg Set.range ?_ ext; simp [IsScalarTower.algebraMap_eq R A B] +variable [Algebra R B] [IsScalarTower R A B] [IsScalarTower R S B] [SMulCommClass S A B] + /-- The map `Ω[A⁄R] →ₗ[A] Ω[B⁄S]` given a square ``` A --→ B diff --git a/Mathlib/RingTheory/Unramified/Finite.lean b/Mathlib/RingTheory/Unramified/Finite.lean index b754c35bf4fee..0afad0fed8c13 100644 --- a/Mathlib/RingTheory/Unramified/Finite.lean +++ b/Mathlib/RingTheory/Unramified/Finite.lean @@ -80,32 +80,6 @@ theorem iff_exists_tensorProduct [EssFiniteType R S] : use 1 ⊗ₜ[R] s - s ⊗ₜ[R] 1 linear_combination ht₁ s -variable [FormallyUnramified R S] [EssFiniteType R S] - -variable (R S) in -/-- -A finite-type `R`-algebra `S` is (formally) unramified iff there exists a `t : S ⊗[R] S` satisfying -1. `t` annihilates every `1 ⊗ s - s ⊗ 1`. -2. the image of `t` is `1` under the map `S ⊗[R] S → S`. -See `Algebra.FormallyUnramified.iff_exists_tensorProduct`. -This is the choice of such a `t`. --/ -noncomputable -def elem : S ⊗[R] S := - (iff_exists_tensorProduct.mp inferInstance).choose - -lemma one_tmul_sub_tmul_one_mul_elem - (s : S) : (1 ⊗ₜ s - s ⊗ₜ 1) * elem R S = 0 := - (iff_exists_tensorProduct.mp inferInstance).choose_spec.1 s - -lemma one_tmul_mul_elem - (s : S) : (1 ⊗ₜ s) * elem R S = (s ⊗ₜ 1) * elem R S := by - rw [← sub_eq_zero, ← sub_mul, one_tmul_sub_tmul_one_mul_elem] - -lemma lmul_elem [EssFiniteType R S] [FormallyUnramified R S] : - TensorProduct.lmul' R (elem R S) = 1 := - (iff_exists_tensorProduct.mp inferInstance).choose_spec.2 - lemma finite_of_free_aux (I) [DecidableEq I] (b : Basis I R S) (f : I →₀ S) (x : S) (a : I → I →₀ R) (ha : a = fun i ↦ b.repr (b i * x)) : (1 ⊗ₜ[R] x * Finsupp.sum f fun i y ↦ y ⊗ₜ[R] b i) = @@ -145,6 +119,33 @@ lemma finite_of_free_aux (I) [DecidableEq I] (b : Basis I R S) simp (config := {contextual := true}) · exact fun _ _ ↦ rfl +variable [FormallyUnramified R S] [EssFiniteType R S] + +variable (R S) in +/-- +A finite-type `R`-algebra `S` is (formally) unramified iff there exists a `t : S ⊗[R] S` satisfying +1. `t` annihilates every `1 ⊗ s - s ⊗ 1`. +2. the image of `t` is `1` under the map `S ⊗[R] S → S`. +See `Algebra.FormallyUnramified.iff_exists_tensorProduct`. +This is the choice of such a `t`. +-/ +noncomputable +def elem : S ⊗[R] S := + (iff_exists_tensorProduct.mp inferInstance).choose + +lemma one_tmul_sub_tmul_one_mul_elem + (s : S) : (1 ⊗ₜ s - s ⊗ₜ 1) * elem R S = 0 := + (iff_exists_tensorProduct.mp inferInstance).choose_spec.1 s + +lemma one_tmul_mul_elem + (s : S) : (1 ⊗ₜ s) * elem R S = (s ⊗ₜ 1) * elem R S := by + rw [← sub_eq_zero, ← sub_mul, one_tmul_sub_tmul_one_mul_elem] + +lemma lmul_elem : + TensorProduct.lmul' R (elem R S) = 1 := + (iff_exists_tensorProduct.mp inferInstance).choose_spec.2 + + variable (R S) /-- An unramified free algebra is finitely generated. Iversen I.2.8 -/ diff --git a/Mathlib/RingTheory/WittVector/Basic.lean b/Mathlib/RingTheory/WittVector/Basic.lean index 7d09dc5c960fb..5630d6a00254e 100644 --- a/Mathlib/RingTheory/WittVector/Basic.lean +++ b/Mathlib/RingTheory/WittVector/Basic.lean @@ -49,7 +49,7 @@ noncomputable section open MvPolynomial Function -variable {p : ℕ} {R S T : Type*} [hp : Fact p.Prime] [CommRing R] [CommRing S] [CommRing T] +variable {p : ℕ} {R S T : Type*} [CommRing R] [CommRing S] [CommRing T] variable {α : Type*} {β : Type*} local notation "𝕎" => WittVector p @@ -76,9 +76,6 @@ theorem surjective (f : α → β) (hf : Surjective f) : Surjective (mapFun f : ⟨mk _ fun n => Classical.choose <| hf <| x.coeff n, by ext n; simp only [mapFun, coeff_mk, comp_apply, Classical.choose_spec (hf (x.coeff n))]⟩ --- Porting note: using `(x y : 𝕎 R)` instead of `(x y : WittVector p R)` produced sorries. -variable (f : R →+* S) (x y : WittVector p R) - /-- Auxiliary tactic for showing that `mapFun` respects the ring operations. -/ -- porting note: a very crude port. macro "map_fun_tac" : tactic => `(tactic| ( @@ -92,6 +89,10 @@ macro "map_fun_tac" : tactic => `(tactic| ( ext ⟨i, k⟩ <;> fin_cases i <;> rfl)) +variable [Fact p.Prime] +-- Porting note: using `(x y : 𝕎 R)` instead of `(x y : WittVector p R)` produced sorries. +variable (f : R →+* S) (x y : WittVector p R) + -- and until `pow`. -- We do not tag these lemmas as `@[simp]` because they will be bundled in `map` later on. theorem zero : mapFun f (0 : 𝕎 R) = 0 := by map_fun_tac @@ -158,8 +159,6 @@ end Tactic section GhostFun -variable (x y : WittVector p R) - -- The following lemmas are not `@[simp]` because they will be bundled in `ghostMap` later on. @[local simp] @@ -167,6 +166,9 @@ theorem matrix_vecEmpty_coeff {R} (i j) : @coeff p R (Matrix.vecEmpty i) j = (Matrix.vecEmpty i : ℕ → R) j := by rcases i with ⟨_ | _ | _ | _ | i_val, ⟨⟩⟩ +variable [Fact p.Prime] +variable (x y : WittVector p R) + private theorem ghostFun_zero : ghostFun (0 : 𝕎 R) = 0 := by ghost_fun_tac 0, ![] @@ -232,6 +234,8 @@ private def ghostEquiv' [Invertible (p : R)] : 𝕎 R ≃ (ℕ → R) where apply_fun aeval x at this simpa only [aeval_bind₁, aeval_X, ghostFun, aeval_wittPolynomial] +variable [Fact p.Prime] + @[local instance] private def comm_ring_aux₁ : CommRing (𝕎 (MvPolynomial R ℚ)) := (ghostEquiv' p (MvPolynomial R ℚ)).injective.commRing ghostFun ghostFun_zero ghostFun_one diff --git a/Mathlib/RingTheory/WittVector/InitTail.lean b/Mathlib/RingTheory/WittVector/InitTail.lean index 63541a917319d..c2c6e3cc9a3a5 100644 --- a/Mathlib/RingTheory/WittVector/InitTail.lean +++ b/Mathlib/RingTheory/WittVector/InitTail.lean @@ -34,7 +34,7 @@ and shows how that polynomial interacts with `MvPolynomial.bind₁`. -/ -variable {p : ℕ} [hp : Fact p.Prime] (n : ℕ) {R : Type*} [CommRing R] +variable {p : ℕ} (n : ℕ) {R : Type*} [CommRing R] -- type as `\bbW` local notation "𝕎" => WittVector p @@ -79,6 +79,8 @@ instance select_isPoly {P : ℕ → Prop} : IsPoly p fun _ _ x => select P x := funext i apply coeff_select +variable [hp : Fact p.Prime] + theorem select_add_select_not : ∀ x : 𝕎 R, select P x + select (fun i => ¬P i) x = x := by -- Porting note: TC search was insufficient to find this instance, even though all required -- instances exist. See zulip: [https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/WittVector.20saga/near/370073526] @@ -127,6 +129,8 @@ theorem coeff_add_of_disjoint (x y : 𝕎 R) (h : ∀ n, x.coeff n = 0 ∨ y.coe end Select +variable [Fact p.Prime] + /-- `WittVector.init n x` is the Witt vector of which the first `n` coefficients are those from `x` and all other coefficients are `0`. See `WittVector.tail` for the complementary part. @@ -186,6 +190,9 @@ theorem init_init (x : 𝕎 R) (n : ℕ) : init n (init n x) = init n x := by simp only [WittVector.init, WittVector.select, WittVector.coeff_mk] by_cases hi : i < n <;> simp [hi] +section +variable [Fact p.Prime] + theorem init_add (x y : 𝕎 R) (n : ℕ) : init n (x + y) = init n (init n x + init n y) := by init_ring using wittAdd_vars @@ -207,6 +214,7 @@ theorem init_zsmul (m : ℤ) (x : 𝕎 R) (n : ℕ) : init n (m • x) = init n theorem init_pow (m : ℕ) (x : 𝕎 R) (n : ℕ) : init n (x ^ m) = init n (init n x ^ m) := by init_ring using fun p [Fact (Nat.Prime p)] n => wittPow_vars p m n +end section variable (p) diff --git a/Mathlib/RingTheory/WittVector/Teichmuller.lean b/Mathlib/RingTheory/WittVector/Teichmuller.lean index a68a5e0c2f286..d90c9006aefba 100644 --- a/Mathlib/RingTheory/WittVector/Teichmuller.lean +++ b/Mathlib/RingTheory/WittVector/Teichmuller.lean @@ -70,14 +70,14 @@ private theorem map_teichmullerFun (f : R →+* S) (r : R) : · rfl · exact f.map_zero -private theorem teichmuller_mul_aux₁ (x y : MvPolynomial R ℚ) : +private theorem teichmuller_mul_aux₁ {R : Type*} (x y : MvPolynomial R ℚ) : teichmullerFun p (x * y) = teichmullerFun p x * teichmullerFun p y := by apply (ghostMap.bijective_of_invertible p (MvPolynomial R ℚ)).1 rw [RingHom.map_mul] ext1 n simp only [Pi.mul_apply, ghostMap_apply, ghostComponent_teichmullerFun, mul_pow] -private theorem teichmuller_mul_aux₂ (x y : MvPolynomial R ℤ) : +private theorem teichmuller_mul_aux₂ {R : Type*} (x y : MvPolynomial R ℤ) : teichmullerFun p (x * y) = teichmullerFun p x * teichmullerFun p y := by refine map_injective (MvPolynomial.map (Int.castRingHom ℚ)) (MvPolynomial.map_injective _ Int.cast_injective) ?_ diff --git a/Mathlib/RingTheory/WittVector/Verschiebung.lean b/Mathlib/RingTheory/WittVector/Verschiebung.lean index def9919ecd7a2..17a60efd46d57 100644 --- a/Mathlib/RingTheory/WittVector/Verschiebung.lean +++ b/Mathlib/RingTheory/WittVector/Verschiebung.lean @@ -21,7 +21,7 @@ namespace WittVector open MvPolynomial -variable {p : ℕ} {R S : Type*} [hp : Fact p.Prime] [CommRing R] [CommRing S] +variable {p : ℕ} {R S : Type*} [CommRing R] [CommRing S] local notation "𝕎" => WittVector p -- type as `\bbW` @@ -49,13 +49,13 @@ theorem verschiebungFun_coeff_succ (x : 𝕎 R) (n : ℕ) : rfl @[ghost_simps] -theorem ghostComponent_zero_verschiebungFun (x : 𝕎 R) : +theorem ghostComponent_zero_verschiebungFun [hp : Fact p.Prime] (x : 𝕎 R) : ghostComponent 0 (verschiebungFun x) = 0 := by rw [ghostComponent_apply, aeval_wittPolynomial, Finset.range_one, Finset.sum_singleton, verschiebungFun_coeff_zero, pow_zero, pow_zero, pow_one, one_mul] @[ghost_simps] -theorem ghostComponent_verschiebungFun (x : 𝕎 R) (n : ℕ) : +theorem ghostComponent_verschiebungFun [hp : Fact p.Prime] (x : 𝕎 R) (n : ℕ) : ghostComponent (n + 1) (verschiebungFun x) = p * ghostComponent n x := by simp only [ghostComponent_apply, aeval_wittPolynomial] rw [Finset.sum_range_succ', verschiebungFun_coeff, if_pos rfl, @@ -96,6 +96,7 @@ example (p : ℕ) (f : ⦃R : Type _⦄ → [CommRing R] → WittVector p R → inferInstance variable {p} +variable [hp : Fact p.Prime] /-- `verschiebung x` shifts the coefficients of `x` up by one, by inserting 0 as the 0th coefficient. @@ -116,7 +117,7 @@ noncomputable def verschiebung : 𝕎 R →+ 𝕎 R where /-- `WittVector.verschiebung` is a polynomial function. -/ @[is_poly] -theorem verschiebung_isPoly : IsPoly p fun R _Rcr => @verschiebung p R hp _Rcr := +theorem verschiebung_isPoly : IsPoly p fun _ _ => verschiebung (p := p) := verschiebungFun_isPoly p /-- verschiebung is a natural transformation -/ diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean index daa03f9b799fb..e4fb9f1ae0d5b 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean @@ -193,32 +193,6 @@ private theorem candidates_lipschitz (fA : f ∈ candidates X Y) : rw [dist_comm] exact candidates_lipschitz_aux fA -/-- candidates give rise to elements of `BoundedContinuousFunction`s -/ -def candidatesBOfCandidates (f : ProdSpaceFun X Y) (fA : f ∈ candidates X Y) : Cb X Y := - BoundedContinuousFunction.mkOfCompact ⟨f, (candidates_lipschitz fA).continuous⟩ - -theorem candidatesBOfCandidates_mem (f : ProdSpaceFun X Y) (fA : f ∈ candidates X Y) : - candidatesBOfCandidates f fA ∈ candidatesB X Y := - fA - -/-- The distance on `X ⊕ Y` is a candidate -/ -private theorem dist_mem_candidates : - (fun p : (X ⊕ Y) × (X ⊕ Y) => dist p.1 p.2) ∈ candidates X Y := by - simp_rw [candidates, Set.mem_setOf_eq, dist_comm, dist_triangle, dist_self, maxVar_bound, - forall_const, and_true] - exact ⟨fun x y => rfl, fun x y => rfl⟩ - -/-- The distance on `X ⊕ Y` as a candidate -/ -def candidatesBDist (X : Type u) (Y : Type v) [MetricSpace X] [CompactSpace X] [Nonempty X] - [MetricSpace Y] [CompactSpace Y] [Nonempty Y] : Cb X Y := - candidatesBOfCandidates _ dist_mem_candidates - -theorem candidatesBDist_mem_candidatesB : candidatesBDist X Y ∈ candidatesB X Y := - candidatesBOfCandidates_mem _ _ - -private theorem candidatesB_nonempty : (candidatesB X Y).Nonempty := - ⟨_, candidatesBDist_mem_candidatesB⟩ - /-- To apply Arzela-Ascoli, we need to check that the set of candidates is closed and equicontinuous. Equicontinuity follows from the Lipschitz control, we check closedness. -/ private theorem closed_candidatesB : IsClosed (candidatesB X Y) := by @@ -249,20 +223,6 @@ private theorem closed_candidatesB : IsClosed (candidatesB X Y) := by |apply isClosed_iInter _ |apply I1 _ _|apply I2 _ _|apply I3 _ _|apply I4 _ _ _|apply I5 _|apply I6 _ _|intro x -/-- Compactness of candidates (in `BoundedContinuousFunction`s) follows. -/ -private theorem isCompact_candidatesB : IsCompact (candidatesB X Y) := by - refine arzela_ascoli₂ - (Icc 0 (maxVar X Y) : Set ℝ) isCompact_Icc (candidatesB X Y) closed_candidatesB ?_ ?_ - · rintro f ⟨x1, x2⟩ hf - simp only [Set.mem_Icc] - exact ⟨candidates_nonneg hf, candidates_le_maxVar hf⟩ - · refine equicontinuous_of_continuity_modulus (fun t => 2 * maxVar X Y * t) ?_ _ ?_ - · have : Tendsto (fun t : ℝ => 2 * (maxVar X Y : ℝ) * t) (𝓝 0) (𝓝 (2 * maxVar X Y * 0)) := - tendsto_const_nhds.mul tendsto_id - simpa using this - · rintro x y ⟨f, hf⟩ - exact (candidates_lipschitz hf).dist_le_mul _ _ - /-- We will then choose the candidate minimizing the Hausdorff distance. Except that we are not in a metric space setting, so we need to define our custom version of Hausdorff distance, called `HD`, and prove its basic properties. -/ @@ -279,7 +239,7 @@ theorem HD_below_aux1 {f : Cb X Y} (C : ℝ) {x : X} : let ⟨cf, hcf⟩ := f.isBounded_range.bddBelow ⟨cf + C, forall_mem_range.2 fun _ => add_le_add_right ((fun x => hcf (mem_range_self x)) _) _⟩ -private theorem HD_bound_aux1 (f : Cb X Y) (C : ℝ) : +private theorem HD_bound_aux1 [Nonempty Y] (f : Cb X Y) (C : ℝ) : BddAbove (range fun x : X => ⨅ y, f (inl x, inr y) + C) := by obtain ⟨Cf, hCf⟩ := f.isBounded_range.bddAbove refine ⟨Cf + C, forall_mem_range.2 fun x => ?_⟩ @@ -292,7 +252,7 @@ theorem HD_below_aux2 {f : Cb X Y} (C : ℝ) {y : Y} : let ⟨cf, hcf⟩ := f.isBounded_range.bddBelow ⟨cf + C, forall_mem_range.2 fun _ => add_le_add_right ((fun x => hcf (mem_range_self x)) _) _⟩ -private theorem HD_bound_aux2 (f : Cb X Y) (C : ℝ) : +private theorem HD_bound_aux2 [Nonempty X] (f : Cb X Y) (C : ℝ) : BddAbove (range fun y : Y => ⨅ x, f (inl x, inr y) + C) := by obtain ⟨Cf, hCf⟩ := f.isBounded_range.bddAbove refine ⟨Cf + C, forall_mem_range.2 fun y => ?_⟩ @@ -300,29 +260,8 @@ private theorem HD_bound_aux2 (f : Cb X Y) (C : ℝ) : ⨅ x, f (inl x, inr y) + C ≤ f (inl default, inr y) + C := ciInf_le (HD_below_aux2 C) default _ ≤ Cf + C := add_le_add ((fun x => hCf (mem_range_self x)) _) le_rfl -/-- Explicit bound on `HD (dist)`. This means that when looking for minimizers it will -be sufficient to look for functions with `HD(f)` bounded by this bound. -/ -theorem HD_candidatesBDist_le : - HD (candidatesBDist X Y) ≤ diam (univ : Set X) + 1 + diam (univ : Set Y) := by - refine max_le (ciSup_le fun x => ?_) (ciSup_le fun y => ?_) - · have A : ⨅ y, candidatesBDist X Y (inl x, inr y) ≤ candidatesBDist X Y (inl x, inr default) := - ciInf_le (by simpa using HD_below_aux1 0) default - have B : dist (inl x) (inr default) ≤ diam (univ : Set X) + 1 + diam (univ : Set Y) := - calc - dist (inl x) (inr (default : Y)) = dist x (default : X) + 1 + dist default default := rfl - _ ≤ diam (univ : Set X) + 1 + diam (univ : Set Y) := by - gcongr <;> - exact dist_le_diam_of_mem isBounded_of_compactSpace (mem_univ _) (mem_univ _) - exact le_trans A B - · have A : ⨅ x, candidatesBDist X Y (inl x, inr y) ≤ candidatesBDist X Y (inl default, inr y) := - ciInf_le (by simpa using HD_below_aux2 0) default - have B : dist (inl default) (inr y) ≤ diam (univ : Set X) + 1 + diam (univ : Set Y) := - calc - dist (inl (default : X)) (inr y) = dist default default + 1 + dist default y := rfl - _ ≤ diam (univ : Set X) + 1 + diam (univ : Set Y) := by - gcongr <;> - exact dist_le_diam_of_mem isBounded_of_compactSpace (mem_univ _) (mem_univ _) - exact le_trans A B +section Nonempty +variable [Nonempty X] [Nonempty Y] /- To check that `HD` is continuous, we check that it is Lipschitz. As `HD` is a max, we prove separately inequalities controlling the two terms (relying too heavily on copy-paste...) -/ @@ -382,7 +321,8 @@ private theorem HD_lipschitz_aux2 (f g : Cb X Y) : -- deduce the result from the above two steps simpa [E2, E1] -private theorem HD_lipschitz_aux3 (f g : Cb X Y) : HD f ≤ HD g + dist f g := +private theorem HD_lipschitz_aux3 (f g : Cb X Y) : + HD f ≤ HD g + dist f g := max_le (le_trans (HD_lipschitz_aux1 f g) (add_le_add_right (le_max_left _ _) _)) (le_trans (HD_lipschitz_aux2 f g) (add_le_add_right (le_max_right _ _) _)) @@ -390,6 +330,77 @@ private theorem HD_lipschitz_aux3 (f g : Cb X Y) : HD f ≤ HD g + dist f g := private theorem HD_continuous : Continuous (HD : Cb X Y → ℝ) := LipschitzWith.continuous (LipschitzWith.of_le_add HD_lipschitz_aux3) +end Nonempty + +variable [CompactSpace X] [CompactSpace Y] + +/-- Compactness of candidates (in `BoundedContinuousFunction`s) follows. -/ +private theorem isCompact_candidatesB : IsCompact (candidatesB X Y) := by + refine arzela_ascoli₂ + (Icc 0 (maxVar X Y) : Set ℝ) isCompact_Icc (candidatesB X Y) closed_candidatesB ?_ ?_ + · rintro f ⟨x1, x2⟩ hf + simp only [Set.mem_Icc] + exact ⟨candidates_nonneg hf, candidates_le_maxVar hf⟩ + · refine equicontinuous_of_continuity_modulus (fun t => 2 * maxVar X Y * t) ?_ _ ?_ + · have : Tendsto (fun t : ℝ => 2 * (maxVar X Y : ℝ) * t) (𝓝 0) (𝓝 (2 * maxVar X Y * 0)) := + tendsto_const_nhds.mul tendsto_id + simpa using this + · rintro x y ⟨f, hf⟩ + exact (candidates_lipschitz hf).dist_le_mul _ _ + +/-- candidates give rise to elements of `BoundedContinuousFunction`s -/ +def candidatesBOfCandidates (f : ProdSpaceFun X Y) (fA : f ∈ candidates X Y) : Cb X Y := + BoundedContinuousFunction.mkOfCompact ⟨f, (candidates_lipschitz fA).continuous⟩ + +theorem candidatesBOfCandidates_mem (f : ProdSpaceFun X Y) (fA : f ∈ candidates X Y) : + candidatesBOfCandidates f fA ∈ candidatesB X Y := + fA + +variable [Nonempty X] [Nonempty Y] + +/-- The distance on `X ⊕ Y` is a candidate -/ +private theorem dist_mem_candidates : + (fun p : (X ⊕ Y) × (X ⊕ Y) => dist p.1 p.2) ∈ candidates X Y := by + simp_rw [candidates, Set.mem_setOf_eq, dist_comm, dist_triangle, dist_self, maxVar_bound, + forall_const, and_true] + exact ⟨fun x y => rfl, fun x y => rfl⟩ + +/-- The distance on `X ⊕ Y` as a candidate -/ +def candidatesBDist (X : Type u) (Y : Type v) [MetricSpace X] [CompactSpace X] [Nonempty X] + [MetricSpace Y] [CompactSpace Y] [Nonempty Y] : Cb X Y := + candidatesBOfCandidates _ dist_mem_candidates + +theorem candidatesBDist_mem_candidatesB : + candidatesBDist X Y ∈ candidatesB X Y := + candidatesBOfCandidates_mem _ _ + +private theorem candidatesB_nonempty : (candidatesB X Y).Nonempty := + ⟨_, candidatesBDist_mem_candidatesB⟩ + +/-- Explicit bound on `HD (dist)`. This means that when looking for minimizers it will +be sufficient to look for functions with `HD(f)` bounded by this bound. -/ +theorem HD_candidatesBDist_le : + HD (candidatesBDist X Y) ≤ diam (univ : Set X) + 1 + diam (univ : Set Y) := by + refine max_le (ciSup_le fun x => ?_) (ciSup_le fun y => ?_) + · have A : ⨅ y, candidatesBDist X Y (inl x, inr y) ≤ candidatesBDist X Y (inl x, inr default) := + ciInf_le (by simpa using HD_below_aux1 0) default + have B : dist (inl x) (inr default) ≤ diam (univ : Set X) + 1 + diam (univ : Set Y) := + calc + dist (inl x) (inr (default : Y)) = dist x (default : X) + 1 + dist default default := rfl + _ ≤ diam (univ : Set X) + 1 + diam (univ : Set Y) := by + gcongr <;> + exact dist_le_diam_of_mem isBounded_of_compactSpace (mem_univ _) (mem_univ _) + exact le_trans A B + · have A : ⨅ x, candidatesBDist X Y (inl x, inr y) ≤ candidatesBDist X Y (inl default, inr y) := + ciInf_le (by simpa using HD_below_aux2 0) default + have B : dist (inl default) (inr y) ≤ diam (univ : Set X) + 1 + diam (univ : Set Y) := + calc + dist (inl (default : X)) (inr y) = dist default default + 1 + dist default y := rfl + _ ≤ diam (univ : Set X) + 1 + diam (univ : Set Y) := by + gcongr <;> + exact dist_le_diam_of_mem isBounded_of_compactSpace (mem_univ _) (mem_univ _) + exact le_trans A B + end Constructions section Consequences From aeb3c05894d02069f1bc626babe6249945362bcf Mon Sep 17 00:00:00 2001 From: Yudai Yamazaki Date: Wed, 7 Aug 2024 11:50:09 +0000 Subject: [PATCH 090/359] feat(Subgroup/IsCommutative): add instance for MonoidHom.range (#15370) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds: - `instance Subgroup.commGroup_isCommutative [CommGroup G] (H : Subgroup G) : H.IsCommutative` for subgroups of an abelian group; and - `instance MonoidHom.range_isCommutative [CommGroup G] (f : G →* N) : f.range.IsCommutative`, for the range of a `MonoidHom` whose domain is abelian. It also adds a lemma `Subgroup.mul_comm_of_mem_isCommutative` to allow applying the commutativity without constructing terms of a `Subgroup`. (This lemma is included in this same PR as I mentioned it in the same Zulip thread, but is independent of the two instances above.) Please refer to the [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/CommGroup.20MonoidHom.2Erange.20if.20the.20domain.20is.20a.20CommGroup) for discussions. --- Mathlib/Algebra/Group/Subgroup/Basic.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 2ca5fe50f09d6..f60917ae388d5 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -1728,6 +1728,11 @@ attribute [class] AddSubgroup.IsCommutative instance IsCommutative.commGroup [h : H.IsCommutative] : CommGroup H := { H.toGroup with mul_comm := h.is_comm.comm } +/-- A subgroup of a commutative group is commutative. -/ +@[to_additive "A subgroup of a commutative group is commutative."] +instance commGroup_isCommutative {G : Type*} [CommGroup G] (H : Subgroup G) : H.IsCommutative := + ⟨CommMagma.to_isCommutative⟩ + @[to_additive] instance map_isCommutative (f : G →* G') [H.IsCommutative] : (H.map f).IsCommutative := ⟨⟨by @@ -1749,6 +1754,11 @@ theorem comap_injective_isCommutative {f : G' →* G} (hf : Injective f) [H.IsCo instance subgroupOf_isCommutative [H.IsCommutative] : (H.subgroupOf K).IsCommutative := H.comap_injective_isCommutative Subtype.coe_injective +@[to_additive] +lemma mul_comm_of_mem_isCommutative [H.IsCommutative] {a b : G} (ha : a ∈ H) (hb : b ∈ H) : + a * b = b * a := by + simpa only [Submonoid.mk_mul_mk, Subtype.mk.injEq] using mul_comm (⟨a, ha⟩ : H) (⟨b, hb⟩ : H) + end Subgroup namespace MulEquiv @@ -1954,6 +1964,11 @@ theorem mem_range {f : G →* N} {y : N} : y ∈ f.range ↔ ∃ x, f x = y := @[to_additive] theorem range_eq_map (f : G →* N) : f.range = (⊤ : Subgroup G).map f := by ext; simp +@[to_additive] +instance range_isCommutative {G : Type*} [CommGroup G] {N : Type*} [Group N] (f : G →* N) : + f.range.IsCommutative := + range_eq_map f ▸ Subgroup.map_isCommutative ⊤ f + @[to_additive (attr := simp)] theorem restrict_range (f : G →* N) : (f.restrict K).range = K.map f := by simp_rw [SetLike.ext_iff, mem_range, mem_map, restrict_apply, SetLike.exists, From 765291c013275e8322559524b582140bf6773c3c Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 7 Aug 2024 12:49:12 +0000 Subject: [PATCH 091/359] feat: have/let linter (#12190) The main PR for the `have` vs `let` linter. The `have` vs `let` linter emits a warning on `have`s introducing a hypothesis whose Type is not `Prop`. There are three settings: * `0` -- inactive; * `1` -- active only on noisy declarations; * `2` or more -- always active. The default value is `1`. The linter lints mathlib, but, by default, it is silent on complete proofs. #12157 is an experiment showing what changes the linter would suggest on mathlib. --- Mathlib.lean | 1 + Mathlib/Tactic.lean | 1 + Mathlib/Tactic/Linter.lean | 1 + Mathlib/Tactic/Linter/HaveLetLinter.lean | 130 +++++++++++++++++++++ test/HaveLetLinter.lean | 139 +++++++++++++++++++++++ 5 files changed, 272 insertions(+) create mode 100644 Mathlib/Tactic/Linter/HaveLetLinter.lean create mode 100644 test/HaveLetLinter.lean diff --git a/Mathlib.lean b/Mathlib.lean index e2132682089cc..8d37716480732 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4121,6 +4121,7 @@ import Mathlib.Tactic.LinearCombination import Mathlib.Tactic.Linter import Mathlib.Tactic.Linter.GlobalAttributeIn import Mathlib.Tactic.Linter.HashCommandLinter +import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Linter.Lint import Mathlib.Tactic.Linter.MinImports import Mathlib.Tactic.Linter.OldObtain diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index c5fdf61f4f3fd..8aa3d3c995b74 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -121,6 +121,7 @@ import Mathlib.Tactic.LinearCombination import Mathlib.Tactic.Linter import Mathlib.Tactic.Linter.GlobalAttributeIn import Mathlib.Tactic.Linter.HashCommandLinter +import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Linter.Lint import Mathlib.Tactic.Linter.MinImports import Mathlib.Tactic.Linter.OldObtain diff --git a/Mathlib/Tactic/Linter.lean b/Mathlib/Tactic/Linter.lean index 5873024a791a8..5ed156753cbd9 100644 --- a/Mathlib/Tactic/Linter.lean +++ b/Mathlib/Tactic/Linter.lean @@ -9,6 +9,7 @@ This file is ignored by `shake`: import Mathlib.Tactic.Linter.GlobalAttributeIn import Mathlib.Tactic.Linter.HashCommandLinter +import Mathlib.Tactic.Linter.HaveLetLinter import Mathlib.Tactic.Linter.Lint import Mathlib.Tactic.Linter.RefineLinter import Mathlib.Tactic.Linter.Style diff --git a/Mathlib/Tactic/Linter/HaveLetLinter.lean b/Mathlib/Tactic/Linter/HaveLetLinter.lean new file mode 100644 index 0000000000000..e59b5417ef447 --- /dev/null +++ b/Mathlib/Tactic/Linter/HaveLetLinter.lean @@ -0,0 +1,130 @@ +/- +Copyright (c) 2024 Damiano Testa. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa +-/ + +import Lean.Elab.Command +import Lean.Server.InfoUtils + +/-! +# The `have` vs `let` linter + +The `have` vs `let` linter flags uses of `have` to introduce a hypothesis whose Type is not `Prop`. + +The option for this linter is a natural number, but really there are only 3 settings: +* `0` -- inactive; +* `1` -- active only on noisy declarations; +* `2` or more -- always active. + +TODO: +* Also lint `let` vs `have`. +* `haveI` may need to change to `let/letI`? +* `replace`, `classical!`, `classical`, `tauto` internally use `have`: + should the linter act on them as well? +-/ + +open Lean Elab Command Meta + +namespace Mathlib.Linter + +/-- The `have` vs `let` linter emits a warning on `have`s introducing a hypothesis whose +Type is not `Prop`. +There are three settings: +* `0` -- inactive; +* `1` -- active only on noisy declarations; +* `2` or more -- always active. + +The default value is `1`. +-/ +register_option linter.haveLet : Nat := { + defValue := 1 + descr := "enable the `have` vs `let` linter:\n\ + * 0 -- inactive;\n\ + * 1 -- active only on noisy declarations;\n\ + * 2 or more -- always active." +} + +namespace haveLet + +/-- find the `have` syntax. -/ +partial +def isHave? : Syntax → Bool + | .node _ ``Lean.Parser.Tactic.tacticHave_ _ => true + |_ => false + +end haveLet + +end Mathlib.Linter + +namespace Mathlib.Linter.haveLet + +/-- a monadic version of `Lean.Elab.InfoTree.foldInfo`. +Used to infer types inside a `CommandElabM`. -/ +def InfoTree.foldInfoM {α m} [Monad m] (f : ContextInfo → Info → α → m α) (init : α) : + InfoTree → m α := + InfoTree.foldInfo (fun ctx i ma => do f ctx i (← ma)) (pure init) + +/-- given a `ContextInfo`, a `LocalContext` and an `Array` of `Expr`essions `es`, +`areProp_toFormat` creates a `MetaM` context, and returns an array of pairs consisting of +* a `Bool`ean, answering the question of whether the Type of `e` is a `Prop` or not, and +* the pretty-printed `Format` of `e` +for each `Expr`ession `e` in `es`. +Concretely, `areProp_toFormat` runs `inferType` in `CommandElabM`. +This is the kind of monadic lift that `nonPropHaves` uses to decide whether the Type of a `have` +is in `Prop` or not. +The output `Format` is just so that the linter displays a better message. -/ +def areProp_toFormat (ctx : ContextInfo) (lc : LocalContext) (es : Array Expr) : + CommandElabM (Array (Bool × Format)) := do + ctx.runMetaM lc do + es.mapM fun e => do + let typ ← inferType (← instantiateMVars e) + return (typ.isProp, ← ppExpr e) + +/-- returns the `have` syntax whose corresponding hypothesis does not have Type `Prop` and +also a `Format`ted version of the corresponding Type. -/ +partial +def nonPropHaves : InfoTree → CommandElabM (Array (Syntax × Format)) := + InfoTree.foldInfoM (init := #[]) fun ctx info args => return args ++ (← do + let .ofTacticInfo i := info | return #[] + let stx := i.stx + let .original .. := stx.getHeadInfo | return #[] + unless isHave? stx do return #[] + let mctx := i.mctxAfter + let mvdecls := (i.goalsAfter.map (mctx.decls.find? ·)).reduceOption + -- we extract the `MetavarDecl` with largest index after a `have`, since this one + -- holds information about the metavariable where `have` introduces the new hypothesis. + let largestIdx := mvdecls.toArray.qsort (·.index > ·.index) + -- the relevant `LocalContext` + let lc := (largestIdx.getD 0 default).lctx + -- we also accumulate all `fvarId`s from all local contexts before the use of `have` + -- so that we can then isolate the `fvarId`s that are created by `have` + let oldMvdecls := (i.goalsBefore.map (mctx.decls.find? ·)).reduceOption + let oldLctx := oldMvdecls.map (·.lctx) + let oldDecls := (oldLctx.map (·.decls.toList.reduceOption)).join + let oldFVars := oldDecls.map (·.fvarId) + -- `newDecls` are the local declarations whose `FVarID` did not exist before the `have` + -- effectively they are the declarations that we want to test for being in `Prop` or not. + let newDecls := lc.decls.toList.reduceOption.filter (! oldFVars.contains ·.fvarId) + -- now, we get the `MetaM` state up and running to find the types of each entry of `newDecls` + let fmts ← areProp_toFormat ctx lc (newDecls.map (·.type)).toArray + let (_propFmts, typeFmts) := (fmts.zip (newDecls.map (·.userName)).toArray).partition (·.1.1) + -- everything that is a Type triggers a warning on `have` + return typeFmts.map fun ((_, fmt), na) => (stx, f!"{na} : {fmt}")) + +/-- The main implementation of the `have` vs `let` linter. -/ +def haveLetLinter : Linter where run := withSetOptionIn fun _stx => do + let gh := linter.haveLet.get (← getOptions) + unless gh != 0 && (← getInfoState).enabled do + return + unless gh == 1 && (← MonadState.get).messages.unreported.isEmpty do + let trees ← getInfoTrees + for t in trees.toArray do + for (s, fmt) in ← nonPropHaves t do + -- Since the linter option is not in `Bool`, the standard `Linter.logLint` does not work. + -- We emulate it with `logWarningAt` + logWarningAt s <| .tagged linter.haveLet.name + m!"'{fmt}' is a Type and not a Prop. Consider using 'let' instead of 'have'.\n\ + You can disable this linter using `set_option linter.haveLet 0`" + +initialize addLinter haveLetLinter diff --git a/test/HaveLetLinter.lean b/test/HaveLetLinter.lean new file mode 100644 index 0000000000000..35accc8e607e7 --- /dev/null +++ b/test/HaveLetLinter.lean @@ -0,0 +1,139 @@ +import Mathlib.Tactic.Linter.HaveLetLinter +import Mathlib.Tactic.Tauto + +/-- +A tactic that adds a vacuous `sorry`. Useful for testing the chattiness of the `haveLet` linter. +-/ +elab "noise" : tactic => do + Lean.Elab.Tactic.evalTactic (← `(tactic| have : 0 = 0 := sorry; clear this )) + +set_option linter.haveLet 2 in +#guard_msgs in +-- check that `tauto`, `replace`, `classical` are ignored +example : True := by + classical + let zero' := 0 + replace _zero := zero' + let eq := (rfl : 0 = 0) + replace _eq := eq + tauto + +/-- +warning: declaration uses 'sorry' +--- +warning: '_zero : ℕ' is a Type and not a Prop. Consider using 'let' instead of 'have'. +You can disable this linter using `set_option linter.haveLet 0` +-/ +#guard_msgs in +example : True := by + noise + have ⟨_zero, _⟩ : Fin 1 := ⟨0, Nat.zero_lt_one⟩ + exact .intro + +/-- +warning: '_zero : ℕ' is a Type and not a Prop. Consider using 'let' instead of 'have'. +You can disable this linter using `set_option linter.haveLet 0` +-/ +#guard_msgs in +set_option linter.haveLet 2 in +example : True := by + have ⟨_zero, _⟩ : Fin 1 := ⟨0, Nat.zero_lt_one⟩ + exact .intro + +#guard_msgs in +example : True := by + have ⟨_zero, _⟩ : Fin 1 := ⟨0, Nat.zero_lt_one⟩ + exact .intro + +/-- +warning: declaration uses 'sorry' +--- +warning: '_zero : ℕ' is a Type and not a Prop. Consider using 'let' instead of 'have'. +You can disable this linter using `set_option linter.haveLet 0` +-/ +#guard_msgs in +example : True := by + have ⟨_zero, _⟩ : Fin 1 := ⟨0, Nat.zero_lt_one⟩ + noise + exact .intro + +/-- +warning: declaration uses 'sorry' +--- +warning: '_zero : ℕ' is a Type and not a Prop. Consider using 'let' instead of 'have'. +You can disable this linter using `set_option linter.haveLet 0` +-/ +#guard_msgs in +example : True := by + have ⟨_zero, _⟩ : Fin 1 := ⟨0, Nat.zero_lt_one⟩ + noise + exact .intro + +/-- +warning: declaration uses 'sorry' +--- +warning: '_a : ℕ' is a Type and not a Prop. Consider using 'let' instead of 'have'. +You can disable this linter using `set_option linter.haveLet 0` +--- +warning: '_b : ℕ' is a Type and not a Prop. Consider using 'let' instead of 'have'. +You can disable this linter using `set_option linter.haveLet 0` +--- +warning: '_oh : ℕ' is a Type and not a Prop. Consider using 'let' instead of 'have'. +You can disable this linter using `set_option linter.haveLet 0` +--- +warning: '_b : ℕ' is a Type and not a Prop. Consider using 'let' instead of 'have'. +You can disable this linter using `set_option linter.haveLet 0` +-/ +#guard_msgs in +example : True := by + noise + have _a := 0 + have _b : Nat := 0 + have _b : 0 = 0 := rfl + have _oh : Nat := 0 + have _b : Nat := 2 + tauto + +set_option linter.haveLet 0 in +set_option linter.haveLet 1 in +/-- +warning: declaration uses 'sorry' +--- +warning: 'this : ℕ' is a Type and not a Prop. Consider using 'let' instead of 'have'. +You can disable this linter using `set_option linter.haveLet 0` +-/ +#guard_msgs in +example : True := by + have := Nat.succ ?_; + noise + exact .intro + exact 0 + +/-- warning: declaration uses 'sorry' -/ +#guard_msgs in +example : True := by + have := And.intro (Nat.add_comm ?_ ?_) (Nat.add_comm ?_ ?_) + apply True.intro + noise + repeat exact 0 + +/-- warning: declaration uses 'sorry' -/ +#guard_msgs in +example (h : False) : True := by + have : False := h + noise + exact .intro + +set_option linter.haveLet 0 in +set_option linter.haveLet 1 in +/-- +warning: declaration uses 'sorry' +--- +warning: 'this : ℕ' is a Type and not a Prop. Consider using 'let' instead of 'have'. +You can disable this linter using `set_option linter.haveLet 0` +-/ +#guard_msgs in +theorem ghi : True := by + noise + have : Nat := Nat.succ 1; + exact .intro From dcd65e93ad52e3e7c377d145d6deab913d9d4e06 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Wed, 7 Aug 2024 13:48:21 +0000 Subject: [PATCH 092/359] chore (Topology.Algebra.Algebra): name `DiscreteTopology.instContinuousSMul` and de-instance (#15339) This instance is reasonable mathematically but creates complexity in the typeclass search. It is currently not used so we de-instance it. --- Mathlib/Topology/Algebra/Algebra.lean | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/Mathlib/Topology/Algebra/Algebra.lean b/Mathlib/Topology/Algebra/Algebra.lean index b4367b4042dc1..5eff70c7f4cc8 100644 --- a/Mathlib/Topology/Algebra/Algebra.lean +++ b/Mathlib/Topology/Algebra/Algebra.lean @@ -68,8 +68,11 @@ theorem algebraMapCLM_toLinearMap : (algebraMapCLM R A).toLinearMap = Algebra.li rfl /-- If `R` is a discrete topological ring, then any topological ring `S` which is an `R`-algebra -is also a topological `R`-algebra. -/ -instance [TopologicalSemiring A] [DiscreteTopology R] : +is also a topological `R`-algebra. + +NB: This could be an instance but the signature makes it very expensive in search. See #15339 +for the regressions caused by making this an instance. -/ +theorem DiscreteTopology.instContinuousSMul [TopologicalSemiring A] [DiscreteTopology R] : ContinuousSMul R A := continuousSMul_of_algebraMap _ _ continuous_of_discreteTopology end TopologicalAlgebra From b9905c84869511ffb448aab2b861edabefc45172 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 7 Aug 2024 13:48:22 +0000 Subject: [PATCH 093/359] chore: backports for leanprover/lean4#4814 (part 26) (#15538) --- Mathlib/Algebra/Jordan/Basic.lean | 7 +- Mathlib/Algebra/Lie/CartanMatrix.lean | 10 +- Mathlib/Algebra/Lie/IdealOperations.lean | 59 ++++++----- Mathlib/Algebra/Lie/Nilpotent.lean | 26 +++-- Mathlib/Algebra/Lie/Quotient.lean | 18 ++-- .../Analysis/InnerProductSpace/Spectrum.lean | 29 ++--- Mathlib/Analysis/Normed/Lp/PiLp.lean | 6 +- Mathlib/FieldTheory/Laurent.lean | 4 +- Mathlib/FieldTheory/Normal.lean | 18 ++-- Mathlib/LinearAlgebra/JordanChevalley.lean | 4 +- Mathlib/NumberTheory/RamificationInertia.lean | 100 ++++++++++-------- Mathlib/RepresentationTheory/Character.lean | 4 +- 12 files changed, 159 insertions(+), 126 deletions(-) diff --git a/Mathlib/Algebra/Jordan/Basic.lean b/Mathlib/Algebra/Jordan/Basic.lean index a1f2aaba68fe4..edd1a4df00c40 100644 --- a/Mathlib/Algebra/Jordan/Basic.lean +++ b/Mathlib/Algebra/Jordan/Basic.lean @@ -148,15 +148,14 @@ theorem commute_rmul_rmul_sq (a : A) : Commute (R a) (R (a * a)) := end Commute -variable {A} [NonUnitalNonAssocCommRing A] [IsCommJordan A] +variable {A} [NonUnitalNonAssocCommRing A] /-! The endomorphisms on an additive monoid `AddMonoid.End` form a `Ring`, and this may be equipped with a Lie Bracket via `Ring.bracket`. -/ - -theorem two_nsmul_lie_lmul_lmul_add_eq_lie_lmul_lmul_add (a b : A) : +theorem two_nsmul_lie_lmul_lmul_add_eq_lie_lmul_lmul_add [IsCommJordan A] (a b : A) : 2 • (⁅L a, L (a * b)⁆ + ⁅L b, L (b * a)⁆) = ⁅L (a * a), L b⁆ + ⁅L (b * b), L a⁆ := by suffices 2 • ⁅L a, L (a * b)⁆ + 2 • ⁅L b, L (b * a)⁆ + ⁅L b, L (a * a)⁆ + ⁅L a, L (b * b)⁆ = 0 by rwa [← sub_eq_zero, ← sub_sub, sub_eq_add_neg, sub_eq_add_neg, lie_skew, lie_skew, nsmul_add] @@ -191,6 +190,8 @@ private theorem aux1 {a b c : A} : rw [add_lie, add_lie] iterate 15 rw [lie_add] +variable [IsCommJordan A] + private theorem aux2 {a b c : A} : ⁅L a, L (a * a)⁆ + ⁅L a, L (b * b)⁆ + ⁅L a, L (c * c)⁆ + ⁅L a, 2 • L (a * b)⁆ + ⁅L a, 2 • L (c * a)⁆ + ⁅L a, 2 • L (b * c)⁆ + diff --git a/Mathlib/Algebra/Lie/CartanMatrix.lean b/Mathlib/Algebra/Lie/CartanMatrix.lean index 85963d26ef498..732a237be0368 100644 --- a/Mathlib/Algebra/Lie/CartanMatrix.lean +++ b/Mathlib/Algebra/Lie/CartanMatrix.lean @@ -83,7 +83,7 @@ universe u v w noncomputable section -variable (R : Type u) {B : Type v} [CommRing R] [DecidableEq B] [Fintype B] +variable (R : Type u) {B : Type v} [CommRing R] variable (A : Matrix B B ℤ) namespace CartanMatrix @@ -116,7 +116,7 @@ def HH : B × B → FreeLieAlgebra R (Generators B) := uncurry fun i j => ⁅H i, H j⁆ /-- The terms corresponding to the `⁅E, F⁆`-relations. -/ -def EF : B × B → FreeLieAlgebra R (Generators B) := +def EF [DecidableEq B] : B × B → FreeLieAlgebra R (Generators B) := uncurry fun i j => if i = j then ⁅E i, F i⁆ - H i else ⁅E i, F j⁆ /-- The terms corresponding to the `⁅H, E⁆`-relations. -/ @@ -150,19 +150,21 @@ private theorem adF_of_eq_eq_zero (i : B) (h : A i i = 2) : adF R A ⟨i, i⟩ = simp [adF, h, h'] /-- The union of all the relations as a subset of the free Lie algebra. -/ -def toSet : Set (FreeLieAlgebra R (Generators B)) := +def toSet [DecidableEq B] : Set (FreeLieAlgebra R (Generators B)) := (Set.range <| HH R) ∪ (Set.range <| EF R) ∪ (Set.range <| HE R A) ∪ (Set.range <| HF R A) ∪ (Set.range <| adE R A) ∪ (Set.range <| adF R A) /-- The ideal of the free Lie algebra generated by the relations. -/ -def toIdeal : LieIdeal R (FreeLieAlgebra R (Generators B)) := +def toIdeal [DecidableEq B] : LieIdeal R (FreeLieAlgebra R (Generators B)) := LieSubmodule.lieSpan R _ <| toSet R A end Relations end CartanMatrix +variable [DecidableEq B] + /-- The Lie algebra corresponding to a Cartan matrix. Note that it is defined for any matrix of integers. Its value for non-Cartan matrices should be diff --git a/Mathlib/Algebra/Lie/IdealOperations.lean b/Mathlib/Algebra/Lie/IdealOperations.lean index fb048de7ef286..0c18a9d49f2d5 100644 --- a/Mathlib/Algebra/Lie/IdealOperations.lean +++ b/Mathlib/Algebra/Lie/IdealOperations.lean @@ -37,13 +37,35 @@ universe u v w w₁ w₂ namespace LieSubmodule variable {R : Type u} {L : Type v} {M : Type w} {M₂ : Type w₁} -variable [CommRing R] [LieRing L] [LieAlgebra R L] -variable [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] -variable [AddCommGroup M₂] [Module R M₂] [LieRingModule L M₂] [LieModule R L M₂] -variable (N N' : LieSubmodule R L M) (I J : LieIdeal R L) (N₂ : LieSubmodule R L M₂) +variable [CommRing R] [LieRing L] +variable [AddCommGroup M] [Module R M] [LieRingModule L M] +variable [AddCommGroup M₂] [Module R M₂] [LieRingModule L M₂] +variable (N N' : LieSubmodule R L M) (N₂ : LieSubmodule R L M₂) +variable (f : M →ₗ⁅R,L⁆ M₂) section LieIdealOperations +theorem map_comap_le : map f (comap f N₂) ≤ N₂ := + (N₂ : Set M₂).image_preimage_subset f + +theorem map_comap_eq (hf : N₂ ≤ f.range) : map f (comap f N₂) = N₂ := by + rw [SetLike.ext'_iff] + exact Set.image_preimage_eq_of_subset hf + +theorem le_comap_map : N ≤ comap f (map f N) := + (N : Set M).subset_preimage_image f + +theorem comap_map_eq (hf : f.ker = ⊥) : comap f (map f N) = N := by + rw [SetLike.ext'_iff] + exact (N : Set M).preimage_image_eq (f.ker_eq_bot.mp hf) + +@[simp] +theorem map_comap_incl : map N.incl (comap N.incl N') = N ⊓ N' := by + rw [← coe_toSubmodule_eq_iff] + exact (N : Submodule R M).map_comap_subtype N' + +variable [LieAlgebra R L] [LieModule R L M₂] (I J : LieIdeal R L) + /-- Given a Lie module `M` over a Lie algebra `L`, the set of Lie ideals of `L` acts on the set of submodules of `M`. -/ instance hasBracket : Bracket (LieIdeal R L) (LieSubmodule R L M) := @@ -55,7 +77,7 @@ theorem lieIdeal_oper_eq_span : /-- See also `LieSubmodule.lieIdeal_oper_eq_linear_span'` and `LieSubmodule.lieIdeal_oper_eq_tensor_map_range`. -/ -theorem lieIdeal_oper_eq_linear_span : +theorem lieIdeal_oper_eq_linear_span [LieModule R L M] : (↑⁅I, N⁆ : Submodule R M) = Submodule.span R { m | ∃ (x : I) (n : N), ⁅(x : L), (n : M)⁆ = m } := by apply le_antisymm @@ -76,7 +98,7 @@ theorem lieIdeal_oper_eq_linear_span : exact Submodule.subset_span · rw [lieIdeal_oper_eq_span]; apply submodule_span_le_lieSpan -theorem lieIdeal_oper_eq_linear_span' : +theorem lieIdeal_oper_eq_linear_span' [LieModule R L M] : (↑⁅I, N⁆ : Submodule R M) = Submodule.span R { m | ∃ x ∈ I, ∃ n ∈ N, ⁅x, n⁆ = m } := by rw [lieIdeal_oper_eq_linear_span] congr @@ -180,9 +202,7 @@ 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] -variable (f : M →ₗ⁅R,L⁆ M₂) - -theorem map_bracket_eq : map f ⁅I, N⁆ = ⁅I, map f N⁆ := by +theorem map_bracket_eq [LieModule R L M] : map f ⁅I, N⁆ = ⁅I, map f N⁆ := by rw [← coe_toSubmodule_eq_iff, coeSubmodule_map, lieIdeal_oper_eq_linear_span, lieIdeal_oper_eq_linear_span, Submodule.map_span] congr @@ -195,30 +215,11 @@ theorem map_bracket_eq : map f ⁅I, N⁆ = ⁅I, map f N⁆ := by obtain ⟨n, hn, rfl⟩ := (mem_map m₂).mp hm₂ exact ⟨⁅x, n⁆, ⟨x, ⟨n, hn⟩, rfl⟩, by simp⟩ -theorem map_comap_le : map f (comap f N₂) ≤ N₂ := - (N₂ : Set M₂).image_preimage_subset f - -theorem map_comap_eq (hf : N₂ ≤ f.range) : map f (comap f N₂) = N₂ := by - rw [SetLike.ext'_iff] - exact Set.image_preimage_eq_of_subset hf - -theorem le_comap_map : N ≤ comap f (map f N) := - (N : Set M).subset_preimage_image f - -theorem comap_map_eq (hf : f.ker = ⊥) : comap f (map f N) = N := by - rw [SetLike.ext'_iff] - exact (N : Set M).preimage_image_eq (f.ker_eq_bot.mp hf) - -theorem comap_bracket_eq (hf₁ : f.ker = ⊥) (hf₂ : N₂ ≤ f.range) : +theorem comap_bracket_eq [LieModule R L M] (hf₁ : f.ker = ⊥) (hf₂ : N₂ ≤ f.range) : comap f ⁅I, N₂⁆ = ⁅I, comap f N₂⁆ := by conv_lhs => rw [← map_comap_eq N₂ f hf₂] rw [← map_bracket_eq, comap_map_eq _ f hf₁] -@[simp] -theorem map_comap_incl : map N.incl (comap N.incl N') = N ⊓ N' := by - rw [← coe_toSubmodule_eq_iff] - exact (N : Submodule R M).map_comap_subtype N' - end LieIdealOperations end LieSubmodule diff --git a/Mathlib/Algebra/Lie/Nilpotent.lean b/Mathlib/Algebra/Lie/Nilpotent.lean index ad9745bddf4de..24358b6aad617 100644 --- a/Mathlib/Algebra/Lie/Nilpotent.lean +++ b/Mathlib/Algebra/Lie/Nilpotent.lean @@ -34,7 +34,7 @@ section NilpotentModules variable {R : Type u} {L : Type v} {M : Type w} variable [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] -variable [LieRingModule L M] [LieModule R L M] +variable [LieRingModule L M] variable (k : ℕ) (N : LieSubmodule R L M) namespace LieSubmodule @@ -99,6 +99,8 @@ theorem lcs_le_self : N.lcs k ≤ N := by · simp only [lcs_succ] exact (LieSubmodule.mono_lie_right ⊤ ih).trans (N.lie_le_right ⊤) +variable [LieModule R L M] + theorem lowerCentralSeries_eq_lcs_comap : lowerCentralSeries R L N k = (N.lcs k).comap N.incl := by induction' k with k ih · simp @@ -106,7 +108,7 @@ theorem lowerCentralSeries_eq_lcs_comap : lowerCentralSeries R L N k = (N.lcs k) have : N.lcs k ≤ N.incl.range := by rw [N.range_incl] apply lcs_le_self - rw [ih, LieSubmodule.comap_bracket_eq _ _ N.incl N.ker_incl this] + rw [ih, LieSubmodule.comap_bracket_eq _ N.incl _ N.ker_incl this] theorem lowerCentralSeries_map_eq_lcs : (lowerCentralSeries R L N k).map N.incl = N.lcs k := by rw [lowerCentralSeries_eq_lcs_comap, LieSubmodule.map_comap_incl, inf_eq_right] @@ -149,6 +151,9 @@ theorem trivial_iff_lower_central_eq_bot : IsTrivial L M ↔ lowerCentralSeries Set.mem_setOf] exact ⟨x, m, rfl⟩ +section +variable [LieModule R L M] + theorem iterate_toEnd_mem_lowerCentralSeries (x : L) (m : M) (k : ℕ) : (toEnd R L M x)^[k] m ∈ lowerCentralSeries R L M k := by induction' k with k ih @@ -186,7 +191,7 @@ lemma map_lowerCentralSeries_eq {f : M →ₗ⁅R,L⁆ M₂} (hf : Function.Surj apply LieSubmodule.mono_lie_right assumption -variable (R L M) +end open LieAlgebra @@ -217,7 +222,9 @@ theorem exists_lowerCentralSeries_eq_bot_of_isNilpotent [IsNilpotent R L M] : theorem isNilpotent_iff : IsNilpotent R L M ↔ ∃ k, lowerCentralSeries R L M k = ⊥ := ⟨fun h => h.nilpotent, fun h => ⟨h⟩⟩ +section variable {R L M} +variable [LieModule R L M] theorem _root_.LieSubmodule.isNilpotent_iff_exists_lcs_eq_bot (N : LieSubmodule R L M) : LieModule.IsNilpotent R L N ↔ ∃ k, N.lcs k = ⊥ := by @@ -292,6 +299,8 @@ theorem iInf_lcs_le_of_isNilpotent_quot (h : IsNilpotent R L (M ⧸ N)) : obtain ⟨k, hk⟩ := (isNilpotent_quotient_iff R L M N).mp h exact iInf_le_of_le k hk +end + /-- Given a nilpotent Lie module `M` with lower central series `M = C₀ ≥ C₁ ≥ ⋯ ≥ Cₖ = ⊥`, this is the natural number `k` (the number of inclusions). @@ -346,7 +355,7 @@ noncomputable def lowerCentralSeriesLast : LieSubmodule R L M := | 0 => ⊥ | k + 1 => lowerCentralSeries R L M k -theorem lowerCentralSeriesLast_le_max_triv : +theorem lowerCentralSeriesLast_le_max_triv [LieModule R L M] : lowerCentralSeriesLast R L M ≤ maxTrivSubmodule R L M := by rw [lowerCentralSeriesLast] cases' h : nilpotencyLength R L M with k @@ -375,6 +384,8 @@ theorem lowerCentralSeriesLast_le_of_not_isTrivial [IsNilpotent R L M] (h : ¬ I · contradiction · exact antitone_lowerCentralSeries _ _ _ (Nat.lt_succ.mp h) +variable [LieModule R L M] + /-- For a nilpotent Lie module `M` of a Lie algebra `L`, the first term in the lower central series of `M` contains a non-zero element on which `L` acts trivially unless the entire action is trivial. @@ -427,6 +438,7 @@ end LieModule namespace LieSubmodule variable {N₁ N₂ : LieSubmodule R L M} +variable [LieModule R L M] /-- The upper (aka ascending) central series. @@ -509,6 +521,7 @@ section Morphisms open LieModule Function +variable [LieModule R L M] variable {L₂ M₂ : Type*} [LieRing L₂] [LieAlgebra R L₂] variable [AddCommGroup M₂] [Module R M₂] [LieRingModule L₂ M₂] [LieModule R L₂ M₂] variable {f : L →ₗ⁅R⁆ L₂} {g : M →ₗ[R] M₂} @@ -709,7 +722,7 @@ namespace LieIdeal open LieModule variable {R L : Type*} [CommRing R] [LieRing L] [LieAlgebra R L] (I : LieIdeal R L) -variable (M : Type*) [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] +variable (M : Type*) [AddCommGroup M] [Module R M] [LieRingModule L M] variable (k : ℕ) /-- Given a Lie module `M` over a Lie algebra `L` together with an ideal `I` of `L`, this is the @@ -733,7 +746,8 @@ theorem lcs_top : (⊤ : LieIdeal R L).lcs M k = lowerCentralSeries R L M k := rfl -- Porting note: added `LieSubmodule.toSubmodule` in the statement -theorem coe_lcs_eq : LieSubmodule.toSubmodule (I.lcs M k) = lowerCentralSeries R I M k := by +theorem coe_lcs_eq [LieModule R L M] : + LieSubmodule.toSubmodule (I.lcs M k) = lowerCentralSeries R I M k := by induction' k with k ih · simp · simp_rw [lowerCentralSeries_succ, lcs_succ, LieSubmodule.lieIdeal_oper_eq_linear_span', ← diff --git a/Mathlib/Algebra/Lie/Quotient.lean b/Mathlib/Algebra/Lie/Quotient.lean index 942582f8857e2..cefd7b350bb49 100644 --- a/Mathlib/Algebra/Lie/Quotient.lean +++ b/Mathlib/Algebra/Lie/Quotient.lean @@ -33,9 +33,9 @@ universe u v w w₁ w₂ namespace LieSubmodule variable {R : Type u} {L : Type v} {M : Type w} -variable [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] -variable [LieRingModule L M] [LieModule R L M] -variable (N N' : LieSubmodule R L M) (I J : LieIdeal R L) +variable [CommRing R] [LieRing L] [AddCommGroup M] [Module R M] +variable [LieRingModule L M] +variable (N N' : LieSubmodule R L M) /-- The quotient of a Lie module by a Lie submodule. It is a Lie module. -/ instance : HasQuotient M (LieSubmodule R L M) := @@ -68,9 +68,16 @@ lie_submodule of the lie_module `N`. -/ abbrev mk : M → M ⧸ N := Submodule.Quotient.mk +-- Porting note: added to replace `mk_eq_zero` as simp lemma. +@[simp] +theorem mk_eq_zero' {m : M} : mk (N := N) m = 0 ↔ m ∈ N := + Submodule.Quotient.mk_eq_zero N.toSubmodule + theorem is_quotient_mk (m : M) : Quotient.mk'' m = (mk m : M ⧸ N) := rfl +variable [LieAlgebra R L] [LieModule R L M] (I J : LieIdeal R L) + /-- Given a Lie module `M` over a Lie algebra `L`, together with a Lie submodule `N ⊆ M`, there is a natural linear map from `L` to the endomorphisms of `M` leaving `N` invariant. -/ def lieSubmoduleInvariant : L →ₗ[R] Submodule.compatibleMaps N.toSubmodule N.toSubmodule := @@ -176,11 +183,6 @@ instance isNoetherian [IsNoetherian R M] : IsNoetherian R (M ⧸ N) := theorem mk_eq_zero {m : M} : mk' N m = 0 ↔ m ∈ N := Submodule.Quotient.mk_eq_zero N.toSubmodule --- Porting note: added to replace `mk_eq_zero` as simp lemma. -@[simp] -theorem mk_eq_zero' {m : M} : mk (N := N) m = 0 ↔ m ∈ N := - Submodule.Quotient.mk_eq_zero N.toSubmodule - @[simp] theorem mk'_ker : (mk' N).ker = N := by ext; simp diff --git a/Mathlib/Analysis/InnerProductSpace/Spectrum.lean b/Mathlib/Analysis/InnerProductSpace/Spectrum.lean index 715ba17b18550..68d1c7ecfc1b4 100644 --- a/Mathlib/Analysis/InnerProductSpace/Spectrum.lean +++ b/Mathlib/Analysis/InnerProductSpace/Spectrum.lean @@ -60,23 +60,24 @@ namespace LinearMap namespace IsSymmetric -variable {T : E →ₗ[𝕜] E} (hT : T.IsSymmetric) +variable {T : E →ₗ[𝕜] E} /-- A self-adjoint operator preserves orthogonal complements of its eigenspaces. -/ -theorem invariant_orthogonalComplement_eigenspace (μ : 𝕜) (v : E) (hv : v ∈ (eigenspace T μ)ᗮ) : - T v ∈ (eigenspace T μ)ᗮ := by +theorem invariant_orthogonalComplement_eigenspace (hT : T.IsSymmetric) (μ : 𝕜) + (v : E) (hv : v ∈ (eigenspace T μ)ᗮ) : T v ∈ (eigenspace T μ)ᗮ := by intro w hw have : T w = (μ : 𝕜) • w := by rwa [mem_eigenspace_iff] at hw simp [← hT w, this, inner_smul_left, hv w hw] /-- The eigenvalues of a self-adjoint operator are real. -/ -theorem conj_eigenvalue_eq_self {μ : 𝕜} (hμ : HasEigenvalue T μ) : conj μ = μ := by +theorem conj_eigenvalue_eq_self (hT : T.IsSymmetric) {μ : 𝕜} (hμ : HasEigenvalue T μ) : + conj μ = μ := by obtain ⟨v, hv₁, hv₂⟩ := hμ.exists_hasEigenvector rw [mem_eigenspace_iff] at hv₁ simpa [hv₂, inner_smul_left, inner_smul_right, hv₁] using hT v v /-- The eigenspaces of a self-adjoint operator are mutually orthogonal. -/ -theorem orthogonalFamily_eigenspaces : +theorem orthogonalFamily_eigenspaces (hT : T.IsSymmetric) : OrthogonalFamily 𝕜 (fun μ => eigenspace T μ) fun μ => (eigenspace T μ).subtypeₗᵢ := by rintro μ ν hμν ⟨v, hv⟩ ⟨w, hw⟩ by_cases hv' : v = 0 @@ -86,21 +87,21 @@ theorem orthogonalFamily_eigenspaces : refine Or.resolve_left ?_ hμν.symm simpa [inner_smul_left, inner_smul_right, hv, hw, H] using (hT v w).symm -theorem orthogonalFamily_eigenspaces' : +theorem orthogonalFamily_eigenspaces' (hT : T.IsSymmetric) : OrthogonalFamily 𝕜 (fun μ : Eigenvalues T => eigenspace T μ) fun μ => (eigenspace T μ).subtypeₗᵢ := hT.orthogonalFamily_eigenspaces.comp Subtype.coe_injective /-- The mutual orthogonal complement of the eigenspaces of a self-adjoint operator on an inner product space is an invariant subspace of the operator. -/ -theorem orthogonalComplement_iSup_eigenspaces_invariant ⦃v : E⦄ (hv : v ∈ (⨆ μ, eigenspace T μ)ᗮ) : - T v ∈ (⨆ μ, eigenspace T μ)ᗮ := by +theorem orthogonalComplement_iSup_eigenspaces_invariant (hT : T.IsSymmetric) + ⦃v : E⦄ (hv : v ∈ (⨆ μ, eigenspace T μ)ᗮ) : T v ∈ (⨆ μ, eigenspace T μ)ᗮ := by rw [← Submodule.iInf_orthogonal] at hv ⊢ exact T.iInf_invariant hT.invariant_orthogonalComplement_eigenspace v hv /-- The mutual orthogonal complement of the eigenspaces of a self-adjoint operator on an inner product space has no eigenvalues. -/ -theorem orthogonalComplement_iSup_eigenspaces (μ : 𝕜) : +theorem orthogonalComplement_iSup_eigenspaces (hT : T.IsSymmetric) (μ : 𝕜) : eigenspace (T.restrict hT.orthogonalComplement_iSup_eigenspaces_invariant) μ = ⊥ := by set p : Submodule 𝕜 E := (⨆ μ, eigenspace T μ)ᗮ refine eigenspace_restrict_eq_bot hT.orthogonalComplement_iSup_eigenspaces_invariant ?_ @@ -114,7 +115,8 @@ variable [FiniteDimensional 𝕜 E] /-- The mutual orthogonal complement of the eigenspaces of a self-adjoint operator on a finite-dimensional inner product space is trivial. -/ -theorem orthogonalComplement_iSup_eigenspaces_eq_bot : (⨆ μ, eigenspace T μ)ᗮ = ⊥ := by +theorem orthogonalComplement_iSup_eigenspaces_eq_bot (hT : T.IsSymmetric) : + (⨆ μ, eigenspace T μ)ᗮ = ⊥ := by have hT' : IsSymmetric _ := hT.restrict_invariant hT.orthogonalComplement_iSup_eigenspaces_invariant -- a self-adjoint operator on a nontrivial inner product space has an eigenvalue @@ -122,7 +124,7 @@ theorem orthogonalComplement_iSup_eigenspaces_eq_bot : (⨆ μ, eigenspace T μ) hT'.subsingleton_of_no_eigenvalue_finiteDimensional hT.orthogonalComplement_iSup_eigenspaces exact Submodule.eq_bot_of_subsingleton -theorem orthogonalComplement_iSup_eigenspaces_eq_bot' : +theorem orthogonalComplement_iSup_eigenspaces_eq_bot' (hT : T.IsSymmetric) : (⨆ μ : Eigenvalues T, eigenspace T μ)ᗮ = ⊥ := show (⨆ μ : { μ // eigenspace T μ ≠ ⊥ }, eigenspace T μ)ᗮ = ⊥ by rw [iSup_ne_bot_subtype, hT.orthogonalComplement_iSup_eigenspaces_eq_bot] @@ -144,10 +146,13 @@ theorem directSum_decompose_apply [_hT : Fact T.IsSymmetric] (x : E) (μ : Eigen /-- The eigenspaces of a self-adjoint operator on a finite-dimensional inner product space `E` gives an internal direct sum decomposition of `E`. -/ -theorem direct_sum_isInternal : DirectSum.IsInternal fun μ : Eigenvalues T => eigenspace T μ := +theorem direct_sum_isInternal (hT : T.IsSymmetric) : + DirectSum.IsInternal fun μ : Eigenvalues T => eigenspace T μ := hT.orthogonalFamily_eigenspaces'.isInternal_iff.mpr hT.orthogonalComplement_iSup_eigenspaces_eq_bot' +variable (hT : T.IsSymmetric) + section Version1 /-- Isometry from an inner product space `E` to the direct sum of the eigenspaces of some diff --git a/Mathlib/Analysis/Normed/Lp/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean index dda1bee40dc60..cabcd6dc2bf05 100644 --- a/Mathlib/Analysis/Normed/Lp/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -458,11 +458,9 @@ theorem continuous_equiv_symm [∀ i, UniformSpace (β i)] : instance bornology [∀ i, Bornology (β i)] : Bornology (PiLp p β) := Pi.instBornology --- throughout the rest of the file, we assume `1 ≤ p` -variable [Fact (1 ≤ p)] section Fintype - +variable [hp : Fact (1 ≤ p)] variable [Fintype ι] /-- pseudoemetric space instance on the product of finitely many pseudoemetric spaces, using the @@ -784,7 +782,7 @@ variable [DecidableEq ι] -- Porting note: added `hp` @[simp] -theorem nnnorm_equiv_symm_single [hp : Fact (1 ≤ p)] (i : ι) (b : β i) : +theorem nnnorm_equiv_symm_single (i : ι) (b : β i) : ‖(WithLp.equiv p (∀ i, β i)).symm (Pi.single i b)‖₊ = ‖b‖₊ := by haveI : Nonempty ι := ⟨i⟩ induction p generalizing hp with diff --git a/Mathlib/FieldTheory/Laurent.lean b/Mathlib/FieldTheory/Laurent.lean index 6336d85e6f5df..20063025146b4 100644 --- a/Mathlib/FieldTheory/Laurent.lean +++ b/Mathlib/FieldTheory/Laurent.lean @@ -32,7 +32,7 @@ open Polynomial open scoped nonZeroDivisors -variable {R : Type u} [CommRing R] [hdomain : IsDomain R] (r s : R) (p q : R[X]) (f : RatFunc R) +variable {R : Type u} [CommRing R] (r s : R) (p q : R[X]) (f : RatFunc R) theorem taylor_mem_nonZeroDivisors (hp : p ∈ R[X]⁰) : taylor r p ∈ R[X]⁰ := by rw [mem_nonZeroDivisors_iff] @@ -58,6 +58,8 @@ theorem laurentAux_ofFractionRing_mk (q : R[X]⁰) : ofFractionRing (.mk (taylor r p) ⟨taylor r q, taylor_mem_nonZeroDivisors r q q.prop⟩) := map_apply_ofFractionRing_mk _ _ _ _ +variable [IsDomain R] + theorem laurentAux_div : laurentAux r (algebraMap _ _ p / algebraMap _ _ q) = algebraMap _ _ (taylor r p) / algebraMap _ _ (taylor r q) := diff --git a/Mathlib/FieldTheory/Normal.lean b/Mathlib/FieldTheory/Normal.lean index 6631e005d1ab9..0b609f7d4d2a1 100644 --- a/Mathlib/FieldTheory/Normal.lean +++ b/Mathlib/FieldTheory/Normal.lean @@ -89,8 +89,8 @@ theorem Normal.tower_top_of_normal [h : Normal F E] : Normal K E := theorem AlgHom.normal_bijective [h : Normal F E] (ϕ : E →ₐ[F] K) : Function.Bijective ϕ := h.toIsAlgebraic.bijective_of_isScalarTower' ϕ --- Porting note: `[Field F] [Field E] [Algebra F E]` added by hand. -variable {F E} {E' : Type*} [Field F] [Field E] [Algebra F E] [Field E'] [Algebra F E'] +variable {E F} +variable {E' : Type*} [Field E'] [Algebra F E'] theorem Normal.of_algEquiv [h : Normal F E] (f : E ≃ₐ[F] E') : Normal F E' := by rw [normal_iff] at h ⊢ @@ -174,9 +174,8 @@ instance normal_sup Normal F (E ⊔ E' : IntermediateField F K) := iSup_bool_eq (f := Bool.rec E' E) ▸ normal_iSup (h := by rintro (_|_) <;> infer_instance) --- Porting note `[Field F] [Field K] [Algebra F K]` added by hand. -variable {F K} {L : Type*} [Field F] [Field K] [Field L] [Algebra F L] [Algebra K L] - [Algebra F K] [IsScalarTower F K L] +variable {F K} +variable {L : Type*} [Field L] [Algebra F L] [Algebra K L] [IsScalarTower F K L] @[simp] theorem restrictScalars_normal {E : IntermediateField K L} : @@ -185,8 +184,8 @@ theorem restrictScalars_normal {E : IntermediateField K L} : end IntermediateField --- Porting note `[Field F]` added by hand. -variable {F} {K} {K₁ K₂ K₃ : Type*} [Field F] [Field K₁] [Field K₂] [Field K₃] [Algebra F K₁] +variable {F} {K} +variable {K₁ K₂ K₃ : Type*} [Field K₁] [Field K₂] [Field K₃] [Algebra F K₁] [Algebra F K₂] [Algebra F K₃] (ϕ : K₁ →ₐ[F] K₂) (χ : K₁ ≃ₐ[F] K₂) (ψ : K₂ →ₐ[F] K₃) (ω : K₂ ≃ₐ[F] K₃) @@ -238,11 +237,8 @@ theorem AlgHom.restrictNormal_comp [Normal F E] : AlgHom.ext fun _ => (algebraMap E K₃).injective (by simp only [AlgHom.comp_apply, AlgHom.restrictNormal_commutes]) --- Porting note `[Algebra F K]` added by hand. -theorem AlgHom.fieldRange_of_normal [Algebra F K] {E : IntermediateField F K} [Normal F E] +theorem AlgHom.fieldRange_of_normal {E : IntermediateField F K} [Normal F E] (f : E →ₐ[F] K) : f.fieldRange = E := by --- Porting note: this was `IsScalarTower F E E := by infer_instance`. - letI : Algebra E E := Algebra.id E let g := f.restrictNormal' E rw [← show E.val.comp ↑g = f from DFunLike.ext_iff.mpr (f.restrictNormal_commutes E), ← AlgHom.map_fieldRange, AlgEquiv.fieldRange_eq_top g, ← AlgHom.fieldRange_eq_map, diff --git a/Mathlib/LinearAlgebra/JordanChevalley.lean b/Mathlib/LinearAlgebra/JordanChevalley.lean index 3361fb5b7e796..76d5f6488464f 100644 --- a/Mathlib/LinearAlgebra/JordanChevalley.lean +++ b/Mathlib/LinearAlgebra/JordanChevalley.lean @@ -37,7 +37,7 @@ open Algebra Polynomial namespace Module.End -variable {K V : Type*} [Field K] [AddCommGroup V] [Module K V] [FiniteDimensional K V] {f : End K V} +variable {K V : Type*} [Field K] [AddCommGroup V] [Module K V] {f : End K V} theorem exists_isNilpotent_isSemisimple_of_separable_of_dvd_pow {P : K[X]} {k : ℕ} (sep : P.Separable) (nil : minpoly K f ∣ P ^ k) : @@ -64,6 +64,8 @@ theorem exists_isNilpotent_isSemisimple_of_separable_of_dvd_pow {P : K[X]} {k : · replace hss : aeval s P = 0 := by rwa [Subtype.ext_iff, coe_aeval_mk_apply] at hss exact isSemisimple_of_squarefree_aeval_eq_zero sep.squarefree hss +variable [FiniteDimensional K V] + /-- **Jordan-Chevalley-Dunford decomposition**: an endomorphism of a finite-dimensional vector space over a perfect field may be written as a sum of nilpotent and semisimple endomorphisms. Moreover these nilpotent and semisimple components are polynomial expressions in the original endomorphism. diff --git a/Mathlib/NumberTheory/RamificationInertia.lean b/Mathlib/NumberTheory/RamificationInertia.lean index c6551c9ef7b46..c059ce200afdd 100644 --- a/Mathlib/NumberTheory/RamificationInertia.lean +++ b/Mathlib/NumberTheory/RamificationInertia.lean @@ -209,7 +209,7 @@ section FinrankQuotientMap open scoped nonZeroDivisors variable [Algebra R S] -variable {K : Type*} [Field K] [Algebra R K] [hRK : IsFractionRing R K] +variable {K : Type*} [Field K] [Algebra R K] variable {L : Type*} [Field L] [Algebra S L] [IsFractionRing S L] variable {V V' V'' : Type*} variable [AddCommGroup V] [Module R V] [Module K V] [IsScalarTower R K V] @@ -217,46 +217,6 @@ variable [AddCommGroup V'] [Module R V'] [Module S V'] [IsScalarTower R S V'] variable [AddCommGroup V''] [Module R V''] variable (K) -/-- Let `V` be a vector space over `K = Frac(R)`, `S / R` a ring extension -and `V'` a module over `S`. If `b`, in the intersection `V''` of `V` and `V'`, -is linear independent over `S` in `V'`, then it is linear independent over `R` in `V`. - -The statement we prove is actually slightly more general: - * it suffices that the inclusion `algebraMap R S : R → S` is nontrivial - * the function `f' : V'' → V'` doesn't need to be injective --/ -theorem FinrankQuotientMap.linearIndependent_of_nontrivial [IsDedekindDomain R] - (hRS : RingHom.ker (algebraMap R S) ≠ ⊤) (f : V'' →ₗ[R] V) (hf : Function.Injective f) - (f' : V'' →ₗ[R] V') {ι : Type*} {b : ι → V''} (hb' : LinearIndependent S (f' ∘ b)) : - LinearIndependent K (f ∘ b) := by - contrapose! hb' with hb - -- Informally, if we have a nontrivial linear dependence with coefficients `g` in `K`, - -- then we can find a linear dependence with coefficients `I.Quotient.mk g'` in `R/I`, - -- where `I = ker (algebraMap R S)`. - -- We make use of the same principle but stay in `R` everywhere. - simp only [linearIndependent_iff', not_forall] at hb ⊢ - obtain ⟨s, g, eq, j', hj's, hj'g⟩ := hb - use s - obtain ⟨a, hag, j, hjs, hgI⟩ := Ideal.exist_integer_multiples_not_mem hRS s g hj's hj'g - choose g'' hg'' using hag - letI := Classical.propDecidable - let g' i := if h : i ∈ s then g'' i h else 0 - have hg' : ∀ i ∈ s, algebraMap _ _ (g' i) = a * g i := by - intro i hi; exact (congr_arg _ (dif_pos hi)).trans (hg'' i hi) - -- Because `R/I` is nontrivial, we can lift `g` to a nontrivial linear dependence in `S`. - have hgI : algebraMap R S (g' j) ≠ 0 := by - simp only [FractionalIdeal.mem_coeIdeal, not_exists, not_and'] at hgI - exact hgI _ (hg' j hjs) - refine ⟨fun i => algebraMap R S (g' i), ?_, j, hjs, hgI⟩ - have eq : f (∑ i ∈ s, g' i • b i) = 0 := by - rw [map_sum, ← smul_zero a, ← eq, Finset.smul_sum] - refine Finset.sum_congr rfl ?_ - intro i hi - rw [LinearMap.map_smul, ← IsScalarTower.algebraMap_smul K, hg' i hi, ← smul_assoc, - smul_eq_mul, Function.comp_apply] - simp only [IsScalarTower.algebraMap_smul, ← map_smul, ← map_sum, - (f.map_eq_zero_iff hf).mp eq, LinearMap.map_zero, (· ∘ ·)] - open scoped Matrix variable {K} @@ -366,7 +326,50 @@ theorem FinrankQuotientMap.span_eq_top [IsDomain R] [IsDomain S] [Algebra K L] [ infer_instance refine IsFractionRing.ideal_span_singleton_map_subset R hRL span_d hx -variable (K L) +variable (K) +variable [hRK : IsFractionRing R K] + +/-- Let `V` be a vector space over `K = Frac(R)`, `S / R` a ring extension +and `V'` a module over `S`. If `b`, in the intersection `V''` of `V` and `V'`, +is linear independent over `S` in `V'`, then it is linear independent over `R` in `V`. + +The statement we prove is actually slightly more general: + * it suffices that the inclusion `algebraMap R S : R → S` is nontrivial + * the function `f' : V'' → V'` doesn't need to be injective +-/ +theorem FinrankQuotientMap.linearIndependent_of_nontrivial [IsDedekindDomain R] + (hRS : RingHom.ker (algebraMap R S) ≠ ⊤) (f : V'' →ₗ[R] V) (hf : Function.Injective f) + (f' : V'' →ₗ[R] V') {ι : Type*} {b : ι → V''} (hb' : LinearIndependent S (f' ∘ b)) : + LinearIndependent K (f ∘ b) := by + contrapose! hb' with hb + -- Informally, if we have a nontrivial linear dependence with coefficients `g` in `K`, + -- then we can find a linear dependence with coefficients `I.Quotient.mk g'` in `R/I`, + -- where `I = ker (algebraMap R S)`. + -- We make use of the same principle but stay in `R` everywhere. + simp only [linearIndependent_iff', not_forall] at hb ⊢ + obtain ⟨s, g, eq, j', hj's, hj'g⟩ := hb + use s + obtain ⟨a, hag, j, hjs, hgI⟩ := Ideal.exist_integer_multiples_not_mem hRS s g hj's hj'g + choose g'' hg'' using hag + letI := Classical.propDecidable + let g' i := if h : i ∈ s then g'' i h else 0 + have hg' : ∀ i ∈ s, algebraMap _ _ (g' i) = a * g i := by + intro i hi; exact (congr_arg _ (dif_pos hi)).trans (hg'' i hi) + -- Because `R/I` is nontrivial, we can lift `g` to a nontrivial linear dependence in `S`. + have hgI : algebraMap R S (g' j) ≠ 0 := by + simp only [FractionalIdeal.mem_coeIdeal, not_exists, not_and'] at hgI + exact hgI _ (hg' j hjs) + refine ⟨fun i => algebraMap R S (g' i), ?_, j, hjs, hgI⟩ + have eq : f (∑ i ∈ s, g' i • b i) = 0 := by + rw [map_sum, ← smul_zero a, ← eq, Finset.smul_sum] + refine Finset.sum_congr rfl ?_ + intro i hi + rw [LinearMap.map_smul, ← IsScalarTower.algebraMap_smul K, hg' i hi, ← smul_assoc, + smul_eq_mul, Function.comp_apply] + simp only [IsScalarTower.algebraMap_smul, ← map_smul, ← map_sum, + (f.map_eq_zero_iff hf).mp eq, LinearMap.map_zero, (· ∘ ·)] + +variable (L) /-- If `p` is a maximal ideal of `R`, and `S` is the integral closure of `R` in `L`, then the dimension `[S/pS : R/p]` is equal to `[Frac(S) : Frac(R)]`. -/ @@ -431,13 +434,12 @@ noncomputable instance Quotient.algebraQuotientPowRamificationIdx : Algebra (R theorem Quotient.algebraMap_quotient_pow_ramificationIdx (x : R) : algebraMap (R ⧸ p) (S ⧸ P ^ e) (Ideal.Quotient.mk p x) = Ideal.Quotient.mk (P ^ e) (f x) := rfl -variable [hfp : NeZero (ramificationIdx f p P)] - /-- If `P` lies over `p`, then `R / p` has a canonical map to `S / P`. This can't be an instance since the map `f : R → S` is generally not inferrable. -/ -def Quotient.algebraQuotientOfRamificationIdxNeZero : Algebra (R ⧸ p) (S ⧸ P) := +def Quotient.algebraQuotientOfRamificationIdxNeZero [hfp : NeZero (ramificationIdx f p P)] : + Algebra (R ⧸ p) (S ⧸ P) := Quotient.algebraQuotientOfLEComap (le_comap_of_ramificationIdx_ne_zero hfp.out) set_option synthInstance.checkSynthOrder false -- Porting note: this is okay by the remark below @@ -445,7 +447,8 @@ set_option synthInstance.checkSynthOrder false -- Porting note: this is okay by attribute [local instance] Ideal.Quotient.algebraQuotientOfRamificationIdxNeZero @[simp] -theorem Quotient.algebraMap_quotient_of_ramificationIdx_neZero (x : R) : +theorem Quotient.algebraMap_quotient_of_ramificationIdx_neZero + [NeZero (ramificationIdx f p P)] (x : R) : algebraMap (R ⧸ p) (S ⧸ P) (Ideal.Quotient.mk p x) = Ideal.Quotient.mk P (f x) := rfl /-- The inclusion `(P^(i + 1) / P^e) ⊂ (P^i / P^e)`. -/ @@ -485,6 +488,9 @@ theorem quotientToQuotientRangePowQuotSuccAux_mk {i : ℕ} {a : S} (a_mem : a Submodule.Quotient.mk ⟨_, Ideal.mem_map_of_mem _ (Ideal.mul_mem_right x _ a_mem)⟩ := by apply Quotient.map'_mk'' +section +variable [hfp : NeZero (ramificationIdx f p P)] + /-- `S ⧸ P` embeds into the quotient by `P^(i+1) ⧸ P^e` as a subspace of `P^i ⧸ P^e`. -/ noncomputable def quotientToQuotientRangePowQuotSucc {i : ℕ} {a : S} (a_mem : a ∈ P ^ i) : S ⧸ P →ₗ[R ⧸ p] @@ -601,6 +607,8 @@ theorem rank_pow_quot [IsDedekindDomain S] [p.IsMaximal] [P.IsPrime] (hP0 : P rw [Nat.sub_self, zero_nsmul, map_quotient_self] exact rank_bot (R ⧸ p) (S ⧸ P ^ e) +end + /-- If `p` is a maximal ideal of `R`, `S` extends `R` and `P^e` lies over `p`, then the dimension `[S/(P^e) : R/p]` is equal to `e * [S/P : R/p]`. -/ theorem rank_prime_pow_ramificationIdx [IsDedekindDomain S] [p.IsMaximal] [P.IsPrime] diff --git a/Mathlib/RepresentationTheory/Character.lean b/Mathlib/RepresentationTheory/Character.lean index 3342fa6cce6db..a0566328c0f4a 100644 --- a/Mathlib/RepresentationTheory/Character.lean +++ b/Mathlib/RepresentationTheory/Character.lean @@ -68,7 +68,9 @@ theorem char_tensor' (V W : FDRep k G) : /-- The character of isomorphic representations is the same. -/ theorem char_iso {V W : FDRep k G} (i : V ≅ W) : V.character = W.character := by - ext g; simp only [character, FDRep.Iso.conj_ρ i]; exact (trace_conj' (V.ρ g) _).symm + ext g + simp only [character, FDRep.Iso.conj_ρ i] + exact (trace_conj' (V.ρ g) _).symm end Monoid From 78791a5454dd1a276f320f1879c259a2f7a38b79 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 7 Aug 2024 14:17:25 +0000 Subject: [PATCH 094/359] feat(RelSeries): RelSeries.take and RelSeries.drop (#15557) from the Carleson project. --- Mathlib/Order/RelSeries.lean | 34 ++++++++++++++++++++++++++++++++++ 1 file changed, 34 insertions(+) diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index 2438ce2f05ea8..0c5da650cc937 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -571,6 +571,40 @@ lemma smash_succ_natAdd {p q : RelSeries r} (h : p.last = q.head) (i : Fin q.len (smash p q h).last = q.last := by delta smash last; aesop +/-- Given the series `a₀ -r→ … -r→ aᵢ -r→ … -r→ aₙ`, the series `a₀ -r→ … -r→ aᵢ`. -/ +@[simps! length] +def take {r : Rel α α} (p : RelSeries r) (i : Fin (p.length + 1)) : RelSeries r where + length := i + toFun := fun ⟨j, h⟩ => p.toFun ⟨j, by omega⟩ + step := fun ⟨j, h⟩ => p.step ⟨j, by omega⟩ + +@[simp] +lemma head_take (p : RelSeries r) (i : Fin (p.length + 1)) : + (p.take i).head = p.head := by simp [take, head] + +@[simp] +lemma last_take (p : RelSeries r) (i : Fin (p.length + 1)) : + (p.take i).last = p i := by simp [take, last, Fin.last] + +/-- Given the series `a₀ -r→ … -r→ aᵢ -r→ … -r→ aₙ`, the series `aᵢ₊₁ -r→ … -r→ aᵢ`. -/ +@[simps! length] +def drop (p : RelSeries r) (i : Fin (p.length + 1)) : RelSeries r where + length := p.length - i + toFun := fun ⟨j, h⟩ => p.toFun ⟨j+i, by omega⟩ + step := fun ⟨j, h⟩ => by + convert p.step ⟨j+i.1, by omega⟩ + simp only [Nat.succ_eq_add_one, Fin.succ_mk]; omega + +@[simp] +lemma head_drop (p : RelSeries r) (i : Fin (p.length + 1)) : (p.drop i).head = p.toFun i := by + simp [drop, head] + +@[simp] +lemma last_drop (p : RelSeries r) (i : Fin (p.length + 1)) : (p.drop i).last = p.last := by + simp only [last, drop, Fin.last] + congr + omega + end RelSeries /-- A type is finite dimensional if its `LTSeries` has bounded length. -/ From e2938e094a7a98a6f18f06ac1bd15cc0d2e89c8e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 7 Aug 2024 14:38:51 +0000 Subject: [PATCH 095/359] chore: backports for leanprover/lean4#4814 (part 29) (#15572) --- .../Geometry/Manifold/ContMDiff/Basic.lean | 35 ++++--- Mathlib/NumberTheory/NumberField/Basic.lean | 12 ++- .../RingTheory/DedekindDomain/Different.lean | 97 ++++++++++-------- Mathlib/RingTheory/Ideal/Norm.lean | 98 ++++++++++--------- .../IntegralClosure/IntegralRestrict.lean | 4 +- 5 files changed, 136 insertions(+), 110 deletions(-) diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index d6a88bd9a46f7..a8f5c4d77931c 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -25,20 +25,21 @@ open Set Filter Function open scoped Topology Manifold variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] - -- declare a smooth manifold `M` over the pair `(E, H)`. + -- declare the prerequisites for a charted space `M` over the pair `(E, H)`. {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] - (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] - [SmoothManifoldWithCorners I M] - -- declare a smooth manifold `M'` over the pair `(E', H')`. + (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] + -- declare the prerequisites for a charted space `M'` over the pair `(E', H')`. {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] - (I' : ModelWithCorners 𝕜 E' H') {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] - [SmoothManifoldWithCorners I' M'] - -- declare a manifold `M''` over the pair `(E'', H'')`. + (I' : ModelWithCorners 𝕜 E' H') {M' : Type*} [TopologicalSpace M'] + -- declare the prerequisites for a charted space `M''` over the pair `(E'', H'')`. {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] - {I'' : ModelWithCorners 𝕜 E'' H''} {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] + {I'' : ModelWithCorners 𝕜 E'' H''} {M'' : Type*} [TopologicalSpace M''] + +section ChartedSpace +variable [ChartedSpace H M] [ChartedSpace H' M'] [ChartedSpace H'' M''] -- declare functions, sets, points and smoothness indices {e : PartialHomeomorph M H} {e' : PartialHomeomorph M' H'} {f f₁ : M → M'} {s s₁ t : Set M} {x : M} {m n : ℕ∞} @@ -349,18 +350,17 @@ theorem smooth_inclusion {U V : Opens M} (h : U ≤ V) : Smooth I I (Set.inclusi end Inclusion +end ChartedSpace + /-! ### Open embeddings and their inverses are smooth -/ section -variable (I) -variable [Nonempty M] {e : M → H} (h : OpenEmbedding e) - [Nonempty M'] {e' : M' → H'} (h' : OpenEmbedding e') - {n : WithTop ℕ} +variable {e : M → H} (h : OpenEmbedding e) {n : WithTop ℕ} /-- If the `ChartedSpace` structure on a manifold `M` is given by an open embedding `e : M → H`, then `e` is smooth. -/ -lemma contMDiff_openEmbedding : +lemma contMDiff_openEmbedding [Nonempty M] : haveI := h.singletonChartedSpace; ContMDiff I I n e := by haveI := h.singleton_smoothManifoldWithCorners I rw [@contMDiff_iff _ _ _ _ _ _ _ _ _ _ h.singletonChartedSpace] @@ -383,14 +383,14 @@ lemma contMDiff_openEmbedding : h.toPartialHomeomorph_target] at this exact this -variable {I} +variable {I I'} /-- If the `ChartedSpace` structure on a manifold `M` is given by an open embedding `e : M → H`, then the inverse of `e` is smooth. -/ -lemma contMDiffOn_openEmbedding_symm : +lemma contMDiffOn_openEmbedding_symm [Nonempty M] : haveI := h.singletonChartedSpace; ContMDiffOn I I n (OpenEmbedding.toPartialHomeomorph e h).symm (range e) := by haveI := h.singleton_smoothManifoldWithCorners I - rw [@contMDiffOn_iff _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ h.singletonChartedSpace] + rw [@contMDiffOn_iff] constructor · rw [← h.toPartialHomeomorph_target] exact (h.toPartialHomeomorph e).continuousOn_symm @@ -407,6 +407,9 @@ lemma contMDiffOn_openEmbedding_symm : apply I.right_inv exact mem_of_subset_of_mem (extChartAt_target_subset_range _ _) hz.1 +variable [ChartedSpace H M] +variable [Nonempty M'] {e' : M' → H'} (h' : OpenEmbedding e') + /-- Let `M'` be a manifold whose chart structure is given by an open embedding `e'` into its model space `H'`. Then the smoothness of `e' ∘ f : M → H'` implies the smoothness of `f`. diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index e1edb1734ec41..e7031b91989cc 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -48,15 +48,15 @@ theorem Int.not_isField : ¬IsField ℤ := fun h => namespace NumberField -variable (K L : Type*) [Field K] [Field L] [nf : NumberField K] +variable (K L : Type*) [Field K] [Field L] -- See note [lower instance priority] attribute [instance] NumberField.to_charZero NumberField.to_finiteDimensional -protected theorem isAlgebraic : Algebra.IsAlgebraic ℚ K := +protected theorem isAlgebraic [NumberField K] : Algebra.IsAlgebraic ℚ K := Algebra.IsAlgebraic.of_finite _ _ -instance [NumberField L] [Algebra K L] : FiniteDimensional K L := +instance [NumberField K] [NumberField L] [Algebra K L] : FiniteDimensional K L := Module.Finite.of_restrictScalars_finite ℚ K L /-- The ring of integers (or number ring) corresponding to a number field @@ -79,7 +79,7 @@ instance : CommRing (𝓞 K) := inferInstanceAs (CommRing (integralClosure _ _)) instance : IsDomain (𝓞 K) := inferInstanceAs (IsDomain (integralClosure _ _)) -instance : CharZero (𝓞 K) := +instance [NumberField K] : CharZero (𝓞 K) := inferInstanceAs (CharZero (integralClosure _ _)) instance : Algebra (𝓞 K) K := inferInstanceAs (Algebra (integralClosure _ _) _) @@ -202,6 +202,8 @@ variable (K) instance [CharZero K] : CharZero (𝓞 K) := CharZero.of_module _ K +variable [NumberField K] + instance : IsNoetherian ℤ (𝓞 K) := IsIntegralClosure.isNoetherian _ ℚ K _ @@ -250,6 +252,8 @@ def restrict_monoidHom [MulOneClass M] (f : M →* K) (h : ∀ x, IsIntegral ℤ end RingOfIntegers +variable [NumberField K] + /-- A basis of `K` over `ℚ` that is also a basis of `𝓞 K` over `ℤ`. -/ noncomputable def integralBasis : Basis (Free.ChooseBasisIndex ℤ (𝓞 K)) ℚ K := Basis.localizationLocalization ℚ (nonZeroDivisors ℤ) K (RingOfIntegers.basis K) diff --git a/Mathlib/RingTheory/DedekindDomain/Different.lean b/Mathlib/RingTheory/DedekindDomain/Different.lean index f77a9914d0a63..e4d4261b6eac5 100644 --- a/Mathlib/RingTheory/DedekindDomain/Different.lean +++ b/Mathlib/RingTheory/DedekindDomain/Different.lean @@ -35,16 +35,11 @@ attribute [local instance] FractionRing.liftAlgebra FractionRing.isScalarTower_l variable (A K : Type*) {L : Type u} {B} [CommRing A] [Field K] [CommRing B] [Field L] variable [Algebra A K] [Algebra B L] [Algebra A B] [Algebra K L] [Algebra A L] variable [IsScalarTower A K L] [IsScalarTower A B L] -variable [IsDomain A] -variable [IsFractionRing A K] [IsIntegralClosure B A L] [IsFractionRing B L] -variable [FiniteDimensional K L] [Algebra.IsSeparable K L] open nonZeroDivisors IsLocalization Matrix Algebra section BIsDomain -variable [IsDomain B] - /-- Under the AKLB setting, `Iᵛ := traceDual A K (I : Submodule B L)` is the `Submodule B L` such that `x ∈ Iᵛ ↔ ∀ y ∈ I, Tr(x, y) ∈ A` -/ noncomputable @@ -105,6 +100,8 @@ lemma traceDual_top' : obtain ⟨c, hc, hc0⟩ := hx' simpa [hc0] using hc (c⁻¹ * b) +variable [IsDomain A] [IsFractionRing A K] [FiniteDimensional K L] [Algebra.IsSeparable K L] + lemma traceDual_top [Decidable (IsField A)] : (⊤ : Submodule B L)ᵛ = if IsField A then ⊤ else ⊥ := by convert traceDual_top' @@ -117,12 +114,43 @@ end Submodule open Submodule +variable [IsFractionRing A K] + +variable (A K) in +lemma map_equiv_traceDual [IsDomain A] [IsFractionRing B L] [IsDomain B] + [NoZeroSMulDivisors A B] (I : Submodule B (FractionRing B)) : + (traceDual A (FractionRing A) I).map (FractionRing.algEquiv B L) = + traceDual A K (I.map (FractionRing.algEquiv B L)) := by + show Submodule.map (FractionRing.algEquiv B L).toLinearEquiv.toLinearMap _ = + traceDual A K (I.map (FractionRing.algEquiv B L).toLinearEquiv.toLinearMap) + rw [Submodule.map_equiv_eq_comap_symm, Submodule.map_equiv_eq_comap_symm] + ext x + simp only [AlgEquiv.toLinearEquiv_symm, AlgEquiv.toLinearEquiv_toLinearMap, + traceDual, traceForm_apply, Submodule.mem_comap, AlgEquiv.toLinearMap_apply, + Submodule.mem_mk, AddSubmonoid.mem_mk, AddSubsemigroup.mem_mk, Set.mem_setOf_eq] + apply (FractionRing.algEquiv B L).forall_congr + simp only [restrictScalars_mem, traceForm_apply, AlgEquiv.toEquiv_eq_coe, + EquivLike.coe_coe, mem_comap, AlgEquiv.toLinearMap_apply, AlgEquiv.symm_apply_apply] + refine fun {y} ↦ (forall_congr' fun hy ↦ ?_) + rw [Algebra.trace_eq_of_equiv_equiv (FractionRing.algEquiv A K).toRingEquiv + (FractionRing.algEquiv B L).toRingEquiv] + swap + · apply IsLocalization.ringHom_ext (M := A⁰); ext + simp only [AlgEquiv.toRingEquiv_eq_coe, AlgEquiv.toRingEquiv_toRingHom, RingHom.coe_comp, + RingHom.coe_coe, Function.comp_apply, AlgEquiv.commutes, ← IsScalarTower.algebraMap_apply] + rw [IsScalarTower.algebraMap_apply A B (FractionRing B), AlgEquiv.commutes, + ← IsScalarTower.algebraMap_apply] + simp only [AlgEquiv.toRingEquiv_eq_coe, _root_.map_mul, AlgEquiv.coe_ringEquiv, + AlgEquiv.apply_symm_apply, ← AlgEquiv.symm_toRingEquiv, mem_one, AlgEquiv.algebraMap_eq_apply] + variable [IsIntegrallyClosed A] lemma Submodule.mem_traceDual_iff_isIntegral {I : Submodule B L} {x} : x ∈ Iᵛ ↔ ∀ a ∈ I, IsIntegral A (traceForm K L x a) := forall₂_congr (fun _ _ ↦ IsIntegrallyClosed.isIntegral_iff.symm) +variable [FiniteDimensional K L] [IsIntegralClosure B A L] + lemma Submodule.one_le_traceDual_one : (1 : Submodule B L) ≤ 1ᵛ := by rw [le_traceDual_iff_map_le_one, mul_one] @@ -132,6 +160,8 @@ lemma Submodule.one_le_traceDual_one : rw [IsIntegralClosure.isIntegral_iff (A := B)] exact ⟨_, rfl⟩ +variable [Algebra.IsSeparable K L] + /-- If `b` is an `A`-integral basis of `L` with discriminant `b`, then `d • a * x` is integral over `A` for all `a ∈ I` and `x ∈ Iᵛ`. -/ lemma isIntegral_discr_mul_of_mem_traceDual @@ -167,33 +197,10 @@ lemma isIntegral_discr_mul_of_mem_traceDual variable (A K) -lemma map_equiv_traceDual [NoZeroSMulDivisors A B] (I : Submodule B (FractionRing B)) : - (traceDual A (FractionRing A) I).map (FractionRing.algEquiv B L) = - traceDual A K (I.map (FractionRing.algEquiv B L)) := by - show Submodule.map (FractionRing.algEquiv B L).toLinearEquiv.toLinearMap _ = - traceDual A K (I.map (FractionRing.algEquiv B L).toLinearEquiv.toLinearMap) - rw [Submodule.map_equiv_eq_comap_symm, Submodule.map_equiv_eq_comap_symm] - ext x - simp only [AlgEquiv.toLinearEquiv_symm, AlgEquiv.toLinearEquiv_toLinearMap, - traceDual, traceForm_apply, Submodule.mem_comap, AlgEquiv.toLinearMap_apply, - Submodule.mem_mk, AddSubmonoid.mem_mk, AddSubsemigroup.mem_mk, Set.mem_setOf_eq] - apply (FractionRing.algEquiv B L).forall_congr - simp only [restrictScalars_mem, traceForm_apply, AlgEquiv.toEquiv_eq_coe, - EquivLike.coe_coe, mem_comap, AlgEquiv.toLinearMap_apply, AlgEquiv.symm_apply_apply] - refine fun {y} ↦ (forall_congr' fun hy ↦ ?_) - rw [Algebra.trace_eq_of_equiv_equiv (FractionRing.algEquiv A K).toRingEquiv - (FractionRing.algEquiv B L).toRingEquiv] - swap - · apply IsLocalization.ringHom_ext (M := A⁰); ext - simp only [AlgEquiv.toRingEquiv_eq_coe, AlgEquiv.toRingEquiv_toRingHom, RingHom.coe_comp, - RingHom.coe_coe, Function.comp_apply, AlgEquiv.commutes, ← IsScalarTower.algebraMap_apply] - rw [IsScalarTower.algebraMap_apply A B (FractionRing B), AlgEquiv.commutes, - ← IsScalarTower.algebraMap_apply] - simp only [AlgEquiv.toRingEquiv_eq_coe, _root_.map_mul, AlgEquiv.coe_ringEquiv, - AlgEquiv.apply_symm_apply, ← AlgEquiv.symm_toRingEquiv, mem_one, AlgEquiv.algebraMap_eq_apply] - open scoped Classical +variable [IsDomain A] [IsFractionRing B L] [Nontrivial B] [NoZeroDivisors B] + namespace FractionalIdeal /-- The dual of a non-zero fractional ideal is the dual of the submodule under the traceform. -/ @@ -219,9 +226,12 @@ end FractionalIdeal end BIsDomain +variable [IsDomain A] [IsFractionRing A K] + [FiniteDimensional K L] [Algebra.IsSeparable K L] [IsIntegralClosure B A L] + namespace FractionalIdeal -variable [IsIntegrallyClosed A] +variable [IsFractionRing B L] [IsIntegrallyClosed A] open Submodule @@ -229,9 +239,9 @@ open scoped Classical local notation:max I:max "ᵛ" => Submodule.traceDual A K I -variable [IsDedekindDomain B] {I J : FractionalIdeal B⁰ L} (hI : I ≠ 0) (hJ : J ≠ 0) +variable [IsDedekindDomain B] {I J : FractionalIdeal B⁰ L} -lemma coe_dual : +lemma coe_dual (hI : I ≠ 0) : (dual A K I : Submodule B L) = Iᵛ := by rw [dual, dif_neg hI, coe_mk] variable (B L) @@ -248,13 +258,13 @@ lemma dual_zero : variable {A K L B} -lemma mem_dual {x} : +lemma mem_dual (hI : I ≠ 0) {x} : x ∈ dual A K I ↔ ∀ a ∈ I, traceForm K L x a ∈ (algebraMap A K).range := by rw [dual, dif_neg hI]; rfl variable (A K) -lemma dual_ne_zero : +lemma dual_ne_zero (hI : I ≠ 0) : dual A K I ≠ 0 := by obtain ⟨b, hb, hb'⟩ := I.prop suffices algebraMap B L b ∈ dual A K I by @@ -284,7 +294,7 @@ lemma dual_ne_zero_iff : variable (A K) -lemma le_dual_inv_aux (hIJ : I * J ≤ 1) : +lemma le_dual_inv_aux (hI : I ≠ 0) (hIJ : I * J ≤ 1) : J ≤ dual A K I := by rw [dual, dif_neg hI] intro x hx y hy @@ -299,7 +309,7 @@ lemma one_le_dual_one : 1 ≤ dual A K (1 : FractionalIdeal B⁰ L) := le_dual_inv_aux A K one_ne_zero (by rw [one_mul]) -lemma le_dual_iff : +lemma le_dual_iff (hJ : J ≠ 0) : I ≤ dual A K J ↔ I * J ≤ dual A K 1 := by by_cases hI : I = 0 · simp [hI, zero_le] @@ -336,11 +346,11 @@ lemma dual_div_dual : · exact inv_div_inv J I · simp only [ne_eq, dual_eq_zero_iff, one_ne_zero, not_false_eq_true] -lemma dual_mul_self : +lemma dual_mul_self (hI : I ≠ 0) : dual A K I * I = dual A K 1 := by rw [dual_eq_mul_inv, mul_assoc, inv_mul_cancel hI, mul_one] -lemma self_mul_dual : +lemma self_mul_dual (hI : I ≠ 0) : I * dual A K I = dual A K 1 := by rw [mul_comm, dual_mul_self A K hI] @@ -360,7 +370,7 @@ lemma dual_dual : variable {I} @[simp] -lemma dual_le_dual : +lemma dual_le_dual (hI : I ≠ 0) (hJ : J ≠ 0) : dual A K I ≤ dual A K J ↔ J ≤ I := by nth_rewrite 2 [← dual_dual A K I] rw [le_dual_iff A K hJ, le_dual_iff A K (I := J) (by rwa [dual_ne_zero_iff]), mul_comm] @@ -404,6 +414,9 @@ lemma coeSubmodule_differentIdeal_fractionRing · exact one_ne_zero · exact one_ne_zero +section +variable [IsFractionRing B L] + lemma coeSubmodule_differentIdeal [NoZeroSMulDivisors A B] : coeSubmodule L (differentIdeal A B) = 1 / Submodule.traceDual A K 1 := by have : (FractionRing.algEquiv B L).toLinearEquiv.comp (Algebra.linearMap B (FractionRing B)) = @@ -503,7 +516,9 @@ lemma traceForm_dualSubmodule_adjoin ← Nat.succ_eq_add_one] at hi exact le_of_not_lt hi -variable (L) +end + +variable (L) {B} open Polynomial Pointwise in lemma conductor_mul_differentIdeal [NoZeroSMulDivisors A B] diff --git a/Mathlib/RingTheory/Ideal/Norm.lean b/Mathlib/RingTheory/Ideal/Norm.lean index fdf93415f5fc1..68a098c40b8ed 100644 --- a/Mathlib/RingTheory/Ideal/Norm.lean +++ b/Mathlib/RingTheory/Ideal/Norm.lean @@ -88,7 +88,7 @@ end Submodule section RingOfIntegers -variable {S : Type*} [CommRing S] [IsDomain S] +variable {S : Type*} [CommRing S] open Submodule @@ -115,12 +115,13 @@ theorem Ideal.mul_add_mem_pow_succ_inj (P : Ideal S) {i : ℕ} (a d d' e e' : S) section PPrime -variable {P : Ideal S} [P_prime : P.IsPrime] (hP : P ≠ ⊥) +variable {P : Ideal S} [P_prime : P.IsPrime] /-- If `a ∈ P^i \ P^(i+1)` and `c ∈ P^i`, then `a * d + e = c` for `e ∈ P^(i+1)`. `Ideal.mul_add_mem_pow_succ_unique` shows the choice of `d` is unique, up to `P`. Inspired by [Neukirch], proposition 6.1 -/ -theorem Ideal.exists_mul_add_mem_pow_succ [IsDedekindDomain S] {i : ℕ} (a c : S) (a_mem : a ∈ P ^ i) +theorem Ideal.exists_mul_add_mem_pow_succ [IsDedekindDomain S] (hP : P ≠ ⊥) + {i : ℕ} (a c : S) (a_mem : a ∈ P ^ i) (a_not_mem : a ∉ P ^ (i + 1)) (c_mem : c ∈ P ^ i) : ∃ d : S, ∃ e ∈ P ^ (i + 1), a * d + e = c := by suffices eq_b : P ^ i = Ideal.span {a} ⊔ P ^ (i + 1) by @@ -143,7 +144,8 @@ theorem Ideal.mem_prime_of_mul_mem_pow [IsDedekindDomain S] {P : Ideal S} [P_pri /-- The choice of `d` in `Ideal.exists_mul_add_mem_pow_succ` is unique, up to `P`. Inspired by [Neukirch], proposition 6.1 -/ -theorem Ideal.mul_add_mem_pow_succ_unique [IsDedekindDomain S] {i : ℕ} (a d d' e e' : S) +theorem Ideal.mul_add_mem_pow_succ_unique [IsDedekindDomain S] (hP : P ≠ ⊥) + {i : ℕ} (a d d' e e' : S) (a_not_mem : a ∉ P ^ (i + 1)) (e_mem : e ∈ P ^ (i + 1)) (e'_mem : e' ∈ P ^ (i + 1)) (h : a * d + e - (a * d' + e') ∈ P ^ (i + 1)) : d - d' ∈ P := by have h' : a * (d - d') ∈ P ^ (i + 1) := by @@ -152,7 +154,7 @@ theorem Ideal.mul_add_mem_pow_succ_unique [IsDedekindDomain S] {i : ℕ} (a d d' exact Ideal.mem_prime_of_mul_mem_pow hP a_not_mem h' /-- Multiplicity of the ideal norm, for powers of prime ideals. -/ -theorem cardQuot_pow_of_prime [IsDedekindDomain S] {i : ℕ} : +theorem cardQuot_pow_of_prime [IsDedekindDomain S] (hP : P ≠ ⊥) {i : ℕ} : cardQuot (P ^ i) = cardQuot P ^ i := by induction' i with i ih · simp @@ -212,7 +214,7 @@ noncomputable def Ideal.absNorm [Nontrivial S] [IsDedekindDomain S] [Module.Free namespace Ideal -variable [Nontrivial S] [IsDedekindDomain S] [Module.Free ℤ S] [Module.Finite ℤ S] +variable [Nontrivial S] [IsDedekindDomain S] [Module.Free ℤ S] theorem absNorm_apply (I : Ideal S) : absNorm I = cardQuot I := rfl @@ -230,6 +232,49 @@ theorem absNorm_ne_zero_iff (I : Ideal S) : Ideal.absNorm I ≠ 0 ↔ Finite (S ⟨fun h => Nat.finite_of_card_ne_zero h, fun h => (@AddSubgroup.finiteIndex_of_finite_quotient _ _ _ h).finiteIndex⟩ +theorem absNorm_dvd_absNorm_of_le {I J : Ideal S} (h : J ≤ I) : Ideal.absNorm I ∣ Ideal.absNorm J := + map_dvd absNorm (dvd_iff_le.mpr h) + +theorem irreducible_of_irreducible_absNorm {I : Ideal S} (hI : Irreducible (Ideal.absNorm I)) : + Irreducible I := + irreducible_iff.mpr + ⟨fun h => + hI.not_unit (by simpa only [Ideal.isUnit_iff, Nat.isUnit_iff, absNorm_eq_one_iff] using h), + by + rintro a b rfl + simpa only [Ideal.isUnit_iff, Nat.isUnit_iff, absNorm_eq_one_iff] using + hI.isUnit_or_isUnit (_root_.map_mul absNorm a b)⟩ + +theorem isPrime_of_irreducible_absNorm {I : Ideal S} (hI : Irreducible (Ideal.absNorm I)) : + I.IsPrime := + isPrime_of_prime + (UniqueFactorizationMonoid.irreducible_iff_prime.mp (irreducible_of_irreducible_absNorm hI)) + +theorem prime_of_irreducible_absNorm_span {a : S} (ha : a ≠ 0) + (hI : Irreducible (Ideal.absNorm (Ideal.span ({a} : Set S)))) : Prime a := + (Ideal.span_singleton_prime ha).mp (isPrime_of_irreducible_absNorm hI) + +theorem absNorm_mem (I : Ideal S) : ↑(Ideal.absNorm I) ∈ I := by + rw [absNorm_apply, cardQuot, ← Ideal.Quotient.eq_zero_iff_mem, map_natCast, + Quotient.index_eq_zero] + +theorem span_singleton_absNorm_le (I : Ideal S) : Ideal.span {(Ideal.absNorm I : S)} ≤ I := by + simp only [Ideal.span_le, Set.singleton_subset_iff, SetLike.mem_coe, Ideal.absNorm_mem I] + +theorem span_singleton_absNorm {I : Ideal S} (hI : (Ideal.absNorm I).Prime) : + Ideal.span (singleton (Ideal.absNorm I : ℤ)) = I.comap (algebraMap ℤ S) := by + have : Ideal.IsPrime (Ideal.span (singleton (Ideal.absNorm I : ℤ))) := by + rwa [Ideal.span_singleton_prime (Int.ofNat_ne_zero.mpr hI.ne_zero), ← Nat.prime_iff_prime_int] + apply (this.isMaximal _).eq_of_le + · exact ((isPrime_of_irreducible_absNorm + ((Nat.irreducible_iff_nat_prime _).mpr hI)).comap (algebraMap ℤ S)).ne_top + · rw [span_singleton_le_iff_mem, mem_comap, algebraMap_int_eq, map_natCast] + exact absNorm_mem I + · rw [Ne, span_singleton_eq_bot] + exact Int.ofNat_ne_zero.mpr hI.ne_zero + +variable [Module.Finite ℤ S] + /-- Let `e : S ≃ I` be an additive isomorphism (therefore a `ℤ`-linear equiv). Then an alternative way to compute the norm of `I` is given by taking the determinant of `e`. See `natAbs_det_basis_change` for a more familiar formulation of this result. -/ @@ -320,9 +365,6 @@ theorem absNorm_span_singleton (r : S) : refine b.ext fun i => ?_ simp -theorem absNorm_dvd_absNorm_of_le {I J : Ideal S} (h : J ≤ I) : Ideal.absNorm I ∣ Ideal.absNorm J := - map_dvd absNorm (dvd_iff_le.mpr h) - theorem absNorm_dvd_norm_of_mem {I : Ideal S} {x : S} (h : x ∈ I) : ↑(Ideal.absNorm I) ∣ Algebra.norm ℤ x := by rw [← Int.dvd_natAbs, ← absNorm_span_singleton x, Int.natCast_dvd_natCast] @@ -352,44 +394,6 @@ theorem absNorm_eq_zero_iff {I : Ideal S} : Ideal.absNorm I = 0 ↔ I = ⊥ := b theorem absNorm_ne_zero_of_nonZeroDivisors (I : (Ideal S)⁰) : Ideal.absNorm (I : Ideal S) ≠ 0 := Ideal.absNorm_eq_zero_iff.not.mpr <| nonZeroDivisors.coe_ne_zero _ -theorem irreducible_of_irreducible_absNorm {I : Ideal S} (hI : Irreducible (Ideal.absNorm I)) : - Irreducible I := - irreducible_iff.mpr - ⟨fun h => - hI.not_unit (by simpa only [Ideal.isUnit_iff, Nat.isUnit_iff, absNorm_eq_one_iff] using h), - by - rintro a b rfl - simpa only [Ideal.isUnit_iff, Nat.isUnit_iff, absNorm_eq_one_iff] using - hI.isUnit_or_isUnit (_root_.map_mul absNorm a b)⟩ - -theorem isPrime_of_irreducible_absNorm {I : Ideal S} (hI : Irreducible (Ideal.absNorm I)) : - I.IsPrime := - isPrime_of_prime - (UniqueFactorizationMonoid.irreducible_iff_prime.mp (irreducible_of_irreducible_absNorm hI)) - -theorem prime_of_irreducible_absNorm_span {a : S} (ha : a ≠ 0) - (hI : Irreducible (Ideal.absNorm (Ideal.span ({a} : Set S)))) : Prime a := - (Ideal.span_singleton_prime ha).mp (isPrime_of_irreducible_absNorm hI) - -theorem absNorm_mem (I : Ideal S) : ↑(Ideal.absNorm I) ∈ I := by - rw [absNorm_apply, cardQuot, ← Ideal.Quotient.eq_zero_iff_mem, map_natCast, - Quotient.index_eq_zero] - -theorem span_singleton_absNorm_le (I : Ideal S) : Ideal.span {(Ideal.absNorm I : S)} ≤ I := by - simp only [Ideal.span_le, Set.singleton_subset_iff, SetLike.mem_coe, Ideal.absNorm_mem I] - -theorem span_singleton_absNorm {I : Ideal S} (hI : (Ideal.absNorm I).Prime) : - Ideal.span (singleton (Ideal.absNorm I : ℤ)) = I.comap (algebraMap ℤ S) := by - have : Ideal.IsPrime (Ideal.span (singleton (Ideal.absNorm I : ℤ))) := by - rwa [Ideal.span_singleton_prime (Int.ofNat_ne_zero.mpr hI.ne_zero), ← Nat.prime_iff_prime_int] - apply (this.isMaximal _).eq_of_le - · exact ((isPrime_of_irreducible_absNorm - ((Nat.irreducible_iff_nat_prime _).mpr hI)).comap (algebraMap ℤ S)).ne_top - · rw [span_singleton_le_iff_mem, mem_comap, algebraMap_int_eq, map_natCast] - exact absNorm_mem I - · rw [Ne, span_singleton_eq_bot] - exact Int.ofNat_ne_zero.mpr hI.ne_zero - theorem finite_setOf_absNorm_eq [CharZero S] {n : ℕ} (hn : 0 < n) : {I : Ideal S | Ideal.absNorm I = n}.Finite := by let f := fun I : Ideal S => Ideal.map (Ideal.Quotient.mk (@Ideal.span S _ {↑n})) I diff --git a/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean b/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean index 74e140516b867..e65a17d8ac054 100644 --- a/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean +++ b/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean @@ -287,13 +287,13 @@ def Algebra.intNormAux [Algebra.IsSeparable K L] : variable {A K L B} -lemma Algebra.map_intNormAux [IsIntegrallyClosed A] [Algebra.IsSeparable K L] (x : B) : +lemma Algebra.map_intNormAux [Algebra.IsSeparable K L] (x : B) : algebraMap A K (Algebra.intNormAux A K L B x) = Algebra.norm K (algebraMap B L x) := by dsimp [Algebra.intNormAux] exact IsIntegralClosure.algebraMap_mk' _ _ _ variable (A B) -variable [IsDomain A] [IsIntegrallyClosed A] [IsDomain B] [IsIntegrallyClosed B] +variable [IsDomain A] [IsDomain B] [IsIntegrallyClosed B] variable [Module.Finite A B] [NoZeroSMulDivisors A B] variable [Algebra.IsSeparable (FractionRing A) (FractionRing B)] -- TODO: remove this From 6c05c5dae0b72501dcea8932ceaafe682b767ee5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Dupuis?= Date: Wed, 7 Aug 2024 14:53:42 +0000 Subject: [PATCH 096/359] feat(Calculus/MeanValue): A function is not differentiable if its derivative tends to infinity (#15200) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR shows that a function `f : ℝ → ℝ` is not differentiable at a point if its derivative tends to infinity at that point. --- Mathlib/Analysis/Calculus/MeanValue.lean | 121 ++++++++++++++++++++++ Mathlib/Topology/Order/LeftRightNhds.lean | 3 + 2 files changed, 124 insertions(+) diff --git a/Mathlib/Analysis/Calculus/MeanValue.lean b/Mathlib/Analysis/Calculus/MeanValue.lean index a52881c81d539..0064eb24ae840 100644 --- a/Mathlib/Analysis/Calculus/MeanValue.lean +++ b/Mathlib/Analysis/Calculus/MeanValue.lean @@ -746,6 +746,127 @@ theorem exists_deriv_eq_slope : ∃ c ∈ Ioo a b, deriv f c = (f b - f a) / (b exists_hasDerivAt_eq_slope f (deriv f) hab hfc fun x hx => ((hfd x hx).differentiableAt <| IsOpen.mem_nhds isOpen_Ioo hx).hasDerivAt +/-- Lagrange's **Mean Value Theorem**, `deriv` version. -/ +theorem exists_deriv_eq_slope' : ∃ c ∈ Ioo a b, deriv f c = slope f a b := by + rw [slope_def_field] + exact exists_deriv_eq_slope f hab hfc hfd + +/-- A real function whose derivative tends to infinity from the right at a point is not +differentiable on the right at that point -/ +theorem not_differentiableWithinAt_of_deriv_tendsto_atTop_Ioi (f : ℝ → ℝ) {a : ℝ} + (hf : Tendsto (deriv f) (𝓝[>] a) atTop) : ¬ DifferentiableWithinAt ℝ f (Ioi a) a := by + replace hf : Tendsto (derivWithin f (Ioi a)) (𝓝[>] a) atTop := by + refine hf.congr' ?_ + filter_upwards [eventually_mem_nhdsWithin] with x hx + have : Ioi a ∈ 𝓝 x := by simp [← mem_interior_iff_mem_nhds, hx] + exact (derivWithin_of_mem_nhds this).symm + by_cases hcont_at_a : ContinuousWithinAt f (Ici a) a + case neg => + intro hcontra + have := hcontra.continuousWithinAt + rw [← ContinuousWithinAt.diff_iff this] at hcont_at_a + simp at hcont_at_a + case pos => + intro hdiff + replace hdiff := hdiff.hasDerivWithinAt + rw [hasDerivWithinAt_iff_tendsto_slope, Set.diff_singleton_eq_self not_mem_Ioi_self] at hdiff + have h₀ : ∀ᶠ b in 𝓝[>] a, + ∀ x ∈ Ioc a b, max (derivWithin f (Ioi a) a + 1) 0 < derivWithin f (Ioi a) x := by + rw [(nhdsWithin_Ioi_basis a).eventually_iff] + rw [(nhdsWithin_Ioi_basis a).tendsto_left_iff] at hf + obtain ⟨b, hab, hb⟩ := hf (Ioi (max (derivWithin f (Ioi a) a + 1) 0)) (Ioi_mem_atTop _) + refine ⟨b, hab, fun x hx z hz => ?_⟩ + simp only [MapsTo, mem_Ioo, mem_Ioi, and_imp] at hb + exact hb hz.1 <| hz.2.trans_lt hx.2 + have h₁ : ∀ᶠ b in 𝓝[>] a, slope f a b < derivWithin f (Ioi a) a + 1 := by + rw [(nhds_basis_Ioo _).tendsto_right_iff] at hdiff + specialize hdiff ⟨derivWithin f (Ioi a) a - 1, derivWithin f (Ioi a) a + 1⟩ <| by simp + filter_upwards [hdiff] with z hz using hz.2 + have hcontra : ∀ᶠ _ in 𝓝[>] a, False := by + filter_upwards [h₀, h₁, eventually_mem_nhdsWithin] with b hb hslope (hab : a < b) + have hdiff' : DifferentiableOn ℝ f (Ioc a b) := fun z hz => by + refine DifferentiableWithinAt.mono (t := Ioi a) ?_ Ioc_subset_Ioi_self + have : derivWithin f (Ioi a) z ≠ 0 := ne_of_gt <| by + simp_all only [mem_Ioo, and_imp, mem_Ioc, max_lt_iff] + exact differentiableWithinAt_of_derivWithin_ne_zero this + have hcont_Ioc : ∀ z ∈ Ioc a b, ContinuousWithinAt f (Icc a b) z := by + intro z hz'' + refine (hdiff'.continuousOn z hz'').mono_of_mem ?_ + have hfinal : 𝓝[Ioc a b] z = 𝓝[Icc a b] z := by + refine nhdsWithin_eq_nhdsWithin' (s := Ioi a) (Ioi_mem_nhds hz''.1) ?_ + simp only [Ioc_inter_Ioi, le_refl, sup_of_le_left] + ext y + exact ⟨fun h => ⟨mem_Icc_of_Ioc h, mem_of_mem_inter_left h⟩, fun ⟨H1, H2⟩ => ⟨H2, H1.2⟩⟩ + rw [← hfinal] + exact self_mem_nhdsWithin + have hcont : ContinuousOn f (Icc a b) := by + intro z hz + by_cases hz' : z = a + · rw [hz'] + exact hcont_at_a.mono Icc_subset_Ici_self + · exact hcont_Ioc z ⟨lt_of_le_of_ne hz.1 (Ne.symm hz'), hz.2⟩ + obtain ⟨x, hx₁, hx₂⟩ := + exists_deriv_eq_slope' f hab hcont (hdiff'.mono (Ioo_subset_Ioc_self)) + specialize hb x ⟨hx₁.1, le_of_lt hx₁.2⟩ + replace hx₂ : derivWithin f (Ioi a) x = slope f a b := by + have : Ioi a ∈ 𝓝 x := by simp [← mem_interior_iff_mem_nhds, hx₁.1] + rwa [derivWithin_of_mem_nhds this] + rw [hx₂, max_lt_iff] at hb + linarith + simp [Filter.eventually_false_iff_eq_bot, ← not_mem_closure_iff_nhdsWithin_eq_bot] at hcontra + +/-- A real function whose derivative tends to minus infinity from the right at a point is not +differentiable on the right at that point -/ +theorem not_differentiableWithinAt_of_deriv_tendsto_atBot_Ioi (f : ℝ → ℝ) {a : ℝ} + (hf : Tendsto (deriv f) (𝓝[>] a) atBot) : ¬ DifferentiableWithinAt ℝ f (Ioi a) a := by + intro h + have hf' : Tendsto (deriv (-f)) (𝓝[>] a) atTop := by + rw [Pi.neg_def, deriv.neg'] + exact tendsto_neg_atBot_atTop.comp hf + exact not_differentiableWithinAt_of_deriv_tendsto_atTop_Ioi (-f) hf' h.neg + +/-- A real function whose derivative tends to minus infinity from the left at a point is not +differentiable on the left at that point -/ +theorem not_differentiableWithinAt_of_deriv_tendsto_atBot_Iio (f : ℝ → ℝ) {a : ℝ} + (hf : Tendsto (deriv f) (𝓝[<] a) atBot) : ¬ DifferentiableWithinAt ℝ f (Iio a) a := by + let f' := f ∘ Neg.neg + have hderiv : deriv f' =ᶠ[𝓝[>] (-a)] -(deriv f ∘ Neg.neg) := by + rw [atBot_basis.tendsto_right_iff] at hf + specialize hf (-1) trivial + rw [(nhdsWithin_Iio_basis a).eventually_iff] at hf + rw [EventuallyEq, (nhdsWithin_Ioi_basis (-a)).eventually_iff] + obtain ⟨b, hb₁, hb₂⟩ := hf + refine ⟨-b, by linarith, fun x hx => ?_⟩ + simp only [Pi.neg_apply, Function.comp_apply] + suffices deriv f' x = deriv f (-x) * deriv (Neg.neg : ℝ → ℝ) x by simpa using this + refine deriv.comp x (differentiableAt_of_deriv_ne_zero ?_) (by fun_prop) + rw [mem_Ioo] at hx + have h₁ : -x ∈ Ioo b a := ⟨by linarith, by linarith⟩ + have h₂ : deriv f (-x) ≤ -1 := hb₂ h₁ + exact ne_of_lt (by linarith) + have hmain : ¬ DifferentiableWithinAt ℝ f' (Ioi (-a)) (-a) := by + refine not_differentiableWithinAt_of_deriv_tendsto_atTop_Ioi f' <| Tendsto.congr' hderiv.symm ?_ + refine Tendsto.comp (g := -deriv f) ?_ tendsto_neg_nhdsWithin_Ioi_neg + exact Tendsto.comp (g := Neg.neg) tendsto_neg_atBot_atTop hf + intro h + have : DifferentiableWithinAt ℝ f' (Ioi (-a)) (-a) := by + refine DifferentiableWithinAt.comp (g := f) (f := Neg.neg) (t := Iio a) (-a) ?_ ?_ ?_ + · simp [h] + · fun_prop + · intro x + simp [neg_lt] + exact hmain this + +/-- A real function whose derivative tends to infinity from the left at a point is not +differentiable on the left at that point -/ +theorem not_differentiableWithinAt_of_deriv_tendsto_atTop_Iio (f : ℝ → ℝ) {a : ℝ} + (hf : Tendsto (deriv f) (𝓝[<] a) atTop) : ¬ DifferentiableWithinAt ℝ f (Iio a) a := by + intro h + have hf' : Tendsto (deriv (-f)) (𝓝[<] a) atBot := by + rw [Pi.neg_def, deriv.neg'] + exact tendsto_neg_atTop_atBot.comp hf + exact not_differentiableWithinAt_of_deriv_tendsto_atBot_Iio (-f) hf' h.neg + end Interval /-- Let `f` be a function continuous on a convex (or, equivalently, connected) subset `D` diff --git a/Mathlib/Topology/Order/LeftRightNhds.lean b/Mathlib/Topology/Order/LeftRightNhds.lean index 5625ca10bb003..626dbfd4cd1b9 100644 --- a/Mathlib/Topology/Order/LeftRightNhds.lean +++ b/Mathlib/Topology/Order/LeftRightNhds.lean @@ -160,6 +160,9 @@ theorem nhdsWithin_Iio_basis' {a : α} (h : ∃ b, b < a) : (𝓝[<] a).HasBasis let ⟨_, h⟩ := h ⟨fun _ => mem_nhdsWithin_Iio_iff_exists_Ioo_subset' h⟩ +theorem nhdsWithin_Iio_basis [NoMinOrder α] (a : α) : (𝓝[<] a).HasBasis (· < a) (Ioo · a) := + nhdsWithin_Iio_basis' <| exists_lt a + theorem nhdsWithin_Iio_eq_bot_iff {a : α} : 𝓝[<] a = ⊥ ↔ IsBot a ∨ ∃ b, b ⋖ a := by convert (config := {preTransparency := .default}) nhdsWithin_Ioi_eq_bot_iff (a := OrderDual.toDual a) using 4 From 5c2ed948ab2bcf566037e033ecd53e0d3b57c397 Mon Sep 17 00:00:00 2001 From: Robert Maxton Date: Wed, 7 Aug 2024 15:24:35 +0000 Subject: [PATCH 097/359] doc(LinearAlgebra/FreeProduct): remove untrue TODOs (#15587) Neither TODO turned out to be mathematically true: `FreeProduct` imposes more structure than `FreeAlgebra`, specifically combining adjacent elements from the same algebra, and my categorical argument for the isomorphism with `PiTensorProduct` in the commutative case was faulty. --- Mathlib/LinearAlgebra/FreeProduct/Basic.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/LinearAlgebra/FreeProduct/Basic.lean b/Mathlib/LinearAlgebra/FreeProduct/Basic.lean index dd2bd58eaa37d..c57193500b619 100644 --- a/Mathlib/LinearAlgebra/FreeProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeProduct/Basic.lean @@ -40,8 +40,6 @@ is just `PiTensorProduct`.) ## TODO - Induction principle for `FreeProduct` -- Equivalence of `FreeProduct` with `PiTensorProduct` in the commutative case and with - `FreeAlgebra R ⨁ (i : ι), A i` -/ universe u v w w' From f0edccd743c8550bc00a713769d249bf8ad0e0d5 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Wed, 7 Aug 2024 15:24:36 +0000 Subject: [PATCH 098/359] chore: remove some useless `simp`s (#15588) They are just instantiating some metavariables that are incidentally there. --- Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean | 1 - Mathlib/MeasureTheory/Order/UpperLower.lean | 1 - Mathlib/NumberTheory/WellApproximable.lean | 1 - 3 files changed, 3 deletions(-) diff --git a/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean b/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean index 434d9339af7b9..2ff8defbbf428 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean @@ -96,7 +96,6 @@ theorem σ_comp_P_eq_zero (X : SimplicialObject C) {n q : ℕ} (i : Fin (n + 1)) decomposition_Q n q, sum_comp, sum_comp, Finset.sum_eq_zero, add_zero, add_neg_eq_zero] intro j hj simp only [true_and_iff, Finset.mem_univ, Finset.mem_filter] at hj - simp only [Nat.succ_eq_add_one] at hi obtain ⟨k, hk⟩ := Nat.le.dest (Nat.lt_succ_iff.mp (Fin.is_lt j)) rw [add_comm] at hk have hi' : i = Fin.castSucc ⟨i, by omega⟩ := by diff --git a/Mathlib/MeasureTheory/Order/UpperLower.lean b/Mathlib/MeasureTheory/Order/UpperLower.lean index 4c8b3268546a2..65f4ad6fa888a 100644 --- a/Mathlib/MeasureTheory/Order/UpperLower.lean +++ b/Mathlib/MeasureTheory/Order/UpperLower.lean @@ -73,7 +73,6 @@ private lemma aux₀ = volume (closedBall (f (ε n) (hε' n)) (ε n / 4)) / volume (closedBall x (ε n)) := ?_ _ ≤ volume (closure s ∩ closedBall x (ε n)) / volume (closedBall x (ε n)) := by gcongr; exact subset_inter ((hf₁ _ $ hε' n).trans interior_subset_closure) $ hf₀ _ $ hε' n - dsimp have := hε' n rw [Real.volume_pi_closedBall, Real.volume_pi_closedBall, ← ENNReal.ofReal_div_of_pos, ← div_pow, mul_div_mul_left _ _ (two_ne_zero' ℝ), div_right_comm, div_self, one_div] diff --git a/Mathlib/NumberTheory/WellApproximable.lean b/Mathlib/NumberTheory/WellApproximable.lean index 0a7f08c466084..69b68a0ebfab8 100644 --- a/Mathlib/NumberTheory/WellApproximable.lean +++ b/Mathlib/NumberTheory/WellApproximable.lean @@ -274,7 +274,6 @@ theorem addWellApproximable_ae_empty_or_univ (δ : ℕ → ℝ) (hδ : Tendsto replace h_div : n / p * p = n := Nat.div_mul_cancel h_div have hf : f = (fun y => x + y) ∘ fun y => p • y := by ext; simp [add_comm x] - simp only at hf simp_rw [Function.comp_apply, le_eq_subset] rw [sSupHom.setImage_toFun, hf, image_comp] have := @monotone_image 𝕊 𝕊 fun y => x + y From 2da5149ffd28286374ac6bda68d10c0fc6db8484 Mon Sep 17 00:00:00 2001 From: Rida Hamadani Date: Wed, 7 Aug 2024 16:11:46 +0000 Subject: [PATCH 099/359] feat(Combinatorics/SimpleGraph): `dist_bot` and `dist_top` (#15465) --- Mathlib/Combinatorics/SimpleGraph/Metric.lean | 23 +++++++++++++++++++ Mathlib/Combinatorics/SimpleGraph/Path.lean | 5 ++++ 2 files changed, 28 insertions(+) diff --git a/Mathlib/Combinatorics/SimpleGraph/Metric.lean b/Mathlib/Combinatorics/SimpleGraph/Metric.lean index 9b21f7d272b05..0e09bff03a186 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Metric.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Metric.lean @@ -70,6 +70,7 @@ theorem edist_eq_zero_iff : G.edist u v = 0 ↔ u = v := by apply Iff.intro <;> simp [edist, ENat.iInf_eq_zero] +@[simp] theorem edist_self : edist G v v = 0 := edist_eq_zero_iff.mpr rfl @@ -128,6 +129,18 @@ theorem edist_eq_one_iff_adj : G.edist u v = 1 ↔ G.Adj u v := by exact w.adj_of_length_eq_one <| Nat.cast_eq_one.mp <| h ▸ hw · exact le_antisymm (edist_le h.toWalk) (ENat.one_le_iff_pos.mpr <| edist_pos_of_ne h.ne) +lemma edist_bot_of_ne (h : u ≠ v) : (⊥ : SimpleGraph V).edist u v = ⊤ := by + rwa [ne_eq, ← reachable_bot.not, ← edist_ne_top_iff_reachable.not, not_not] at h + +lemma edist_bot [DecidableEq V] : (⊥ : SimpleGraph V).edist u v = (if u = v then 0 else ⊤) := by + by_cases h : u = v <;> simp [h, edist_bot_of_ne] + +lemma edist_top_of_ne (h : u ≠ v) : (⊤ : SimpleGraph V).edist u v = 1 := by + simp [h] + +lemma edist_top [DecidableEq V] : (⊤ : SimpleGraph V).edist u v = (if u = v then 0 else 1) := by + by_cases h : u = v <;> simp [h] + /-- Supergraphs have smaller or equal extended distances to their subgraphs. -/ theorem edist_le_subgraph_edist {G' : SimpleGraph V} (h : G ≤ G') : G'.edist u v ≤ G.edist u v := by @@ -243,6 +256,16 @@ lemma Connected.exists_path_of_dist (hconn : G.Connected) (u v : V) : obtain ⟨p, h⟩ := hconn.exists_walk_length_eq_dist u v exact ⟨p, p.isPath_of_length_eq_dist h, h⟩ +@[simp] +lemma dist_bot : (⊥ : SimpleGraph V).dist u v = 0 := by + by_cases h : u = v <;> simp [h] + +lemma dist_top_of_ne (h : u ≠ v) : (⊤ : SimpleGraph V).dist u v = 1 := by + simp [h] + +lemma dist_top [DecidableEq V] : (⊤ : SimpleGraph V).dist u v = (if u = v then 0 else 1) := by + by_cases h : u = v <;> simp [h] + /-- Supergraphs have smaller or equal distances to their subgraphs. -/ theorem dist_le_subgraph_dist {G' : SimpleGraph V} (h : G ≤ G') (hr : G.Reachable u v) : G'.dist u v ≤ G.dist u v := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index 7dd0f77cbd996..0f5aaeef7169e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -700,6 +700,11 @@ variable (G) theorem reachable_is_equivalence : Equivalence G.Reachable := Equivalence.mk (@Reachable.refl _ G) (@Reachable.symm _ G) (@Reachable.trans _ G) +/-- Distinct vertices are not reachable in the empty graph. -/ +@[simp] +lemma reachable_bot {u v : V} : (⊥ : SimpleGraph V).Reachable u v ↔ u = v := + ⟨fun h ↦ h.elim fun p ↦ match p with | .nil => rfl, fun h ↦ h ▸ .rfl⟩ + /-- The equivalence relation on vertices given by `SimpleGraph.Reachable`. -/ def reachableSetoid : Setoid V := Setoid.mk _ G.reachable_is_equivalence From 8ad482ab86615753b9bab7c79e95efebbc47a6da Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Wed, 7 Aug 2024 16:22:37 +0000 Subject: [PATCH 100/359] chore: update Mathlib dependencies 2024-08-07 (#15590) 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 6fc3bcbf513a7..64809f172e29e 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -55,7 +55,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "68cd8ae0f5b996176d1243d94c56e17de570e3bf", + "rev": "57bd2065f1dbea5e9235646fb836c7cea9ab03b6", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", From a88cdb667fd69c027152ed86e7909e9c31437b5c Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 7 Aug 2024 16:33:42 +0000 Subject: [PATCH 101/359] feat(RelSeries): head_singleton, last_singleton (#15384) from the Carleson project. --- Mathlib/Order/RelSeries.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index 0c5da650cc937..c3676e955e9ca 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -219,6 +219,12 @@ lemma head_mem (x : RelSeries r) : x.head ∈ x := ⟨_, rfl⟩ lemma last_mem (x : RelSeries r) : x.last ∈ x := ⟨_, rfl⟩ +@[simp] +lemma head_singleton {r : Rel α α} (x : α) : (singleton r x).head = x := by simp [singleton, head] + +@[simp] +lemma last_singleton {r : Rel α α} (x : α) : (singleton r x).last = x := by simp [singleton, last] + end variable {r s} From 339689fb16756e4a0b52ddb36b83e0a81f327d0f Mon Sep 17 00:00:00 2001 From: grunweg Date: Wed, 7 Aug 2024 16:33:44 +0000 Subject: [PATCH 102/359] chore: follow-up to #15498 (#15499) Make the same replacements elsewhere: blocks_fun -> blocksFun, size_up_to -> sizeUpTo --- Mathlib/Analysis/Analytic/Inverse.lean | 4 ++-- Mathlib/Combinatorics/Enumerative/Composition.lean | 14 +++++++------- 2 files changed, 9 insertions(+), 9 deletions(-) diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index 76ec01632ad96..dc8f10b6918b0 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -399,8 +399,8 @@ theorem radius_right_inv_pos_of_radius_pos_aux1 (n : ℕ) (p : ℕ → ℝ) (hp _ = ∑ e ∈ compPartialSumSource 2 (n + 1) n, ∏ j : Fin e.1, r * (a ^ e.2 j * p (e.2 j)) := by symm apply compChangeOfVariables_sum - rintro ⟨k, blocks_fun⟩ H - have K : (compChangeOfVariables 2 (n + 1) n ⟨k, blocks_fun⟩ H).snd.length = k := by simp + rintro ⟨k, blocksFun⟩ H + have K : (compChangeOfVariables 2 (n + 1) n ⟨k, blocksFun⟩ H).snd.length = k := by simp congr 2 <;> try rw [K] rw [Fin.heq_fun_iff K.symm] intro j diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index 401c788269c57..ac4b0794cb1d6 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -37,11 +37,11 @@ between the two types. Let `c : Composition n` be a composition of `n`. Then * `c.blocks` is the list of blocks in `c`. * `c.length` is the number of blocks in the composition. -* `c.blocks_fun : Fin c.length → ℕ` is the realization of `c.blocks` as a function on +* `c.blocksFun : Fin c.length → ℕ` is the realization of `c.blocks` as a function on `Fin c.length`. This is the main object when using compositions to understand the composition of analytic functions. * `c.sizeUpTo : ℕ → ℕ` is the sum of the size of the blocks up to `i`.; -* `c.embedding i : Fin (c.blocks_fun i) → Fin n` is the increasing embedding of the `i`-th block in +* `c.embedding i : Fin (c.blocksFun i) → Fin n` is the increasing embedding of the `i`-th block in `Fin n`; * `c.index j`, for `j : Fin n`, is the index of the block containing `j`. @@ -253,7 +253,7 @@ theorem orderEmbOfFin_boundaries : refine (Finset.orderEmbOfFin_unique' _ ?_).symm exact fun i => (Finset.mem_map' _).2 (Finset.mem_univ _) -/-- Embedding the `i`-th block of a composition (identified with `Fin (c.blocks_fun i)`) into +/-- Embedding the `i`-th block of a composition (identified with `Fin (c.blocksFun i)`) into `Fin n` at the relevant position. -/ def embedding (i : Fin c.length) : Fin (c.blocksFun i) ↪o Fin n := (Fin.natAddOrderEmb <| c.sizeUpTo i).trans <| Fin.castLEOrderEmb <| @@ -267,7 +267,7 @@ theorem coe_embedding (i : Fin c.length) (j : Fin (c.blocksFun i)) : (c.embedding i j : ℕ) = c.sizeUpTo i + j := rfl -/-- `index_exists` asserts there is some `i` with `j < c.size_up_to (i+1)`. +/-- `index_exists` asserts there is some `i` with `j < c.sizeUpTo (i+1)`. In the next definition `index` we use `Nat.find` to produce the minimal such index. -/ theorem index_exists {j : ℕ} (h : j < n) : ∃ i : ℕ, j < c.sizeUpTo (i + 1) ∧ i < c.length := by @@ -301,7 +301,7 @@ theorem sizeUpTo_index_le (j : Fin n) : c.sizeUpTo (c.index j) ≤ j := by exact Nat.lt_le_asymm H this /-- Mapping an element `j` of `Fin n` to the element in the block containing it, identified with -`Fin (c.blocks_fun (c.index j))` through the canonical increasing bijection. -/ +`Fin (c.blocksFun (c.index j))` through the canonical increasing bijection. -/ def invEmbedding (j : Fin n) : Fin (c.blocksFun (c.index j)) := ⟨j - c.sizeUpTo (c.index j), by rw [tsub_lt_iff_right, add_comm, ← sizeUpTo_succ'] @@ -377,7 +377,7 @@ theorem invEmbedding_comp (i : Fin c.length) (j : Fin (c.blocksFun i)) : simp_rw [coe_invEmbedding, index_embedding, coe_embedding, add_tsub_cancel_left] /-- Equivalence between the disjoint union of the blocks (each of them seen as -`Fin (c.blocks_fun i)`) with `Fin n`. -/ +`Fin (c.blocksFun i)`) with `Fin n`. -/ def blocksFinEquiv : (Σi : Fin c.length, Fin (c.blocksFun i)) ≃ Fin n where toFun x := c.embedding x.1 x.2 invFun j := ⟨c.index j, c.invEmbedding j⟩ @@ -547,7 +547,7 @@ end Composition ### Splitting a list Given a list of length `n` and a composition `c` of `n`, one can split `l` into `c.length` sublists -of respective lengths `c.blocks_fun 0`, ..., `c.blocks_fun (c.length-1)`. This is inverse to the +of respective lengths `c.blocksFun 0`, ..., `c.blocksFun (c.length-1)`. This is inverse to the join operation. -/ From bdfc2b33b94a81651385eaa46c1e410ed9e2fb97 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Wed, 7 Aug 2024 18:10:50 +0000 Subject: [PATCH 103/359] chore: update Mathlib dependencies 2024-08-07 (#15594) 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 64809f172e29e..607daf7296fe1 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "8ff1f48157c719e04bebc729f12af55cc22aabe4", + "rev": "0444234b4216e944d5be2ce42a25d7410c67876f", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", From e860e3b8a9fe291681487172a0cb1f99303e3c8d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 7 Aug 2024 19:09:01 +0000 Subject: [PATCH 104/359] chore: reduce imports in Order.Ring.Hom (#15581) --- Archive/ZagierTwoSquares.lean | 1 + Counterexamples/MapFloor.lean | 1 + Mathlib.lean | 3 +- Mathlib/Algebra/AddConstMap/Basic.lean | 2 +- .../Computation/ApproximationCorollaries.lean | 2 +- .../Basic.lean} | 0 Mathlib/Algebra/Order/Archimedean/Hom.lean | 47 +++++++++++++++++++ .../Algebra/Order/CauSeq/BigOperators.lean | 2 +- Mathlib/Algebra/Order/Chebyshev.lean | 1 - Mathlib/Algebra/Order/CompleteField.lean | 2 +- Mathlib/Algebra/Order/Hom/Ring.lean | 42 +---------------- Mathlib/Algebra/Order/ToIntervalMod.lean | 2 +- Mathlib/Algebra/Periodic.lean | 2 +- Mathlib/Algebra/Polynomial/HasseDeriv.lean | 1 + Mathlib/Algebra/Ring/Subring/Order.lean | 1 + .../Additive/PluenneckeRuzsa.lean | 1 - Mathlib/Data/Real/Archimedean.lean | 2 +- Mathlib/Data/ZMod/Basic.lean | 2 + Mathlib/GroupTheory/Archimedean.lean | 2 +- Mathlib/GroupTheory/Coxeter/Basic.lean | 2 + Mathlib/GroupTheory/Coxeter/Inversion.lean | 1 + Mathlib/GroupTheory/Coxeter/Length.lean | 1 + Mathlib/GroupTheory/Exponent.lean | 1 + Mathlib/ModelTheory/Fraisse.lean | 1 + Mathlib/NumberTheory/LucasLehmer.lean | 2 +- Mathlib/Order/Filter/Archimedean.lean | 2 +- Mathlib/RingTheory/Binomial.lean | 2 + Mathlib/RingTheory/Noetherian.lean | 1 + Mathlib/RingTheory/Valuation/Basic.lean | 1 - Mathlib/Topology/Algebra/Group/Basic.lean | 1 + .../Topology/Algebra/InfiniteSum/Order.lean | 2 +- scripts/noshake.json | 22 ++++----- 32 files changed, 89 insertions(+), 66 deletions(-) rename Mathlib/Algebra/Order/{Archimedean.lean => Archimedean/Basic.lean} (100%) create mode 100644 Mathlib/Algebra/Order/Archimedean/Hom.lean diff --git a/Archive/ZagierTwoSquares.lean b/Archive/ZagierTwoSquares.lean index 6577bb99be719..e9f7d1bcb5b30 100644 --- a/Archive/ZagierTwoSquares.lean +++ b/Archive/ZagierTwoSquares.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Tan, Thomas Browning -/ import Mathlib.GroupTheory.Perm.Cycle.Type +import Mathlib.Tactic.Linarith /-! # Zagier's "one-sentence proof" of Fermat's theorem on sums of two squares diff --git a/Counterexamples/MapFloor.lean b/Counterexamples/MapFloor.lean index 6b5e97fe6ebed..5ddc5c17af16d 100644 --- a/Counterexamples/MapFloor.lean +++ b/Counterexamples/MapFloor.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import Mathlib.Algebra.Order.Floor import Mathlib.Algebra.Order.Group.PiLex import Mathlib.Algebra.Polynomial.Reverse diff --git a/Mathlib.lean b/Mathlib.lean index 8d37716480732..94fc9f82bcb87 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -523,7 +523,8 @@ import Mathlib.Algebra.Order.Algebra import Mathlib.Algebra.Order.Antidiag.Finsupp import Mathlib.Algebra.Order.Antidiag.Pi import Mathlib.Algebra.Order.Antidiag.Prod -import Mathlib.Algebra.Order.Archimedean +import Mathlib.Algebra.Order.Archimedean.Basic +import Mathlib.Algebra.Order.Archimedean.Hom import Mathlib.Algebra.Order.BigOperators.Group.Finset import Mathlib.Algebra.Order.BigOperators.Group.List import Mathlib.Algebra.Order.BigOperators.Group.Multiset diff --git a/Mathlib/Algebra/AddConstMap/Basic.lean b/Mathlib/Algebra/AddConstMap/Basic.lean index bcd78b97109b0..cf6f7d5f8a91a 100644 --- a/Mathlib/Algebra/AddConstMap/Basic.lean +++ b/Mathlib/Algebra/AddConstMap/Basic.lean @@ -6,7 +6,7 @@ Authors: Yury Kudryashov import Mathlib.Algebra.Group.Action.Pi import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.Module.Defs -import Mathlib.Algebra.Order.Archimedean +import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Algebra.Order.Group.Instances import Mathlib.Logic.Function.Iterate diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean b/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean index 6700e7e2320cb..22eed92dcfa54 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/ApproximationCorollaries.lean @@ -5,7 +5,7 @@ Authors: Kevin Kappelmann -/ import Mathlib.Algebra.ContinuedFractions.Computation.Approximations import Mathlib.Algebra.ContinuedFractions.ConvergentsEquiv -import Mathlib.Algebra.Order.Archimedean +import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Tactic.GCongr import Mathlib.Topology.Order.LeftRightNhds diff --git a/Mathlib/Algebra/Order/Archimedean.lean b/Mathlib/Algebra/Order/Archimedean/Basic.lean similarity index 100% rename from Mathlib/Algebra/Order/Archimedean.lean rename to Mathlib/Algebra/Order/Archimedean/Basic.lean diff --git a/Mathlib/Algebra/Order/Archimedean/Hom.lean b/Mathlib/Algebra/Order/Archimedean/Hom.lean new file mode 100644 index 0000000000000..92a89b61ee68e --- /dev/null +++ b/Mathlib/Algebra/Order/Archimedean/Hom.lean @@ -0,0 +1,47 @@ +/- +Copyright (c) 2022 Alex J. Best, Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alex J. Best, Yaël Dillies +-/ +import Mathlib.Algebra.Order.Archimedean.Basic +import Mathlib.Algebra.Order.Hom.Ring + +/-! +### Uniqueness of ring homomorphisms to archimedean fields. + +There is at most one ordered ring homomorphism from a linear ordered field to an archimedean linear +ordered field. Reciprocally, such an ordered ring homomorphism exists when the codomain is further +conditionally complete. +-/ + +variable {α β : Type*} + +/-- There is at most one ordered ring homomorphism from a linear ordered field to an archimedean +linear ordered field. -/ +instance OrderRingHom.subsingleton [LinearOrderedField α] [LinearOrderedField β] [Archimedean β] : + Subsingleton (α →+*o β) := + ⟨fun f g => by + ext x + by_contra! h' : f x ≠ g x + wlog h : f x < g x generalizing α β with h₂ + -- Porting note: had to add the `generalizing` as there are random variables + -- `F γ δ` flying around in context. + · exact h₂ g f x (Ne.symm h') (h'.lt_or_lt.resolve_left h) + obtain ⟨q, hf, hg⟩ := exists_rat_btwn h + rw [← map_ratCast f] at hf + rw [← map_ratCast g] at hg + exact + (lt_asymm ((OrderHomClass.mono g).reflect_lt hg) <| + (OrderHomClass.mono f).reflect_lt hf).elim⟩ + +/-- There is at most one ordered ring isomorphism between a linear ordered field and an archimedean +linear ordered field. -/ +instance OrderRingIso.subsingleton_right [LinearOrderedField α] [LinearOrderedField β] + [Archimedean β] : Subsingleton (α ≃+*o β) := + OrderRingIso.toOrderRingHom_injective.subsingleton + +/-- There is at most one ordered ring isomorphism between an archimedean linear ordered field and a +linear ordered field. -/ +instance OrderRingIso.subsingleton_left [LinearOrderedField α] [Archimedean α] + [LinearOrderedField β] : Subsingleton (α ≃+*o β) := + OrderRingIso.symm_bijective.injective.subsingleton diff --git a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean index 1e545d67b0704..129e4899f3ace 100644 --- a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean +++ b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yaël Dillies -/ import Mathlib.Algebra.GeomSum -import Mathlib.Algebra.Order.Archimedean +import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Algebra.Order.CauSeq.Basic /-! diff --git a/Mathlib/Algebra/Order/Chebyshev.lean b/Mathlib/Algebra/Order/Chebyshev.lean index db3067b750a85..ebbe5643ee00c 100644 --- a/Mathlib/Algebra/Order/Chebyshev.lean +++ b/Mathlib/Algebra/Order/Chebyshev.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mantas Bakšys, Yaël Dillies -/ import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Algebra.Order.Group.Basic import Mathlib.Algebra.Order.Rearrangement import Mathlib.Algebra.Order.Ring.Basic import Mathlib.GroupTheory.Perm.Cycle.Basic diff --git a/Mathlib/Algebra/Order/CompleteField.lean b/Mathlib/Algebra/Order/CompleteField.lean index 6b540009fb273..2c8942e53c15c 100644 --- a/Mathlib/Algebra/Order/CompleteField.lean +++ b/Mathlib/Algebra/Order/CompleteField.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Alex J. Best. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best, Yaël Dillies -/ -import Mathlib.Algebra.Order.Hom.Ring +import Mathlib.Algebra.Order.Archimedean.Hom import Mathlib.Algebra.Order.Pointwise import Mathlib.Analysis.SpecialFunctions.Pow.Real diff --git a/Mathlib/Algebra/Order/Hom/Ring.lean b/Mathlib/Algebra/Order/Hom/Ring.lean index 52a204af77c29..72c23ef5ca89e 100644 --- a/Mathlib/Algebra/Order/Hom/Ring.lean +++ b/Mathlib/Algebra/Order/Hom/Ring.lean @@ -3,9 +3,7 @@ Copyright (c) 2022 Alex J. Best, Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best, Yaël Dillies -/ -import Mathlib.Algebra.Order.Archimedean import Mathlib.Algebra.Order.Hom.Monoid -import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Algebra.Ring.Equiv /-! @@ -35,6 +33,8 @@ making some typeclasses and instances irrelevant. ordered ring homomorphism, order homomorphism -/ +assert_not_exists FloorRing +assert_not_exists Archimedean open Function @@ -452,41 +452,3 @@ theorem toOrderRingHom_injective : Injective (toOrderRingHom : α ≃+*o β → end NonAssocSemiring end OrderRingIso - -/-! -### Uniqueness - -There is at most one ordered ring homomorphism from a linear ordered field to an archimedean linear -ordered field. Reciprocally, such an ordered ring homomorphism exists when the codomain is further -conditionally complete. --/ - -/-- There is at most one ordered ring homomorphism from a linear ordered field to an archimedean -linear ordered field. -/ -instance OrderRingHom.subsingleton [LinearOrderedField α] [LinearOrderedField β] [Archimedean β] : - Subsingleton (α →+*o β) := - ⟨fun f g => by - ext x - by_contra! h' : f x ≠ g x - wlog h : f x < g x generalizing α β with h₂ - -- Porting note: had to add the `generalizing` as there are random variables - -- `F γ δ` flying around in context. - · exact h₂ g f x (Ne.symm h') (h'.lt_or_lt.resolve_left h) - obtain ⟨q, hf, hg⟩ := exists_rat_btwn h - rw [← map_ratCast f] at hf - rw [← map_ratCast g] at hg - exact - (lt_asymm ((OrderHomClass.mono g).reflect_lt hg) <| - (OrderHomClass.mono f).reflect_lt hf).elim⟩ - -/-- There is at most one ordered ring isomorphism between a linear ordered field and an archimedean -linear ordered field. -/ -instance OrderRingIso.subsingleton_right [LinearOrderedField α] [LinearOrderedField β] - [Archimedean β] : Subsingleton (α ≃+*o β) := - OrderRingIso.toOrderRingHom_injective.subsingleton - -/-- There is at most one ordered ring isomorphism between an archimedean linear ordered field and a -linear ordered field. -/ -instance OrderRingIso.subsingleton_left [LinearOrderedField α] [Archimedean α] - [LinearOrderedField β] : Subsingleton (α ≃+*o β) := - OrderRingIso.symm_bijective.injective.subsingleton diff --git a/Mathlib/Algebra/Order/ToIntervalMod.lean b/Mathlib/Algebra/Order/ToIntervalMod.lean index c77210693d946..706b51d23eb81 100644 --- a/Mathlib/Algebra/Order/ToIntervalMod.lean +++ b/Mathlib/Algebra/Order/ToIntervalMod.lean @@ -5,7 +5,7 @@ Authors: Joseph Myers -/ import Mathlib.Algebra.ModEq import Mathlib.Algebra.Module.Defs -import Mathlib.Algebra.Order.Archimedean +import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Algebra.Periodic import Mathlib.Data.Int.SuccPred import Mathlib.GroupTheory.QuotientGroup diff --git a/Mathlib/Algebra/Periodic.lean b/Mathlib/Algebra/Periodic.lean index b3ff7880c4e2b..2a4c1ad2350fe 100644 --- a/Mathlib/Algebra/Periodic.lean +++ b/Mathlib/Algebra/Periodic.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Field.Opposite import Mathlib.Algebra.Group.Subgroup.ZPowers import Mathlib.Algebra.Group.Submonoid.Membership import Mathlib.Algebra.Ring.NegOnePow -import Mathlib.Algebra.Order.Archimedean +import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.GroupTheory.Coset /-! diff --git a/Mathlib/Algebra/Polynomial/HasseDeriv.lean b/Mathlib/Algebra/Polynomial/HasseDeriv.lean index a9f6382c9f6d7..526b25659bb93 100644 --- a/Mathlib/Algebra/Polynomial/HasseDeriv.lean +++ b/Mathlib/Algebra/Polynomial/HasseDeriv.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Polynomial.Derivative import Mathlib.Data.Nat.Choose.Cast import Mathlib.Data.Nat.Choose.Vandermonde import Mathlib.Tactic.FieldSimp +import Mathlib.Tactic.Positivity /-! # Hasse derivative of polynomials diff --git a/Mathlib/Algebra/Ring/Subring/Order.lean b/Mathlib/Algebra/Ring/Subring/Order.lean index 0fbd9feacceed..3f640625b8d8f 100644 --- a/Mathlib/Algebra/Ring/Subring/Order.lean +++ b/Mathlib/Algebra/Ring/Subring/Order.lean @@ -5,6 +5,7 @@ Authors: Damiano Testa -/ import Mathlib.Algebra.Ring.Subring.Basic import Mathlib.Algebra.Order.Hom.Ring +import Mathlib.Algebra.Order.Ring.InjSurj /-! diff --git a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean index aadd4b09f1f38..b1c058d95cf66 100644 --- a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean +++ b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 Yaël Dillies, George Shakan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, George Shakan -/ -import Mathlib.Algebra.Order.Group.Basic import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Combinatorics.Enumerative.DoubleCounting import Mathlib.Data.Finset.Pointwise diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index db4cfc3bee867..1d868592b239f 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.Bounds -import Mathlib.Algebra.Order.Archimedean +import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Data.Real.Basic import Mathlib.Order.Interval.Set.Disjoint diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 0f50b93aa8261..dfdf5727a1b17 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -6,6 +6,8 @@ Authors: Chris Hughes import Mathlib.Algebra.Ring.Prod import Mathlib.GroupTheory.OrderOfElement import Mathlib.Tactic.FinCases +import Mathlib.Tactic.Linarith + /-! # Integers mod `n` diff --git a/Mathlib/GroupTheory/Archimedean.lean b/Mathlib/GroupTheory/Archimedean.lean index 84300c7017e45..77957e91b3685 100644 --- a/Mathlib/GroupTheory/Archimedean.lean +++ b/Mathlib/GroupTheory/Archimedean.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth, Patrick Massot -/ import Mathlib.Algebra.Group.Subgroup.Order -import Mathlib.Algebra.Order.Archimedean +import Mathlib.Algebra.Order.Archimedean.Basic /-! # Archimedean groups diff --git a/Mathlib/GroupTheory/Coxeter/Basic.lean b/Mathlib/GroupTheory/Coxeter/Basic.lean index ee7b5537ffb66..5fc6d5c513a62 100644 --- a/Mathlib/GroupTheory/Coxeter/Basic.lean +++ b/Mathlib/GroupTheory/Coxeter/Basic.lean @@ -6,6 +6,8 @@ Authors: Newell Jensen, Mitchell Lee import Mathlib.Algebra.Ring.Int import Mathlib.GroupTheory.PresentedGroup import Mathlib.GroupTheory.Coxeter.Matrix +import Mathlib.Tactic.Ring +import Mathlib.Tactic.Use /-! # Coxeter groups and Coxeter systems diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index 90a3048bbb6de..cdc9cb64b7dab 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -5,6 +5,7 @@ Authors: Mitchell Lee -/ import Mathlib.GroupTheory.Coxeter.Length import Mathlib.Data.ZMod.Parity +import Mathlib.Data.List.GetD /-! # Reflections, inversions, and inversion sequences diff --git a/Mathlib/GroupTheory/Coxeter/Length.lean b/Mathlib/GroupTheory/Coxeter/Length.lean index 06893e436cc1a..da1191305642b 100644 --- a/Mathlib/GroupTheory/Coxeter/Length.lean +++ b/Mathlib/GroupTheory/Coxeter/Length.lean @@ -5,6 +5,7 @@ Authors: Mitchell Lee -/ import Mathlib.Data.ZMod.Basic import Mathlib.GroupTheory.Coxeter.Basic +import Mathlib.Tactic.Zify /-! # The length function, reduced words, and descents diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index 92d9f9da0567e..dcdd5fdff4dee 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.GCDMonoid.Finset import Mathlib.Algebra.GCDMonoid.Nat import Mathlib.Data.Nat.Factorization.Basic import Mathlib.Tactic.Peel +import Mathlib.Algebra.Order.Archimedean.Basic /-! # Exponent of a group diff --git a/Mathlib/ModelTheory/Fraisse.lean b/Mathlib/ModelTheory/Fraisse.lean index 599d3e5bafc19..ca502e44d67e8 100644 --- a/Mathlib/ModelTheory/Fraisse.lean +++ b/Mathlib/ModelTheory/Fraisse.lean @@ -6,6 +6,7 @@ Authors: Aaron Anderson import Mathlib.ModelTheory.FinitelyGenerated import Mathlib.ModelTheory.DirectLimit import Mathlib.ModelTheory.Bundled +import Mathlib.Algebra.Order.Archimedean.Basic /-! # Fraïssé Classes and Fraïssé Limits diff --git a/Mathlib/NumberTheory/LucasLehmer.lean b/Mathlib/NumberTheory/LucasLehmer.lean index 9e9d7f058799f..13cbbcc14f54e 100644 --- a/Mathlib/NumberTheory/LucasLehmer.lean +++ b/Mathlib/NumberTheory/LucasLehmer.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Scott Morrison, Ainsley Pahljina -/ -import Mathlib.Algebra.Order.Group.Basic import Mathlib.Algebra.Order.Ring.Abs import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Algebra.Ring.Nat @@ -11,6 +10,7 @@ import Mathlib.Data.ZMod.Basic import Mathlib.GroupTheory.OrderOfElement import Mathlib.RingTheory.Fintype import Mathlib.Tactic.IntervalCases +import Mathlib.Tactic.Zify /-! # The Lucas-Lehmer test for Mersenne primes. diff --git a/Mathlib/Order/Filter/Archimedean.lean b/Mathlib/Order/Filter/Archimedean.lean index abd9d7f0176e8..7b5d1fe77631c 100644 --- a/Mathlib/Order/Filter/Archimedean.lean +++ b/Mathlib/Order/Filter/Archimedean.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ -import Mathlib.Algebra.Order.Archimedean +import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Order.Filter.AtTopBot import Mathlib.Tactic.GCongr diff --git a/Mathlib/RingTheory/Binomial.lean b/Mathlib/RingTheory/Binomial.lean index 5becb23233943..17b424f495698 100644 --- a/Mathlib/RingTheory/Binomial.lean +++ b/Mathlib/RingTheory/Binomial.lean @@ -4,8 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Carnahan -/ import Mathlib.Algebra.Polynomial.Smeval +import Mathlib.Algebra.Order.Floor import Mathlib.GroupTheory.GroupAction.Ring import Mathlib.RingTheory.Polynomial.Pochhammer +import Mathlib.Tactic.FieldSimp /-! # Binomial rings diff --git a/Mathlib/RingTheory/Noetherian.lean b/Mathlib/RingTheory/Noetherian.lean index 1b9760d77848f..5115056fb0474 100644 --- a/Mathlib/RingTheory/Noetherian.lean +++ b/Mathlib/RingTheory/Noetherian.lean @@ -8,6 +8,7 @@ import Mathlib.Order.PartialSups import Mathlib.Algebra.Module.Submodule.IterateMapComap import Mathlib.RingTheory.OrzechProperty import Mathlib.RingTheory.Nilpotent.Lemmas +import Mathlib.Algebra.Order.Archimedean.Basic /-! # Noetherian rings and modules diff --git a/Mathlib/RingTheory/Valuation/Basic.lean b/Mathlib/RingTheory/Valuation/Basic.lean index 1008370dd53ef..9dfcfbe3ccddc 100644 --- a/Mathlib/RingTheory/Valuation/Basic.lean +++ b/Mathlib/RingTheory/Valuation/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Buzzard, Johan Commelin, Patrick Massot -/ -import Mathlib.Algebra.Order.Group.Basic import Mathlib.Algebra.Order.Ring.Basic import Mathlib.RingTheory.Ideal.Maps import Mathlib.Tactic.TFAE diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 8a5349a3e7a28..aa02b48710111 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -8,6 +8,7 @@ import Mathlib.GroupTheory.GroupAction.Quotient import Mathlib.GroupTheory.QuotientGroup import Mathlib.Topology.Algebra.Monoid import Mathlib.Topology.Algebra.Constructions +import Mathlib.Algebra.Order.Archimedean.Basic /-! # Topological groups diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean index 414a2aa7f83c2..e0e206d0522bb 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ -import Mathlib.Algebra.Order.Archimedean +import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Topology.Algebra.InfiniteSum.NatInt import Mathlib.Topology.Algebra.Order.Field import Mathlib.Topology.Order.MonotoneConvergence diff --git a/scripts/noshake.json b/scripts/noshake.json index c89c42678e386..e2defe13f3e81 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -196,8 +196,6 @@ "Mathlib.Topology.Order.LeftRightNhds": ["Mathlib.Data.Set.Pointwise.Basic"], "Mathlib.Topology.Germ": ["Mathlib.Analysis.Normed.Module.Basic"], "Mathlib.Topology.Defs.Basic": ["Mathlib.Tactic.FunProp"], - "Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital": - ["Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital"], "Mathlib.Topology.Category.UniformSpace": ["Mathlib.CategoryTheory.Monad.Limits"], "Mathlib.Topology.Algebra.Nonarchimedean.AdicTopology": @@ -281,18 +279,16 @@ "Mathlib.Tactic.CategoryTheory.BicategoricalComp"], "Mathlib.Tactic.CC.Addition": ["Mathlib.Logic.Basic", "Mathlib.Tactic.CC.Lemmas"], - "Mathlib.Tactic.Basic": ["Mathlib.Tactic.Linter.OldObtain"], + "Mathlib.Tactic.Bound.Init": ["Aesop.Frontend.Command"], + "Mathlib.Tactic.Bound.Attribute": + ["Mathlib.Algebra.Group.ZeroOne", "Mathlib.Tactic.Bound.Init"], "Mathlib.Tactic.Bound": - ["Mathlib.Analysis.Analytic.Basic", - "Mathlib.Algebra.Order.AbsoluteValue", + ["Mathlib.Algebra.Order.AbsoluteValue", "Mathlib.Algebra.Order.BigOperators.Group.Finset", "Mathlib.Algebra.Order.Floor", - "Mathlib.Tactic.Bound.Attribute", + "Mathlib.Analysis.Analytic.Basic", "Mathlib.Tactic.Bound.Init"], - "Mathlib.Tactic.Bound.Attribute": - ["Mathlib.Algebra.Group.ZeroOne", - "Mathlib.Tactic.Bound.Init"], - "Mathlib.Tactic.Bound.Init": ["Aesop.Frontend.Command"], + "Mathlib.Tactic.Basic": ["Mathlib.Tactic.Linter.OldObtain"], "Mathlib.Tactic.Attr.Register": ["Lean.Meta.Tactic.Simp.SimpTheorems"], "Mathlib.Tactic.ArithMult": ["Mathlib.Tactic.ArithMult.Init"], "Mathlib.RingTheory.PowerSeries.Basic": @@ -300,6 +296,7 @@ "Mathlib.RingTheory.PolynomialAlgebra": ["Mathlib.Data.Matrix.DMatrix"], "Mathlib.RingTheory.MvPolynomial.Homogeneous": ["Mathlib.Algebra.DirectSum.Internal"], + "Mathlib.RingTheory.Binomial": ["Mathlib.Algebra.Order.Floor"], "Mathlib.Probability.Notation": ["Mathlib.MeasureTheory.Decomposition.Lebesgue", "Mathlib.MeasureTheory.Function.ConditionalExpectation.Basic"], @@ -366,12 +363,15 @@ ["Mathlib.CategoryTheory.Limits.Shapes.Pullback.HasPullback", "Mathlib.CategoryTheory.Limits.Shapes.Pullbacks"], "Mathlib.CategoryTheory.Limits.IsLimit": ["Batteries.Tactic.Congr"], - "Mathlib.Analysis.Normed.Operator.LinearIsometry": ["Mathlib.Algebra.Star.Basic"], + "Mathlib.Analysis.Normed.Operator.LinearIsometry": + ["Mathlib.Algebra.Star.Basic"], "Mathlib.Analysis.InnerProductSpace.Basic": ["Mathlib.Algebra.Module.LinearMap.Basic"], "Mathlib.Analysis.Distribution.SchwartzSpace": ["Mathlib.Tactic.MoveAdd"], "Mathlib.Analysis.Convex.Basic": ["Mathlib.Algebra.Order.BigOperators.Ring.Finset"], + "Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital": + ["Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unital"], "Mathlib.AlgebraicTopology.DoldKan.Notations": ["Mathlib.AlgebraicTopology.AlternatingFaceMapComplex"], "Mathlib.Algebra.Star.Module": ["Mathlib.Algebra.Module.LinearMap.Star"], From f29176d1c90742c1183ae30dbe08622272588e57 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 8 Aug 2024 01:59:27 +0000 Subject: [PATCH 105/359] chore: backports for leanprover/lean4#4814 (part 31) (#15575) --- .../NormedSpace/HahnBanach/Separation.lean | 3 +- .../CategoryTheory/Triangulated/Yoneda.lean | 1 + Mathlib/MeasureTheory/Integral/Marginal.lean | 41 ++++++++++--------- 3 files changed, 24 insertions(+), 21 deletions(-) diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean index c14eafb2b6a82..e35d4c0f03deb 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean @@ -205,8 +205,7 @@ theorem iInter_halfspaces_eq (hs₁ : Convex ℝ s) (hs₂ : IsClosed s) : namespace RCLike -variable [RCLike 𝕜] [TopologicalSpace E] [AddCommGroup E] [TopologicalAddGroup E] - [Module 𝕜 E] [Module ℝ E] [ContinuousSMul 𝕜 E] [IsScalarTower ℝ 𝕜 E] +variable [RCLike 𝕜] [Module 𝕜 E] [ContinuousSMul 𝕜 E] [IsScalarTower ℝ 𝕜 E] /--Real linear extension of continuous extension of `LinearMap.extendTo𝕜'` -/ noncomputable def extendTo𝕜'ₗ : (E →L[ℝ] ℝ) →ₗ[ℝ] (E →L[𝕜] 𝕜) := diff --git a/Mathlib/CategoryTheory/Triangulated/Yoneda.lean b/Mathlib/CategoryTheory/Triangulated/Yoneda.lean index c8c8d28da654c..3d97552da7183 100644 --- a/Mathlib/CategoryTheory/Triangulated/Yoneda.lean +++ b/Mathlib/CategoryTheory/Triangulated/Yoneda.lean @@ -49,6 +49,7 @@ noncomputable instance (A : Cᵒᵖ) : (preadditiveCoyoneda.obj A).ShiftSequence Functor.ShiftSequence.tautological _ _ lemma preadditiveCoyoneda_homologySequenceδ_apply + {C : Type*} [Category C] [Preadditive C] [HasShift C ℤ] (A : Cᵒᵖ) (T : Triangle C) (n₀ n₁ : ℤ) (h : n₀ + 1 = n₁) (x : A.unop ⟶ T.obj₃⟦n₀⟧) : (preadditiveCoyoneda.obj A).homologySequenceδ T n₀ n₁ h x = x ≫ T.mor₃⟦n₀⟧' ≫ (shiftFunctorAdd' C 1 n₀ n₁ (by omega)).inv.app _ := by diff --git a/Mathlib/MeasureTheory/Integral/Marginal.lean b/Mathlib/MeasureTheory/Integral/Marginal.lean index 60cf5d485cb95..76fcaa682df92 100644 --- a/Mathlib/MeasureTheory/Integral/Marginal.lean +++ b/Mathlib/MeasureTheory/Integral/Marginal.lean @@ -65,7 +65,7 @@ namespace MeasureTheory section LMarginal variable {δ δ' : Type*} {π : δ → Type*} [∀ x, MeasurableSpace (π x)] -variable {μ : ∀ i, Measure (π i)} [∀ i, SigmaFinite (μ i)] [DecidableEq δ] +variable {μ : ∀ i, Measure (π i)} [DecidableEq δ] variable {s t : Finset δ} {f g : (∀ i, π i) → ℝ≥0∞} {x y : ∀ i, π i} {i : δ} /-- Integrate `f(x₁,…,xₙ)` over all variables `xᵢ` where `i ∈ s`. Return a function in the @@ -85,7 +85,8 @@ notation "∫⋯∫⁻_" s ", " f => lmarginal (fun _ ↦ volume) s f variable (μ) -theorem _root_.Measurable.lmarginal (hf : Measurable f) : Measurable (∫⋯∫⁻_s, f ∂μ) := by +theorem _root_.Measurable.lmarginal [∀ i, SigmaFinite (μ i)] (hf : Measurable f) : + Measurable (∫⋯∫⁻_s, f ∂μ) := by refine Measurable.lintegral_prod_right ?_ refine hf.comp ?_ rw [measurable_pi_iff]; intro i @@ -113,6 +114,25 @@ theorem lmarginal_update_of_mem {i : δ} (hi : i ∈ s) have : j ≠ i := by rintro rfl; exact hj hi apply update_noteq this +variable {μ} in +theorem lmarginal_singleton (f : (∀ i, π i) → ℝ≥0∞) (i : δ) : + ∫⋯∫⁻_{i}, f ∂μ = fun x => ∫⁻ xᵢ, f (Function.update x i xᵢ) ∂μ i := by + let α : Type _ := ({i} : Finset δ) + let e := (MeasurableEquiv.piUnique fun j : α ↦ π j).symm + ext1 x + calc (∫⋯∫⁻_{i}, f ∂μ) x + = ∫⁻ (y : π (default : α)), f (updateFinset x {i} (e y)) ∂μ (default : α) := by + simp_rw [lmarginal, measurePreserving_piUnique (fun j : ({i} : Finset δ) ↦ μ j) |>.symm _ + |>.lintegral_map_equiv] + _ = ∫⁻ xᵢ, f (Function.update x i xᵢ) ∂μ i := by simp [update_eq_updateFinset]; rfl + +variable {μ} in +@[gcongr] +theorem lmarginal_mono {f g : (∀ i, π i) → ℝ≥0∞} (hfg : f ≤ g) : ∫⋯∫⁻_s, f ∂μ ≤ ∫⋯∫⁻_s, g ∂μ := + fun _ => lintegral_mono fun _ => hfg _ + +variable [∀ i, SigmaFinite (μ i)] + theorem lmarginal_union (f : (∀ i, π i) → ℝ≥0∞) (hf : Measurable f) (hst : Disjoint s t) : ∫⋯∫⁻_s ∪ t, f ∂μ = ∫⋯∫⁻_s, ∫⋯∫⁻_t, f ∂μ ∂μ := by ext1 x @@ -138,17 +158,6 @@ theorem lmarginal_union' (f : (∀ i, π i) → ℝ≥0∞) (hf : Measurable f) variable {μ} -theorem lmarginal_singleton (f : (∀ i, π i) → ℝ≥0∞) (i : δ) : - ∫⋯∫⁻_{i}, f ∂μ = fun x => ∫⁻ xᵢ, f (Function.update x i xᵢ) ∂μ i := by - let α : Type _ := ({i} : Finset δ) - let e := (MeasurableEquiv.piUnique fun j : α ↦ π j).symm - ext1 x - calc (∫⋯∫⁻_{i}, f ∂μ) x - = ∫⁻ (y : π (default : α)), f (updateFinset x {i} (e y)) ∂μ (default : α) := by - simp_rw [lmarginal, measurePreserving_piUnique (fun j : ({i} : Finset δ) ↦ μ j) |>.symm _ - |>.lintegral_map_equiv] - _ = ∫⁻ xᵢ, f (Function.update x i xᵢ) ∂μ i := by simp [update_eq_updateFinset]; rfl - /-- Peel off a single integral from a `lmarginal` integral at the beginning (compare with `lmarginal_insert'`, which peels off an integral at the end). -/ theorem lmarginal_insert (f : (∀ i, π i) → ℝ≥0∞) (hf : Measurable f) {i : δ} @@ -179,12 +188,6 @@ theorem lmarginal_erase' (f : (∀ i, π i) → ℝ≥0∞) (hf : Measurable f) ∫⋯∫⁻_s, f ∂μ = ∫⋯∫⁻_(erase s i), (fun x ↦ ∫⁻ xᵢ, f (Function.update x i xᵢ) ∂μ i) ∂μ := by simpa [insert_erase hi] using lmarginal_insert' _ hf (not_mem_erase i s) -open Filter - -@[gcongr] -theorem lmarginal_mono {f g : (∀ i, π i) → ℝ≥0∞} (hfg : f ≤ g) : ∫⋯∫⁻_s, f ∂μ ≤ ∫⋯∫⁻_s, g ∂μ := - fun _ => lintegral_mono fun _ => hfg _ - @[simp] theorem lmarginal_univ [Fintype δ] {f : (∀ i, π i) → ℝ≥0∞} : ∫⋯∫⁻_univ, f ∂μ = fun _ => ∫⁻ x, f x ∂Measure.pi μ := by let e : { j // j ∈ Finset.univ } ≃ δ := Equiv.subtypeUnivEquiv mem_univ From f786e149de0df265c1ffd26d3cc5dfb3767440e5 Mon Sep 17 00:00:00 2001 From: Jeremy Tan Jie Rui Date: Thu, 8 Aug 2024 02:10:37 +0000 Subject: [PATCH 106/359] chore: backports for leanprover/lean4#4814 (part 32) (#15577) --- .../Algebra/Algebra/NonUnitalSubalgebra.lean | 113 +++++++------ Mathlib/Algebra/DirectLimit.lean | 25 ++- Mathlib/Algebra/DirectSum/Module.lean | 16 +- Mathlib/Algebra/Lie/Abelian.lean | 3 +- Mathlib/Algebra/Lie/Rank.lean | 37 ++-- Mathlib/Algebra/Module/PID.lean | 2 +- .../BoxIntegral/Partition/Filter.lean | 6 +- .../Triangulated/HomologicalFunctor.lean | 25 ++- Mathlib/Geometry/Manifold/ContMDiff/Defs.lean | 119 +++++++------ Mathlib/LinearAlgebra/FreeModule/PID.lean | 10 +- .../TensorProduct/Graded/External.lean | 22 ++- Mathlib/MeasureTheory/Measure/Regular.lean | 159 ++++++++++-------- Mathlib/RingTheory/Generators.lean | 40 +++-- .../MvPolynomial/NewtonIdentities.lean | 10 +- 14 files changed, 343 insertions(+), 244 deletions(-) diff --git a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean index 9a7700afdd395..1bc3dc1438938 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean @@ -509,10 +509,13 @@ lemma span_eq_toSubmodule (s : NonUnitalSubalgebra R A) : Submodule.span R (s : Set A) = s.toSubmodule := by simp [SetLike.ext'_iff, Submodule.coe_span_eq_self] -variable [IsScalarTower R A A] [SMulCommClass R A A] -variable [NonUnitalNonAssocSemiring B] [Module R B] [IsScalarTower R B B] [SMulCommClass R B B] +variable [NonUnitalNonAssocSemiring B] [Module R B] variable [FunLike F A B] [NonUnitalAlgHomClass F R A B] +section IsScalarTower + +variable [IsScalarTower R A A] [SMulCommClass R A A] + /-- The minimal non-unital subalgebra that includes `s`. -/ def adjoin (s : Set A) : NonUnitalSubalgebra R A := { Submodule.span R (NonUnitalSubsemiring.closure s : Set A) with @@ -657,15 +660,15 @@ theorem adjoin_univ : adjoin R (Set.univ : Set A) = ⊤ := eq_top_iff.2 fun _x hx => subset_adjoin R hx open NonUnitalSubalgebra in -lemma _root_.NonUnitalAlgHom.map_adjoin (f : F) (s : Set A) : - map f (adjoin R s) = adjoin R (f '' s) := +lemma _root_.NonUnitalAlgHom.map_adjoin [IsScalarTower R B B] [SMulCommClass R B B] + (f : F) (s : Set A) : map f (adjoin R s) = adjoin R (f '' s) := Set.image_preimage.l_comm_of_u_comm (gc_map_comap f) NonUnitalAlgebra.gi.gc NonUnitalAlgebra.gi.gc fun _t => rfl open NonUnitalSubalgebra in @[simp] -lemma _root_.NonUnitalAlgHom.map_adjoin_singleton (f : F) (x : A) : - map f (adjoin R {x}) = adjoin R {f x} := by +lemma _root_.NonUnitalAlgHom.map_adjoin_singleton [IsScalarTower R B B] [SMulCommClass R B B] + (f : F) (x : A) : map f (adjoin R {x}) = adjoin R {f x} := by simp [NonUnitalAlgHom.map_adjoin] variable {R A} @@ -718,7 +721,8 @@ theorem mul_mem_sup {S T : NonUnitalSubalgebra R A} {x y : A} (hx : x ∈ S) (hy x * y ∈ S ⊔ T := mul_mem (mem_sup_left hx) (mem_sup_right hy) -theorem map_sup (f : F) (S T : NonUnitalSubalgebra R A) : +theorem map_sup [IsScalarTower R B B] [SMulCommClass R B B] + (f : F) (S T : NonUnitalSubalgebra R A) : ((S ⊔ T).map f : NonUnitalSubalgebra R B) = S.map f ⊔ T.map f := (NonUnitalSubalgebra.gc_map_comap f).l_sup @@ -788,10 +792,6 @@ theorem coe_bot : ((⊥ : NonUnitalSubalgebra R A) : Set A) = {0} := by theorem eq_top_iff {S : NonUnitalSubalgebra R A} : S = ⊤ ↔ ∀ x : A, x ∈ S := ⟨fun h x => by rw [h]; exact mem_top, fun h => by ext x; exact ⟨fun _ => mem_top, fun _ => h x⟩⟩ -theorem range_top_iff_surjective (f : A →ₙₐ[R] B) : - NonUnitalAlgHom.range f = (⊤ : NonUnitalSubalgebra R B) ↔ Function.Surjective f := - NonUnitalAlgebra.eq_top_iff - @[simp] theorem range_id : NonUnitalAlgHom.range (NonUnitalAlgHom.id R A) = ⊤ := SetLike.coe_injective Set.range_id @@ -801,17 +801,25 @@ theorem map_top (f : A →ₙₐ[R] B) : (⊤ : NonUnitalSubalgebra R A).map f = SetLike.coe_injective Set.image_univ @[simp] -theorem map_bot (f : A →ₙₐ[R] B) : (⊥ : NonUnitalSubalgebra R A).map f = ⊥ := +theorem map_bot [IsScalarTower R B B] [SMulCommClass R B B] + (f : A →ₙₐ[R] B) : (⊥ : NonUnitalSubalgebra R A).map f = ⊥ := SetLike.coe_injective <| by simp [NonUnitalAlgebra.coe_bot, NonUnitalSubalgebra.coe_map] @[simp] -theorem comap_top (f : A →ₙₐ[R] B) : (⊤ : NonUnitalSubalgebra R B).comap f = ⊤ := +theorem comap_top [IsScalarTower R B B] [SMulCommClass R B B] + (f : A →ₙₐ[R] B) : (⊤ : NonUnitalSubalgebra R B).comap f = ⊤ := eq_top_iff.2 fun _ => mem_top /-- `NonUnitalAlgHom` to `⊤ : NonUnitalSubalgebra R A`. -/ def toTop : A →ₙₐ[R] (⊤ : NonUnitalSubalgebra R A) := NonUnitalAlgHom.codRestrict (NonUnitalAlgHom.id R A) ⊤ fun _ => mem_top +end IsScalarTower + +theorem range_top_iff_surjective [IsScalarTower R B B] [SMulCommClass R B B] (f : A →ₙₐ[R] B) : + NonUnitalAlgHom.range f = (⊤ : NonUnitalSubalgebra R B) ↔ Function.Surjective f := + NonUnitalAlgebra.eq_top_iff + end NonUnitalAlgebra namespace NonUnitalSubalgebra @@ -831,9 +839,48 @@ theorem range_val : NonUnitalAlgHom.range (NonUnitalSubalgebraClass.subtype S) = instance subsingleton_of_subsingleton [Subsingleton A] : Subsingleton (NonUnitalSubalgebra R A) := ⟨fun B C => ext fun x => by simp only [Subsingleton.elim x 0, zero_mem B, zero_mem C]⟩ -variable [IsScalarTower R A A] [SMulCommClass R A A] variable [NonUnitalNonAssocSemiring B] [Module R B] +section Prod + +variable (S₁ : NonUnitalSubalgebra R B) + +/-- The product of two non-unital subalgebras is a non-unital subalgebra. -/ +def prod : NonUnitalSubalgebra R (A × B) := + { S.toNonUnitalSubsemiring.prod S₁.toNonUnitalSubsemiring with + carrier := S ×ˢ S₁ + smul_mem' := fun r _x hx => ⟨SMulMemClass.smul_mem r hx.1, SMulMemClass.smul_mem r hx.2⟩ } + +@[simp] +theorem coe_prod : (prod S S₁ : Set (A × B)) = (S : Set A) ×ˢ S₁ := + rfl + +theorem prod_toSubmodule : (S.prod S₁).toSubmodule = S.toSubmodule.prod S₁.toSubmodule := + rfl + +@[simp] +theorem mem_prod {S : NonUnitalSubalgebra R A} {S₁ : NonUnitalSubalgebra R B} {x : A × B} : + x ∈ prod S S₁ ↔ x.1 ∈ S ∧ x.2 ∈ S₁ := + Set.mem_prod + +variable [IsScalarTower R A A] [SMulCommClass R A A] [IsScalarTower R B B] [SMulCommClass R B B] + +@[simp] +theorem prod_top : (prod ⊤ ⊤ : NonUnitalSubalgebra R (A × B)) = ⊤ := by ext; simp + +theorem prod_mono {S T : NonUnitalSubalgebra R A} {S₁ T₁ : NonUnitalSubalgebra R B} : + S ≤ T → S₁ ≤ T₁ → prod S S₁ ≤ prod T T₁ := + Set.prod_mono + +@[simp] +theorem prod_inf_prod {S T : NonUnitalSubalgebra R A} {S₁ T₁ : NonUnitalSubalgebra R B} : + S.prod S₁ ⊓ T.prod T₁ = (S ⊓ T).prod (S₁ ⊓ T₁) := + SetLike.coe_injective Set.prod_inter_prod + +end Prod + +variable [IsScalarTower R A A] [SMulCommClass R A A] + instance _root_.NonUnitalAlgHom.subsingleton [Subsingleton (NonUnitalSubalgebra R A)] : Subsingleton (A →ₙₐ[R] B) := ⟨fun f g => @@ -880,44 +927,6 @@ theorem coe_inclusion {S T : NonUnitalSubalgebra R A} (h : S ≤ T) (s : S) : (inclusion h s : A) = s := rfl -section Prod - -variable (S₁ : NonUnitalSubalgebra R B) - -/-- The product of two non-unital subalgebras is a non-unital subalgebra. -/ -def prod : NonUnitalSubalgebra R (A × B) := - { S.toNonUnitalSubsemiring.prod S₁.toNonUnitalSubsemiring with - carrier := S ×ˢ S₁ - smul_mem' := fun r _x hx => ⟨SMulMemClass.smul_mem r hx.1, SMulMemClass.smul_mem r hx.2⟩ } - -@[simp] -theorem coe_prod : (prod S S₁ : Set (A × B)) = (S : Set A) ×ˢ S₁ := - rfl - -theorem prod_toSubmodule : (S.prod S₁).toSubmodule = S.toSubmodule.prod S₁.toSubmodule := - rfl - -@[simp] -theorem mem_prod {S : NonUnitalSubalgebra R A} {S₁ : NonUnitalSubalgebra R B} {x : A × B} : - x ∈ prod S S₁ ↔ x.1 ∈ S ∧ x.2 ∈ S₁ := - Set.mem_prod - -variable [IsScalarTower R B B] [SMulCommClass R B B] - -@[simp] -theorem prod_top : (prod ⊤ ⊤ : NonUnitalSubalgebra R (A × B)) = ⊤ := by ext; simp - -theorem prod_mono {S T : NonUnitalSubalgebra R A} {S₁ T₁ : NonUnitalSubalgebra R B} : - S ≤ T → S₁ ≤ T₁ → prod S S₁ ≤ prod T T₁ := - Set.prod_mono - -@[simp] -theorem prod_inf_prod {S T : NonUnitalSubalgebra R A} {S₁ T₁ : NonUnitalSubalgebra R B} : - S.prod S₁ ⊓ T.prod T₁ = (S ⊓ T).prod (S₁ ⊓ T₁) := - SetLike.coe_injective Set.prod_inter_prod - -end Prod - section SuprLift variable {ι : Type*} diff --git a/Mathlib/Algebra/DirectLimit.lean b/Mathlib/Algebra/DirectLimit.lean index 747314ff6af4a..29b9817a14331 100644 --- a/Mathlib/Algebra/DirectLimit.lean +++ b/Mathlib/Algebra/DirectLimit.lean @@ -79,10 +79,9 @@ nonrec theorem DirectedSystem.map_map [DirectedSystem G fun i j h => f i j h] {i DirectedSystem.map_map (fun i j h => f i j h) hij hjk x variable (G) -variable [DecidableEq ι] /-- The direct limit of a directed system is the modules glued together along the maps. -/ -def DirectLimit : Type max v w := +def DirectLimit [DecidableEq ι] : Type max v w := DirectSum ι G ⧸ (span R <| { a | @@ -91,6 +90,10 @@ def DirectLimit : Type max v w := namespace DirectLimit +section Basic + +variable [DecidableEq ι] + instance addCommGroup : AddCommGroup (DirectLimit G f) := Quotient.addCommGroup _ @@ -247,10 +250,9 @@ lemma congr_symm_apply_of [IsDirected ι (· ≤ ·)] end functorial -section Totalize - +end Basic -variable (G f) +section Totalize open Classical in /-- `totalize G f i j` is a linear map from `G i` to `G j`, for *every* `i` and `j`. @@ -269,7 +271,8 @@ theorem totalize_of_not_le {i j} (h : ¬i ≤ j) : totalize G f i j = 0 := end Totalize -variable [DirectedSystem G fun i j h => f i j h] +variable [DecidableEq ι] [DirectedSystem G fun i j h => f i j h] +variable {G f} theorem toModule_totalize_of_le [∀ i (k : G i), Decidable (k ≠ 0)] {x : DirectSum ι G} {i j : ι} (hij : i ≤ j) (hx : ∀ k ∈ x.support, k ≤ i) : @@ -343,10 +346,10 @@ end Module namespace AddCommGroup -variable [DecidableEq ι] [∀ i, AddCommGroup (G i)] +variable [∀ i, AddCommGroup (G i)] /-- The direct limit of a directed system is the abelian groups glued together along the maps. -/ -def DirectLimit (f : ∀ i j, i ≤ j → G i →+ G j) : Type _ := +def DirectLimit [DecidableEq ι] (f : ∀ i j, i ≤ j → G i →+ G j) : Type _ := @Module.DirectLimit ℤ _ ι _ G _ _ (fun i j hij => (f i j hij).toIntLinearMap) _ namespace DirectLimit @@ -359,6 +362,8 @@ protected theorem directedSystem [h : DirectedSystem G fun i j h => f i j h] : attribute [local instance] DirectLimit.directedSystem +variable [DecidableEq ι] + instance : AddCommGroup (DirectLimit G f) := Module.DirectLimit.addCommGroup G fun i j hij => (f i j hij).toIntLinearMap @@ -842,7 +847,7 @@ variable {G' : ι → Type v'} [∀ i, CommRing (G' i)] variable {f' : ∀ i j, i ≤ j → G' i →+* G' j} variable {G'' : ι → Type v''} [∀ i, CommRing (G'' i)] variable {f'' : ∀ i j, i ≤ j → G'' i →+* G'' j} -variable [Nonempty ι] + /-- Consider direct limits `lim G` and `lim G'` with direct system `f` and `f'` respectively, any family of ring homomorphisms `gᵢ : Gᵢ ⟶ G'ᵢ` such that `g ∘ f = f' ∘ g` induces a ring @@ -862,6 +867,8 @@ def map (g : (i : ι) → G i →+* G' i) map g hg (of G _ _ x) = of G' (fun _ _ h ↦ f' _ _ h) i (g i x) := lift_of _ _ _ _ _ +variable [Nonempty ι] + @[simp] lemma map_id [IsDirected ι (· ≤ ·)] : map (fun i ↦ RingHom.id _) (fun _ _ _ ↦ rfl) = RingHom.id (DirectLimit G fun _ _ h ↦ f _ _ h) := diff --git a/Mathlib/Algebra/DirectSum/Module.lean b/Mathlib/Algebra/DirectSum/Module.lean index 8321707132911..69dfbc7dc4fe0 100644 --- a/Mathlib/Algebra/DirectSum/Module.lean +++ b/Mathlib/Algebra/DirectSum/Module.lean @@ -54,6 +54,9 @@ theorem smul_apply (b : R) (v : ⨁ i, M i) (i : ι) : (b • v) i = b • v i : DFinsupp.smul_apply _ _ _ variable (R ι M) + +section DecidableEq + variable [DecidableEq ι] /-- Create the direct sum given a family `M` of `R` modules indexed over `ι`. -/ @@ -158,6 +161,8 @@ theorem linearEquivFunOnFintype_symm_single [Fintype ι] (i : ι) (m : M i) : rw [DFinsupp.equivFunOnFintype_symm_single i m] rfl +end DecidableEq + @[simp] theorem linearEquivFunOnFintype_symm_coe [Fintype ι] (f : ⨁ i, M i) : (linearEquivFunOnFintype R ι M).symm f = f := by @@ -168,8 +173,6 @@ protected def lid (M : Type v) (ι : Type* := PUnit) [AddCommMonoid M] [Module R (⨁ _ : ι, M) ≃ₗ[R] M := { DirectSum.id M ι, toModule R ι M fun _ ↦ LinearMap.id with } -variable (ι M) - /-- The projection map onto one component, as a linear map. -/ def component (i : ι) : (⨁ i, M i) →ₗ[R] M i := DFinsupp.lapply i @@ -187,14 +190,15 @@ theorem ext_iff {f g : ⨁ i, M i} : f = g ↔ ∀ i, component R ι M i f = com ⟨fun h _ ↦ by rw [h], ext R⟩ @[simp] -theorem lof_apply (i : ι) (b : M i) : ((lof R ι M i) b) i = b := +theorem lof_apply [DecidableEq ι] (i : ι) (b : M i) : ((lof R ι M i) b) i = b := DFinsupp.single_eq_same @[simp] -theorem component.lof_self (i : ι) (b : M i) : component R ι M i ((lof R ι M i) b) = b := +theorem component.lof_self [DecidableEq ι] (i : ι) (b : M i) : + component R ι M i ((lof R ι M i) b) = b := lof_apply R i b -theorem component.of (i j : ι) (b : M j) : +theorem component.of [DecidableEq ι] (i j : ι) (b : M j) : component R ι M i ((lof R ι M j) b) = if h : j = i then Eq.recOn h b else 0 := DFinsupp.single_apply @@ -216,7 +220,7 @@ end CongrLeft section Sigma variable {α : ι → Type*} {δ : ∀ i, α i → Type w} -variable [∀ i j, AddCommMonoid (δ i j)] [∀ i j, Module R (δ i j)] +variable [DecidableEq ι] [∀ i j, AddCommMonoid (δ i j)] [∀ i j, Module R (δ i j)] /-- `curry` as a linear map. -/ def sigmaLcurry : (⨁ i : Σi, _, δ i.1 i.2) →ₗ[R] ⨁ (i) (j), δ i j := diff --git a/Mathlib/Algebra/Lie/Abelian.lean b/Mathlib/Algebra/Lie/Abelian.lean index c5c44b1c83217..af06765422898 100644 --- a/Mathlib/Algebra/Lie/Abelian.lean +++ b/Mathlib/Algebra/Lie/Abelian.lean @@ -271,8 +271,7 @@ open LieSubmodule LieSubalgebra variable {R : Type u} {L : Type v} {M : Type w} variable [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] -variable [LieRingModule L M] [LieModule R L M] -variable (N N' : LieSubmodule R L M) (I J : LieIdeal R L) +variable [LieRingModule L M] (N N' : LieSubmodule R L M) (I J : LieIdeal R L) @[simp] theorem LieSubmodule.trivial_lie_oper_zero [LieModule.IsTrivial L M] : ⁅I, N⁆ = ⊥ := by diff --git a/Mathlib/Algebra/Lie/Rank.lean b/Mathlib/Algebra/Lie/Rank.lean index bafbba08fddc2..e4e3f5c3f6738 100644 --- a/Mathlib/Algebra/Lie/Rank.lean +++ b/Mathlib/Algebra/Lie/Rank.lean @@ -31,13 +31,13 @@ if the `n`-th coefficient of the characteristic polynomial of `ad R L x` is non- -/ variable {R A L M ι ιₘ : Type*} -variable [CommRing R] [Nontrivial R] +variable [CommRing R] variable [CommRing A] [Algebra R A] variable [LieRing L] [LieAlgebra R L] [Module.Finite R L] [Module.Free R L] variable [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] variable [Module.Finite R M] [Module.Free R M] -variable [Fintype ι] [DecidableEq ι] -variable [Fintype ιₘ] [DecidableEq ιₘ] +variable [Fintype ι] +variable [Fintype ιₘ] variable (b : Basis ι R L) (bₘ : Basis ιₘ R M) (x : L) namespace LieModule @@ -57,25 +57,26 @@ The *rank* of `M` is the smallest `n` for which the `n`-th coefficient is not th noncomputable def rank : ℕ := nilRank φ -lemma polyCharpoly_coeff_rank_ne_zero : +lemma polyCharpoly_coeff_rank_ne_zero [Nontrivial R] [DecidableEq ι] : (polyCharpoly φ b).coeff (rank R L M) ≠ 0 := polyCharpoly_coeff_nilRank_ne_zero _ _ -lemma rank_eq_natTrailingDegree : +lemma rank_eq_natTrailingDegree [Nontrivial R] [DecidableEq ι] : rank R L M = (polyCharpoly φ b).natTrailingDegree := by apply nilRank_eq_polyCharpoly_natTrailingDegree open FiniteDimensional -lemma rank_le_card : rank R L M ≤ Fintype.card ιₘ := + +lemma rank_le_card [Nontrivial R] : rank R L M ≤ Fintype.card ιₘ := nilRank_le_card _ bₘ open FiniteDimensional -lemma rank_le_finrank : rank R L M ≤ finrank R M := +lemma rank_le_finrank [Nontrivial R] : rank R L M ≤ finrank R M := nilRank_le_finrank _ variable {L} -lemma rank_le_natTrailingDegree_charpoly_ad : +lemma rank_le_natTrailingDegree_charpoly_ad [Nontrivial R] : rank R L M ≤ (toEnd R L M x).charpoly.natTrailingDegree := nilRank_le_natTrailingDegree_charpoly _ _ @@ -87,13 +88,13 @@ def IsRegular (x : L) : Prop := LinearMap.IsNilRegular φ x lemma isRegular_def : IsRegular R M x ↔ (toEnd R L M x).charpoly.coeff (rank R L M) ≠ 0 := Iff.rfl -lemma isRegular_iff_coeff_polyCharpoly_rank_ne_zero : +lemma isRegular_iff_coeff_polyCharpoly_rank_ne_zero [DecidableEq ι] : IsRegular R M x ↔ MvPolynomial.eval (b.repr x) ((polyCharpoly φ b).coeff (rank R L M)) ≠ 0 := LinearMap.isNilRegular_iff_coeff_polyCharpoly_nilRank_ne_zero _ _ _ -lemma isRegular_iff_natTrailingDegree_charpoly_eq_rank : +lemma isRegular_iff_natTrailingDegree_charpoly_eq_rank [Nontrivial R] : IsRegular R M x ↔ (toEnd R L M x).charpoly.natTrailingDegree = rank R L M := LinearMap.isNilRegular_iff_natTrailingDegree_charpoly_eq_nilRank _ _ section IsDomain @@ -128,25 +129,25 @@ The *rank* of `L` is the smallest `n` for which the `n`-th coefficient is not th noncomputable abbrev rank : ℕ := LieModule.rank R L L -lemma polyCharpoly_coeff_rank_ne_zero : +lemma polyCharpoly_coeff_rank_ne_zero [Nontrivial R] [DecidableEq ι] : (polyCharpoly (ad R L).toLinearMap b).coeff (rank R L) ≠ 0 := polyCharpoly_coeff_nilRank_ne_zero _ _ -lemma rank_eq_natTrailingDegree : +lemma rank_eq_natTrailingDegree [Nontrivial R] [DecidableEq ι] : rank R L = (polyCharpoly (ad R L).toLinearMap b).natTrailingDegree := by apply nilRank_eq_polyCharpoly_natTrailingDegree open FiniteDimensional -lemma rank_le_card : rank R L ≤ Fintype.card ι := + +lemma rank_le_card [Nontrivial R] : rank R L ≤ Fintype.card ι := nilRank_le_card _ b -open FiniteDimensional -lemma rank_le_finrank : rank R L ≤ finrank R L := +lemma rank_le_finrank [Nontrivial R] : rank R L ≤ finrank R L := nilRank_le_finrank _ variable {L} -lemma rank_le_natTrailingDegree_charpoly_ad : +lemma rank_le_natTrailingDegree_charpoly_ad [Nontrivial R] : rank R L ≤ (ad R L x).charpoly.natTrailingDegree := nilRank_le_natTrailingDegree_charpoly _ _ @@ -158,13 +159,13 @@ abbrev IsRegular (x : L) : Prop := LieModule.IsRegular R L x lemma isRegular_def : IsRegular R x ↔ (Polynomial.coeff (ad R L x).charpoly (rank R L) ≠ 0) := Iff.rfl -lemma isRegular_iff_coeff_polyCharpoly_rank_ne_zero : +lemma isRegular_iff_coeff_polyCharpoly_rank_ne_zero [DecidableEq ι] : IsRegular R x ↔ MvPolynomial.eval (b.repr x) ((polyCharpoly (ad R L).toLinearMap b).coeff (rank R L)) ≠ 0 := LinearMap.isNilRegular_iff_coeff_polyCharpoly_nilRank_ne_zero _ _ _ -lemma isRegular_iff_natTrailingDegree_charpoly_eq_rank : +lemma isRegular_iff_natTrailingDegree_charpoly_eq_rank [Nontrivial R] : IsRegular R x ↔ (ad R L x).charpoly.natTrailingDegree = rank R L := LinearMap.isNilRegular_iff_natTrailingDegree_charpoly_eq_nilRank _ _ section IsDomain diff --git a/Mathlib/Algebra/Module/PID.lean b/Mathlib/Algebra/Module/PID.lean index 83c608ab901a1..50be6d05c621a 100644 --- a/Mathlib/Algebra/Module/PID.lean +++ b/Mathlib/Algebra/Module/PID.lean @@ -264,7 +264,7 @@ theorem equiv_free_prod_directSum [h' : Module.Finite R N] : haveI := Module.Finite.of_surjective _ (torsion R N).mkQ_surjective obtain ⟨I, fI, p, hp, e, ⟨h⟩⟩ := equiv_directSum_of_isTorsion.{u, v} (@torsion_isTorsion R N _ _ _) - obtain ⟨n, ⟨g⟩⟩ := @Module.basisOfFiniteTypeTorsionFree' R _ _ _ (N ⧸ torsion R N) _ _ _ _ + obtain ⟨n, ⟨g⟩⟩ := @Module.basisOfFiniteTypeTorsionFree' R _ (N ⧸ torsion R N) _ _ _ _ _ _ haveI : Module.Projective R (N ⧸ torsion R N) := Module.Projective.of_basis ⟨g⟩ obtain ⟨f, hf⟩ := Module.projective_lifting_property _ LinearMap.id (torsion R N).mkQ_surjective refine diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean index 03f25141416a6..b0558a5db054c 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean @@ -328,15 +328,17 @@ theorem toFilter_inf_iUnion_eq (l : IntegrationParams) (I : Box ι) (π₀ : Pre variable {r r₁ r₂ : (ι → ℝ) → Ioi (0 : ℝ)} {π π₁ π₂ : TaggedPrepartition I} -theorem MemBaseSet.mono' (I : Box ι) (h : l₁ ≤ l₂) (hc : c₁ ≤ c₂) {π : TaggedPrepartition I} +variable (I) in +theorem MemBaseSet.mono' (h : l₁ ≤ l₂) (hc : c₁ ≤ c₂) (hr : ∀ J ∈ π, r₁ (π.tag J) ≤ r₂ (π.tag J)) (hπ : l₁.MemBaseSet I c₁ r₁ π) : l₂.MemBaseSet I c₂ r₂ π := ⟨hπ.1.mono' hr, fun h₂ => hπ.2 (le_iff_imp.1 h.2.1 h₂), fun hD => (hπ.3 (le_iff_imp.1 h.2.2 hD)).trans hc, fun hD => (hπ.4 (le_iff_imp.1 h.2.2 hD)).imp fun _ hπ => ⟨hπ.1, hπ.2.trans hc⟩⟩ +variable (I) in @[mono] -theorem MemBaseSet.mono (I : Box ι) (h : l₁ ≤ l₂) (hc : c₁ ≤ c₂) {π : TaggedPrepartition I} +theorem MemBaseSet.mono (h : l₁ ≤ l₂) (hc : c₁ ≤ c₂) (hr : ∀ x ∈ Box.Icc I, r₁ x ≤ r₂ x) (hπ : l₁.MemBaseSet I c₁ r₁ π) : l₂.MemBaseSet I c₂ r₂ π := hπ.mono' I h hc fun J _ => hr _ <| π.tag_mem_Icc J diff --git a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean index ee62de01bca66..3c73fda709c13 100644 --- a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean +++ b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean @@ -52,16 +52,20 @@ namespace CategoryTheory open Category Limits Pretriangulated ZeroObject Preadditive -variable {C D A : Type*} [Category C] [HasZeroObject C] [HasShift C ℤ] [Preadditive C] - [∀ (n : ℤ), (CategoryTheory.shiftFunctor C n).Additive] [Pretriangulated C] +variable {C D A : Type*} [Category C] [HasShift C ℤ] [Category D] [HasZeroObject D] [HasShift D ℤ] [Preadditive D] [∀ (n : ℤ), (CategoryTheory.shiftFunctor D n).Additive] [Pretriangulated D] - [Category A] [Abelian A] + [Category A] namespace Functor variable (F : C ⥤ A) +section Pretriangulated + +variable [HasZeroObject C] [Preadditive C] [∀ (n : ℤ), (CategoryTheory.shiftFunctor C n).Additive] + [Pretriangulated C] [Abelian A] + /-- A functor from a pretriangulated category to an abelian category is an homological functor if it sends distinguished triangles to exact sequences. -/ class IsHomological extends F.PreservesZeroMorphisms : Prop where @@ -151,27 +155,32 @@ lemma isHomological_of_localization (L : C ⥤ D) obtain ⟨T₀, e, hT₀⟩ := hT exact ⟨L.mapTriangle.obj T₀, e, (L ⋙ F).map_distinguished_exact _ hT₀⟩ -section +end Pretriangulated -variable [F.IsHomological] [F.ShiftSequence ℤ] (T T' : Triangle C) (hT : T ∈ distTriang C) - (hT' : T' ∈ distTriang C) (φ : T ⟶ T') (n₀ n₁ : ℤ) (h : n₀ + 1 = n₁) +section /-- The connecting homomorphism in the long exact sequence attached to an homological functor and a distinguished triangle. -/ -noncomputable def homologySequenceδ : +noncomputable def homologySequenceδ + [F.ShiftSequence ℤ] (T : Triangle C) (n₀ n₁ : ℤ) (h : n₀ + 1 = n₁) : (F.shift n₀).obj T.obj₃ ⟶ (F.shift n₁).obj T.obj₁ := F.shiftMap T.mor₃ n₀ n₁ (by rw [add_comm 1, h]) variable {T T'} @[reassoc] -lemma homologySequenceδ_naturality : +lemma homologySequenceδ_naturality + [F.ShiftSequence ℤ] (T T' : Triangle C) (φ : T ⟶ T') (n₀ n₁ : ℤ) (h : n₀ + 1 = n₁) : (F.shift n₀).map φ.hom₃ ≫ F.homologySequenceδ T' n₀ n₁ h = F.homologySequenceδ T n₀ n₁ h ≫ (F.shift n₁).map φ.hom₁ := by dsimp only [homologySequenceδ] rw [← shiftMap_comp', ← φ.comm₃, shiftMap_comp] variable (T) +variable [HasZeroObject C] [Preadditive C] [∀ (n : ℤ), (CategoryTheory.shiftFunctor C n).Additive] + [Pretriangulated C] [Abelian A] [F.IsHomological] +variable [F.ShiftSequence ℤ] (T T' : Triangle C) (hT : T ∈ distTriang C) + (hT' : T' ∈ distTriang C) (φ : T ⟶ T') (n₀ n₁ : ℤ) (h : n₀ + 1 = n₁) @[reassoc] lemma comp_homologySequenceδ : diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index 2c13c76476527..3028f6318c3cd 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -52,17 +52,15 @@ open scoped Topology Manifold variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] - -- declare a smooth manifold `M` over the pair `(E, H)`. + -- Prerequisite typeclasses to say that `M` is a smooth manifold over the pair `(E, H)` {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] - [SmoothManifoldWithCorners I M] - -- declare a smooth manifold `M'` over the pair `(E', H')`. + -- Prerequisite typeclasses to say that `M'` is a smooth manifold over the pair `(E', H')` {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] (I' : ModelWithCorners 𝕜 E' H') {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] - [SmoothManifoldWithCorners I' M'] - -- declare a manifold `M''` over the pair `(E'', H'')`. + -- Prerequisite typeclasses to say that `M''` is a smooth manifold over the pair `(E'', H'')` {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] {I'' : ModelWithCorners 𝕜 E'' H''} {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] @@ -342,6 +340,58 @@ 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 +section SmoothManifoldWithCorners + +theorem contMDiffWithinAt_iff_source_of_mem_maximalAtlas + [SmoothManifoldWithCorners I M] (he : e ∈ maximalAtlas I M) (hx : x ∈ e.source) : + ContMDiffWithinAt I I' n f s x ↔ + ContMDiffWithinAt 𝓘(𝕜, E) I' n (f ∘ (e.extend I).symm) ((e.extend I).symm ⁻¹' s ∩ range I) + (e.extend I x) := by + have h2x := hx; rw [← e.extend_source I] at h2x + simp_rw [ContMDiffWithinAt, + (contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_indep_chart_source he hx, + StructureGroupoid.liftPropWithinAt_self_source, + e.extend_symm_continuousWithinAt_comp_right_iff, contDiffWithinAtProp_self_source, + ContDiffWithinAtProp, Function.comp, e.left_inv hx, (e.extend I).left_inv h2x] + rfl + +theorem contMDiffWithinAt_iff_source_of_mem_source + [SmoothManifoldWithCorners I M] {x' : M} (hx' : x' ∈ (chartAt H x).source) : + ContMDiffWithinAt I I' n f s x' ↔ + ContMDiffWithinAt 𝓘(𝕜, E) I' n (f ∘ (extChartAt I x).symm) + ((extChartAt I x).symm ⁻¹' s ∩ range I) (extChartAt I x x') := + contMDiffWithinAt_iff_source_of_mem_maximalAtlas (chart_mem_maximalAtlas I x) hx' + +theorem contMDiffAt_iff_source_of_mem_source + [SmoothManifoldWithCorners I M] {x' : M} (hx' : x' ∈ (chartAt H x).source) : + ContMDiffAt I I' n f x' ↔ + ContMDiffWithinAt 𝓘(𝕜, E) I' n (f ∘ (extChartAt I x).symm) (range I) (extChartAt I x x') := by + simp_rw [ContMDiffAt, contMDiffWithinAt_iff_source_of_mem_source hx', preimage_univ, univ_inter] + +theorem contMDiffWithinAt_iff_target_of_mem_source + [SmoothManifoldWithCorners I' M'] {x : M} {y : M'} (hy : f x ∈ (chartAt H' y).source) : + ContMDiffWithinAt I I' n f s x ↔ + ContinuousWithinAt f s x ∧ ContMDiffWithinAt I 𝓘(𝕜, E') n (extChartAt I' y ∘ f) s x := by + simp_rw [ContMDiffWithinAt] + rw [(contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_indep_chart_target + (chart_mem_maximalAtlas I' y) hy, + and_congr_right] + intro hf + simp_rw [StructureGroupoid.liftPropWithinAt_self_target] + simp_rw [((chartAt H' y).continuousAt hy).comp_continuousWithinAt hf] + rw [← extChartAt_source I'] at hy + simp_rw [(continuousAt_extChartAt' I' hy).comp_continuousWithinAt hf] + rfl + +theorem contMDiffAt_iff_target_of_mem_source + [SmoothManifoldWithCorners I' M'] {x : M} {y : M'} (hy : f x ∈ (chartAt H' y).source) : + ContMDiffAt I I' n f x ↔ + ContinuousAt f x ∧ ContMDiffAt I 𝓘(𝕜, E') n (extChartAt I' y ∘ f) x := by + rw [ContMDiffAt, contMDiffWithinAt_iff_target_of_mem_source hy, continuousWithinAt_univ, + ContMDiffAt] + +variable [SmoothManifoldWithCorners I M] [SmoothManifoldWithCorners I' M'] + theorem contMDiffWithinAt_iff_of_mem_maximalAtlas {x : M} (he : e ∈ maximalAtlas I M) (he' : e' ∈ maximalAtlas I' M') (hx : x ∈ e.source) (hy : f x ∈ e'.source) : ContMDiffWithinAt I I' n f s x ↔ @@ -399,51 +449,6 @@ theorem contMDiffAt_iff_of_mem_source {x' : M} {y : M'} (hx : x' ∈ (chartAt H (contMDiffWithinAt_iff_of_mem_source hx hy).trans <| by rw [continuousWithinAt_univ, preimage_univ, univ_inter] -theorem contMDiffWithinAt_iff_target_of_mem_source {x : M} {y : M'} - (hy : f x ∈ (chartAt H' y).source) : - ContMDiffWithinAt I I' n f s x ↔ - ContinuousWithinAt f s x ∧ ContMDiffWithinAt I 𝓘(𝕜, E') n (extChartAt I' y ∘ f) s x := by - simp_rw [ContMDiffWithinAt] - rw [(contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_indep_chart_target - (chart_mem_maximalAtlas I' y) hy, - and_congr_right] - intro hf - simp_rw [StructureGroupoid.liftPropWithinAt_self_target] - simp_rw [((chartAt H' y).continuousAt hy).comp_continuousWithinAt hf] - rw [← extChartAt_source I'] at hy - simp_rw [(continuousAt_extChartAt' I' hy).comp_continuousWithinAt hf] - rfl - -theorem contMDiffAt_iff_target_of_mem_source {x : M} {y : M'} (hy : f x ∈ (chartAt H' y).source) : - ContMDiffAt I I' n f x ↔ - ContinuousAt f x ∧ ContMDiffAt I 𝓘(𝕜, E') n (extChartAt I' y ∘ f) x := by - rw [ContMDiffAt, contMDiffWithinAt_iff_target_of_mem_source hy, continuousWithinAt_univ, - ContMDiffAt] - -theorem contMDiffWithinAt_iff_source_of_mem_maximalAtlas (he : e ∈ maximalAtlas I M) - (hx : x ∈ e.source) : - ContMDiffWithinAt I I' n f s x ↔ - ContMDiffWithinAt 𝓘(𝕜, E) I' n (f ∘ (e.extend I).symm) ((e.extend I).symm ⁻¹' s ∩ range I) - (e.extend I x) := by - have h2x := hx; rw [← e.extend_source I] at h2x - simp_rw [ContMDiffWithinAt, - (contDiffWithinAt_localInvariantProp I I' n).liftPropWithinAt_indep_chart_source he hx, - StructureGroupoid.liftPropWithinAt_self_source, - e.extend_symm_continuousWithinAt_comp_right_iff, contDiffWithinAtProp_self_source, - ContDiffWithinAtProp, Function.comp, e.left_inv hx, (e.extend I).left_inv h2x] - rfl - -theorem contMDiffWithinAt_iff_source_of_mem_source {x' : M} (hx' : x' ∈ (chartAt H x).source) : - ContMDiffWithinAt I I' n f s x' ↔ - ContMDiffWithinAt 𝓘(𝕜, E) I' n (f ∘ (extChartAt I x).symm) - ((extChartAt I x).symm ⁻¹' s ∩ range I) (extChartAt I x x') := - contMDiffWithinAt_iff_source_of_mem_maximalAtlas (chart_mem_maximalAtlas I x) hx' - -theorem contMDiffAt_iff_source_of_mem_source {x' : M} (hx' : x' ∈ (chartAt H x).source) : - ContMDiffAt I I' n f x' ↔ - ContMDiffWithinAt 𝓘(𝕜, E) I' n (f ∘ (extChartAt I x).symm) (range I) (extChartAt I x x') := by - simp_rw [ContMDiffAt, contMDiffWithinAt_iff_source_of_mem_source hx', preimage_univ, univ_inter] - theorem contMDiffOn_iff_of_mem_maximalAtlas (he : e ∈ maximalAtlas I M) (he' : e' ∈ maximalAtlas I' M') (hs : s ⊆ e.source) (h2s : MapsTo f s e'.source) : ContMDiffOn I I' n f s ↔ @@ -580,6 +585,8 @@ theorem smooth_iff_target : ∀ y : M', SmoothOn I 𝓘(𝕜, E') (extChartAt I' y ∘ f) (f ⁻¹' (extChartAt I' y).source) := contMDiff_iff_target +end SmoothManifoldWithCorners + /-! ### Deducing smoothness from smoothness one step beyond -/ @@ -703,7 +710,8 @@ theorem ContMDiffOn.contMDiffAt (h : ContMDiffOn I I' n f s) (hx : s ∈ 𝓝 x) theorem SmoothOn.smoothAt (h : SmoothOn I I' f s) (hx : s ∈ 𝓝 x) : SmoothAt I I' f x := h.contMDiffAt hx -theorem contMDiffOn_iff_source_of_mem_maximalAtlas (he : e ∈ maximalAtlas I M) (hs : s ⊆ e.source) : +theorem contMDiffOn_iff_source_of_mem_maximalAtlas [SmoothManifoldWithCorners I M] + (he : e ∈ maximalAtlas I M) (hs : s ⊆ e.source) : ContMDiffOn I I' n f s ↔ ContMDiffOn 𝓘(𝕜, E) I' n (f ∘ (e.extend I).symm) (e.extend I '' s) := by simp_rw [ContMDiffOn, Set.forall_mem_image] @@ -716,7 +724,8 @@ theorem contMDiffOn_iff_source_of_mem_maximalAtlas (he : e ∈ maximalAtlas I M) -- Porting note: didn't compile; fixed by golfing the proof and moving parts to lemmas /-- A function is `C^n` within a set at a point, for `n : ℕ`, if and only if it is `C^n` on a neighborhood of this point. -/ -theorem contMDiffWithinAt_iff_contMDiffOn_nhds {n : ℕ} : +theorem contMDiffWithinAt_iff_contMDiffOn_nhds + [SmoothManifoldWithCorners I M] [SmoothManifoldWithCorners I' M'] {n : ℕ} : ContMDiffWithinAt I I' n f s x ↔ ∃ u ∈ 𝓝[insert x s] x, ContMDiffOn I I' n f u := by -- WLOG, `x ∈ s`, otherwise we add `x` to `s` wlog hxs : x ∈ s generalizing s @@ -747,13 +756,15 @@ theorem contMDiffWithinAt_iff_contMDiffOn_nhds {n : ℕ} : /-- A function is `C^n` at a point, for `n : ℕ`, if and only if it is `C^n` on a neighborhood of this point. -/ -theorem contMDiffAt_iff_contMDiffOn_nhds {n : ℕ} : +theorem contMDiffAt_iff_contMDiffOn_nhds + [SmoothManifoldWithCorners I M] [SmoothManifoldWithCorners I' M'] {n : ℕ} : ContMDiffAt I I' n f x ↔ ∃ u ∈ 𝓝 x, ContMDiffOn I I' n f u := by simp [← contMDiffWithinAt_univ, contMDiffWithinAt_iff_contMDiffOn_nhds, nhdsWithin_univ] /-- Note: This does not hold for `n = ∞`. `f` being `C^∞` at `x` means that for every `n`, `f` is `C^n` on some neighborhood of `x`, but this neighborhood can depend on `n`. -/ -theorem contMDiffAt_iff_contMDiffAt_nhds {n : ℕ} : +theorem contMDiffAt_iff_contMDiffAt_nhds + [SmoothManifoldWithCorners I M] [SmoothManifoldWithCorners I' M'] {n : ℕ} : ContMDiffAt I I' n f x ↔ ∀ᶠ x' in 𝓝 x, ContMDiffAt I I' n f x' := by refine ⟨?_, fun h => h.self_of_nhds⟩ rw [contMDiffAt_iff_contMDiffOn_nhds] diff --git a/Mathlib/LinearAlgebra/FreeModule/PID.lean b/Mathlib/LinearAlgebra/FreeModule/PID.lean index d15998c9d710a..4d837a3b56384 100644 --- a/Mathlib/LinearAlgebra/FreeModule/PID.lean +++ b/Mathlib/LinearAlgebra/FreeModule/PID.lean @@ -99,9 +99,13 @@ section PrincipalIdealDomain open Submodule.IsPrincipal Set Submodule -variable {ι : Type*} {R : Type*} [CommRing R] [IsDomain R] [IsPrincipalIdealRing R] +variable {ι : Type*} {R : Type*} [CommRing R] variable {M : Type*} [AddCommGroup M] [Module R M] {b : ι → M} +section StrongRankCondition + +variable [IsDomain R] [IsPrincipalIdealRing R] + open Submodule.IsPrincipal theorem generator_maximal_submoduleImage_dvd {N O : Submodule R M} (hNO : N ≤ O) {ϕ : O →ₗ[R] R} @@ -401,6 +405,8 @@ theorem Module.free_iff_noZeroSMulDivisors [Module.Finite R M] : Module.Free R M ↔ NoZeroSMulDivisors R M := ⟨fun _ ↦ inferInstance, fun _ ↦ inferInstance⟩ +end StrongRankCondition + section SmithNormal /-- A Smith normal form basis for a submodule `N` of a module `M` consists of @@ -475,6 +481,8 @@ lemma toMatrix_restrict_eq_toMatrix [Fintype ι] [DecidableEq ι] end Basis.SmithNormalForm +variable [IsDomain R] [IsPrincipalIdealRing R] + /-- If `M` is finite free over a PID `R`, then any submodule `N` is free and we can find a basis for `M` and `N` such that the inclusion map is a diagonal matrix in Smith normal form. diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean index b8af1ec61709a..64465d3ce42e3 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/External.lean @@ -56,8 +56,6 @@ variable (𝒜 : ι → Type*) (ℬ : ι → Type*) variable [CommRing R] variable [∀ i, AddCommGroup (𝒜 i)] [∀ i, AddCommGroup (ℬ i)] variable [∀ i, Module R (𝒜 i)] [∀ i, Module R (ℬ i)] -variable [DirectSum.GRing 𝒜] [DirectSum.GRing ℬ] -variable [DirectSum.GAlgebra R 𝒜] [DirectSum.GAlgebra R ℬ] -- this helps with performance instance (i : ι × ι) : Module R (𝒜 (Prod.fst i) ⊗[R] ℬ (Prod.snd i)) := @@ -144,30 +142,38 @@ theorem gradedComm_of_zero_tmul (a : 𝒜 0) (b : ⨁ i, ℬ i) : dsimp rw [gradedComm_of_tmul_of, mul_zero, uzpow_zero, one_smul] -theorem gradedComm_tmul_one (a : ⨁ i, 𝒜 i) : gradedComm R 𝒜 ℬ (a ⊗ₜ 1) = 1 ⊗ₜ a := +theorem gradedComm_tmul_one [DirectSum.GRing ℬ] (a : ⨁ i, 𝒜 i) : + gradedComm R 𝒜 ℬ (a ⊗ₜ 1) = 1 ⊗ₜ a := gradedComm_tmul_of_zero _ _ _ _ _ -theorem gradedComm_one_tmul (b : ⨁ i, ℬ i) : gradedComm R 𝒜 ℬ (1 ⊗ₜ b) = b ⊗ₜ 1 := +theorem gradedComm_one_tmul [DirectSum.GRing 𝒜] (b : ⨁ i, ℬ i) : + gradedComm R 𝒜 ℬ (1 ⊗ₜ b) = b ⊗ₜ 1 := gradedComm_of_zero_tmul _ _ _ _ _ @[simp, nolint simpNF] -- linter times out -theorem gradedComm_one : gradedComm R 𝒜 ℬ 1 = 1 := +theorem gradedComm_one [DirectSum.GRing 𝒜] [DirectSum.GRing ℬ] : gradedComm R 𝒜 ℬ 1 = 1 := gradedComm_one_tmul _ _ _ _ -theorem gradedComm_tmul_algebraMap (a : ⨁ i, 𝒜 i) (r : R) : +theorem gradedComm_tmul_algebraMap [DirectSum.GRing ℬ] [DirectSum.GAlgebra R ℬ] + (a : ⨁ i, 𝒜 i) (r : R) : gradedComm R 𝒜 ℬ (a ⊗ₜ algebraMap R _ r) = algebraMap R _ r ⊗ₜ a := gradedComm_tmul_of_zero _ _ _ _ _ -theorem gradedComm_algebraMap_tmul (r : R) (b : ⨁ i, ℬ i) : +theorem gradedComm_algebraMap_tmul [DirectSum.GRing 𝒜] [DirectSum.GAlgebra R 𝒜] + (r : R) (b : ⨁ i, ℬ i) : gradedComm R 𝒜 ℬ (algebraMap R _ r ⊗ₜ b) = b ⊗ₜ algebraMap R _ r := gradedComm_of_zero_tmul _ _ _ _ _ -theorem gradedComm_algebraMap (r : R) : +theorem gradedComm_algebraMap [DirectSum.GRing 𝒜] [DirectSum.GRing ℬ] + [DirectSum.GAlgebra R 𝒜] [DirectSum.GAlgebra R ℬ] (r : R) : gradedComm R 𝒜 ℬ (algebraMap R _ r) = algebraMap R _ r := (gradedComm_algebraMap_tmul R 𝒜 ℬ r 1).trans (Algebra.TensorProduct.algebraMap_apply' r).symm end gradedComm +variable [DirectSum.GRing 𝒜] [DirectSum.GRing ℬ] +variable [DirectSum.GAlgebra R 𝒜] [DirectSum.GAlgebra R ℬ] + open TensorProduct (assoc map) in /-- The multiplication operation for tensor products of externally `ι`-graded algebras. -/ noncomputable irreducible_def gradedMul : diff --git a/Mathlib/MeasureTheory/Measure/Regular.lean b/Mathlib/MeasureTheory/Measure/Regular.lean index c983a4d2caa82..da222d3e2a93c 100644 --- a/Mathlib/MeasureTheory/Measure/Regular.lean +++ b/Mathlib/MeasureTheory/Measure/Regular.lean @@ -269,7 +269,11 @@ theorem mono {p' q' : Set α → Prop} (H : InnerRegularWRT μ p q) end InnerRegularWRT -variable {α β : Type*} [MeasurableSpace α] [TopologicalSpace α] {μ : Measure α} +variable {α β : Type*} [MeasurableSpace α] {μ : Measure α} + +section Classes + +variable [TopologicalSpace α] /-- A measure `μ` is outer regular if `μ(A) = inf {μ(U) | A ⊆ U open}` for a measurable set `A`. @@ -315,8 +319,12 @@ instance (priority := 100) Regular.weaklyRegular [R1Space α] [Regular μ] : ⟨closure K, K_comp.closure_subset_of_isOpen hU KU, isClosed_closure, hK.trans_le (measure_mono subset_closure)⟩ +end Classes + namespace OuterRegular +variable [TopologicalSpace α] + instance zero : OuterRegular (0 : Measure α) := ⟨fun A _ _r hr => ⟨univ, subset_univ A, isOpen_univ, hr⟩⟩ @@ -429,7 +437,8 @@ end OuterRegular /-- If a measure `μ` admits finite spanning open sets such that the restriction of `μ` to each set is outer regular, then the original measure is outer regular as well. -/ -protected theorem FiniteSpanningSetsIn.outerRegular [OpensMeasurableSpace α] {μ : Measure α} +protected theorem FiniteSpanningSetsIn.outerRegular + [TopologicalSpace α] [OpensMeasurableSpace α] {μ : Measure α} (s : μ.FiniteSpanningSetsIn { U | IsOpen U ∧ OuterRegular (μ.restrict U) }) : OuterRegular μ := OuterRegular.of_restrict (s := fun n ↦ s.set n) (fun n ↦ (s.set_mem n).2) @@ -439,6 +448,76 @@ namespace InnerRegularWRT variable {p q : Set α → Prop} {U s : Set α} {ε r : ℝ≥0∞} +/-- 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 +inner regular. -/ +lemma of_restrict {μ : Measure α} {s : ℕ → Set α} + (h : ∀ n, InnerRegularWRT (μ.restrict (s n)) p MeasurableSet) + (hs : univ ⊆ ⋃ n, s n) (hmono : Monotone s) : InnerRegularWRT μ p MeasurableSet := by + intro F hF r hr + have hBU : ⋃ n, F ∩ s n = F := by rw [← inter_iUnion, univ_subset_iff.mp hs, inter_univ] + have : μ F = ⨆ n, μ (F ∩ s n) := by + rw [← measure_iUnion_eq_iSup, hBU] + exact Monotone.directed_le fun m n h ↦ inter_subset_inter_right _ (hmono h) + rw [this] at hr + rcases lt_iSup_iff.1 hr with ⟨n, hn⟩ + rw [← restrict_apply hF] at hn + rcases h n hF _ hn with ⟨K, KF, hKp, hK⟩ + exact ⟨K, KF, hKp, hK.trans_le (restrict_apply_le _ _)⟩ + +/-- If `μ` is inner regular for measurable finite measure sets with respect to some class of sets, +then its restriction to any set is also inner regular for measurable finite measure sets, with +respect to the same class of sets. -/ +lemma restrict (h : InnerRegularWRT μ p (fun s ↦ MeasurableSet s ∧ μ s ≠ ∞)) (A : Set α) : + InnerRegularWRT (μ.restrict A) p (fun s ↦ MeasurableSet s ∧ μ.restrict A s ≠ ∞) := by + rintro s ⟨s_meas, hs⟩ r hr + rw [restrict_apply s_meas] at hs + obtain ⟨K, K_subs, pK, rK⟩ : ∃ K, K ⊆ (toMeasurable μ (s ∩ A)) ∩ s ∧ p K ∧ r < μ K := by + have : r < μ ((toMeasurable μ (s ∩ A)) ∩ s) := by + apply hr.trans_le + rw [restrict_apply s_meas] + exact measure_mono <| subset_inter (subset_toMeasurable μ (s ∩ A)) inter_subset_left + refine h ⟨(measurableSet_toMeasurable _ _).inter s_meas, ?_⟩ _ this + apply (lt_of_le_of_lt _ hs.lt_top).ne + rw [← measure_toMeasurable (s ∩ A)] + exact measure_mono inter_subset_left + refine ⟨K, K_subs.trans inter_subset_right, pK, ?_⟩ + calc + r < μ K := rK + _ = μ.restrict (toMeasurable μ (s ∩ A)) K := by + rw [restrict_apply' (measurableSet_toMeasurable μ (s ∩ A))] + congr + apply (inter_eq_left.2 ?_).symm + exact K_subs.trans inter_subset_left + _ = μ.restrict (s ∩ A) K := by rwa [restrict_toMeasurable] + _ ≤ μ.restrict A K := Measure.le_iff'.1 (restrict_mono inter_subset_right le_rfl) K + +/-- If `μ` is inner regular for measurable finite measure sets with respect to some class of sets, +then its restriction to any finite measure set is also inner regular for measurable sets with +respect to the same class of sets. -/ +lemma restrict_of_measure_ne_top (h : InnerRegularWRT μ p (fun s ↦ MeasurableSet s ∧ μ s ≠ ∞)) + {A : Set α} (hA : μ A ≠ ∞) : + InnerRegularWRT (μ.restrict A) p (fun s ↦ MeasurableSet s) := by + have : Fact (μ A < ∞) := ⟨hA.lt_top⟩ + exact (restrict h A).trans (of_imp (fun s hs ↦ ⟨hs, measure_ne_top _ _⟩)) + +/-- Given a σ-finite measure, any measurable set can be approximated from inside by a measurable +set of finite measure. -/ +lemma of_sigmaFinite [SigmaFinite μ] : + InnerRegularWRT μ (fun s ↦ MeasurableSet s ∧ μ s ≠ ∞) (fun s ↦ MeasurableSet s) := by + intro s hs r hr + set B : ℕ → Set α := spanningSets μ + have hBU : ⋃ n, s ∩ B n = s := by rw [← inter_iUnion, iUnion_spanningSets, inter_univ] + have : μ s = ⨆ n, μ (s ∩ B n) := by + rw [← measure_iUnion_eq_iSup, hBU] + exact Monotone.directed_le fun m n h => inter_subset_inter_right _ (monotone_spanningSets μ h) + rw [this] at hr + rcases lt_iSup_iff.1 hr with ⟨n, hn⟩ + refine ⟨s ∩ B n, inter_subset_left, ⟨hs.inter (measurable_spanningSets μ n), ?_⟩, hn⟩ + exact ((measure_mono inter_subset_right).trans_lt (measure_spanningSets_lt_top μ n)).ne + +variable [TopologicalSpace α] + /-- If a measure is inner regular (using closed or compact sets) for open sets, then every measurable set of finite measure can be approximated by a (closed or compact) subset. -/ theorem measurableSet_of_isOpen [OuterRegular μ] (H : InnerRegularWRT μ p IsOpen) @@ -531,23 +610,6 @@ theorem weaklyRegular_of_finite [BorelSpace α] (μ : Measure α) [IsFiniteMeasu _ = μ (⋃ n, s n) + ∑' n, δ n := by rw [measure_iUnion hsd hsm, ENNReal.tsum_add] _ ≤ μ (⋃ n, s n) + ε := add_le_add_left (hδε.le.trans ENNReal.half_le_self) _ -/-- 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 -inner regular. -/ -lemma of_restrict {μ : Measure α} {s : ℕ → Set α} - (h : ∀ n, InnerRegularWRT (μ.restrict (s n)) p MeasurableSet) - (hs : univ ⊆ ⋃ n, s n) (hmono : Monotone s) : InnerRegularWRT μ p MeasurableSet := by - intro F hF r hr - have hBU : ⋃ n, F ∩ s n = F := by rw [← inter_iUnion, univ_subset_iff.mp hs, inter_univ] - have : μ F = ⨆ n, μ (F ∩ s n) := by - rw [← measure_iUnion_eq_iSup, hBU] - exact Monotone.directed_le fun m n h ↦ inter_subset_inter_right _ (hmono h) - rw [this] at hr - rcases lt_iSup_iff.1 hr with ⟨n, hn⟩ - rw [← restrict_apply hF] at hn - rcases h n hF _ hn with ⟨K, KF, hKp, hK⟩ - exact ⟨K, KF, hKp, hK.trans_le (restrict_apply_le _ _)⟩ - /-- In a metrizable space (or even a pseudo metrizable space), an open set can be approximated from inside by closed sets. -/ theorem of_pseudoMetrizableSpace {X : Type*} [TopologicalSpace X] [PseudoMetrizableSpace X] @@ -573,62 +635,11 @@ theorem isCompact_isClosed {X : Type*} [TopologicalSpace X] [SigmaCompactSpace X rcases lt_iSup_iff.1 hr with ⟨n, hn⟩ exact ⟨_, inter_subset_left, hBc n, hn⟩ -/-- If `μ` is inner regular for measurable finite measure sets with respect to some class of sets, -then its restriction to any set is also inner regular for measurable finite measure sets, with -respect to the same class of sets. -/ -lemma restrict (h : InnerRegularWRT μ p (fun s ↦ MeasurableSet s ∧ μ s ≠ ∞)) (A : Set α) : - InnerRegularWRT (μ.restrict A) p (fun s ↦ MeasurableSet s ∧ μ.restrict A s ≠ ∞) := by - rintro s ⟨s_meas, hs⟩ r hr - rw [restrict_apply s_meas] at hs - obtain ⟨K, K_subs, pK, rK⟩ : ∃ K, K ⊆ (toMeasurable μ (s ∩ A)) ∩ s ∧ p K ∧ r < μ K := by - have : r < μ ((toMeasurable μ (s ∩ A)) ∩ s) := by - apply hr.trans_le - rw [restrict_apply s_meas] - exact measure_mono <| subset_inter (subset_toMeasurable μ (s ∩ A)) inter_subset_left - refine h ⟨(measurableSet_toMeasurable _ _).inter s_meas, ?_⟩ _ this - apply (lt_of_le_of_lt _ hs.lt_top).ne - rw [← measure_toMeasurable (s ∩ A)] - exact measure_mono inter_subset_left - refine ⟨K, K_subs.trans inter_subset_right, pK, ?_⟩ - calc - r < μ K := rK - _ = μ.restrict (toMeasurable μ (s ∩ A)) K := by - rw [restrict_apply' (measurableSet_toMeasurable μ (s ∩ A))] - congr - apply (inter_eq_left.2 ?_).symm - exact K_subs.trans inter_subset_left - _ = μ.restrict (s ∩ A) K := by rwa [restrict_toMeasurable] - _ ≤ μ.restrict A K := Measure.le_iff'.1 (restrict_mono inter_subset_right le_rfl) K - -/-- If `μ` is inner regular for measurable finite measure sets with respect to some class of sets, -then its restriction to any finite measure set is also inner regular for measurable sets with -respect to the same class of sets. -/ -lemma restrict_of_measure_ne_top (h : InnerRegularWRT μ p (fun s ↦ MeasurableSet s ∧ μ s ≠ ∞)) - {A : Set α} (hA : μ A ≠ ∞) : - InnerRegularWRT (μ.restrict A) p (fun s ↦ MeasurableSet s) := by - have : Fact (μ A < ∞) := ⟨hA.lt_top⟩ - exact (restrict h A).trans (of_imp (fun s hs ↦ ⟨hs, measure_ne_top _ _⟩)) - -/-- Given a σ-finite measure, any measurable set can be approximated from inside by a measurable -set of finite measure. -/ -lemma of_sigmaFinite [SigmaFinite μ] : - InnerRegularWRT μ (fun s ↦ MeasurableSet s ∧ μ s ≠ ∞) (fun s ↦ MeasurableSet s) := by - intro s hs r hr - set B : ℕ → Set α := spanningSets μ - have hBU : ⋃ n, s ∩ B n = s := by rw [← inter_iUnion, iUnion_spanningSets, inter_univ] - have : μ s = ⨆ n, μ (s ∩ B n) := by - rw [← measure_iUnion_eq_iSup, hBU] - exact Monotone.directed_le fun m n h => inter_subset_inter_right _ (monotone_spanningSets μ h) - rw [this] at hr - rcases lt_iSup_iff.1 hr with ⟨n, hn⟩ - refine ⟨s ∩ B n, inter_subset_left, ⟨hs.inter (measurable_spanningSets μ n), ?_⟩, hn⟩ - exact ((measure_mono inter_subset_right).trans_lt (measure_spanningSets_lt_top μ n)).ne - end InnerRegularWRT namespace InnerRegular -variable {U : Set α} {ε : ℝ≥0∞} +variable {U : Set α} {ε : ℝ≥0∞} [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) @@ -687,6 +698,8 @@ end InnerRegular namespace InnerRegularCompactLTTop +variable [TopologicalSpace α] + /-- If `μ` is inner regular for finite measure sets with respect to compact sets, then any measurable set of finite measure can be approximated by a compact subset. See also `MeasurableSet.exists_lt_isCompact_of_ne_top`. -/ @@ -861,6 +874,8 @@ end InnerRegularCompactLTTop namespace WeaklyRegular +variable [TopologicalSpace α] + instance zero : WeaklyRegular (0 : Measure α) := ⟨fun _ _ _r hr => ⟨∅, empty_subset _, isClosed_empty, hr⟩⟩ @@ -947,6 +962,8 @@ end WeaklyRegular namespace Regular +variable [TopologicalSpace α] + instance zero : Regular (0 : Measure α) := ⟨fun _ _ _r hr => ⟨∅, empty_subset _, isCompact_empty, hr⟩⟩ diff --git a/Mathlib/RingTheory/Generators.lean b/Mathlib/RingTheory/Generators.lean index c2204df972a59..d55ec361f5c74 100644 --- a/Mathlib/RingTheory/Generators.lean +++ b/Mathlib/RingTheory/Generators.lean @@ -215,17 +215,15 @@ def baseChange {T} [CommRing T] [Algebra R T] (P : Generators R S) : Generators end Construction variable {R' S'} [CommRing R'] [CommRing S'] [Algebra R' S'] (P' : Generators R' S') -variable [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] - variable {R'' S''} [CommRing R''] [CommRing S''] [Algebra R'' S''] (P'' : Generators R'' S'') -variable [Algebra R R''] [Algebra S S''] [Algebra R S''] - [IsScalarTower R R'' S''] [IsScalarTower R S S''] -variable [Algebra R' R''] [Algebra S' S''] [Algebra R' S''] - [IsScalarTower R' R'' S''] [IsScalarTower R' S' S''] -variable [IsScalarTower R R' R''] [IsScalarTower S S' S''] section Hom +section + +variable [Algebra R R'] [Algebra R' R''] [Algebra R' S''] +variable [Algebra S S'] [Algebra S' S''] [Algebra S S''] + /-- Given a commuting square R --→ P = R[X] ---→ S | | @@ -249,6 +247,7 @@ an algebra homomorphism between the polynomial rings. -/ noncomputable def Hom.toAlgHom (f : Hom P P') : P.Ring →ₐ[R] P'.Ring := MvPolynomial.aeval f.val +variable [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] in @[simp] lemma Hom.algebraMap_toAlgHom (f : Hom P P') (x) : MvPolynomial.aeval P'.val (f.toAlgHom x) = algebraMap S S' (MvPolynomial.aeval P.val x) := by @@ -270,6 +269,7 @@ lemma Hom.toAlgHom_monomial (f : Generators.Hom P P') (v r) : f.toAlgHom (monomial v r) = r • v.prod (f.val · ^ ·) := by rw [toAlgHom, aeval_monomial, Algebra.smul_def] +variable [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] in /-- Giving a hom between two families of generators is equivalent to giving an algebra homomorphism between the polynomial rings. -/ @[simps] @@ -301,7 +301,8 @@ variable {P P' P''} /-- The composition of two homs. -/ @[simps] -noncomputable def Hom.comp (f : Hom P' P'') (g : Hom P P') : Hom P P'' where +noncomputable def Hom.comp [IsScalarTower R' R'' S''] [IsScalarTower R' S' S''] + [IsScalarTower S S' S''] (f : Hom P' P'') (g : Hom P P') : Hom P P'' where val x := aeval f.val (g.val x) aeval_val x := by simp only @@ -312,13 +313,23 @@ noncomputable def Hom.comp (f : Hom P' P'') (g : Hom P P') : Hom P P'' where | h_X p i hp => simp only [_root_.map_mul, hp, aeval_X, aeval_val] @[simp] -lemma Hom.comp_id (f : Hom P P') : f.comp (Hom.id P) = f := by ext; simp +lemma Hom.comp_id [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] (f : Hom P P') : + f.comp (Hom.id P) = f := by ext; simp + +end @[simp] -lemma Hom.id_comp (f : Hom P P') : (Hom.id P').comp f = f := by ext; simp [Hom.id, aeval_X_left] +lemma Hom.id_comp [Algebra S S'] (f : Hom P P') : (Hom.id P').comp f = f := by + ext; simp [Hom.id, aeval_X_left] + +variable [Algebra R R'] [Algebra R' R''] [Algebra R' S''] +variable [Algebra S S'] [Algebra S' S''] [Algebra S S''] @[simp] -lemma Hom.toAlgHom_comp_apply (f : Hom P P') (g : Hom P' P'') (x) : +lemma Hom.toAlgHom_comp_apply + [Algebra R R''] [IsScalarTower R R' R''] [IsScalarTower R' R'' S''] + [IsScalarTower R' S' S''] [IsScalarTower S S' S''] + (f : Hom P P') (g : Hom P' P'') (x) : (g.comp f).toAlgHom x = g.toAlgHom (f.toAlgHom x) := by induction x using MvPolynomial.induction_on with | h_C r => simp only [← MvPolynomial.algebraMap_eq, AlgHom.map_algebraMap] @@ -457,6 +468,9 @@ lemma Cotangent.mk_surjective : Function.Surjective (mk (P := P)) := fun x ↦ Ideal.toCotangent_surjective P.ker x.val variable {P'} +variable [Algebra R R'] [Algebra R' R''] [Algebra R' S''] +variable [Algebra S S'] [Algebra S' S''] [Algebra S S''] +variable [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] /-- A hom between families of generators induce a map between cotangent spaces. -/ noncomputable @@ -490,6 +504,10 @@ lemma Cotangent.map_id : simp only [map_mk, Hom.toAlgHom_id, AlgHom.coe_id, id_eq, Subtype.coe_eta, val_mk, LinearMap.id_coe] +variable [Algebra R R''] [IsScalarTower R R' R''] [IsScalarTower R' R'' S''] + [IsScalarTower R' S' S''] [IsScalarTower S S' S''] + [Algebra R S''] [IsScalarTower R R'' S''] [IsScalarTower R S S''] + lemma Cotangent.map_comp (f : Hom P P') (g : Hom P' P'') : Cotangent.map (g.comp f) = (map g).restrictScalars S ∘ₗ map f := by ext x diff --git a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean index e9b925a2e5ed6..5d1767881cad0 100644 --- a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean @@ -52,7 +52,11 @@ open Finset Nat namespace NewtonIdentities -variable (σ : Type*) [DecidableEq σ] (R : Type*) [CommRing R] +variable (σ : Type*) (R : Type*) [CommRing R] + +section DecidableEq + +variable [DecidableEq σ] private def pairMap (t : Finset σ × σ) : Finset σ × σ := if h : t.snd ∈ t.fst then (t.fst.erase t.snd, t.snd) else (t.fst.cons t.snd h, t.snd) @@ -187,6 +191,10 @@ private theorem disjUnion_filter_pairs_eq_pairs (k : ℕ) : have hacard := le_iff_lt_or_eq.mp ha.2.1 tauto +end DecidableEq + +variable [Fintype σ] + private theorem esymm_summand_to_weight (k : ℕ) (A : Finset σ) (h : A ∈ powersetCard k univ) : ∑ j ∈ A, weight σ R k (A, j) = k * (-1) ^ k * (∏ i ∈ A, X i : MvPolynomial σ R) := by simp [weight, mem_powersetCard_univ.mp h, mul_assoc] From 15720c78a8142f7d35418f15b528bdb3f184989a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 8 Aug 2024 06:02:38 +0000 Subject: [PATCH 107/359] chore: backports for leanprover/lean4#4814 (part 20) (#15440) see https://github.com/leanprover-community/mathlib4/pull/15245 --- .../HomotopyCategory/Pretriangulated.lean | 9 +- .../HomotopyCategory/Triangulated.lean | 4 +- Mathlib/Algebra/Module/Torsion.lean | 14 ++- .../Analysis/Normed/Group/CocompactMap.lean | 12 +- .../Normed/Operator/LinearIsometry.lean | 8 -- Mathlib/Analysis/Oscillation.lean | 6 +- Mathlib/CategoryTheory/Sites/Discrete.lean | 18 +-- .../CategoryTheory/Sites/InducedTopology.lean | 4 +- Mathlib/GroupTheory/Frattini.lean | 5 +- .../Dimension/Constructions.lean | 14 +-- Mathlib/LinearAlgebra/Dimension/Finite.lean | 53 ++++++--- .../HomogeneousLocalization.lean | 112 +++++++++--------- .../IntegralClosure/IsIntegral/Basic.lean | 14 +-- .../RingTheory/Valuation/ValuationRing.lean | 8 +- Mathlib/Topology/Algebra/Group/Basic.lean | 3 +- Mathlib/Topology/Instances/Discrete.lean | 4 +- .../Sheaves/SheafCondition/Sites.lean | 10 +- .../UniformConvergenceTopology.lean | 2 +- 18 files changed, 168 insertions(+), 132 deletions(-) diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean index 7f171cc03ba0c..b6a0a567c16d5 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean @@ -33,8 +33,8 @@ complexes. Here, we follow the original definitions in [Verdiers's thesis, I.3][ open CategoryTheory Category Limits CochainComplex.HomComplex Pretriangulated variable {C D : Type*} [Category C] [Category D] - [Preadditive C] [HasZeroObject C] [HasBinaryBiproducts C] - [Preadditive D] [HasZeroObject D] [HasBinaryBiproducts D] + [Preadditive C] [HasBinaryBiproducts C] + [Preadditive D] [HasBinaryBiproducts D] {K L : CochainComplex C ℤ} (φ : K ⟶ L) namespace CochainComplex @@ -420,6 +420,7 @@ lemma isomorphic_distinguished (T₁ : Triangle (HomotopyCategory C (ComplexShap obtain ⟨X, Y, f, ⟨e'⟩⟩ := hT₁ exact ⟨X, Y, f, ⟨e ≪≫ e'⟩⟩ +variable [HasZeroObject C] in lemma contractible_distinguished (X : HomotopyCategory C (ComplexShape.up ℤ)) : Pretriangulated.contractibleTriangle X ∈ distinguishedTriangles C := by obtain ⟨X⟩ := X @@ -491,6 +492,8 @@ lemma complete_distinguished_triangle_morphism end Pretriangulated +variable [HasZeroObject C] + instance : Pretriangulated (HomotopyCategory C (ComplexShape.up ℤ)) where distinguishedTriangles := Pretriangulated.distinguishedTriangles C isomorphic_distinguished := Pretriangulated.isomorphic_distinguished @@ -506,6 +509,8 @@ lemma mappingCone_triangleh_distinguished {X Y : CochainComplex C ℤ} (f : X CochainComplex.mappingCone.triangleh f ∈ distTriang (HomotopyCategory _ _) := ⟨_, _, f, ⟨Iso.refl _⟩⟩ +variable [HasZeroObject D] + instance (G : C ⥤ D) [G.Additive] : (G.mapHomotopyCategory (ComplexShape.up ℤ)).IsTriangulated where map_distinguished := by diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean index 58ad6bfd825ff..9c7b600049093 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean @@ -16,7 +16,7 @@ the pretriangulated category `HomotopyCategory C (ComplexShape.up ℤ)` is trian open CategoryTheory Category Limits Pretriangulated ComposableArrows -variable {C : Type*} [Category C] [Preadditive C] [HasZeroObject C] [HasBinaryBiproducts C] +variable {C : Type*} [Category C] [Preadditive C] [HasBinaryBiproducts C] {X₁ X₂ X₃ : CochainComplex C ℤ} (f : X₁ ⟶ X₂) (g : X₂ ⟶ X₃) namespace CochainComplex @@ -176,6 +176,8 @@ end CochainComplex namespace HomotopyCategory +variable [HasZeroObject C] + lemma mappingConeCompTriangleh_distinguished : (CochainComplex.mappingConeCompTriangleh f g) ∈ distTriang (HomotopyCategory C (ComplexShape.up ℤ)) := by diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index 4e7d09b893f8f..e0b05d0b86998 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -357,10 +357,10 @@ variable {R M} section Coprime variable {ι : Type*} {p : ι → Ideal R} {S : Finset ι} -variable (hp : (S : Set ι).Pairwise fun i j => p i ⊔ p j = ⊤) -- Porting note: mem_iSup_finset_iff_exists_sum now requires DecidableEq ι -theorem iSup_torsionBySet_ideal_eq_torsionBySet_iInf : +theorem iSup_torsionBySet_ideal_eq_torsionBySet_iInf + (hp : (S : Set ι).Pairwise fun i j => p i ⊔ p j = ⊤) : ⨆ i ∈ S, torsionBySet R M (p i) = torsionBySet R M ↑(⨅ i ∈ S, p i) := by rcases S.eq_empty_or_nonempty with h | h · simp only [h] @@ -396,7 +396,8 @@ theorem iSup_torsionBySet_ideal_eq_torsionBySet_iInf : · rw [← Finset.sum_smul, hμ, one_smul] -- Porting note: iSup_torsionBySet_ideal_eq_torsionBySet_iInf now requires DecidableEq ι -theorem supIndep_torsionBySet_ideal : S.SupIndep fun i => torsionBySet R M <| p i := +theorem supIndep_torsionBySet_ideal (hp : (S : Set ι).Pairwise fun i j => p i ⊔ p j = ⊤) : + S.SupIndep fun i => torsionBySet R M <| p i := fun T hT i hi hiT => by rw [disjoint_iff, Finset.sup_eq_iSup, iSup_torsionBySet_ideal_eq_torsionBySet_iInf fun i hi j hj ij => hp (hT hi) (hT hj) ij] @@ -406,9 +407,9 @@ theorem supIndep_torsionBySet_ideal : S.SupIndep fun i => torsionBySet R M <| p rw [← this, Ideal.sup_iInf_eq_top, top_coe, torsionBySet_univ] intro j hj; apply hp hi (hT hj); rintro rfl; exact hiT hj -variable {q : ι → R} (hq : (S : Set ι).Pairwise <| (IsCoprime on q)) +variable {q : ι → R} -theorem iSup_torsionBy_eq_torsionBy_prod : +theorem iSup_torsionBy_eq_torsionBy_prod (hq : (S : Set ι).Pairwise <| (IsCoprime on q)) : ⨆ i ∈ S, torsionBy R M (q i) = torsionBy R M (∏ i ∈ S, q i) := by rw [← torsionBySet_span_singleton_eq, Ideal.submodule_span_eq, ← Ideal.finset_inf_span_singleton _ _ hq, Finset.inf_eq_iInf, ← @@ -420,7 +421,8 @@ theorem iSup_torsionBy_eq_torsionBy_prod : exact (torsionBySet_span_singleton_eq _).symm exact fun i hi j hj ij => (Ideal.sup_eq_top_iff_isCoprime _ _).mpr (hq hi hj ij) -theorem supIndep_torsionBy : S.SupIndep fun i => torsionBy R M <| q i := by +theorem supIndep_torsionBy (hq : (S : Set ι).Pairwise <| (IsCoprime on q)) : + S.SupIndep fun i => torsionBy R M <| q i := by convert supIndep_torsionBySet_ideal (M := M) fun i hi j hj ij => (Ideal.sup_eq_top_iff_isCoprime (q i) _).mpr <| hq hi hj ij exact (torsionBySet_span_singleton_eq (R := R) (M := M) _).symm diff --git a/Mathlib/Analysis/Normed/Group/CocompactMap.lean b/Mathlib/Analysis/Normed/Group/CocompactMap.lean index 0a85b81d0c3be..ff947f47393ef 100644 --- a/Mathlib/Analysis/Normed/Group/CocompactMap.lean +++ b/Mathlib/Analysis/Normed/Group/CocompactMap.lean @@ -24,11 +24,11 @@ map is cocompact. open Filter Metric variable {𝕜 E F 𝓕 : Type*} -variable [NormedAddCommGroup E] [NormedAddCommGroup F] [ProperSpace E] [ProperSpace F] +variable [NormedAddCommGroup E] [NormedAddCommGroup F] variable {f : 𝓕} -theorem CocompactMapClass.norm_le [FunLike 𝓕 E F] [CocompactMapClass 𝓕 E F] (ε : ℝ) : - ∃ r : ℝ, ∀ x : E, r < ‖x‖ → ε < ‖f x‖ := by +theorem CocompactMapClass.norm_le [ProperSpace F] [FunLike 𝓕 E F] [CocompactMapClass 𝓕 E F] + (ε : ℝ) : ∃ r : ℝ, ∀ x : E, r < ‖x‖ → ε < ‖f x‖ := by have h := cocompact_tendsto f rw [tendsto_def] at h specialize h (Metric.closedBall 0 ε)ᶜ (mem_cocompact_of_closedBall_compl_subset 0 ⟨ε, rfl.subset⟩) @@ -39,7 +39,7 @@ theorem CocompactMapClass.norm_le [FunLike 𝓕 E F] [CocompactMapClass 𝓕 E F apply hr simp [hx] -theorem Filter.tendsto_cocompact_cocompact_of_norm {f : E → F} +theorem Filter.tendsto_cocompact_cocompact_of_norm [ProperSpace E] {f : E → F} (h : ∀ ε : ℝ, ∃ r : ℝ, ∀ x : E, r < ‖x‖ → ε < ‖f x‖) : Tendsto f (cocompact E) (cocompact F) := by rw [tendsto_def] @@ -53,7 +53,7 @@ theorem Filter.tendsto_cocompact_cocompact_of_norm {f : E → F} apply hε simp [hr x hx] -theorem ContinuousMapClass.toCocompactMapClass_of_norm [FunLike 𝓕 E F] [ContinuousMapClass 𝓕 E F] - (h : ∀ (f : 𝓕) (ε : ℝ), ∃ r : ℝ, ∀ x : E, r < ‖x‖ → ε < ‖f x‖) : +theorem ContinuousMapClass.toCocompactMapClass_of_norm [ProperSpace E] [FunLike 𝓕 E F] + [ContinuousMapClass 𝓕 E F] (h : ∀ (f : 𝓕) (ε : ℝ), ∃ r : ℝ, ∀ x : E, r < ‖x‖ → ε < ‖f x‖) : CocompactMapClass 𝓕 E F where cocompact_tendsto := (tendsto_cocompact_cocompact_of_norm <| h ·) diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index 087c79c8eb84c..6c8da81a65094 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -170,14 +170,6 @@ theorem ext {f g : E →ₛₗᵢ[σ₁₂] E₂} (h : ∀ x, f x = g x) : f = g variable [FunLike 𝓕 E E₂] -protected theorem congr_arg {f : 𝓕} : - ∀ {x x' : E}, x = x' → f x = f x' - | _, _, rfl => rfl - -protected theorem congr_fun {f g : 𝓕} (h : f = g) (x : E) : - f x = g x := - h ▸ rfl - -- @[simp] -- Porting note (#10618): simp can prove this protected theorem map_zero : f 0 = 0 := f.toLinearMap.map_zero diff --git a/Mathlib/Analysis/Oscillation.lean b/Mathlib/Analysis/Oscillation.lean index 54f9794a0f235..48e3533c9ef10 100644 --- a/Mathlib/Analysis/Oscillation.lean +++ b/Mathlib/Analysis/Oscillation.lean @@ -93,12 +93,12 @@ end Oscillation namespace IsCompact -variable [PseudoEMetricSpace E] {K : Set E} (comp : IsCompact K) +variable [PseudoEMetricSpace E] {K : Set E} variable {f : E → F} {D : Set E} {ε : ENNReal} /-- If `oscillationWithin f D x < ε` at every `x` in a compact set `K`, then there exists `δ > 0` such that the oscillation of `f` on `ball x δ ∩ D` is less than `ε` for every `x` in `K`. -/ -theorem uniform_oscillationWithin (hK : ∀ x ∈ K, oscillationWithin f D x < ε) : +theorem uniform_oscillationWithin (comp : IsCompact K) (hK : ∀ x ∈ K, oscillationWithin f D x < ε) : ∃ δ > 0, ∀ x ∈ K, diam (f '' (ball x (ENNReal.ofReal δ) ∩ D)) ≤ ε := by let S := fun r ↦ { x : E | ∃ (a : ℝ), (a > r ∧ diam (f '' (ball x (ENNReal.ofReal a) ∩ D)) ≤ ε) } have S_open : ∀ r > 0, IsOpen (S r) := by @@ -144,7 +144,7 @@ theorem uniform_oscillationWithin (hK : ∀ x ∈ K, oscillationWithin f D x < /-- If `oscillation f x < ε` at every `x` in a compact set `K`, then there exists `δ > 0` such that the oscillation of `f` on `ball x δ` is less than `ε` for every `x` in `K`. -/ -theorem uniform_oscillation [PseudoEMetricSpace E] {K : Set E} (comp : IsCompact K) +theorem uniform_oscillation {K : Set E} (comp : IsCompact K) {f : E → F} {ε : ENNReal} (hK : ∀ x ∈ K, oscillation f x < ε) : ∃ δ > 0, ∀ x ∈ K, diam (f '' (ball x (ENNReal.ofReal δ))) ≤ ε := by simp only [← oscillationWithin_univ_eq_oscillation] at hK diff --git a/Mathlib/CategoryTheory/Sites/Discrete.lean b/Mathlib/CategoryTheory/Sites/Discrete.lean index 2b59807d4ad46..e2a50fdc19dd8 100644 --- a/Mathlib/CategoryTheory/Sites/Discrete.lean +++ b/Mathlib/CategoryTheory/Sites/Discrete.lean @@ -39,8 +39,10 @@ open CategoryTheory Limits Functor Adjunction Opposite Category Functor namespace CategoryTheory.Sheaf variable {C : Type*} [Category C] (J : GrothendieckTopology C) {A : Type*} [Category A] - [HasWeakSheafify J A] [(constantSheaf J A).Faithful] [(constantSheaf J A).Full] - {t : C} (ht : IsTerminal t) + [HasWeakSheafify J A] {t : C} (ht : IsTerminal t) + +section +variable [(constantSheaf J A).Faithful] [(constantSheaf J A).Full] /-- A sheaf is discrete if it is a discrete object of the "underlying object" functor from the sheaf @@ -80,8 +82,6 @@ variable (G : C ⥤ D) [∀ (X : Dᵒᵖ), HasLimitsOfShape (StructuredArrow X G.op) A] [G.IsDenseSubsite J K] (ht' : IsTerminal (G.obj t)) -variable [(constantSheaf J A).Faithful] [(constantSheaf J A).Full] - open Functor.IsDenseSubsite noncomputable example : @@ -152,12 +152,11 @@ lemma isDiscrete_iff_of_equivalence (F : Sheaf K A) : end Equivalence +end + section Forget -variable {B : Type*} [Category B] (U : A ⥤ B) - [HasWeakSheafify J A] [HasWeakSheafify J B] - [(constantSheaf J A).Faithful] [(constantSheaf J A).Full] - [(constantSheaf J B).Faithful] [(constantSheaf J B).Full] +variable {B : Type*} [Category B] (U : A ⥤ B) [HasWeakSheafify J B] [J.PreservesSheafification U] [J.HasSheafCompose U] (F : Sheaf J A) open Limits @@ -263,6 +262,9 @@ lemma sheafCompose_reflects_discrete [(sheafCompose J U).ReflectsIsomorphisms] apply ReflectsIsomorphisms.reflects (sheafToPresheaf J B) _ apply ReflectsIsomorphisms.reflects (sheafCompose J U) _ +variable [(constantSheaf J A).Full] [(constantSheaf J A).Faithful] + [(constantSheaf J B).Full] [(constantSheaf J B).Faithful] + instance [h : F.IsDiscrete J ht] : ((sheafCompose J U).obj F).IsDiscrete J ht := by rw [isDiscrete_iff_mem_essImage] at h ⊢ diff --git a/Mathlib/CategoryTheory/Sites/InducedTopology.lean b/Mathlib/CategoryTheory/Sites/InducedTopology.lean index aadc6c8e2f4fb..8c1aae2362041 100644 --- a/Mathlib/CategoryTheory/Sites/InducedTopology.lean +++ b/Mathlib/CategoryTheory/Sites/InducedTopology.lean @@ -49,7 +49,7 @@ class LocallyCoverDense : Prop where functorPushforward_functorPullback_mem : ∀ ⦃X : C⦄ (T : K (G.obj X)), (T.val.functorPullback G).functorPushforward G ∈ K (G.obj X) -variable [G.LocallyCoverDense K] [G.IsLocallyFull K] [G.IsLocallyFaithful K] +variable [G.LocallyCoverDense K] theorem pushforward_cover_iff_cover_pullback [G.Full] [G.Faithful] {X : C} (S : Sieve X) : K _ (S.functorPushforward G) ↔ ∃ T : K (G.obj X), T.val.functorPullback G = S := by @@ -59,6 +59,8 @@ theorem pushforward_cover_iff_cover_pullback [G.Full] [G.Faithful] {X : C} (S : · rintro ⟨T, rfl⟩ exact LocallyCoverDense.functorPushforward_functorPullback_mem T +variable [G.IsLocallyFull K] [G.IsLocallyFaithful K] + /-- If a functor `G : C ⥤ (D, K)` is fully faithful and locally dense, then the set `{ T ∩ mor(C) | T ∈ K }` is a grothendieck topology of `C`. -/ diff --git a/Mathlib/GroupTheory/Frattini.lean b/Mathlib/GroupTheory/Frattini.lean index 8b5bfdfbc3bed..131e0c34f2ae1 100644 --- a/Mathlib/GroupTheory/Frattini.lean +++ b/Mathlib/GroupTheory/Frattini.lean @@ -20,14 +20,15 @@ the Frattini subgroup consists of the non-generating elements of the group. def frattini (G : Type*) [Group G] : Subgroup G := Order.radical (Subgroup G) -variable {G H : Type*} [Group G] [Group H] {φ : G →* H} (hφ : Function.Surjective φ) +variable {G H : Type*} [Group G] [Group H] {φ : G →* H} lemma frattini_le_coatom {K : Subgroup G} (h : IsCoatom K) : frattini G ≤ K := Order.radical_le_coatom h open Subgroup -lemma frattini_le_comap_frattini_of_surjective : frattini G ≤ (frattini H).comap φ := by +lemma frattini_le_comap_frattini_of_surjective (hφ : Function.Surjective φ) : + frattini G ≤ (frattini H).comap φ := by simp_rw [frattini, Order.radical, comap_iInf, le_iInf_iff] intro M hM apply biInf_le diff --git a/Mathlib/LinearAlgebra/Dimension/Constructions.lean b/Mathlib/LinearAlgebra/Dimension/Constructions.lean index 6edb9f787e61d..3724ba1ee3ccd 100644 --- a/Mathlib/LinearAlgebra/Dimension/Constructions.lean +++ b/Mathlib/LinearAlgebra/Dimension/Constructions.lean @@ -38,7 +38,7 @@ variable {ι : Type w} {ι' : Type w'} {η : Type u₁'} {φ : η → Type*} open Cardinal Basis Submodule Function Set FiniteDimensional DirectSum variable [Ring R] [CommRing S] [AddCommGroup M] [AddCommGroup M'] [AddCommGroup M₁] -variable [Module R M] [Module R M'] [Module R M₁] +variable [Module R M] section Quotient @@ -106,6 +106,7 @@ end ULift section Prod variable (R M M') +variable [Module R M₁] [Module R M'] open LinearMap in theorem lift_rank_add_lift_rank_le_rank_prod [Nontrivial R] : @@ -150,7 +151,7 @@ end Prod section Finsupp variable (R M M') -variable [StrongRankCondition R] [Module.Free R M] [Module.Free R M'] +variable [StrongRankCondition R] [Module.Free R M] [Module R M'] [Module.Free R M'] open Module.Free @@ -206,8 +207,6 @@ theorem rank_matrix' (m n : Type v) [Finite m] [Finite n] : theorem rank_matrix'' (m n : Type u) [Finite m] [Finite n] : Module.rank R (Matrix m n R) = #m * #n := by simp -variable [Module.Finite R M] [Module.Finite R M'] - open Fintype namespace FiniteDimensional @@ -322,9 +321,9 @@ section TensorProduct open TensorProduct variable [StrongRankCondition R] [StrongRankCondition S] -variable [Module S M] [Module.Free S M] [Module S M'] [Module.Free S M'] +variable [Module S M] [Module S M'] [Module.Free S M'] variable [Module S M₁] [Module.Free S M₁] -variable [Algebra S R] [Module R M] [IsScalarTower S R M] [Module.Free R M] +variable [Algebra S R] [IsScalarTower S R M] [Module.Free R M] open Module.Free @@ -382,7 +381,8 @@ theorem Submodule.finrank_quotient_le [Module.Finite R M] (s : Submodule R M) : (rank_lt_aleph0 _ _) /-- Pushforwards of finite submodules have a smaller finrank. -/ -theorem Submodule.finrank_map_le (f : M →ₗ[R] M') (p : Submodule R M) [Module.Finite R p] : +theorem Submodule.finrank_map_le + [Module R M'] (f : M →ₗ[R] M') (p : Submodule R M) [Module.Finite R p] : finrank R (p.map f) ≤ finrank R p := finrank_le_finrank_of_rank_le_rank (lift_rank_map_le _ _) (rank_lt_aleph0 _ _) diff --git a/Mathlib/LinearAlgebra/Dimension/Finite.lean b/Mathlib/LinearAlgebra/Dimension/Finite.lean index 28691d5396df9..75052d2a62473 100644 --- a/Mathlib/LinearAlgebra/Dimension/Finite.lean +++ b/Mathlib/LinearAlgebra/Dimension/Finite.lean @@ -61,6 +61,8 @@ lemma rank_eq_zero_iff : simpa using DFunLike.congr_fun (linearIndependent_iff.mp hs (Finsupp.single i a) (by simpa)) i variable [Nontrivial R] + +section variable [NoZeroSMulDivisors R M] theorem rank_zero_iff_forall_zero : @@ -80,14 +82,16 @@ theorem rank_pos_iff_exists_ne_zero : 0 < Module.rank R M ↔ ∃ x : M, x ≠ 0 theorem rank_pos_iff_nontrivial : 0 < Module.rank R M ↔ Nontrivial M := rank_pos_iff_exists_ne_zero.trans (nontrivial_iff_exists_ne 0).symm +theorem rank_pos [Nontrivial M] : 0 < Module.rank R M := + rank_pos_iff_nontrivial.mpr ‹_› + +end + lemma rank_eq_zero_iff_isTorsion {R M} [CommRing R] [IsDomain R] [AddCommGroup M] [Module R M] : Module.rank R M = 0 ↔ Module.IsTorsion R M := by rw [Module.IsTorsion, rank_eq_zero_iff] simp [mem_nonZeroDivisors_iff_ne_zero] -theorem rank_pos [Nontrivial M] : 0 < Module.rank R M := - rank_pos_iff_nontrivial.mpr ‹_› - variable (R M) /-- See `rank_subsingleton` that assumes `Subsingleton R` instead. -/ @@ -129,6 +133,7 @@ theorem Module.finite_of_rank_eq_one [Module.Free R M] (h : Module.rank R M = 1) Module.Finite R M := Module.finite_of_rank_eq_nat <| h.trans Nat.cast_one.symm +section variable [StrongRankCondition R] /-- If a module has a finite dimension, all bases are indexed by a finite type. -/ @@ -147,7 +152,10 @@ theorem Basis.finite_index_of_rank_lt_aleph0 {ι : Type*} {s : Set ι} (b : Basi (h : Module.rank R M < ℵ₀) : s.Finite := finite_def.2 (b.nonempty_fintype_index_of_rank_lt_aleph0 h) +end + namespace LinearIndependent +variable [StrongRankCondition R] theorem cardinal_mk_le_finrank [Module.Finite R M] {ι : Type w} {b : ι → M} (h : LinearIndependent R b) : #ι ≤ finrank R M := by @@ -217,12 +225,12 @@ lemma exists_linearIndependent_of_le_rank {n : ℕ} (hn : n ≤ Module.rank R M) have ⟨_, hs, hs'⟩ := exists_finset_linearIndependent_of_le_rank hn ⟨_, (linearIndependent_equiv (Finset.equivFinOfCardEq hs).symm).mpr hs'⟩ -lemma natCast_le_rank_iff {n : ℕ} : +lemma natCast_le_rank_iff [Nontrivial R] {n : ℕ} : n ≤ Module.rank R M ↔ ∃ f : Fin n → M, LinearIndependent R f := ⟨exists_linearIndependent_of_le_rank, fun H ↦ by simpa using H.choose_spec.cardinal_lift_le_rank⟩ -lemma natCast_le_rank_iff_finset {n : ℕ} : +lemma natCast_le_rank_iff_finset [Nontrivial R] {n : ℕ} : n ≤ Module.rank R M ↔ ∃ s : Finset M, s.card = n ∧ LinearIndependent R ((↑) : s → M) := ⟨exists_finset_linearIndependent_of_le_rank, fun ⟨s, h₁, h₂⟩ ↦ by simpa [h₁] using h₂.cardinal_le_rank⟩ @@ -240,11 +248,11 @@ lemma exists_linearIndependent_of_le_finrank {n : ℕ} (hn : n ≤ finrank R M) have ⟨_, hs, hs'⟩ := exists_finset_linearIndependent_of_le_finrank hn ⟨_, (linearIndependent_equiv (Finset.equivFinOfCardEq hs).symm).mpr hs'⟩ -variable [Module.Finite R M] - +variable [Module.Finite R M] [StrongRankCondition R] in theorem Module.Finite.not_linearIndependent_of_infinite {ι : Type*} [Infinite ι] (v : ι → M) : ¬LinearIndependent R v := mt LinearIndependent.finite <| @not_finite _ _ +section variable [NoZeroSMulDivisors R M] theorem CompleteLattice.Independent.subtype_ne_bot_le_rank [Nontrivial R] @@ -259,6 +267,8 @@ 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 +variable [Module.Finite R M] [StrongRankCondition R] + theorem CompleteLattice.Independent.subtype_ne_bot_le_finrank_aux {p : ι → Submodule R M} (hp : CompleteLattice.Independent p) : #{ i // p i ≠ ⊥ } ≤ (finrank R M : Cardinal.{w}) := by @@ -290,6 +300,10 @@ theorem CompleteLattice.Independent.subtype_ne_bot_le_finrank {p : ι → Submodule R M} (hp : CompleteLattice.Independent p) [Fintype { i // p i ≠ ⊥ }] : Fintype.card { i // p i ≠ ⊥ } ≤ finrank R M := by simpa using hp.subtype_ne_bot_le_finrank_aux +end + +variable [Module.Finite R M] [StrongRankCondition R] + section open Finset @@ -344,7 +358,20 @@ end Finite section FinrankZero -variable [Nontrivial R] [NoZeroSMulDivisors R M] +section +variable [Nontrivial R] + +/-- A (finite dimensional) space that is a subsingleton has zero `finrank`. -/ +@[nontriviality] +theorem FiniteDimensional.finrank_zero_of_subsingleton [Subsingleton M] : + finrank R M = 0 := by + rw [finrank, rank_subsingleton', _root_.map_zero] + +lemma LinearIndependent.finrank_eq_zero_of_infinite {ι} [Infinite ι] {v : ι → M} + (hv : LinearIndependent R v) : finrank R M = 0 := toNat_eq_zero.mpr <| .inr hv.aleph0_le_rank + +section +variable [NoZeroSMulDivisors R M] /-- A finite dimensional space is nontrivial if it has positive `finrank`. -/ theorem FiniteDimensional.nontrivial_of_finrank_pos (h : 0 < finrank R M) : Nontrivial M := @@ -356,13 +383,7 @@ theorem FiniteDimensional.nontrivial_of_finrank_eq_succ {n : ℕ} (hn : finrank R M = n.succ) : Nontrivial M := nontrivial_of_finrank_pos (by rw [hn]; exact n.succ_pos) -/-- A (finite dimensional) space that is a subsingleton has zero `finrank`. -/ -@[nontriviality] -theorem FiniteDimensional.finrank_zero_of_subsingleton [Subsingleton M] : finrank R M = 0 := by - rw [finrank, rank_subsingleton', _root_.map_zero] - -lemma LinearIndependent.finrank_eq_zero_of_infinite {ι} [Nontrivial R] [Infinite ι] {v : ι → M} - (hv : LinearIndependent R v) : finrank R M = 0 := toNat_eq_zero.mpr <| .inr hv.aleph0_le_rank +end variable (R M) @@ -370,7 +391,7 @@ variable (R M) theorem finrank_bot : finrank R (⊥ : Submodule R M) = 0 := finrank_eq_of_rank_eq (rank_bot _ _) -variable {R M} +end section StrongRankCondition diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index e38a2983a5ab2..3822701b81a5c 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -73,9 +73,8 @@ open DirectSum Pointwise open DirectSum SetLike variable {ι R A : Type*} -variable [AddCommMonoid ι] [DecidableEq ι] variable [CommRing R] [CommRing A] [Algebra R A] -variable (𝒜 : ι → Submodule R A) [GradedAlgebra 𝒜] +variable (𝒜 : ι → Submodule R A) variable (x : Submonoid A) local notation "at " x => Localization x @@ -111,6 +110,44 @@ theorem ext {c1 c2 : NumDenSameDeg 𝒜 x} (hdeg : c1.deg = c2.deg) (hnum : (c1. subst hdeg hnum hden congr +instance : Neg (NumDenSameDeg 𝒜 x) where + neg c := ⟨c.deg, ⟨-c.num, neg_mem c.num.2⟩, c.den, c.den_mem⟩ + +@[simp] +theorem deg_neg (c : NumDenSameDeg 𝒜 x) : (-c).deg = c.deg := + rfl + +@[simp] +theorem num_neg (c : NumDenSameDeg 𝒜 x) : ((-c).num : A) = -c.num := + rfl + +@[simp] +theorem den_neg (c : NumDenSameDeg 𝒜 x) : ((-c).den : A) = c.den := + rfl + +section SMul + +variable {α : Type*} [SMul α R] [SMul α A] [IsScalarTower α R A] + +instance : SMul α (NumDenSameDeg 𝒜 x) where + smul m c := ⟨c.deg, m • c.num, c.den, c.den_mem⟩ + +@[simp] +theorem deg_smul (c : NumDenSameDeg 𝒜 x) (m : α) : (m • c).deg = c.deg := + rfl + +@[simp] +theorem num_smul (c : NumDenSameDeg 𝒜 x) (m : α) : ((m • c).num : A) = m • c.num := + rfl + +@[simp] +theorem den_smul (c : NumDenSameDeg 𝒜 x) (m : α) : ((m • c).den : A) = c.den := + rfl + +end SMul + +variable [AddCommMonoid ι] [DecidableEq ι] [GradedAlgebra 𝒜] + instance : One (NumDenSameDeg 𝒜 x) where one := { deg := 0 @@ -188,21 +225,6 @@ theorem num_add (c1 c2 : NumDenSameDeg 𝒜 x) : theorem den_add (c1 c2 : NumDenSameDeg 𝒜 x) : ((c1 + c2).den : A) = c1.den * c2.den := rfl -instance : Neg (NumDenSameDeg 𝒜 x) where - neg c := ⟨c.deg, ⟨-c.num, neg_mem c.num.2⟩, c.den, c.den_mem⟩ - -@[simp] -theorem deg_neg (c : NumDenSameDeg 𝒜 x) : (-c).deg = c.deg := - rfl - -@[simp] -theorem num_neg (c : NumDenSameDeg 𝒜 x) : ((-c).num : A) = -c.num := - rfl - -@[simp] -theorem den_neg (c : NumDenSameDeg 𝒜 x) : ((-c).den : A) = c.den := - rfl - instance : CommMonoid (NumDenSameDeg 𝒜 x) where one := 1 mul := (· * ·) @@ -231,27 +253,6 @@ theorem num_pow (c : NumDenSameDeg 𝒜 x) (n : ℕ) : ((c ^ n).num : A) = (c.nu theorem den_pow (c : NumDenSameDeg 𝒜 x) (n : ℕ) : ((c ^ n).den : A) = (c.den : A) ^ n := rfl -section SMul - -variable {α : Type*} [SMul α R] [SMul α A] [IsScalarTower α R A] - -instance : SMul α (NumDenSameDeg 𝒜 x) where - smul m c := ⟨c.deg, m • c.num, c.den, c.den_mem⟩ - -@[simp] -theorem deg_smul (c : NumDenSameDeg 𝒜 x) (m : α) : (m • c).deg = c.deg := - rfl - -@[simp] -theorem num_smul (c : NumDenSameDeg 𝒜 x) (m : α) : ((m • c).num : A) = m • c.num := - rfl - -@[simp] -theorem den_smul (c : NumDenSameDeg 𝒜 x) (m : α) : ((m • c).den : A) = c.den := - rfl - -end SMul - variable (𝒜) /-- For `x : prime ideal of A` and any `p : NumDenSameDeg 𝒜 x`, or equivalent a numerator and a @@ -308,17 +309,6 @@ lemma subsingleton (hx : 0 ∈ x) : Subsingleton (HomogeneousLocalization 𝒜 x have := IsLocalization.subsingleton (S := at x) hx (HomogeneousLocalization.val_injective (𝒜 := 𝒜) (x := x)).subsingleton -instance hasPow : Pow (HomogeneousLocalization 𝒜 x) ℕ where - pow z n := - (Quotient.map' (· ^ n) fun c1 c2 (h : Localization.mk _ _ = Localization.mk _ _) => by - change Localization.mk _ _ = Localization.mk _ _ - simp only [num_pow, den_pow] - convert congr_arg (fun z : at x => z ^ n) h <;> erw [Localization.mk_pow] <;> rfl : - HomogeneousLocalization 𝒜 x → HomogeneousLocalization 𝒜 x) - z - -@[simp] lemma mk_pow (i : NumDenSameDeg 𝒜 x) (n : ℕ) : mk (i ^ n) = mk i ^ n := rfl - section SMul variable {α : Type*} [SMul α R] [SMul α A] [IsScalarTower α R A] @@ -352,6 +342,23 @@ instance : Neg (HomogeneousLocalization 𝒜 x) where @[simp] lemma mk_neg (i : NumDenSameDeg 𝒜 x) : mk (-i) = -mk i := rfl +@[simp] +theorem val_neg {x} : ∀ y : HomogeneousLocalization 𝒜 x, (-y).val = -y.val := + Quotient.ind' fun y ↦ by rw [← mk_neg, val_mk, val_mk, Localization.neg_mk]; rfl + +variable [AddCommMonoid ι] [DecidableEq ι] [GradedAlgebra 𝒜] + +instance hasPow : Pow (HomogeneousLocalization 𝒜 x) ℕ where + pow z n := + (Quotient.map' (· ^ n) fun c1 c2 (h : Localization.mk _ _ = Localization.mk _ _) => by + change Localization.mk _ _ = Localization.mk _ _ + simp only [num_pow, den_pow] + convert congr_arg (fun z : at x => z ^ n) h <;> erw [Localization.mk_pow] <;> rfl : + HomogeneousLocalization 𝒜 x → HomogeneousLocalization 𝒜 x) + z + +@[simp] lemma mk_pow (i : NumDenSameDeg 𝒜 x) (n : ℕ) : mk (i ^ n) = mk i ^ n := rfl + instance : Add (HomogeneousLocalization 𝒜 x) where add := Quotient.map₂' (· + ·) @@ -408,10 +415,6 @@ theorem val_add : ∀ y1 y2 : HomogeneousLocalization 𝒜 x, (y1 + y2).val = y1 theorem val_mul : ∀ y1 y2 : HomogeneousLocalization 𝒜 x, (y1 * y2).val = y1.val * y2.val := Quotient.ind₂' fun y1 y2 ↦ by rw [← mk_mul, val_mk, val_mk, val_mk, Localization.mk_mul]; rfl -@[simp] -theorem val_neg : ∀ y : HomogeneousLocalization 𝒜 x, (-y).val = -y.val := - Quotient.ind' fun y ↦ by rw [← mk_neg, val_mk, val_mk, Localization.neg_mk]; rfl - @[simp] theorem val_sub (y1 y2 : HomogeneousLocalization 𝒜 x) : (y1 - y2).val = y1.val - y2.val := by rw [sub_eq_add_neg, ← val_neg, ← val_add]; rfl @@ -504,6 +507,7 @@ theorem ext_iff_val (f g : HomogeneousLocalization 𝒜 x) : f = g ↔ f.val = g section +variable [AddCommMonoid ι] [DecidableEq ι] [GradedAlgebra 𝒜] variable (𝒜) (𝔭 : Ideal A) [Ideal.IsPrime 𝔭] /-- Localizing a ring homogeneously at a prime ideal. -/ @@ -544,6 +548,7 @@ variable (𝒜) (f : A) abbrev Away := HomogeneousLocalization 𝒜 (Submonoid.powers f) +variable [AddCommMonoid ι] [DecidableEq ι] [GradedAlgebra 𝒜] variable {𝒜} {f} theorem Away.eventually_smul_mem {m} (hf : f ∈ 𝒜 m) (z : Away 𝒜 f) : @@ -565,6 +570,7 @@ end section +variable [AddCommMonoid ι] [DecidableEq ι] [GradedAlgebra 𝒜] variable (𝒜) variable {B : Type*} [CommRing B] [Algebra R B] variable (ℬ : ι → Submodule R B) [GradedAlgebra ℬ] diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean index c67dd1d640062..4950bcffb0f99 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegral/Basic.lean @@ -33,7 +33,7 @@ section variable {R A B S : Type*} variable [CommRing R] [CommRing A] [Ring B] [CommRing S] -variable [Algebra R A] [Algebra R B] (f : R →+* S) +variable [Algebra R A] (f : R →+* S) theorem IsIntegral.map {B C F : Type*} [Ring B] [Ring C] [Algebra R B] [Algebra A B] [Algebra R C] [IsScalarTower R A B] [Algebra A C] [IsScalarTower R A C] {b : B} @@ -47,9 +47,9 @@ theorem IsIntegral.map {B C F : Type*} [Ring B] [Ring C] [Algebra R B] [Algebra section variable {A B : Type*} [Ring A] [Ring B] [Algebra R A] [Algebra R B] -variable (f : A →ₐ[R] B) (hf : Function.Injective f) -theorem isIntegral_algHom_iff {x : A} : IsIntegral R (f x) ↔ IsIntegral R x := by +theorem isIntegral_algHom_iff (f : A →ₐ[R] B) (hf : Function.Injective f) {x : A} : + IsIntegral R (f x) ↔ IsIntegral R x := by refine ⟨fun ⟨p, hp, hx⟩ ↦ ⟨p, hp, ?_⟩, IsIntegral.map f⟩ rwa [← f.comp_algebraMap, ← AlgHom.coe_toRingHom, ← hom_eval₂, AlgHom.coe_toRingHom, map_eq_zero_iff f hf] at hx @@ -75,7 +75,7 @@ theorem Submodule.span_range_natDegree_eq_adjoin {R A} [CommRing R] [Semiring A] exact smul_mem _ _ (subset_span <| Finset.mem_image_of_mem _ <| Finset.mem_range.mpr <| (le_natDegree_of_mem_supp _ hkq).trans_lt <| natDegree_modByMonic_lt p hf hf1) -theorem IsIntegral.fg_adjoin_singleton {x : B} (hx : IsIntegral R x) : +theorem IsIntegral.fg_adjoin_singleton [Algebra R B] {x : B} (hx : IsIntegral R x) : (Algebra.adjoin R {x}).toSubmodule.FG := by classical rcases hx with ⟨f, hfm, hfx⟩ @@ -87,18 +87,18 @@ variable (f : R →+* B) theorem RingHom.isIntegralElem_zero : f.IsIntegralElem 0 := f.map_zero ▸ f.isIntegralElem_map -theorem isIntegral_zero : IsIntegral R (0 : B) := +theorem isIntegral_zero [Algebra R B] : IsIntegral R (0 : B) := (algebraMap R B).isIntegralElem_zero theorem RingHom.isIntegralElem_one : f.IsIntegralElem 1 := f.map_one ▸ f.isIntegralElem_map -theorem isIntegral_one : IsIntegral R (1 : B) := +theorem isIntegral_one [Algebra R B] : IsIntegral R (1 : B) := (algebraMap R B).isIntegralElem_one variable (f : R →+* S) -theorem IsIntegral.of_pow {x : B} {n : ℕ} (hn : 0 < n) (hx : IsIntegral R <| x ^ n) : +theorem IsIntegral.of_pow [Algebra R B] {x : B} {n : ℕ} (hn : 0 < n) (hx : IsIntegral R <| x ^ n) : IsIntegral R x := by rcases hx with ⟨p, hmonic, heval⟩ exact ⟨expand R n p, hmonic.expand hn, by rwa [← aeval_def, expand_aeval]⟩ diff --git a/Mathlib/RingTheory/Valuation/ValuationRing.lean b/Mathlib/RingTheory/Valuation/ValuationRing.lean index bcbcd0db5ea08..976b2a387a2a7 100644 --- a/Mathlib/RingTheory/Valuation/ValuationRing.lean +++ b/Mathlib/RingTheory/Valuation/ValuationRing.lean @@ -381,11 +381,11 @@ theorem _root_.Function.Surjective.valuationRing {R S : Type*} [CommRing R] [IsD section variable {𝒪 : Type u} {K : Type v} {Γ : Type w} [CommRing 𝒪] [IsDomain 𝒪] [Field K] [Algebra 𝒪 K] - [LinearOrderedCommGroupWithZero Γ] (v : Valuation K Γ) (hh : v.Integers 𝒪) + [LinearOrderedCommGroupWithZero Γ] /-- If `𝒪` satisfies `v.integers 𝒪` where `v` is a valuation on a field, then `𝒪` is a valuation ring. -/ -theorem of_integers : ValuationRing 𝒪 := by +theorem of_integers (v : Valuation K Γ) (hh : v.Integers 𝒪) : ValuationRing 𝒪 := by constructor intro a b rcases le_total (v (algebraMap 𝒪 K a)) (v (algebraMap 𝒪 K b)) with h | h @@ -394,7 +394,7 @@ theorem of_integers : ValuationRing 𝒪 := by · obtain ⟨c, hc⟩ := Valuation.Integers.dvd_of_le hh h use c; exact Or.inl hc.symm -instance instValuationRingInteger : ValuationRing v.integer := +instance instValuationRingInteger (v : Valuation K Γ) : ValuationRing v.integer := of_integers (v := v) (Valuation.integer.integers v) theorem isFractionRing_iff [ValuationRing 𝒪] : @@ -419,7 +419,7 @@ theorem isFractionRing_iff [ValuationRing 𝒪] : · intro _ _ hab exact ⟨1, by simp only [OneMemClass.coe_one, h.2 hab, one_mul]⟩ -instance instIsFractionRingInteger : IsFractionRing v.integer K := +instance instIsFractionRingInteger (v : Valuation K Γ) : IsFractionRing v.integer K := ValuationRing.isFractionRing_iff.mpr ⟨Valuation.Integers.eq_algebraMap_or_inv_eq_algebraMap (Valuation.integer.integers v), Subtype.coe_injective⟩ diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index aa02b48710111..fda640ec9b168 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -1243,7 +1243,8 @@ theorem QuotientGroup.isClosedMap_coe {H : Subgroup G} (hH : IsCompact (H : Set rfl @[to_additive] -lemma subset_mul_closure_one (s : Set G) : s ⊆ s * (closure {1} : Set G) := by +lemma subset_mul_closure_one {G} [MulOneClass G] [TopologicalSpace G] (s : Set G) : + s ⊆ s * (closure {1} : Set G) := by have : s ⊆ s * ({1} : Set G) := by simp exact this.trans (smul_subset_smul_left subset_closure) diff --git a/Mathlib/Topology/Instances/Discrete.lean b/Mathlib/Topology/Instances/Discrete.lean index 6eda72d280c32..b937256b9520c 100644 --- a/Mathlib/Topology/Instances/Discrete.lean +++ b/Mathlib/Topology/Instances/Discrete.lean @@ -43,7 +43,7 @@ theorem DiscreteTopology.secondCountableTopology_of_encodable {α : Type*} [TopologicalSpace α] [DiscreteTopology α] [Countable α] : SecondCountableTopology α := DiscreteTopology.secondCountableTopology_of_countable -theorem bot_topologicalSpace_eq_generateFrom_of_pred_succOrder [PartialOrder α] [PredOrder α] +theorem bot_topologicalSpace_eq_generateFrom_of_pred_succOrder {α} [PartialOrder α] [PredOrder α] [SuccOrder α] [NoMinOrder α] [NoMaxOrder α] : (⊥ : TopologicalSpace α) = generateFrom { s | ∃ a, s = Ioi a ∨ s = Iio a } := by refine (eq_bot_of_singletons_open fun a => ?_).symm @@ -69,7 +69,7 @@ instance (priority := 100) DiscreteTopology.orderTopology_of_pred_succ' [h : Dis [PartialOrder α] [PredOrder α] [SuccOrder α] [NoMinOrder α] [NoMaxOrder α] : OrderTopology α := discreteTopology_iff_orderTopology_of_pred_succ'.1 h -theorem LinearOrder.bot_topologicalSpace_eq_generateFrom [LinearOrder α] [PredOrder α] +theorem LinearOrder.bot_topologicalSpace_eq_generateFrom {α} [LinearOrder α] [PredOrder α] [SuccOrder α] : (⊥ : TopologicalSpace α) = generateFrom { s | ∃ a, s = Ioi a ∨ s = Iio a } := by refine (eq_bot_of_singletons_open fun a => ?_).symm have h_singleton_eq_inter : {a} = Iic a ∩ Ici a := by rw [inter_comm, Ici_inter_Iic, Icc_self a] diff --git a/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean b/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean index 09348e7e2cbaa..07a4029efeb34 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean @@ -211,7 +211,7 @@ open TopCat Opposite variable {C : Type u} [Category.{v} C] variable {X : TopCat.{w}} {ι : Type*} {B : ι → Opens X} -variable (F : X.Presheaf C) (F' : Sheaf C X) (h : Opens.IsBasis (Set.range B)) +variable (F : X.Presheaf C) (F' : Sheaf C X) /-- The empty component of a sheaf is terminal. -/ def isTerminalOfEmpty (F : Sheaf C X) : Limits.IsTerminal (F.val.obj (op ⊥)) := @@ -226,18 +226,20 @@ def isTerminalOfEqEmpty (F : X.Sheaf C) {U : Opens X} (h : U = ⊥) : is a sheaf on `X`, then a homomorphism between a presheaf `F` on `X` and `F'` is equivalent to a homomorphism between their restrictions to the indexing type `ι` of `B`, with the induced category structure on `ι`. -/ -def restrictHomEquivHom : +def restrictHomEquivHom (h : Opens.IsBasis (Set.range B)) : ((inducedFunctor B).op ⋙ F ⟶ (inducedFunctor B).op ⋙ F'.1) ≃ (F ⟶ F'.1) := @Functor.IsCoverDense.restrictHomEquivHom _ _ _ _ _ _ _ _ (Opens.coverDense_inducedFunctor h) _ F F' @[simp] -theorem extend_hom_app (α : (inducedFunctor B).op ⋙ F ⟶ (inducedFunctor B).op ⋙ F'.1) (i : ι) : +theorem extend_hom_app (h : Opens.IsBasis (Set.range B)) + (α : (inducedFunctor B).op ⋙ F ⟶ (inducedFunctor B).op ⋙ F'.1) (i : ι) : (restrictHomEquivHom F F' h α).app (op (B i)) = α.app (op i) := by nth_rw 2 [← (restrictHomEquivHom F F' h).left_inv α] rfl -theorem hom_ext {α β : F ⟶ F'.1} (he : ∀ i, α.app (op (B i)) = β.app (op (B i))) : α = β := by +theorem hom_ext (h : Opens.IsBasis (Set.range B)) + {α β : F ⟶ F'.1} (he : ∀ i, α.app (op (B i)) = β.app (op (B i))) : α = β := by apply (restrictHomEquivHom F F' h).symm.injective ext i exact he i.unop diff --git a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean index bc5bf5abab8fd..60512c6535c4d 100644 --- a/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean +++ b/Mathlib/Topology/UniformSpace/UniformConvergenceTopology.lean @@ -333,7 +333,7 @@ theorem uniformContinuous_eval (x : α) : variable {β} @[simp] -protected lemma mem_gen {f g : α →ᵤ β} {V : Set (β × β)} : +protected lemma mem_gen {β} {f g : α →ᵤ β} {V : Set (β × β)} : (f, g) ∈ UniformFun.gen α β V ↔ ∀ x, (toFun f x, toFun g x) ∈ V := .rfl From 2edef7d509a7aeba89793a84905121a9f19171b6 Mon Sep 17 00:00:00 2001 From: Jeremy Tan Jie Rui Date: Thu, 8 Aug 2024 10:31:48 +0000 Subject: [PATCH 108/359] chore: backports for leanprover/lean4#4814 (part 36) (#15612) --- Archive/Imo/Imo1975Q1.lean | 7 +-- Archive/Imo/Imo2019Q2.lean | 16 ++++-- Archive/Imo/Imo2024Q2.lean | 4 ++ Archive/Wiedijk100Theorems/BuffonsNeedle.lean | 4 +- Archive/Wiedijk100Theorems/CubingACube.lean | 6 ++ .../Wiedijk100Theorems/SolutionOfCubic.lean | 7 +-- Archive/ZagierTwoSquares.lean | 4 +- Counterexamples/QuadraticForm.lean | 8 ++- .../ContinuousFunctionalCalculus/Order.lean | 12 +++- .../ContinuousFunctionalCalculus/Unique.lean | 20 ++++++- Mathlib/Analysis/CStarAlgebra/Module.lean | 33 +++++++---- .../Fourier/FourierTransformDeriv.lean | 10 +++- .../FunctionalSpaces/SobolevInequality.lean | 19 ++++--- Mathlib/Geometry/Manifold/BumpFunction.lean | 18 +++++- .../Geometry/Manifold/VectorBundle/Basic.lean | 2 +- .../Function/StronglyMeasurable/Basic.lean | 3 +- .../MeasureTheory/Integral/SetIntegral.lean | 13 +++-- Mathlib/NumberTheory/Cyclotomic/Rat.lean | 13 ++++- Mathlib/NumberTheory/Cyclotomic/Three.lean | 13 +++-- Mathlib/NumberTheory/FLT/Three.lean | 55 ++++++++++++++++--- .../NumberTheory/NumberField/Embeddings.lean | 23 ++++++-- .../NumberField/EquivReindex.lean | 8 +-- Mathlib/NumberTheory/NumberField/House.lean | 29 +++++----- .../NumberField/Units/DirichletTheorem.lean | 10 +++- 24 files changed, 249 insertions(+), 88 deletions(-) diff --git a/Archive/Imo/Imo1975Q1.lean b/Archive/Imo/Imo1975Q1.lean index 81306d8a6eea2..ebfba2f144703 100644 --- a/Archive/Imo/Imo1975Q1.lean +++ b/Archive/Imo/Imo1975Q1.lean @@ -25,11 +25,10 @@ by the Rearrangement Inequality /- Let `n` be a natural number, `x` and `y` be as in the problem statement and `σ` be the permutation of natural numbers such that `z = y ∘ σ` -/ -variable (n : ℕ) (σ : Equiv.Perm ℕ) (hσ : {x | σ x ≠ x} ⊆ Finset.Icc 1 n) (x y : ℕ → ℝ) -variable (hx : AntitoneOn x (Finset.Icc 1 n)) -variable (hy : AntitoneOn y (Finset.Icc 1 n)) +variable (n : ℕ) (σ : Equiv.Perm ℕ) (x y : ℕ → ℝ) -theorem imo1975_q1 : +theorem imo1975_q1 (hσ : {x | σ x ≠ x} ⊆ Finset.Icc 1 n) + (hx : AntitoneOn x (Finset.Icc 1 n)) (hy : AntitoneOn y (Finset.Icc 1 n)) : ∑ i ∈ Finset.Icc 1 n, (x i - y i) ^ 2 ≤ ∑ i ∈ Finset.Icc 1 n, (x i - y (σ i)) ^ 2 := by simp only [sub_sq, Finset.sum_add_distrib, Finset.sum_sub_distrib] -- a finite sum is invariant if we permute the order of summation diff --git a/Archive/Imo/Imo2019Q2.lean b/Archive/Imo/Imo2019Q2.lean index 2ec513036ce01..db1650a221391 100644 --- a/Archive/Imo/Imo2019Q2.lean +++ b/Archive/Imo/Imo2019Q2.lean @@ -66,7 +66,7 @@ attribute [local instance] FiniteDimensional.of_fact_finrank_eq_two variable (V : Type*) (Pt : Type*) variable [NormedAddCommGroup V] [InnerProductSpace ℝ V] [MetricSpace Pt] -variable [NormedAddTorsor V Pt] [hd2 : Fact (finrank ℝ V = 2)] +variable [NormedAddTorsor V Pt] namespace Imo2019Q2 @@ -96,7 +96,7 @@ structure Imo2019q2Cfg where C_ne_Q₁ : C ≠ Q₁ /-- A default choice of orientation, for lemmas that need to pick one. -/ -def someOrientation : Module.Oriented ℝ V (Fin 2) := +def someOrientation [hd2 : Fact (finrank ℝ V = 2)] : Module.Oriented ℝ V (Fin 2) := ⟨Basis.orientation (finBasisOfFinrankEq _ _ hd2.out)⟩ variable {V Pt} @@ -249,7 +249,7 @@ section Oriented variable [Module.Oriented ℝ V (Fin 2)] -theorem oangle_CQ₁Q_sign_eq_oangle_CBA_sign : +theorem oangle_CQ₁Q_sign_eq_oangle_CBA_sign [Fact (finrank ℝ V = 2)] : (∡ cfg.C cfg.Q₁ cfg.Q).sign = (∡ cfg.C cfg.B cfg.A).sign := by rw [← cfg.sbtw_Q_A₁_Q₁.symm.oangle_eq_right, cfg.sOppSide_CB_Q_Q₁.oangle_sign_eq_neg (left_mem_affineSpan_pair ℝ cfg.C cfg.B) @@ -259,7 +259,8 @@ theorem oangle_CQ₁Q_sign_eq_oangle_CBA_sign : cfg.wbtw_B_Q_B₁.oangle_eq_right cfg.Q_ne_B, cfg.wbtw_A_B₁_C.symm.oangle_sign_eq_of_ne_left cfg.B cfg.B₁_ne_C.symm] -theorem oangle_CQ₁Q_eq_oangle_CBA : ∡ cfg.C cfg.Q₁ cfg.Q = ∡ cfg.C cfg.B cfg.A := +theorem oangle_CQ₁Q_eq_oangle_CBA [Fact (finrank ℝ V = 2)] : + ∡ cfg.C cfg.Q₁ cfg.Q = ∡ cfg.C cfg.B cfg.A := oangle_eq_of_angle_eq_of_sign_eq cfg.angle_CQ₁Q_eq_angle_CBA cfg.oangle_CQ₁Q_sign_eq_oangle_CBA_sign @@ -267,6 +268,9 @@ end Oriented /-! ### More obvious configuration properties -/ +section + +variable [hd2 : Fact (finrank ℝ V = 2)] theorem A₁_ne_B : cfg.A₁ ≠ cfg.B := by intro h @@ -568,6 +572,8 @@ theorem result : Concyclic ({cfg.P, cfg.Q, cfg.P₁, cfg.Q₁} : Set Pt) := by simp only [Set.insert_subset_iff, Set.singleton_subset_iff] exact ⟨cfg.P_mem_ω, cfg.Q_mem_ω, cfg.P₁_mem_ω, cfg.Q₁_mem_ω⟩ +end + end Imo2019q2Cfg end @@ -576,7 +582,7 @@ end Imo2019Q2 open Imo2019Q2 -theorem imo2019_q2 (A B C A₁ B₁ P Q P₁ Q₁ : Pt) +theorem imo2019_q2 [Fact (finrank ℝ V = 2)] (A B C A₁ B₁ P Q P₁ Q₁ : Pt) (affine_independent_ABC : AffineIndependent ℝ ![A, B, C]) (wbtw_B_A₁_C : Wbtw ℝ B A₁ C) (wbtw_A_B₁_C : Wbtw ℝ A B₁ C) (wbtw_A_P_A₁ : Wbtw ℝ A P A₁) (wbtw_B_Q_B₁ : Wbtw ℝ B Q B₁) (PQ_parallel_AB : line[ℝ, P, Q] ∥ line[ℝ, A, B]) (P_ne_Q : P ≠ Q) diff --git a/Archive/Imo/Imo2024Q2.lean b/Archive/Imo/Imo2024Q2.lean index 0801f9425db8f..048b384ebbf9f 100644 --- a/Archive/Imo/Imo2024Q2.lean +++ b/Archive/Imo/Imo2024Q2.lean @@ -47,6 +47,8 @@ namespace Condition variable {a b : ℕ} (h : Condition a b) +section + lemma a_pos : 0 < a := h.1 lemma b_pos : 0 < b := h.2.1 @@ -79,6 +81,8 @@ lemma dvd_g_of_le_N_of_dvd {n : ℕ} (hn : h.N ≤ n) {d : ℕ} (hab : d ∣ a ^ rw [← h.gcd_eq_g hn, Nat.dvd_gcd_iff] exact ⟨hab, hba⟩ +end + lemma a_coprime_ab_add_one : a.Coprime (a * b + 1) := by simp diff --git a/Archive/Wiedijk100Theorems/BuffonsNeedle.lean b/Archive/Wiedijk100Theorems/BuffonsNeedle.lean index ccb8a3b6357e8..3017493e369aa 100644 --- a/Archive/Wiedijk100Theorems/BuffonsNeedle.lean +++ b/Archive/Wiedijk100Theorems/BuffonsNeedle.lean @@ -76,7 +76,7 @@ namespace BuffonsNeedle variable /- Probability theory variables. -/ - {Ω : Type*} [MeasureSpace Ω] [IsProbabilityMeasure (ℙ : Measure Ω)] + {Ω : Type*} [MeasureSpace Ω] /- Buffon's needle variables. -/ @@ -124,7 +124,7 @@ noncomputable def N : Ω → ℝ := needleCrossesIndicator l ∘ B /-- The possible x-positions and angle relative to the y-axis of a needle. -/ -abbrev needleSpace: Set (ℝ × ℝ) := Set.Icc (-d / 2) (d / 2) ×ˢ Set.Icc 0 π +abbrev needleSpace : Set (ℝ × ℝ) := Set.Icc (-d / 2) (d / 2) ×ˢ Set.Icc 0 π lemma volume_needleSpace : ℙ (needleSpace d) = ENNReal.ofReal (d * π) := by simp_rw [MeasureTheory.Measure.volume_eq_prod, MeasureTheory.Measure.prod_prod, Real.volume_Icc, diff --git a/Archive/Wiedijk100Theorems/CubingACube.lean b/Archive/Wiedijk100Theorems/CubingACube.lean index ce0ceacc81d8d..9d7164a59e35c 100644 --- a/Archive/Wiedijk100Theorems/CubingACube.lean +++ b/Archive/Wiedijk100Theorems/CubingACube.lean @@ -260,6 +260,8 @@ theorem t_le_t (hi : i ∈ bcubes cs c) (j : Fin n) : · exact h'.2 · simp [hw] +section + /-- Every cube in the valley must be smaller than it -/ theorem w_lt_w (hi : i ∈ bcubes cs c) : (cs i).w < c.w := by apply lt_of_le_of_ne _ (v.2.2 i hi.1) @@ -508,6 +510,10 @@ noncomputable def sequenceOfCubes : ℕ → { i : ι // Valley cs (cs i).shiftUp def decreasingSequence (k : ℕ) : ℝ := (cs (sequenceOfCubes h k).1).w +end + +variable [Finite ι] [Nontrivial ι] + theorem strictAnti_sequenceOfCubes : StrictAnti <| decreasingSequence h := strictAnti_nat_of_succ_lt fun k => by let v := (sequenceOfCubes h k).2; dsimp only [decreasingSequence, sequenceOfCubes] diff --git a/Archive/Wiedijk100Theorems/SolutionOfCubic.lean b/Archive/Wiedijk100Theorems/SolutionOfCubic.lean index 5162c4b4038c7..b780b1926bc5e 100644 --- a/Archive/Wiedijk100Theorems/SolutionOfCubic.lean +++ b/Archive/Wiedijk100Theorems/SolutionOfCubic.lean @@ -41,10 +41,7 @@ section Field open Polynomial -variable {K : Type*} [Field K] -variable [Invertible (2 : K)] [Invertible (3 : K)] -variable (a b c d : K) -variable {ω p q r s t : K} +variable {K : Type*} [Field K] (a b c d : K) {ω p q r s t : K} theorem cube_root_of_unity_sum (hω : IsPrimitiveRoot ω 3) : 1 + ω + ω ^ 2 = 0 := by simpa [cyclotomic_prime, Finset.sum_range_succ] using hω.isRoot_cyclotomic (by decide) @@ -67,6 +64,8 @@ theorem cubic_basic_eq_zero_iff (hω : IsPrimitiveRoot ω 3) (hp_nonzero : p ≠ (x ^ 2 * (s - t) + x * (-ω * (s ^ 2 + t ^ 2) + s * t * (3 + ω ^ 2 - ω)) - (-(s ^ 3 - t ^ 3) * (ω - 1) + s ^ 2 * t * ω ^ 2 - s * t ^ 2 * ω ^ 2)) * s ^ 3 * H +variable [Invertible (2 : K)] [Invertible (3 : K)] + /-- Roots of a monic cubic whose discriminant is nonzero. -/ theorem cubic_monic_eq_zero_iff (hω : IsPrimitiveRoot ω 3) (hp : p = (3 * c - b ^ 2) / 9) (hp_nonzero : p ≠ 0) (hq : q = (9 * b * c - 2 * b ^ 3 - 27 * d) / 54) diff --git a/Archive/ZagierTwoSquares.lean b/Archive/ZagierTwoSquares.lean index e9f7d1bcb5b30..894cf667b4e0c 100644 --- a/Archive/ZagierTwoSquares.lean +++ b/Archive/ZagierTwoSquares.lean @@ -76,7 +76,7 @@ section Involutions open Function -variable (k : ℕ) [hk : Fact (4 * k + 1).Prime] +variable (k : ℕ) /-- The obvious involution `(x, y, z) ↦ (x, z, y)`. -/ def obvInvo : Function.End (zagierSet k) := fun ⟨⟨x, y, z⟩, h⟩ => ⟨⟨x, z, y⟩, by @@ -111,6 +111,8 @@ def complexInvo : Function.End (zagierSet k) := fun ⟨⟨x, y, z⟩, h⟩ => · -- middle: `x` is neither less than `y - z` or more than `2 * y` push_neg at less more; zify [less, more] at h ⊢; linarith [h]⟩ +variable [hk : Fact (4 * k + 1).Prime] + /-- `complexInvo k` is indeed an involution. -/ theorem complexInvo_sq : complexInvo k ^ 2 = 1 := by change complexInvo k ∘ complexInvo k = id diff --git a/Counterexamples/QuadraticForm.lean b/Counterexamples/QuadraticForm.lean index fd174fc9ffa69..c691445111fa7 100644 --- a/Counterexamples/QuadraticForm.lean +++ b/Counterexamples/QuadraticForm.lean @@ -16,7 +16,7 @@ The counterexample we use is $B (x, y) (x', y') ↦ xy' + x'y$ where `x y x' y' -/ -variable (F : Type*) [Nontrivial F] [CommRing F] [CharP F 2] +variable (F : Type*) [CommRing F] open LinearMap open LinearMap.BilinForm @@ -36,9 +36,11 @@ theorem B_apply (x y : F × F) : B F x y = x.1 * y.2 + x.2 * y.1 := theorem isSymm_B : (B F).IsSymm := fun x y => by simp [mul_comm, add_comm] -theorem isAlt_B : (B F).IsAlt := fun x => by simp [mul_comm, CharTwo.add_self_eq_zero (x.1 * x.2)] +theorem isAlt_B [CharP F 2] : (B F).IsAlt := fun x => by + simp [mul_comm, CharTwo.add_self_eq_zero (x.1 * x.2)] -theorem B_ne_zero : B F ≠ 0 := fun h => by simpa using LinearMap.congr_fun₂ h (1, 0) (1, 1) +theorem B_ne_zero [Nontrivial F] : B F ≠ 0 := fun h => by + simpa using LinearMap.congr_fun₂ h (1, 0) (1, 1) /-- `LinearMap.BilinForm.toQuadraticForm` is not injective on symmetric bilinear forms. diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean index c16f8ee372b0d..2bebd79a58986 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Order.lean @@ -37,7 +37,7 @@ open scoped NNReal namespace Unitization -variable {A : Type*} [NonUnitalNormedRing A] [CompleteSpace A] [Nontrivial A] +variable {A : Type*} [NonUnitalNormedRing A] [CompleteSpace A] [PartialOrder A] [StarRing A] [StarOrderedRing A] [CStarRing A] [NormedSpace ℂ A] [StarModule ℂ A] [SMulCommClass ℂ A A] [IsScalarTower ℂ A A] @@ -67,7 +67,11 @@ end Unitization section CStar_unital variable {A : Type*} [NormedRing A] [StarRing A] [CStarRing A] [CompleteSpace A] -variable [NormedAlgebra ℂ A] [StarModule ℂ A] [PartialOrder A] [StarOrderedRing A] +variable [NormedAlgebra ℂ A] [StarModule ℂ A] + +section StarOrderedRing + +variable [PartialOrder A] [StarOrderedRing A] lemma IsSelfAdjoint.le_algebraMap_norm_self {a : A} (ha : IsSelfAdjoint a := by cfc_tac) : a ≤ algebraMap ℝ A ‖a‖ := by @@ -93,6 +97,8 @@ lemma CStarRing.star_mul_le_algebraMap_norm_sq {a : A} : star a * a ≤ algebraM have : star a * a ≤ algebraMap ℝ A ‖star a * a‖ := IsSelfAdjoint.le_algebraMap_norm_self rwa [CStarRing.norm_star_mul_self, ← pow_two] at this +end StarOrderedRing + lemma IsSelfAdjoint.toReal_spectralRadius_eq_norm {a : A} (ha : IsSelfAdjoint a) : (spectralRadius ℝ a).toReal = ‖a‖ := by simp [ha.spectrumRestricts.spectralRadius_eq, ha.spectralRadius_eq_nnnorm] @@ -103,6 +109,8 @@ lemma CStarRing.norm_or_neg_norm_mem_spectrum [Nontrivial A] {a : A} rw [← ha.toReal_spectralRadius_eq_norm] exact Real.spectralRadius_mem_spectrum_or (ha'.image ▸ (spectrum.nonempty a).image _) +variable [PartialOrder A] [StarOrderedRing A] + lemma CStarRing.nnnorm_mem_spectrum_of_nonneg [Nontrivial A] {a : A} (ha : 0 ≤ a := by cfc_tac) : ‖a‖₊ ∈ spectrum ℝ≥0 a := by have : IsSelfAdjoint a := .of_nonneg ha diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean index 49ccc1221a710..830edf4006647 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unique.lean @@ -92,10 +92,14 @@ lemma toNNReal_neg_one : (-1 : C(X, ℝ)).toNNReal = 0 := toNNReal_neg_algebraMa end ContinuousMap -variable {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] [Algebra ℝ A] [TopologicalRing A] +variable {A : Type*} [Ring A] [StarRing A] [Algebra ℝ A] namespace StarAlgHom +section TopologicalRing + +variable [TopologicalSpace A] [TopologicalRing A] + /-- Given a star `ℝ≥0`-algebra homomorphism `φ` from `C(X, ℝ≥0)` into an `ℝ`-algebra `A`, this is the unique extension of `φ` from `C(X, ℝ)` to `A` as a star `ℝ`-algebra homomorphism. -/ @[simps] @@ -135,6 +139,8 @@ lemma continuous_realContinuousMapOfNNReal (φ : C(X, ℝ≥0) →⋆ₐ[ℝ≥0 simp [realContinuousMapOfNNReal] fun_prop +end TopologicalRing + @[simp high] lemma realContinuousMapOfNNReal_apply_comp_toReal (φ : C(X, ℝ≥0) →⋆ₐ[ℝ≥0] A) (f : C(X, ℝ≥0)) : @@ -155,6 +161,8 @@ lemma realContinuousMapOfNNReal_injective : end StarAlgHom +variable [TopologicalSpace A] [TopologicalRing A] + instance NNReal.instUniqueContinuousFunctionalCalculus [UniqueContinuousFunctionalCalculus ℝ A] : UniqueContinuousFunctionalCalculus ℝ≥0 A where compactSpace_spectrum a := by @@ -283,13 +291,16 @@ lemma toNNReal_add_add_neg_add_neg_eq (f g : C(X, ℝ)₀) : end ContinuousMapZero -variable {A : Type*} [TopologicalSpace A] [NonUnitalRing A] [StarRing A] [Module ℝ A] -variable [IsScalarTower ℝ A A] [SMulCommClass ℝ A A] [TopologicalRing A] +variable {A : Type*} [NonUnitalRing A] [StarRing A] [Module ℝ A] namespace NonUnitalStarAlgHom open ContinuousMapZero +section TopologicalRing + +variable [TopologicalSpace A] [TopologicalRing A] + /-- Given a non-unital star `ℝ≥0`-algebra homomorphism `φ` from `C(X, ℝ≥0)₀` into a non-unital `ℝ`-algebra `A`, this is the unique extension of `φ` from `C(X, ℝ)₀` to `A` as a non-unital star `ℝ`-algebra homomorphism. -/ @@ -331,6 +342,8 @@ lemma continuous_realContinuousMapZeroOfNNReal (φ : C(X, ℝ≥0)₀ →⋆ₙ simp [realContinuousMapZeroOfNNReal] fun_prop +end TopologicalRing + @[simp high] lemma realContinuousMapZeroOfNNReal_apply_comp_toReal (φ : C(X, ℝ≥0)₀ →⋆ₙₐ[ℝ≥0] A) (f : C(X, ℝ≥0)₀) : @@ -355,6 +368,7 @@ end NonUnitalStarAlgHom open ContinuousMapZero instance NNReal.instUniqueNonUnitalContinuousFunctionalCalculus + [TopologicalSpace A] [TopologicalRing A] [IsScalarTower ℝ A A] [SMulCommClass ℝ A A] [UniqueNonUnitalContinuousFunctionalCalculus ℝ A] : UniqueNonUnitalContinuousFunctionalCalculus ℝ≥0 A where compactSpace_quasispectrum a := by diff --git a/Mathlib/Analysis/CStarAlgebra/Module.lean b/Mathlib/Analysis/CStarAlgebra/Module.lean index d9abc7498c893..c524080d6161f 100644 --- a/Mathlib/Analysis/CStarAlgebra/Module.lean +++ b/Mathlib/Analysis/CStarAlgebra/Module.lean @@ -68,8 +68,7 @@ namespace CStarModule section general variable {A E : Type*} [NonUnitalRing A] [StarRing A] [AddCommGroup E] [Module ℂ A] - [Module ℂ E] [PartialOrder A] [SMul Aᵐᵒᵖ E] [StarModule ℂ A] [Norm A] [Norm E] - [CStarModule A E] + [Module ℂ E] [PartialOrder A] [SMul Aᵐᵒᵖ E] [Norm A] [Norm E] [CStarModule A E] local notation "⟪" x ", " y "⟫" => inner (𝕜 := A) x y @@ -82,6 +81,10 @@ lemma inner_add_left {x y z : E} : ⟪x + y, z⟫ = ⟪x, z⟫ + ⟪y, z⟫ := b lemma inner_op_smul_left {a : A} {x y : E} : ⟪x <• a, y⟫ = star a * ⟪x, y⟫ := by rw [← star_inner]; simp +section StarModule + +variable [StarModule ℂ A] + @[simp] lemma inner_smul_left_complex {z : ℂ} {x y : E} : ⟪z • x, y⟫ = star z • ⟪x, y⟫ := by rw [← star_inner] @@ -128,6 +131,8 @@ lemma inner_sum_left {ι : Type*} {s : Finset ι} {x : ι → E} {y : E} : ⟪∑ i ∈ s, x i, y⟫ = ∑ i ∈ s, ⟪x i, y⟫ := map_sum (innerₛₗ.flip y) .. +end StarModule + @[simp] lemma isSelfAdjoint_inner_self {x : E} : IsSelfAdjoint ⟪x, x⟫ := star_inner _ _ @@ -135,9 +140,8 @@ end general section norm -variable {A E : Type*} [NonUnitalNormedRing A] [StarRing A] [CStarRing A] [PartialOrder A] - [StarOrderedRing A] [AddCommGroup E] [NormedSpace ℂ A] [Module ℂ E] [SMul Aᵐᵒᵖ E] [Norm E] - [StarModule ℂ A] [CStarModule A E] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] +variable {A E : Type*} [NonUnitalNormedRing A] [StarRing A] [PartialOrder A] + [AddCommGroup E] [NormedSpace ℂ A] [Module ℂ E] [SMul Aᵐᵒᵖ E] [Norm E] [CStarModule A E] local notation "⟪" x ", " y "⟫" => inner (𝕜 := A) x y @@ -149,11 +153,7 @@ noncomputable def norm : Norm E where lemma inner_self_eq_norm_sq {x : E} : ‖⟪x, x⟫‖ = ‖x‖ ^ 2 := by simp [norm_eq_sqrt_norm_inner_self] -protected lemma norm_zero : ‖(0 : E)‖ = 0 := by simp [norm_eq_sqrt_norm_inner_self] - -lemma norm_zero_iff (x : E) : ‖x‖ = 0 ↔ x = 0 := - ⟨fun h => by simpa [norm_eq_sqrt_norm_inner_self, inner_self] using h, - fun h => by simp [norm, h]; rw [CStarModule.norm_zero] ⟩ +section protected lemma norm_nonneg {x : E} : 0 ≤ ‖x‖ := by simp [norm_eq_sqrt_norm_inner_self]; positivity @@ -163,8 +163,21 @@ protected lemma norm_pos {x : E} (hx : x ≠ 0) : 0 < ‖x‖ := by rw [inner_self] at H exact hx H +variable [StarModule ℂ A] + +protected lemma norm_zero : ‖(0 : E)‖ = 0 := by simp [norm_eq_sqrt_norm_inner_self] + +lemma norm_zero_iff (x : E) : ‖x‖ = 0 ↔ x = 0 := + ⟨fun h => by simpa [norm_eq_sqrt_norm_inner_self, inner_self] using h, + fun h => by simp [norm, h]; rw [CStarModule.norm_zero] ⟩ + +end + lemma norm_sq_eq {x : E} : ‖x‖ ^ 2 = ‖⟪x, x⟫‖ := by simp [norm_eq_sqrt_norm_inner_self] +variable [CStarRing A] [StarOrderedRing A] [StarModule ℂ A] + [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] + /-- The C⋆-algebra-valued Cauchy-Schwarz inequality for Hilbert C⋆-modules. -/ lemma inner_mul_inner_swap_le [CompleteSpace A] {x y : E} : ⟪y, x⟫ * ⟪x, y⟫ ≤ ‖x‖ ^ 2 • ⟪y, y⟫ := by rcases eq_or_ne x 0 with h|h diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 4ffbc633d5466..226ce61296678 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -420,7 +420,11 @@ lemma norm_iteratedFDeriv_fourierPowSMulRight simp only [← Finset.sum_mul, ← Nat.cast_sum, Nat.sum_range_choose, mul_one, ← mul_assoc, Nat.cast_pow, Nat.cast_ofNat, Nat.cast_add, Nat.cast_one, ← mul_pow, mul_add] -variable [SecondCountableTopology V] [MeasurableSpace V] [BorelSpace V] {μ : Measure V} +variable [MeasurableSpace V] [BorelSpace V] {μ : Measure V} + +section SecondCountableTopology + +variable [SecondCountableTopology V] lemma _root_.MeasureTheory.AEStronglyMeasurable.fourierPowSMulRight (hf : AEStronglyMeasurable f μ) (n : ℕ) : @@ -506,6 +510,8 @@ lemma iteratedFDeriv_fourierIntegral {N : ℕ∞} ext w : 1 exact ((hasFTaylorSeriesUpTo_fourierIntegral L hf h'f).eq_iteratedFDeriv hn w).symm +end SecondCountableTopology + /-- The Fourier integral of the `n`-th derivative of a function is obtained by multiplying the Fourier integral of the original function by `(2πI L w ⬝ )^n`. -/ theorem fourierIntegral_iteratedFDeriv [FiniteDimensional ℝ V] @@ -702,7 +708,7 @@ theorem fourierIntegral_iteratedFDeriv {N : ℕ∞} (hf : ContDiff ℝ N f) /-- One can bound `‖w‖^n * ‖D^k (𝓕 f) w‖` in terms of integrals of the derivatives of `f` (or order at most `n`) multiplied by powers of `v` (of order at most `k`). -/ -lemma pow_mul_norm_iteratedFDeriv_fourierIntegral_le [FiniteDimensional ℝ V] +lemma pow_mul_norm_iteratedFDeriv_fourierIntegral_le {K N : ℕ∞} (hf : ContDiff ℝ N f) (h'f : ∀ (k n : ℕ), k ≤ K → n ≤ N → Integrable (fun v ↦ ‖v‖^k * ‖iteratedFDeriv ℝ n f v‖)) {k n : ℕ} (hk : k ≤ K) (hn : n ≤ N) (w : V) : diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 5acc22f6991c2..6ef37580057c0 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -54,14 +54,14 @@ open Set Function Finset MeasureTheory Measure Filter noncomputable section -variable {ι : Type*} [Fintype ι] +variable {ι : Type*} local prefix:max "#" => Fintype.card /-! ## The grid-lines lemma -/ variable {A : ι → Type*} [∀ i, MeasurableSpace (A i)] - (μ : ∀ i, Measure (A i)) [∀ i, SigmaFinite (μ i)] + (μ : ∀ i, Measure (A i)) namespace MeasureTheory @@ -91,7 +91,7 @@ def T (p : ℝ) (f : (∀ i, A i) → ℝ≥0∞) (s : Finset ι) : (∀ i, A i) variable {p : ℝ} -@[simp] lemma T_univ (f : (∀ i, A i) → ℝ≥0∞) (x : ∀ i, A i) : +@[simp] lemma T_univ [Fintype ι] [∀ i, SigmaFinite (μ i)] (f : (∀ i, A i) → ℝ≥0∞) (x : ∀ i, A i) : T μ p f univ x = ∫⁻ (x : ∀ i, A i), (f x ^ (1 - (#ι - 1 : ℝ) * p) * ∏ i : ι, (∫⁻ t : A i, f (update x i t) ∂(μ i)) ^ p) ∂(.pi μ) := by @@ -106,7 +106,7 @@ variable {p : ℝ} The grid-lines operation `GridLines.T` on a nonnegative function on a finitary product type is less than or equal to the grid-lines operation of its partial integral in one co-ordinate (the latter intuitively considered as a function on a space "one dimension down"). -/ -theorem T_insert_le_T_lmarginal_singleton (hp₀ : 0 ≤ p) (s : Finset ι) +theorem T_insert_le_T_lmarginal_singleton [∀ i, SigmaFinite (μ i)] (hp₀ : 0 ≤ p) (s : Finset ι) (hp : (s.card : ℝ) * p ≤ 1) (i : ι) (hi : i ∉ s) {f : (∀ i, A i) → ℝ≥0∞} (hf : Measurable f) : T μ p f (insert i s) ≤ T μ p (∫⋯∫⁻_{i}, f ∂μ) s := by @@ -214,8 +214,8 @@ type indexed by `ι`, and a set `s` in `ι`, consider partially integrating over `sᶜ` and performing the "grid-lines operation" (see `GridLines.T`) to the resulting function in the variables `s`. This theorem states that this operation decreases as the number of grid-lines taken increases. -/ -theorem T_lmarginal_antitone (hp₀ : 0 ≤ p) (hp : (#ι - 1 : ℝ) * p ≤ 1) - {f : (∀ i, A i) → ℝ≥0∞} (hf : Measurable f) : +theorem T_lmarginal_antitone [Fintype ι] [∀ i, SigmaFinite (μ i)] + (hp₀ : 0 ≤ p) (hp : (#ι - 1 : ℝ) * p ≤ 1) {f : (∀ i, A i) → ℝ≥0∞} (hf : Measurable f) : Antitone (fun s ↦ T μ p (∫⋯∫⁻_sᶜ, f ∂μ) s) := by -- Reformulate (by induction): a function is decreasing on `Finset ι` if it decreases under the -- insertion of any element to any set. @@ -254,7 +254,8 @@ direction through `x`. This lemma bounds the Lebesgue integral of the grid-lines quantity by a power of the Lebesgue integral of `f`. -/ -theorem lintegral_mul_prod_lintegral_pow_le {p : ℝ} (hp₀ : 0 ≤ p) +theorem lintegral_mul_prod_lintegral_pow_le + [Fintype ι] [∀ i, SigmaFinite (μ i)] {p : ℝ} (hp₀ : 0 ≤ p) (hp : (#ι - 1 : ℝ) * p ≤ 1) {f : (∀ i : ι, A i) → ℝ≥0∞} (hf : Measurable f) : ∫⁻ x, f x ^ (1 - (#ι - 1 : ℝ) * p) * ∏ i, (∫⁻ xᵢ, f (update x i xᵢ) ∂μ i) ^ p ∂.pi μ ≤ (∫⁻ x, f x ∂.pi μ) ^ (1 + p) := by @@ -266,7 +267,7 @@ theorem lintegral_mul_prod_lintegral_pow_le {p : ℝ} (hp₀ : 0 ≤ p) /-- Special case of the grid-lines lemma `lintegral_mul_prod_lintegral_pow_le`, taking the extremal exponent `p = (#ι - 1)⁻¹`. -/ -theorem lintegral_prod_lintegral_pow_le +theorem lintegral_prod_lintegral_pow_le [Fintype ι] [∀ i, SigmaFinite (μ i)] {p : ℝ} (hp : Real.IsConjExponent #ι p) {f} (hf : Measurable f) : ∫⁻ x, ∏ i, (∫⁻ xᵢ, f (update x i xᵢ) ∂μ i) ^ ((1 : ℝ) / (#ι - 1 : ℝ)) ∂.pi μ @@ -295,7 +296,7 @@ expression `|u x| ^ (n / (n - 1))` is bounded above by the `n / (n - 1)`-th powe integral of the Fréchet derivative of `u`. For a basis-free version, see `lintegral_pow_le_pow_lintegral_fderiv`. -/ -theorem lintegral_pow_le_pow_lintegral_fderiv_aux +theorem lintegral_pow_le_pow_lintegral_fderiv_aux [Fintype ι] {p : ℝ} (hp : Real.IsConjExponent #ι p) {u : (ι → ℝ) → F} (hu : ContDiff ℝ 1 u) (h2u : HasCompactSupport u) : diff --git a/Mathlib/Geometry/Manifold/BumpFunction.lean b/Mathlib/Geometry/Manifold/BumpFunction.lean index 0cc7636a1ddc0..455a203e54316 100644 --- a/Mathlib/Geometry/Manifold/BumpFunction.lean +++ b/Mathlib/Geometry/Manifold/BumpFunction.lean @@ -30,7 +30,7 @@ manifold, smooth bump function universe uE uF uH uM -variable {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] +variable {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] {H : Type uH} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} {M : Type uM} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] @@ -64,6 +64,10 @@ structure SmoothBumpFunction (c : M) extends ContDiffBump (extChartAt I c c) whe namespace SmoothBumpFunction +section FiniteDimensional + +variable [FiniteDimensional ℝ E] + variable {c : M} (f : SmoothBumpFunction I c) {x : M} /-- The function defined by `f : SmoothBumpFunction c`. Use automatic coercion to function @@ -77,6 +81,10 @@ instance : CoeFun (SmoothBumpFunction I c) fun _ => M → ℝ := theorem coe_def : ⇑f = indicator (chartAt H c).source (f.toContDiffBump ∘ extChartAt I c) := rfl +end FiniteDimensional + +variable {c : M} (f : SmoothBumpFunction I c) {x : M} + theorem rOut_pos : 0 < f.rOut := f.toContDiffBump.rOut_pos @@ -89,6 +97,10 @@ theorem ball_inter_range_eq_ball_inter_target : (subset_inter inter_subset_left f.ball_subset).antisymm <| inter_subset_inter_right _ <| extChartAt_target_subset_range _ _ +section FiniteDimensional + +variable [FiniteDimensional ℝ E] + theorem eqOn_source : EqOn f (f.toContDiffBump ∘ extChartAt I c) (chartAt H c).source := eqOn_indicator @@ -173,6 +185,8 @@ theorem isCompact_symm_image_closedBall : ((isCompact_closedBall _ _).inter_right I.isClosed_range).image_of_continuousOn <| (continuousOn_extChartAt_symm _ _).mono f.closedBall_subset +end FiniteDimensional + /-- Given a smooth bump function `f : SmoothBumpFunction I c`, the closed ball of radius `f.R` is known to include the support of `f`. These closed balls (in the model normed space `E`) intersected with `Set.range I` form a basis of `𝓝[range I] (extChartAt I c c)`. -/ @@ -186,6 +200,8 @@ theorem nhdsWithin_range_basis : · exact fun f _ => inter_mem (mem_nhdsWithin_of_mem_nhds <| closedBall_mem_nhds _ f.rOut_pos) self_mem_nhdsWithin +variable [FiniteDimensional ℝ E] + theorem isClosed_image_of_isClosed {s : Set M} (hsc : IsClosed s) (hs : s ⊆ support f) : IsClosed (extChartAt I c '' s) := by rw [f.image_eq_inter_preimage_of_subset_support hs] diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean index 3be66afa23f67..91703740e9643 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean @@ -462,7 +462,7 @@ instance SmoothFiberwiseLinear.hasGroupoid : exact (e.apply_symm_apply_eq_coordChangeL e' hb.1 v).symm /-- A smooth vector bundle `E` is naturally a smooth manifold. -/ -instance Bundle.TotalSpace.smoothManifoldWithCorners : +instance Bundle.TotalSpace.smoothManifoldWithCorners [SmoothManifoldWithCorners IB B] : SmoothManifoldWithCorners (IB.prod 𝓘(𝕜, F)) (TotalSpace F E) := by refine { StructureGroupoid.HasGroupoid.comp (smoothFiberwiseLinear B F IB) ?_ with } intro e he diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index fd82c96b14643..d84856b7e6f41 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -900,7 +900,8 @@ theorem exists_spanning_measurableSet_norm_le [SeminormedAddCommGroup β] {m m0 ∃ s : ℕ → Set α, (∀ n, MeasurableSet[m] (s n) ∧ μ (s n) < ∞ ∧ ∀ x ∈ s n, ‖f x‖ ≤ n) ∧ ⋃ i, s i = Set.univ := by - obtain ⟨s, hs, hs_univ⟩ := exists_spanning_measurableSet_le hf.nnnorm.measurable (μ.trim hm) + obtain ⟨s, hs, hs_univ⟩ := + @exists_spanning_measurableSet_le _ m _ hf.nnnorm.measurable (μ.trim hm) _ refine ⟨s, fun n ↦ ⟨(hs n).1, (le_trim hm).trans_lt (hs n).2.1, fun x hx ↦ ?_⟩, hs_univ⟩ have hx_nnnorm : ‖f x‖₊ ≤ n := (hs n).2.2 x hx rw [← coe_nnnorm] diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 0e612aee0fb86..ff79440e0ab5c 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -57,10 +57,12 @@ open Set Filter TopologicalSpace MeasureTheory Function RCLike open scoped Classical Topology ENNReal NNReal -variable {X Y E F : Type*} [MeasurableSpace X] +variable {X Y E F : Type*} namespace MeasureTheory +variable [MeasurableSpace X] + section NormedAddCommGroup variable [NormedAddCommGroup E] [NormedSpace ℝ E] @@ -1032,7 +1034,8 @@ section OpenPos open Measure -variable [TopologicalSpace X] [OpensMeasurableSpace X] {μ : Measure X} [IsOpenPosMeasure μ] +variable [MeasurableSpace X] [TopologicalSpace X] [OpensMeasurableSpace X] + {μ : Measure X} [IsOpenPosMeasure μ] theorem Continuous.integral_pos_of_hasCompactSupport_nonneg_nonzero [IsFiniteMeasureOnCompacts μ] {f : X → ℝ} {x : X} (f_cont : Continuous f) (f_comp : HasCompactSupport f) (f_nonneg : 0 ≤ f) @@ -1047,7 +1050,7 @@ section FTC open MeasureTheory Asymptotics Metric -variable {ι : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +variable [MeasurableSpace X] {ι : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] /-- Fundamental theorem of calculus for set integrals: if `μ` is a measure that is finite at a filter `l` and @@ -1140,6 +1143,8 @@ end FTC section +variable [MeasurableSpace X] + /-! ### Continuous linear maps composed with integration The goal of this section is to prove that integration commutes with continuous linear maps. @@ -1418,7 +1423,7 @@ end section thickenedIndicator -variable [PseudoEMetricSpace X] +variable [MeasurableSpace X] [PseudoEMetricSpace X] theorem measure_le_lintegral_thickenedIndicatorAux (μ : Measure X) {E : Set X} (E_mble : MeasurableSet E) (δ : ℝ) : μ E ≤ ∫⁻ x, (thickenedIndicatorAux δ E x : ℝ≥0∞) ∂μ := by diff --git a/Mathlib/NumberTheory/Cyclotomic/Rat.lean b/Mathlib/NumberTheory/Cyclotomic/Rat.lean index 88d03d54e67f4..5c6aee417a142 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Rat.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Rat.lean @@ -28,10 +28,12 @@ open Algebra IsCyclotomicExtension Polynomial NumberField open scoped Cyclotomic Nat -variable {p : ℕ+} {k : ℕ} {K : Type u} [Field K] [CharZero K] {ζ : K} [hp : Fact (p : ℕ).Prime] +variable {p : ℕ+} {k : ℕ} {K : Type u} [Field K] {ζ : K} [hp : Fact (p : ℕ).Prime] namespace IsCyclotomicExtension.Rat +variable [CharZero K] + /-- The discriminant of the power basis given by `ζ - 1`. -/ theorem discr_prime_pow_ne_two' [IsCyclotomicExtension {p ^ (k + 1)} ℚ K] (hζ : IsPrimitiveRoot ζ ↑(p ^ (k + 1))) (hk : p ^ (k + 1) ≠ 2) : @@ -145,6 +147,10 @@ open IsCyclotomicExtension.Rat namespace IsPrimitiveRoot +section CharZero + +variable [CharZero K] + /-- The algebra isomorphism `adjoin ℤ {ζ} ≃ₐ[ℤ] (𝓞 K)`, where `ζ` is a primitive `p ^ k`-th root of unity and `K` is a `p ^ k`-th cyclotomic extension of `ℚ`. -/ @[simps!] @@ -169,6 +175,8 @@ noncomputable def integralPowerBasis [IsCyclotomicExtension {p ^ k} ℚ K] /-- Abbreviation to see a primitive root of unity as a member of the ring of integers. -/ abbrev toInteger {k : ℕ+} (hζ : IsPrimitiveRoot ζ k) : 𝓞 K := ⟨ζ, hζ.isIntegral k.pos⟩ +end CharZero + lemma coe_toInteger {k : ℕ+} (hζ : IsPrimitiveRoot ζ k) : hζ.toInteger.1 = ζ := rfl /-- `𝓞 K ⧸ Ideal.span {ζ - 1}` is finite. -/ @@ -194,6 +202,8 @@ lemma toInteger_isPrimitiveRoot {k : ℕ+} (hζ : IsPrimitiveRoot ζ k) : IsPrimitiveRoot hζ.toInteger k := IsPrimitiveRoot.of_map_of_injective (by exact hζ) RingOfIntegers.coe_injective +variable [CharZero K] + -- Porting note: the proof changed because `simp` unfolds too much. @[simp] theorem integralPowerBasis_gen [hcycl : IsCyclotomicExtension {p ^ k} ℚ K] @@ -547,6 +557,7 @@ namespace IsCyclotomicExtension.Rat open nonZeroDivisors IsPrimitiveRoot variable (K p k) +variable [CharZero K] /-- We compute the absolute discriminant of a `p ^ k`-th cyclotomic field. Beware that in the cases `p ^ k = 1` and `p ^ k = 2` the formula uses `1 / 2 = 0` and `0 - 1 = 0`. diff --git a/Mathlib/NumberTheory/Cyclotomic/Three.lean b/Mathlib/NumberTheory/Cyclotomic/Three.lean index a54ac2dce6955..034ec4ce85a40 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Three.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Three.lean @@ -31,7 +31,7 @@ open NumberField Units InfinitePlace nonZeroDivisors Polynomial namespace IsCyclotomicExtension.Rat.Three -variable {K : Type*} [Field K] [NumberField K] [IsCyclotomicExtension {3} ℚ K] +variable {K : Type*} [Field K] variable {ζ : K} (hζ : IsPrimitiveRoot ζ ↑(3 : ℕ+)) (u : (𝓞 K)ˣ) local notation3 "η" => (IsPrimitiveRoot.isUnit (hζ.toInteger_isPrimitiveRoot) (by decide)).unit local notation3 "λ" => hζ.toInteger - 1 @@ -44,7 +44,8 @@ lemma _root_.IsPrimitiveRoot.toInteger_cube_eq_one : hζ.toInteger ^ 3 = 1 := /-- Let `u` be a unit in `(𝓞 K)ˣ`, then `u ∈ [1, -1, η, -η, η^2, -η^2]`. -/ -- Here `List` is more convenient than `Finset`, even if further from the informal statement. -- For example, `fin_cases` below does not work with a `Finset`. -theorem Units.mem : u ∈ [1, -1, η, -η, η ^ 2, -η ^ 2] := by +theorem Units.mem [NumberField K] [IsCyclotomicExtension {3} ℚ K] : + u ∈ [1, -1, η, -η, η ^ 2, -η ^ 2] := by have hrank : rank K = 0 := by dsimp only [rank] rw [card_eq_nrRealPlaces_add_nrComplexPlaces, nrRealPlaces_eq_zero (n := 3) K (by decide), @@ -89,7 +90,8 @@ lemma eta_sq : (η ^ 2 : 𝓞 K) = - η - 1 := by /-- If a unit `u` is congruent to an integer modulo `λ ^ 2`, then `u = 1` or `u = -1`. This is a special case of the so-called *Kummer's lemma*. -/ -theorem eq_one_or_neg_one_of_unit_of_congruent (hcong : ∃ n : ℤ, λ ^ 2 ∣ (u - n : 𝓞 K)) : +theorem eq_one_or_neg_one_of_unit_of_congruent + [NumberField K] [IsCyclotomicExtension {3} ℚ K] (hcong : ∃ n : ℤ, λ ^ 2 ∣ (u - n : 𝓞 K)) : u = 1 ∨ u = -1 := by replace hcong : ∃ n : ℤ, (3 : 𝓞 K) ∣ (↑u - n : 𝓞 K) := by obtain ⟨n, x, hx⟩ := hcong @@ -120,7 +122,8 @@ theorem eq_one_or_neg_one_of_unit_of_congruent (hcong : ∃ n : ℤ, λ ^ 2 ∣ variable (x : 𝓞 K) /-- Let `(x : 𝓞 K)`. Then we have that `λ` divides one amongst `x`, `x - 1` and `x + 1`. -/ -lemma lambda_dvd_or_dvd_sub_one_or_dvd_add_one : λ ∣ x ∨ λ ∣ x - 1 ∨ λ ∣ x + 1 := by +lemma lambda_dvd_or_dvd_sub_one_or_dvd_add_one [NumberField K] [IsCyclotomicExtension {3} ℚ K] : + λ ∣ x ∨ λ ∣ x - 1 ∨ λ ∣ x + 1 := by classical have := hζ.finite_quotient_toInteger_sub_one (by decide) let _ := Fintype.ofFinite (𝓞 K ⧸ Ideal.span {λ}) @@ -156,6 +159,8 @@ lemma cube_sub_one_eq_mul : x ^ 3 - 1 = (x - 1) * (x - η) * (x - η ^ 2) := by simp [hζ.toInteger_cube_eq_one] _ = x ^ 3 - 1 := by rw [eta_sq_add_eta_add_one hζ]; ring +variable [NumberField K] [IsCyclotomicExtension {3} ℚ K] + /-- We have that `λ` divides `x * (x - 1) * (x - (η + 1))`. -/ lemma lambda_dvd_mul_sub_one_mul_sub_eta_add_one : λ ∣ x * (x - 1) * (x - (η + 1)) := by rcases lambda_dvd_or_dvd_sub_one_or_dvd_add_one hζ x with h | h | h diff --git a/Mathlib/NumberTheory/FLT/Three.lean b/Mathlib/NumberTheory/FLT/Three.lean index 55759a13c44a7..09088e17b9fee 100644 --- a/Mathlib/NumberTheory/FLT/Three.lean +++ b/Mathlib/NumberTheory/FLT/Three.lean @@ -146,7 +146,7 @@ section eisenstein open NumberField IsCyclotomicExtension.Rat.Three -variable {K : Type*} [Field K] [NumberField K] [IsCyclotomicExtension {3} ℚ K] +variable {K : Type*} [Field K] variable {ζ : K} (hζ : IsPrimitiveRoot ζ (3 : ℕ+)) local notation3 "η" => (IsPrimitiveRoot.isUnit (hζ.toInteger_isPrimitiveRoot) (by decide)).unit @@ -160,7 +160,8 @@ def FermatLastTheoremForThreeGen : Prop := a ^ 3 + b ^ 3 ≠ u * c ^ 3 /-- To prove `FermatLastTheoremFor 3`, it is enough to prove `FermatLastTheoremForThreeGen`. -/ -lemma FermatLastTheoremForThree_of_FermatLastTheoremThreeGen : +lemma FermatLastTheoremForThree_of_FermatLastTheoremThreeGen + [NumberField K] [IsCyclotomicExtension {3} ℚ K] : FermatLastTheoremForThreeGen hζ → FermatLastTheoremFor 3 := by intro H refine fermatLastTheoremThree_of_three_dvd_only_c (fun a b c hc ha hb ⟨x, hx⟩ hcoprime h ↦ ?_) @@ -201,13 +202,21 @@ structure Solution extends Solution' hζ where hab : λ ^ 2 ∣ a + b variable {hζ} -variable (S : Solution hζ) (S' : Solution' hζ) [DecidableRel fun (a b : 𝓞 K) ↦ a ∣ b] +variable (S : Solution hζ) (S' : Solution' hζ) + +section IsCyclotomicExtension + +variable [NumberField K] [IsCyclotomicExtension {3} ℚ K] /-- For any `S' : Solution'`, the multiplicity of `λ` in `S'.c` is finite. -/ lemma Solution'.multiplicity_lambda_c_finite : multiplicity.Finite (hζ.toInteger - 1) S'.c := multiplicity.finite_of_not_isUnit hζ.zeta_sub_one_prime'.not_unit S'.hc +section DecidableRel + +variable [DecidableRel fun (a b : 𝓞 K) ↦ a ∣ b] + /-- Given `S' : Solution'`, `S'.multiplicity` is the multiplicity of `λ` in `S'.c`, as a natural number. -/ def Solution'.multiplicity := @@ -228,6 +237,8 @@ lemma Solution.exists_minimal : ∃ (S₁ : Solution hζ), S₁.isMinimal := by rcases Nat.find_spec (⟨S.multiplicity, ⟨S, rfl⟩⟩ : T.Nonempty) with ⟨S₁, hS₁⟩ exact ⟨S₁, fun S'' ↦ hS₁ ▸ Nat.find_min' _ ⟨S'', rfl⟩⟩ +end DecidableRel + /-- Given `S' : Solution'`, then `S'.a` and `S'.b` are both congruent to `1` modulo `λ ^ 4` or are both congruent to `-1`. -/ lemma a_cube_b_cube_congr_one_or_neg_one : @@ -267,6 +278,10 @@ lemma lambda_pow_four_dvd_c_cube : λ ^ 4 ∣ S'.c ^ 3 := by _ = S'.u⁻¹ * (S'.u * S'.c ^ 3) := by rw [S'.H] _ = S'.c ^ 3 := by simp +section DecidableRel + +variable [DecidableRel fun (a b : 𝓞 K) ↦ a ∣ b] + /-- Given `S' : Solution'`, we have that `λ ^ 2` divides `S'.c`. -/ lemma lambda_sq_dvd_c : λ ^ 2 ∣ S'.c := by have hm := S'.multiplicity_lambda_c_finite @@ -293,6 +308,10 @@ lemma Solution'.two_le_multiplicity : 2 ≤ S'.multiplicity := by lemma Solution.two_le_multiplicity : 2 ≤ S.multiplicity := S.toSolution'.two_le_multiplicity +end DecidableRel + +end IsCyclotomicExtension + /-- Given `S' : Solution'`, the key factorization of `S'.a ^ 3 + S'.b ^ 3`. -/ lemma a_cube_add_b_cube_eq_mul : S'.a ^ 3 + S'.b ^ 3 = (S'.a + S'.b) * (S'.a + η * S'.b) * (S'.a + η ^ 2 * S'.b) := by @@ -302,6 +321,10 @@ lemma a_cube_add_b_cube_eq_mul : simp [hζ.toInteger_cube_eq_one] _ = S'.a ^ 3 + S'.b ^ 3 := by rw [eta_sq]; ring +section DecidableRel + +variable [NumberField K] [IsCyclotomicExtension {3} ℚ K] [DecidableRel fun (a b : 𝓞 K) ↦ a ∣ b] + open PartENat in /-- Given `S' : Solution'`, we have that `λ ^ 2` divides one amongst `S'.a + S'.b`, `S'.a + η * S'.b` and `S'.a + η ^ 2 * S'.b`. -/ @@ -370,6 +393,8 @@ lemma exists_Solution_of_Solution' : ∃ (S₁ : Solution hζ), S₁.multiplicit H := H hab := hab }, rfl⟩ +end DecidableRel + namespace Solution lemma a_add_eta_mul_b : S.a + η * S.b = (S.a + S.b) + λ * S.b := by rw [coe_eta]; ring @@ -384,6 +409,10 @@ lemma lambda_dvd_a_add_eta_sq_mul_b : λ ∣ (S.a + η ^ 2 * S.b) := by exact dvd_add (dvd_add (dvd_trans (dvd_pow_self _ (by decide)) S.hab) ⟨λ * S.b, by ring⟩) ⟨2 * S.b, by ring⟩ +section IsCyclotomicExtension + +variable [NumberField K] [IsCyclotomicExtension {3} ℚ K] + /-- Given `(S : Solution)`, we have that `λ ^ 2` does not divide `S.a + η * S.b`. -/ lemma lambda_sq_not_dvd_a_add_eta_mul_b : ¬ λ ^ 2 ∣ (S.a + η * S.b) := by simp_rw [a_add_eta_mul_b, dvd_add_right S.hab, pow_two, mul_dvd_mul_iff_left @@ -445,6 +474,8 @@ lemma associated_of_dvd_a_add_eta_mul_b_of_dvd_a_add_eta_sq_mul_b {p : 𝓞 K} ( _ = 2 * η.1 ^ 2 + 3 * η + 1 := by ring _ = λ := by rw [eta_sq, coe_eta]; ring +end IsCyclotomicExtension + /-- Given `S : Solution`, we let `S.y` be any element such that `S.a + η * S.b = λ * S.y` -/ private noncomputable def y := (lambda_dvd_a_add_eta_mul_b S).choose private lemma y_spec : S.a + η * S.b = λ * S.y := @@ -455,6 +486,8 @@ private noncomputable def z := (lambda_dvd_a_add_eta_sq_mul_b S).choose private lemma z_spec : S.a + η ^ 2 * S.b = λ * S.z := (lambda_dvd_a_add_eta_sq_mul_b S).choose_spec +variable [NumberField K] [IsCyclotomicExtension {3} ℚ K] + private lemma lambda_not_dvd_y : ¬ λ ∣ S.y := fun h ↦ by replace h := mul_dvd_mul_left ((η : 𝓞 K) - 1) h rw [coe_eta, ← y_spec, ← pow_two] at h @@ -465,6 +498,10 @@ private lemma lambda_not_dvd_z : ¬ λ ∣ S.z := fun h ↦ by rw [coe_eta, ← z_spec, ← pow_two] at h exact lambda_sq_not_dvd_a_add_eta_sq_mul_b _ h +section DecidableRel + +variable [DecidableRel fun (a b : 𝓞 K) ↦ a ∣ b] + /-- We have that `λ ^ (3*S.multiplicity-2)` divides `S.a + S.b`. -/ private lemma lambda_pow_dvd_a_add_b : λ ^ (3 * S.multiplicity - 2) ∣ S.a + S.b := by have h : λ ^ S.multiplicity ∣ S.c := multiplicity.pow_multiplicity_dvd _ @@ -510,6 +547,10 @@ private lemma lambda_not_dvd_x : ¬ λ ∣ S.x := fun h ↦ by · exact lambda_not_dvd_w _ <| hζ.zeta_sub_one_prime'.dvd_of_dvd_pow h · simp [hζ.zeta_sub_one_prime'.ne_zero] +end DecidableRel + +attribute [local instance] IsCyclotomicExtension.Rat.three_pid + private lemma isCoprime_helper {r s t w : 𝓞 K} (hr : ¬ λ ∣ r) (hs : ¬ λ ∣ s) (Hp : ∀ {p}, Prime p → p ∣ t → p ∣ w → Associated p λ) (H₁ : ∀ {q}, q ∣ r → q ∣ t) (H₂ : ∀ {q}, q ∣ s → q ∣ w) : IsCoprime r s := by @@ -517,12 +558,12 @@ private lemma isCoprime_helper {r s t w : 𝓞 K} (hr : ¬ λ ∣ r) (hs : ¬ λ (fun p hp p_dvd_r p_dvd_s ↦ hr ?_) rwa [← Associated.dvd_iff_dvd_left <| Hp hp (H₁ p_dvd_r) (H₂ p_dvd_s)] -private lemma isCoprime_x_y : IsCoprime S.x S.y := +private lemma isCoprime_x_y [DecidableRel fun (a b : 𝓞 K) ↦ a ∣ b] : IsCoprime S.x S.y := isCoprime_helper (lambda_not_dvd_x S) (lambda_not_dvd_y S) (associated_of_dvd_a_add_b_of_dvd_a_add_eta_mul_b S) (fun hq ↦ x_spec S ▸ hq.mul_left _) (fun hq ↦ y_spec S ▸ hq.mul_left _) -private lemma isCoprime_x_z : IsCoprime S.x S.z := +private lemma isCoprime_x_z [DecidableRel fun (a b : 𝓞 K) ↦ a ∣ b] : IsCoprime S.x S.z := isCoprime_helper (lambda_not_dvd_x S) (lambda_not_dvd_z S) (associated_of_dvd_a_add_b_of_dvd_a_add_eta_sq_mul_b S) (fun hq ↦ x_spec S ▸ hq.mul_left _) (fun hq ↦ z_spec S ▸ hq.mul_left _) @@ -532,6 +573,8 @@ private lemma isCoprime_y_z : IsCoprime S.y S.z := (associated_of_dvd_a_add_eta_mul_b_of_dvd_a_add_eta_sq_mul_b S) (fun hq ↦ y_spec S ▸ hq.mul_left _) (fun hq ↦ z_spec S ▸ hq.mul_left _) +variable [DecidableRel fun (a b : 𝓞 K) ↦ a ∣ b] + private lemma x_mul_y_mul_z_eq_u_mul_w_cube : S.x * S.y * S.z = S.u * S.w ^ 3 := by suffices hh : λ ^ (3 * S.multiplicity - 2) * S.x * λ * S.y * λ * S.z = S.u * λ ^ (3 * S.multiplicity) * S.w ^ 3 by @@ -638,8 +681,6 @@ private lemma lambda_sq_div_u₅_mul : λ ^ 2 ∣ S.u₅ * (λ ^ (S.multiplicity _ = λ^2*λ^(3*S.multiplicity-5)*S.u₅*S.X^3 := by rw [this, pow_add] _ = λ^2*(λ^(3*S.multiplicity-5)*S.u₅*S.X^3) := by ring -variable [DecidableEq (𝓞 K)] - private lemma u₄_eq_one_or_neg_one : S.u₄ = 1 ∨ S.u₄ = -1 := by have : λ^2 ∣ λ^4 := ⟨λ^2, by ring⟩ have h := S.lambda_sq_div_u₅_mul diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 9c9cc81cf80ee..7df0582059f8b 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -478,6 +478,8 @@ theorem mkReal_coe (φ : { φ : K →+* ℂ // ComplexEmbedding.IsReal φ }) : theorem mkComplex_coe (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ }) : (mkComplex φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl +section NumberField + variable [NumberField K] /-- The infinite part of the product formula : for `x ∈ K`, we have `Π_w ‖x‖_w = |norm(x)|` where @@ -542,10 +544,16 @@ theorem _root_.NumberField.adjoin_eq_top_of_infinitePlace_lt {x : 𝓞 K} {w : I exact congr_arg IntermediateField.toSubalgebra <| NumberField.is_primitive_element_of_infinitePlace_lt h₁ h₂ h₃ +end NumberField + open Fintype FiniteDimensional variable (K) +section NumberField + +variable [NumberField K] + /-- The number of infinite real places of the number field `K`. -/ noncomputable abbrev NrRealPlaces := card { w : InfinitePlace K // IsReal w } @@ -599,6 +607,10 @@ def comap (w : InfinitePlace K) (f : k →+* K) : InfinitePlace k := ⟨w.1.comp f.injective, w.embedding.comp f, by { ext x; show _ = w.1 (f x); rw [← w.2.choose_spec]; rfl }⟩ +end NumberField + +variable {K} + @[simp] lemma comap_mk (φ : K →+* ℂ) (f : k →+* K) : (mk φ).comap f = mk (φ.comp f) := rfl @@ -630,8 +642,7 @@ lemma mult_comap_le (f : k →+* K) (w : InfinitePlace K) : mult (w.comap f) ≤ · exact (h₁ (h₂.comap _)).elim all_goals decide -variable [Algebra k K] [Algebra k F] [Algebra K F] [IsScalarTower k K F] -variable (σ : K ≃ₐ[k] K) (w : InfinitePlace K) +variable [Algebra k K] (σ : K ≃ₐ[k] K) (w : InfinitePlace K) variable (k K) lemma card_mono [NumberField k] [NumberField K] : @@ -663,7 +674,7 @@ lemma isReal_smul_iff : IsReal (σ • w) ↔ IsReal w := isReal_comap_iff (f := lemma isComplex_smul_iff : IsComplex (σ • w) ↔ IsComplex w := by rw [← not_isReal_iff_isComplex, ← not_isReal_iff_isComplex, isReal_smul_iff] -lemma ComplexEmbedding.exists_comp_symm_eq_of_comp_eq [Algebra k K] [IsGalois k K] (φ ψ : K →+* ℂ) +lemma ComplexEmbedding.exists_comp_symm_eq_of_comp_eq [IsGalois k K] (φ ψ : K →+* ℂ) (h : φ.comp (algebraMap k K) = ψ.comp (algebraMap k K)) : ∃ σ : K ≃ₐ[k] K, φ.comp σ.symm = ψ := by letI := (φ.comp (algebraMap k K)).toAlgebra @@ -674,7 +685,7 @@ lemma ComplexEmbedding.exists_comp_symm_eq_of_comp_eq [Algebra k K] [IsGalois k ext1 x exact AlgHom.restrictNormal_commutes ψ' K x -lemma exists_smul_eq_of_comap_eq [Algebra k K] [IsGalois k K] {w w' : InfinitePlace K} +lemma exists_smul_eq_of_comap_eq [IsGalois k K] {w w' : InfinitePlace K} (h : w.comap (algebraMap k K) = w'.comap (algebraMap k K)) : ∃ σ : K ≃ₐ[k] K, σ • w = w' := by rw [← mk_embedding w, ← mk_embedding w', comap_mk, comap_mk, mk_eq_iff] at h cases h with @@ -732,12 +743,15 @@ lemma isUnramified_iff_mult_le : rw [IsUnramified, le_antisymm_iff, and_iff_right] exact mult_comap_le _ _ +variable [Algebra k F] + lemma IsUnramified.comap_algHom {w : InfinitePlace F} (h : IsUnramified k w) (f : K →ₐ[k] F) : IsUnramified k (w.comap (f : K →+* F)) := by rw [InfinitePlace.isUnramified_iff_mult_le, ← InfinitePlace.comap_comp, f.comp_algebraMap, h.eq] exact InfinitePlace.mult_comap_le _ _ variable (K) +variable [Algebra K F] [IsScalarTower k K F] lemma IsUnramified.of_restrictScalars {w : InfinitePlace F} (h : IsUnramified k w) : IsUnramified K w := by @@ -912,6 +926,7 @@ lemma even_finrank_of_not_isUnramifiedIn exact even_finrank_of_not_isUnramified hw variable (k K) +variable [NumberField K] open Finset in lemma card_isUnramified [NumberField k] [IsGalois k K] : diff --git a/Mathlib/NumberTheory/NumberField/EquivReindex.lean b/Mathlib/NumberTheory/NumberField/EquivReindex.lean index 49dbcfc3ef6ec..67fc4926a8307 100644 --- a/Mathlib/NumberTheory/NumberField/EquivReindex.lean +++ b/Mathlib/NumberTheory/NumberField/EquivReindex.lean @@ -35,9 +35,7 @@ abbrev equivReindex : (K →+* ℂ) ≃ (ChooseBasisIndex ℤ (𝓞 K)) := abbrev basisMatrix : Matrix (K →+* ℂ) (K →+* ℂ) ℂ := (Matrix.of fun i ↦ latticeBasis K (equivReindex K i)) -variable [DecidableEq (K →+* ℂ)] - -theorem det_of_basisMatrix_non_zero : (basisMatrix K).det ≠ 0 := by +theorem det_of_basisMatrix_non_zero [DecidableEq (K →+* ℂ)] : (basisMatrix K).det ≠ 0 := by let e : (K →+* ℂ) ≃ ChooseBasisIndex ℤ (𝓞 K) := equivReindex K let N := Algebra.embeddingsMatrixReindex ℚ ℂ (fun i => integralBasis K (e i)) RingHom.equivRatAlgHom @@ -50,7 +48,7 @@ theorem det_of_basisMatrix_non_zero : (basisMatrix K).det ≠ 0 := by exact (Algebra.discr_eq_det_embeddingsMatrixReindex_pow_two ℚ ℂ (fun _ => integralBasis K (e _)) RingHom.equivRatAlgHom).symm -instance : Invertible (basisMatrix K) := invertibleOfIsUnitDet _ +instance [DecidableEq (K →+* ℂ)] : Invertible (basisMatrix K) := invertibleOfIsUnitDet _ (Ne.isUnit (det_of_basisMatrix_non_zero K)) variable {K} @@ -64,7 +62,7 @@ theorem canonicalEmbedding_eq_basisMatrix_mulVec (α : K) : transpose_apply, of_apply, Fintype.sum_apply, mul_comm, Basis.repr_reindex, Finsupp.mapDomain_equiv_apply, Equiv.symm_symm, Pi.smul_apply, smul_eq_mul] -theorem inverse_basisMatrix_mulVec_eq_repr (α : 𝓞 K) : +theorem inverse_basisMatrix_mulVec_eq_repr [DecidableEq (K →+* ℂ)] (α : 𝓞 K) : ∀ i, ((basisMatrix K).transpose)⁻¹.mulVec (fun j => canonicalEmbedding K (algebraMap (𝓞 K) K α) j) i = ((integralBasis K).reindex (equivReindex K).symm).repr α i := fun i => by diff --git a/Mathlib/NumberTheory/NumberField/House.lean b/Mathlib/NumberTheory/NumberField/House.lean index d3aa73c74a420..87d342f35542d 100644 --- a/Mathlib/NumberTheory/NumberField/House.lean +++ b/Mathlib/NumberTheory/NumberField/House.lean @@ -62,12 +62,14 @@ noncomputable section variable (K) -variable [DecidableEq (K →+* ℂ)] - open Module.Free FiniteDimensional canonicalEmbedding Matrix Finset attribute [local instance] Matrix.seminormedAddCommGroup +section DecidableEq + +variable [DecidableEq (K →+* ℂ)] + /-- `c` is defined as the product of the maximum absolute value of the entries of the inverse of the matrix `basisMatrix` and `finrank ℚ K`. -/ private def c := (finrank ℚ K) * ‖((basisMatrix K).transpose)⁻¹‖ @@ -107,13 +109,13 @@ private def newBasis := (RingOfIntegers.basis K).reindex (equivReindex K).symm private def supOfBasis : ℝ := univ.sup' univ_nonempty fun r ↦ house (algebraMap (𝓞 K) K (newBasis K r)) +end DecidableEq + private theorem supOfBasis_nonneg : 0 ≤ supOfBasis K := by simp only [supOfBasis, le_sup'_iff, mem_univ, and_self, exists_const, house_nonneg] -variable {α : Type*} {β : Type*} [Fintype α] [Fintype β] - -variable (a : Matrix α β (𝓞 K)) +variable {α : Type*} {β : Type*} (a : Matrix α β (𝓞 K)) /-- `a' K a` returns the integer coefficients of the basis vector in the expansion of the product of an algebraic integer and a basis vectors. -/ @@ -144,13 +146,7 @@ private theorem asiegel_ne_0 : asiegel K a ≠ 0 := by simp only [mul_eq_zero] at this exact this.resolve_right (Basis.ne_zero (newBasis K) b) -variable {p q : ℕ} - (cardα : Fintype.card α = p) (cardβ : Fintype.card β = q) - (h0p : 0 < p) (hpq : p < q) - (x : β × (K →+* ℂ) → ℤ) - (hxl : x ≠ 0) - (hmulvec0 : asiegel K a *ᵥ x = 0) - (hxbound : ‖x‖ ≤ (q * finrank ℚ K * ‖asiegel K a‖) ^ ((p : ℝ) / (q - p))) +variable {p q : ℕ} (h0p : 0 < p) (hpq : p < q) (x : β × (K →+* ℂ) → ℤ) (hxl : x ≠ 0) /-- `ξ` is the the product of `x (l, r)` and the `r`-th basis element of the newBasis of `K`. -/ private def ξ : β → 𝓞 K := fun l => ∑ r : K →+* ℂ, x (l, r) * (newBasis K r) @@ -168,6 +164,8 @@ private theorem lin_1 (l k r) : a k l * (newBasis K) r = ∑ u, (a' K a k l r u) * (newBasis K) u := by simp only [Basis.sum_repr (newBasis K) (a k l * (newBasis K) r), a', ← zsmul_eq_mul] +variable [Fintype β] (cardβ : Fintype.card β = q) (hmulvec0 : asiegel K a *ᵥ x = 0) + private theorem ξ_mulVec_eq_0 : a *ᵥ ξ K x = 0 := by funext k; simp only [Pi.zero_apply]; rw [eq_comm] @@ -200,13 +198,16 @@ private theorem ξ_mulVec_eq_0 : a *ᵥ ξ K x = 0 := by variable {A : ℝ} (habs : ∀ k l, (house ((algebraMap (𝓞 K) K) (a k l))) ≤ A) +variable [DecidableEq (K →+* ℂ)] + /-- `c₂` is the product of the maximum of `1` and `c`, and `supOfBasis`. -/ private abbrev c₂ := max 1 (c K) * (supOfBasis K) private theorem c₂_nonneg : 0 ≤ c₂ K := mul_nonneg (le_trans zero_le_one (le_max_left ..)) (supOfBasis_nonneg _) -variable (Apos : 0 ≤ A) +variable [Fintype α] (cardα : Fintype.card α = p) (Apos : 0 ≤ A) + (hxbound : ‖x‖ ≤ (q * finrank ℚ K * ‖asiegel K a‖) ^ ((p : ℝ) / (q - p))) private theorem asiegel_remark : ‖asiegel K a‖ ≤ c₂ K * A := by rw [Matrix.norm_le_iff] @@ -305,7 +306,7 @@ theorem exists_ne_zero_int_vec_house_le : have ⟨l⟩ := Fintype.card_pos_iff.1 (cardβ ▸ h0p.trans hpq) exact le_trans (house_nonneg _) (habs k l) use ξ K x, ξ_ne_0 K x hxl, ξ_mulVec_eq_0 K a x hxl hmulvec0, - house_le_bound K a hpq x hxbound habs Apos + house_le_bound K a hpq x habs Apos hxbound end diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 868816ee37ad9..5a609f303cc0d 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -42,7 +42,7 @@ noncomputable section open NumberField NumberField.InfinitePlace NumberField.Units BigOperators -variable (K : Type*) [Field K] [NumberField K] +variable (K : Type*) [Field K] namespace NumberField.Units.dirichletUnitTheorem @@ -64,6 +64,10 @@ open Finset variable {K} +section NumberField + +variable [NumberField K] + /-- The distinguished infinite place. -/ def w₀ : InfinitePlace K := (inferInstance : Nonempty (InfinitePlace K)).some @@ -95,6 +99,8 @@ theorem sum_logEmbedding_component (x : (𝓞 K)ˣ) : · norm_num · exact fun w _ => pow_ne_zero _ (AbsoluteValue.ne_zero _ (coe_ne_zero x)) +end NumberField + theorem mult_log_place_eq_zero {x : (𝓞 K)ˣ} {w : InfinitePlace K} : mult w * Real.log (w x) = 0 ↔ w x = 1 := by rw [mul_eq_zero, or_iff_right, Real.log_eq_zero, or_iff_right, or_iff_left] @@ -103,6 +109,8 @@ theorem mult_log_place_eq_zero {x : (𝓞 K)ˣ} {w : InfinitePlace K} : · refine (ne_of_gt ?_) rw [mult]; split_ifs <;> norm_num +variable [NumberField K] + theorem logEmbedding_eq_zero_iff {x : (𝓞 K)ˣ} : logEmbedding K x = 0 ↔ x ∈ torsion K := by rw [mem_torsion] From a180e024b196dd867dd317c18b03cbc638d3d67e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 8 Aug 2024 13:48:14 +0000 Subject: [PATCH 109/359] chore: removing unused Mathport files (#15189) This does not yet remove `Mathport.Syntax`. We may want to preserve the information about unported tactics implicitly stored there (although this information is quite messy, in large part out of date, and in any case preserved in the git history if anyone wants to do the archaeology). --- Mathlib.lean | 2 - Mathlib/Init/Logic.lean | 1 - Mathlib/Mathport/Attributes.lean | 16 -- Mathlib/Mathport/README.md | 6 - Mathlib/Mathport/Rename.lean | 274 ------------------------------- 5 files changed, 299 deletions(-) delete mode 100644 Mathlib/Mathport/Attributes.lean delete mode 100644 Mathlib/Mathport/Rename.lean diff --git a/Mathlib.lean b/Mathlib.lean index 94fc9f82bcb87..81517a645910c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3070,9 +3070,7 @@ import Mathlib.Logic.Small.Ring import Mathlib.Logic.Small.Set import Mathlib.Logic.Unique import Mathlib.Logic.UnivLE -import Mathlib.Mathport.Attributes import Mathlib.Mathport.Notation -import Mathlib.Mathport.Rename import Mathlib.Mathport.Syntax import Mathlib.MeasureTheory.Category.MeasCat import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic diff --git a/Mathlib/Init/Logic.lean b/Mathlib/Init/Logic.lean index 7f19a88f52634..0496c9b280296 100644 --- a/Mathlib/Init/Logic.lean +++ b/Mathlib/Init/Logic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Floris van Doorn -/ import Mathlib.Tactic.Lemma -import Mathlib.Mathport.Attributes import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.ProjectionNotation import Batteries.Tactic.Alias diff --git a/Mathlib/Mathport/Attributes.lean b/Mathlib/Mathport/Attributes.lean deleted file mode 100644 index 1fabfc5668c4a..0000000000000 --- a/Mathlib/Mathport/Attributes.lean +++ /dev/null @@ -1,16 +0,0 @@ -/- -Copyright (c) 2021 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro --/ -import Lean.Attributes - -/-! -# Defines the "substitution" attribute for mathport - -This has to be defined in a separate file. --/ - -namespace Lean.Attr - -initialize substAttr : TagAttribute ← registerTagAttribute `subst "substitution" diff --git a/Mathlib/Mathport/README.md b/Mathlib/Mathport/README.md index 7dc2c3c33481d..4a25af5bc8789 100644 --- a/Mathlib/Mathport/README.md +++ b/Mathlib/Mathport/README.md @@ -5,12 +5,6 @@ This directory contains instructions for `mathport`, helping it to translate `mathlib` and align declarations and tactics with `mathlib4`. (These files were formerly part of `mathport`, in the directory `Mathport/Prelude/`.) -`SpecialNames.lean` -: Contains `#align X Y` statements, where `X` is an identifier from mathlib3 - which should be aligned with the identifier `Y` from mathlib4. - Sometimes we need `#align` statements just to handle exceptions to casing rules, - but there are also many exceptional cases. - `Syntax.lean` : Contains unimplemented stubs of tactics which need to be migrated from Lean3 to Lean4. When porting tactics, you can move the relevant stubs to a new file and diff --git a/Mathlib/Mathport/Rename.lean b/Mathlib/Mathport/Rename.lean deleted file mode 100644 index 6cd246b84d732..0000000000000 --- a/Mathlib/Mathport/Rename.lean +++ /dev/null @@ -1,274 +0,0 @@ -/- -Copyright (c) 2021 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Daniel Selsam --/ -import Lean.Elab.Command -import Lean.Linter.Util - -/-! -# Mathport infrastructure for tracking renaming Lean 3 to Lean 4 names - -This file defines mathport infrastructure for tracking renaming of Lean 3 declarations to -their Lean 4 counterparts. This correspondence is declared in the ported file using the -`#align` command (and its variants with `ₓ` and `#noalign`). --/ - -namespace Mathlib.Prelude.Rename - -open Lean -open System (FilePath) -open Lean (HashMap) - -/-- This structure keeps track of alignments from lean 3 names to lean 4 names and vice versa. -/ -structure RenameMap where - /-- This maps `n3 ↦ (dubious, n4)` where `n3` is the lean 3 name and `n4` is the corresponding - lean 4 name. `dubious` is either empty, or a warning message to be displayed when `n3` is - translated, which indicates that the translation from `n3` to `n4` is approximate and may cause - downstream errors. -/ - toLean4 : NameMap (String × Name) := {} - /-- This maps `n4 ↦ (n3, clashes)` where `n4` is the lean 4 name and `n3::clashes` is the list of - all (non-`synthetic`) declarations that map to `n4`. (That is, we do not assume the mapping - from lean 3 to lean 4 name is injective.) -/ - toLean3 : NameMap (Name × List Name) := {} - deriving Inhabited - -/-- An `olean` entry for the rename extension. -/ -structure NameEntry where - /-- The lean 3 name. -/ - n3 : Name - /-- The lean 4 name, or `.anonymous` for a `#noalign`. -/ - n4 : Name - /-- If true, this lean 3 -> lean 4 mapping will not be entered into the converse map. - This is used for "internal" definitions that should never be referred to in the source syntax. -/ - synthetic := false - /-- A dubious translation is one where there is a type mismatch - from the translated lean 3 definition to a pre-existing lean 4 definition. - Type errors are likely in downstream theorems. - The string stored here is printed in the event that `n3` is encountered by synport. -/ - dubious := "" - -/-- Insert a name entry into the `RenameMap`. -/ -def RenameMap.insert (m : RenameMap) (e : NameEntry) : RenameMap := - let ⟨to4, to3⟩ := m - let to4 := to4.insert e.n3 (e.dubious, e.n4) - let to3 := if e.synthetic || e.n4.isAnonymous then to3 else - match to3.find? e.n4 with - | none => to3.insert e.n4 (e.n3, []) - | some (a, l) => if (a::l).contains e.n3 then to3 else to3.insert e.n4 (a, e.n3 :: l) - ⟨to4, to3⟩ - -/-- Look up a lean 4 name from the lean 3 name. Also return the `dubious` error message. -/ -def RenameMap.find? (m : RenameMap) : Name → Option (String × Name) := m.toLean4.find? - -universe u in --- TODO: upstream into core/std -instance {α : Type u} [Inhabited α] : Inhabited (Thunk α) where - default := .pure default - -/-- This extension stores the lookup data generated from `#align` commands. -/ --- wrap state in `Thunk` as downstream projects rarely actually query it; it only --- gets queried when new `#align`s are added. -initialize renameExtension : SimplePersistentEnvExtension NameEntry (Thunk RenameMap) ← - registerSimplePersistentEnvExtension { - addEntryFn := fun t e => t.map (·.insert e) - addImportedFn := fun es => ⟨fun _ => mkStateFromImportedEntries (·.insert) {} es⟩ - } - -/-- Insert a new name alignment into the rename extension. -/ -def addNameAlignment (n3 : Name) (n4 : Name) (synthetic := false) (dubious := "") : CoreM Unit := do - modifyEnv fun env => renameExtension.addEntry env { n3, n4, synthetic, dubious } - -/-- The `@[binport]` attribute should not be added manually, it is added automatically by mathport -to definitions that it created based on a lean 3 definition (as opposed to pre-existing -definitions). -/ -initialize binportTag : TagAttribute ← - registerTagAttribute `binport "this definition was autogenerated by mathport" - -/-- -Removes all occurrences of `ₓ` from the name. -This is the same processing used by mathport to generate name references, -and declarations with `ₓ` are used to align declarations that do not defeq match the originals. --/ -def removeX : Name → Name - | .anonymous => .anonymous - | .str p s => - let s := if s.contains 'ₓ' then - s.foldl (fun acc c => if c = 'ₓ' then acc else acc.push c) "" - else s - .str (removeX p) s - | .num p n => .num (removeX p) n - -open Lean.Elab Lean.Elab.Command - -/-- Because lean 3 uses a lowercase snake case convention, it is expected that all lean 3 -declaration names should use lowercase, with a few rare exceptions for categories and the set union -operator. This linter warns if you use uppercase in the lean 3 part of an `#align` statement, -because this is most likely a typo. But if the declaration actually uses capitals it is not unusual -to disable this lint locally or at file scope. -/ -register_option linter.uppercaseLean3 : Bool := { - defValue := true - descr := "enable the lean 3 casing lint" -} - -/-- Check that the referenced lean 4 definition exists in an `#align` directive. -/ -register_option align.precheck : Bool := { - defValue := true - descr := "Check that the referenced lean 4 definition exists in an `#align` directive." -} - -/-- -`#align lean_3.def_name Lean4.defName` will record an "alignment" from the lean 3 name -to the corresponding lean 4 name. This information is used by the -[mathport](https://github.com/leanprover-community/mathport) utility to translate later uses of -the definition. - -If there is no alignment for a given definition, mathport will attempt to convert -from the lean 3 `snake_case` style to `UpperCamelCase` for namespaces and types and -`lowerCamelCase` for definitions, and `snake_case` for theorems. But for various reasons, -it may fail either to determine whether it is a type, definition, or theorem, or to determine -that a specific definition is in fact being called. Or a specific definition may need a different -name altogether because the existing name is already taken in lean 4 for something else. For -these reasons, you should use `#align` on any theorem that needs to be renamed from the default. --/ -syntax (name := align) "#align " ident ppSpace ident : command - -/-- Checks that `id` has not already been `#align`ed or `#noalign`ed. -/ -def ensureUnused {m : Type → Type} [Monad m] [MonadEnv m] [MonadError m] (id : Name) : m Unit := do - if let some (_, n) := (renameExtension.getState (← getEnv)).get.toLean4.find? id then - if n.isAnonymous then - throwError "{id} has already been no-aligned" - else - throwError "{id} has already been aligned (to {n})" - -/-- -Purported Lean 3 names containing capital letters are suspicious. -However, we disregard capital letters occurring in a few common names. --/ -def suspiciousLean3Name (s : String) : Bool := Id.run do - let allowed : List String := - ["Prop", "Type", "Pi", "Exists", "End", - "Inf", "Sup", "Union", "Inter", - "Hausdorff", "is_R_or_C", - "Ioo", "Ico", "Iio", "Icc", "Iic", "Ioc", "Ici", "Ioi", "Ixx"] - let mut s := s - for a in allowed do - s := s.replace a "" - return s.any (·.isUpper) - -/-- Elaborate an `#align` command. -/ -@[command_elab align] def elabAlign : CommandElab - | `(#align $id3:ident $id4:ident) => do - if (← getInfoState).enabled then - addCompletionInfo <| CompletionInfo.id id4 id4.getId (danglingDot := false) {} none - let c := removeX id4.getId - if (← getEnv).contains c then - addConstInfo id4 c none - else if align.precheck.get (← getOptions) then - let note := "(add `set_option align.precheck false` to suppress this message)" - let inner := match ← - try some <$> (liftCoreM <| realizeGlobalConstWithInfos id4) - catch _ => pure none with - | none => m!"" - | some cs => m!" Did you mean:\n\n\ - {("\n":MessageData).joinSep (cs.map fun c' => m!" #align {id3} {c'}")}\n\n\ - #align inputs have to be fully qualified. \ - (Double check the lean 3 name too, we can't check that!)" - throwErrorAt id4 "Declaration {c} not found.{inner}\n{note}" - if Linter.getLinterValue linter.uppercaseLean3 (← getOptions) then - if id3.getId.anyS suspiciousLean3Name then - Linter.logLint linter.uppercaseLean3 id3 - "Lean 3 names are usually lowercase. This might be a typo.\n\ - If the Lean 3 name is correct, then above this line, add:\n\ - set_option linter.uppercaseLean3 false in\n" - withRef id3 <| ensureUnused id3.getId - liftCoreM <| addNameAlignment id3.getId id4.getId - | _ => throwUnsupportedSyntax - -/-- -`#noalign lean_3.def_name` will record that `lean_3.def_name` has been marked for non-porting. -This information is used by the [mathport](https://github.com/leanprover-community/mathport) -utility, which will remove the declaration from the corresponding mathport file, and later -uses of the definition will be replaced by `sorry`. --/ -syntax (name := noalign) "#noalign " ident : command - -/-- Elaborate a `#noalign` command. -/ -@[command_elab noalign] def elabNoAlign : CommandElab - | `(#noalign $id3:ident) => do - withRef id3 <| ensureUnused id3.getId - liftCoreM <| addNameAlignment id3.getId .anonymous - | _ => throwUnsupportedSyntax - -/-- Show information about the alignment status of a lean 3 definition. -/ -syntax (name := lookup3) "#lookup3 " ident : command - -/-- Elaborate a `#lookup3` command. -/ -@[command_elab lookup3] def elabLookup3 : CommandElab - | `(#lookup3%$tk $id3:ident) => do - let n3 := id3.getId - let m := renameExtension.getState (← getEnv) |>.get - match m.find? n3 with - | none => logInfoAt tk s!"name `{n3} not found" - | some (dubious, n4) => do - if n4.isAnonymous then - logInfoAt tk m!"{n3} has been no-aligned" - else - let mut msg := m!"{n4}" - if !dubious.isEmpty then - msg := msg ++ s!" (dubious: {dubious})" - logInfoAt tk <| - match m.toLean3.find? n4 with - | none | some (_, []) => msg - | some (n, l) => m!"{msg} (aliases {n :: l})" - | _ => throwUnsupportedSyntax - -open Lean Lean.Parser Lean.PrettyPrinter - -/-- An entry in the `#align_import` extension, containing all the data in the command. -/ -structure ImportEntry where - /-- The lean 3 name of the module. -/ - mod3 : Name - /-- The original repository from which this module was derived. -/ - origin : Option (String × String) - -/-- The data for `#align_import` that is stored in memory while processing a lean file. -/ -structure ImportState where - /-- This is the same as `(← getEnv).header.mainModule`, - but we need access to it in `exportEntriesFn` where the environment is not available. -/ - mod4 : Name := .anonymous - /-- The original list of import entries from imported files. We do not process these because - only mathport actually uses it. -/ - extern : Array (Array (Name × ImportEntry)) := #[] - /-- The import entries from the current file. -/ - entries : List ImportEntry := [] - deriving Inhabited - -/-- This extension stores the lookup data generated from `#align_import` commands. -/ -initialize renameImportExtension : - PersistentEnvExtension (Name × ImportEntry) (Name × ImportEntry) ImportState ← - registerPersistentEnvExtension { - mkInitial := pure {} - addEntryFn := fun s (mod4, e) => { s with mod4, entries := e :: s.entries } - addImportedFn := fun extern => return { mod4 := (← read).env.header.mainModule, extern } - exportEntriesFn := fun s => s.entries.reverse.foldl (fun a b => a.push (s.mod4, b)) #[] - } - -/-- Declare the corresponding mathlib3 module for the current mathlib4 module. -/ -syntax (name := alignImport) "#align_import " ident (" from " str "@" str)? : command - -/-- Elaborate a `#align_import` command. -/ -@[command_elab alignImport] def elabAlignImport : CommandElab - | `(#align_import $mod3 $[from $repo @ $sha]?) => do - let origin ← repo.mapM fun repo => do - let sha := sha.get! - let shaStr := sha.getString - if !shaStr.all ("abcdef0123456789".contains) then - throwErrorAt sha "not a valid hex sha, bad digits" - else if shaStr.length ≠ 40 then - throwErrorAt sha "must be a full sha" - else - pure (repo.getString, shaStr) - modifyEnv fun env => - renameImportExtension.addEntry env (env.header.mainModule, { mod3 := mod3.getId, origin }) - | _ => throwUnsupportedSyntax From 76dca296ce5765e9bc81b45221a65c0ba1e67c19 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Thu, 8 Aug 2024 15:03:31 +0000 Subject: [PATCH 110/359] perf(LinearAlgebra.Quotient): de-simp theorems with weak keys (#15591) For an example, the simp keys for `mk''_eq_mk` are `Quotient.mk'' _ _ _` meaning it will be tried on *anything* whose left-hand side reduces to `Quotient.mk''` regardless of the context. These are used a total of 6 times. --- Mathlib/Algebra/Lie/Weights/Basic.lean | 3 ++- Mathlib/LinearAlgebra/Isomorphisms.lean | 5 +++-- Mathlib/LinearAlgebra/Quotient.lean | 7 ++----- Mathlib/LinearAlgebra/Semisimple.lean | 1 + 4 files changed, 8 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index cfab261a0e7f0..404329ebc837e 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -455,7 +455,8 @@ lemma posFittingComp_le_iInf_lowerCentralSeries : simp_rw [← LieSubmodule.mem_coeSubmodule, posFittingCompOf, hk k (le_refl k)] apply LinearMap.mem_range_self suffices (toEnd R L (M ⧸ F) x ^ k) (LieSubmodule.Quotient.mk (N := F) m) = - LieSubmodule.Quotient.mk (N := F) ((toEnd R L M x ^ k) m) by simpa [this] + LieSubmodule.Quotient.mk (N := F) ((toEnd R L M x ^ k) m) + by simpa [Submodule.Quotient.quot_mk_eq_mk, this] have := LinearMap.congr_fun (LinearMap.commute_pow_left_of_commute (LieSubmodule.Quotient.toEnd_comp_mk' F x) k) m simpa using this diff --git a/Mathlib/LinearAlgebra/Isomorphisms.lean b/Mathlib/LinearAlgebra/Isomorphisms.lean index 0ca2f4302dd79..7a9bb31cd43fa 100644 --- a/Mathlib/LinearAlgebra/Isomorphisms.lean +++ b/Mathlib/LinearAlgebra/Isomorphisms.lean @@ -161,8 +161,9 @@ def quotientQuotientEquivQuotient : ((M ⧸ S) ⧸ T.map S.mkQ) ≃ₗ[R] M ⧸ { quotientQuotientEquivQuotientAux S T h with toFun := quotientQuotientEquivQuotientAux S T h invFun := mapQ _ _ (mkQ S) (le_comap_map _ _) - left_inv := fun x => Quotient.inductionOn' x fun x => Quotient.inductionOn' x fun x => by simp - right_inv := fun x => Quotient.inductionOn' x fun x => by simp } + left_inv := fun x => Quotient.inductionOn' x fun x => Quotient.inductionOn' x fun x => + by simp [Quotient.mk''_eq_mk] + right_inv := fun x => Quotient.inductionOn' x fun x => by simp [Quotient.mk''_eq_mk] } /-- Essentially the same equivalence as in the third isomorphism theorem, except restated in terms of suprema/addition of submodules instead of `≤`. -/ diff --git a/Mathlib/LinearAlgebra/Quotient.lean b/Mathlib/LinearAlgebra/Quotient.lean index 3d20b104b0b09..f19649d0da829 100644 --- a/Mathlib/LinearAlgebra/Quotient.lean +++ b/Mathlib/LinearAlgebra/Quotient.lean @@ -53,16 +53,13 @@ def mk {p : Submodule R M} : M → M ⧸ p := /- porting note: here and throughout elaboration is sped up *tremendously* (in some cases even avoiding timeouts) by providing type ascriptions to `mk` (or `mk x`) and its variants. Lean 3 didn't need this help. -/ -@[simp] theorem mk'_eq_mk' {p : Submodule R M} (x : M) : @Quotient.mk' _ (quotientRel p) x = (mk : M → M ⧸ p) x := rfl -@[simp] theorem mk''_eq_mk {p : Submodule R M} (x : M) : (Quotient.mk'' x : M ⧸ p) = (mk : M → M ⧸ p) x := rfl -@[simp] theorem quot_mk_eq_mk {p : Submodule R M} (x : M) : (Quot.mk _ x : M ⧸ p) = (mk : M → M ⧸ p) x := rfl @@ -463,8 +460,8 @@ def Quotient.equiv {N : Type*} [AddCommGroup N] [Module R N] (P : Submodule R M) rw [← hf, Submodule.mem_map] at hx obtain ⟨y, hy, rfl⟩ := hx simpa - left_inv := fun x => Quotient.inductionOn' x (by simp) - right_inv := fun x => Quotient.inductionOn' x (by simp) } + left_inv := fun x => Quotient.inductionOn' x (by simp [mk''_eq_mk]) + right_inv := fun x => Quotient.inductionOn' x (by simp [mk''_eq_mk]) } @[simp] theorem Quotient.equiv_symm {R M N : Type*} [CommRing R] [AddCommGroup M] [Module R M] diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index bdfaa811a6e1d..32882cdc8f657 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -167,6 +167,7 @@ section PerfectField variable [PerfectField K] (comm : Commute f g) (hf : f.IsSemisimple) (hg : g.IsSemisimple) +attribute [local simp] Submodule.Quotient.quot_mk_eq_mk in theorem IsSemisimple.of_mem_adjoin_pair {a : End K M} (ha : a ∈ Algebra.adjoin K {f, g}) : a.IsSemisimple := by let R := K[X] ⧸ Ideal.span {minpoly K f} From 54ff6aaa0586de2ec9dd30d78de9da39d8331f09 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 8 Aug 2024 15:42:07 +0000 Subject: [PATCH 111/359] chore: avoid import Cardinal unnecessarily (#15580) --- Counterexamples/MapFloor.lean | 1 + Mathlib.lean | 3 +- Mathlib/Algebra/Algebra/Operations.lean | 2 +- Mathlib/Algebra/Algebra/Subalgebra/Basic.lean | 1 - .../Algebra/Subalgebra/Operations.lean | 1 + Mathlib/Algebra/Group/UniqueProds.lean | 2 +- Mathlib/Algebra/MonoidAlgebra/Support.lean | 2 +- Mathlib/Algebra/Order/Antidiag/Pi.lean | 2 +- Mathlib/Algebra/Pointwise/Stabilizer.lean | 2 +- .../Polynomial/Degree/Definitions.lean | 1 + .../Polynomial/Degree/TrailingDegree.lean | 1 + Mathlib/Algebra/Vertex/HVertexOperator.lean | 2 + .../Combinatorics/Additive/ETransform.lean | 2 +- Mathlib/Combinatorics/Additive/Energy.lean | 2 +- .../Additive/PluenneckeRuzsa.lean | 4 +- .../Combinatorics/Additive/RuzsaCovering.lean | 2 +- Mathlib/Data/Complex/Module.lean | 1 + Mathlib/Data/DFinsupp/Interval.lean | 2 +- Mathlib/Data/Finset/Finsupp.lean | 2 +- Mathlib/Data/Finset/NatDivisors.lean | 3 +- .../{Pointwise.lean => Pointwise/Basic.lean} | 49 ++------------- Mathlib/Data/Finset/Pointwise/Card.lean | 62 +++++++++++++++++++ Mathlib/Data/Finset/Pointwise/Interval.lean | 3 +- Mathlib/Data/Multiset/Interval.lean | 1 + .../RingTheory/HahnSeries/Multiplication.lean | 1 - scripts/style-exceptions.txt | 5 +- 26 files changed, 96 insertions(+), 63 deletions(-) rename Mathlib/Data/Finset/{Pointwise.lean => Pointwise/Basic.lean} (98%) create mode 100644 Mathlib/Data/Finset/Pointwise/Card.lean diff --git a/Counterexamples/MapFloor.lean b/Counterexamples/MapFloor.lean index 5ddc5c17af16d..08097ff2b0b83 100644 --- a/Counterexamples/MapFloor.lean +++ b/Counterexamples/MapFloor.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Algebra.Order.Floor import Mathlib.Algebra.Order.Group.PiLex +import Mathlib.Algebra.Order.Hom.Ring import Mathlib.Algebra.Polynomial.Reverse /-! diff --git a/Mathlib.lean b/Mathlib.lean index 81517a645910c..c4fa6e74e8a44 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2097,7 +2097,8 @@ import Mathlib.Data.Finset.Pairwise import Mathlib.Data.Finset.Pi import Mathlib.Data.Finset.PiInduction import Mathlib.Data.Finset.Piecewise -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic +import Mathlib.Data.Finset.Pointwise.Card import Mathlib.Data.Finset.Pointwise.Interval import Mathlib.Data.Finset.Powerset import Mathlib.Data.Finset.Preimage diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 50b8453d2021e..6f96922e603a9 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -11,7 +11,7 @@ import Mathlib.Algebra.Module.Opposites import Mathlib.Algebra.Module.Submodule.Bilinear import Mathlib.Algebra.Module.Submodule.Pointwise import Mathlib.Algebra.Order.Kleene -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic import Mathlib.Data.Set.Pointwise.BigOperators import Mathlib.Data.Set.Semiring import Mathlib.GroupTheory.GroupAction.SubMulAction.Pointwise diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index 5359885be61ba..81ea05a0347ff 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -13,7 +13,6 @@ In this file we define `Subalgebra`s and the usual operations on them (`map`, `c More lemmas about `adjoin` can be found in `RingTheory.Adjoin`. -/ - universe u u' v w w' /-- A subalgebra is a sub(semi)ring that includes the range of `algebraMap`. -/ diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Operations.lean b/Mathlib/Algebra/Algebra/Subalgebra/Operations.lean index b96fa1964bacb..8d05a87504817 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Operations.lean @@ -15,6 +15,7 @@ and `Mathlib.RingTheory.Ideal.Operations` are somewhat of a grab-bag of definiti whatever ends up in the intersection. -/ +assert_not_exists Cardinal namespace AlgHom diff --git a/Mathlib/Algebra/Group/UniqueProds.lean b/Mathlib/Algebra/Group/UniqueProds.lean index 6e1f39960f5be..6ac2374fea1d6 100644 --- a/Mathlib/Algebra/Group/UniqueProds.lean +++ b/Mathlib/Algebra/Group/UniqueProds.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ import Mathlib.Data.DFinsupp.Basic -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic import Mathlib.LinearAlgebra.Basis.VectorSpace /-! diff --git a/Mathlib/Algebra/MonoidAlgebra/Support.lean b/Mathlib/Algebra/MonoidAlgebra/Support.lean index d50cce52c12e4..18069fab871f8 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Support.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Support.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ import Mathlib.Algebra.MonoidAlgebra.Basic -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic /-! # Lemmas about the support of a finitely supported function diff --git a/Mathlib/Algebra/Order/Antidiag/Pi.lean b/Mathlib/Algebra/Order/Antidiag/Pi.lean index 8cc4c25cbbcbf..7af083a0c2b08 100644 --- a/Mathlib/Algebra/Order/Antidiag/Pi.lean +++ b/Mathlib/Algebra/Order/Antidiag/Pi.lean @@ -5,7 +5,7 @@ Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández, Eric Wieser, Yaël Dillies -/ import Mathlib.Algebra.Order.BigOperators.Group.Finset -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic import Mathlib.Data.Fin.Tuple.NatAntidiagonal /-! diff --git a/Mathlib/Algebra/Pointwise/Stabilizer.lean b/Mathlib/Algebra/Pointwise/Stabilizer.lean index 14d5840e635a6..a63e174396fa1 100644 --- a/Mathlib/Algebra/Pointwise/Stabilizer.lean +++ b/Mathlib/Algebra/Pointwise/Stabilizer.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic import Mathlib.GroupTheory.QuotientGroup /-! diff --git a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean index a5e3b5b388592..70d9b2bb4743e 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean @@ -10,6 +10,7 @@ import Mathlib.Data.Fintype.BigOperators import Mathlib.Data.Nat.WithBot import Mathlib.Data.Nat.Cast.WithTop import Mathlib.Data.Nat.SuccPred +import Mathlib.Algebra.Order.Ring.WithTop /-! # Theory of univariate polynomials diff --git a/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean b/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean index 5217a632caa26..af0a168b4849f 100644 --- a/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean +++ b/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean @@ -5,6 +5,7 @@ Authors: Damiano Testa -/ import Mathlib.Algebra.Polynomial.Degree.Definitions import Mathlib.Data.ENat.Basic +import Mathlib.Data.ENat.Lattice /-! # Trailing degree of univariate polynomials diff --git a/Mathlib/Algebra/Vertex/HVertexOperator.lean b/Mathlib/Algebra/Vertex/HVertexOperator.lean index 229a327c1952f..fdb943fb8ad04 100644 --- a/Mathlib/Algebra/Vertex/HVertexOperator.lean +++ b/Mathlib/Algebra/Vertex/HVertexOperator.lean @@ -29,6 +29,8 @@ here allows us to consider composites and scalar-multiply by multivariable Laure -/ +assert_not_exists Cardinal + noncomputable section variable {Γ : Type*} [PartialOrder Γ] {R : Type*} {V W : Type*} [CommRing R] diff --git a/Mathlib/Combinatorics/Additive/ETransform.lean b/Mathlib/Combinatorics/Additive/ETransform.lean index a782fbe10f426..c146fce4176ae 100644 --- a/Mathlib/Combinatorics/Additive/ETransform.lean +++ b/Mathlib/Combinatorics/Additive/ETransform.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic /-! # e-transforms diff --git a/Mathlib/Combinatorics/Additive/Energy.lean b/Mathlib/Combinatorics/Additive/Energy.lean index 8414e0d08328d..fdc2bce807bcc 100644 --- a/Mathlib/Combinatorics/Additive/Energy.lean +++ b/Mathlib/Combinatorics/Additive/Energy.lean @@ -6,7 +6,7 @@ Authors: Yaël Dillies, Ella Yu import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Data.Finset.Prod import Mathlib.Data.Fintype.Prod -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic /-! # Additive energy diff --git a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean index b1c058d95cf66..c959973f91a54 100644 --- a/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean +++ b/Mathlib/Combinatorics/Additive/PluenneckeRuzsa.lean @@ -5,8 +5,10 @@ Authors: Yaël Dillies, George Shakan -/ import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Combinatorics.Enumerative.DoubleCounting -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic import Mathlib.Tactic.GCongr +import Mathlib.Algebra.Order.Field.Basic +import Mathlib.Algebra.Order.Field.Rat /-! # The Plünnecke-Ruzsa inequality diff --git a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean index 4846ae040a625..06670fc204bb7 100644 --- a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean +++ b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic import Mathlib.SetTheory.Cardinal.Finite /-! diff --git a/Mathlib/Data/Complex/Module.lean b/Mathlib/Data/Complex/Module.lean index ab8d178d1afd7..d79c2f4a80d70 100644 --- a/Mathlib/Data/Complex/Module.lean +++ b/Mathlib/Data/Complex/Module.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.CharP.Invertible import Mathlib.Data.Complex.Basic import Mathlib.LinearAlgebra.Matrix.ToLin import Mathlib.Data.Real.Star +import Mathlib.Data.ZMod.Defs /-! # Complex number as a vector space over `ℝ` diff --git a/Mathlib/Data/DFinsupp/Interval.lean b/Mathlib/Data/DFinsupp/Interval.lean index 0df18bd559127..460f8e07408e8 100644 --- a/Mathlib/Data/DFinsupp/Interval.lean +++ b/Mathlib/Data/DFinsupp/Interval.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 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.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic import Mathlib.Data.Fintype.BigOperators import Mathlib.Data.DFinsupp.Order import Mathlib.Order.Interval.Finset.Basic diff --git a/Mathlib/Data/Finset/Finsupp.lean b/Mathlib/Data/Finset/Finsupp.lean index 8bbc8f7059f02..08de380f2d4b9 100644 --- a/Mathlib/Data/Finset/Finsupp.lean +++ b/Mathlib/Data/Finset/Finsupp.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.BigOperators.Finsupp -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic import Mathlib.Data.Finsupp.Indicator import Mathlib.Data.Fintype.BigOperators diff --git a/Mathlib/Data/Finset/NatDivisors.lean b/Mathlib/Data/Finset/NatDivisors.lean index ef1bc660839eb..07fa594f4b6a3 100644 --- a/Mathlib/Data/Finset/NatDivisors.lean +++ b/Mathlib/Data/Finset/NatDivisors.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa, Yury Kudryashov -/ import Mathlib.NumberTheory.Divisors -import Mathlib.Data.Nat.Order.Lemmas -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic /-! # `Nat.divisors` as a multiplicative homomorpism diff --git a/Mathlib/Data/Finset/Pointwise.lean b/Mathlib/Data/Finset/Pointwise/Basic.lean similarity index 98% rename from Mathlib/Data/Finset/Pointwise.lean rename to Mathlib/Data/Finset/Pointwise/Basic.lean index 96525b4a505bb..a20237588d38e 100644 --- a/Mathlib/Data/Finset/Pointwise.lean +++ b/Mathlib/Data/Finset/Pointwise/Basic.lean @@ -9,7 +9,9 @@ import Mathlib.Data.Finset.Preimage import Mathlib.Data.Set.Pointwise.Finite import Mathlib.Data.Set.Pointwise.SMul import Mathlib.Data.Set.Pointwise.ListOfFn -import Mathlib.SetTheory.Cardinal.Finite +import Mathlib.Data.ULift +import Mathlib.Algebra.BigOperators.Group.Finset +import Mathlib.Algebra.Order.Ring.Nat /-! # Pointwise operations of finsets @@ -54,6 +56,8 @@ pointwise subtraction -/ +assert_not_exists Cardinal + open Function MulOpposite open scoped Pointwise @@ -1985,49 +1989,6 @@ theorem Finite.toFinset_vsub (hs : s.Finite) (ht : t.Finite) (hf := hs.vsub ht) end VSub -section MulAction -variable [Group α] [MulAction α β] - -@[to_additive (attr := simp)] -lemma card_smul_set (a : α) (s : Set β) : Nat.card ↥(a • s) = Nat.card s := - Nat.card_image_of_injective (MulAction.injective a) _ - -end MulAction - -section IsCancelMul -variable [Mul α] [IsCancelMul α] {s t : Set α} - -@[to_additive] -lemma card_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by - classical - obtain h | h := (s * t).infinite_or_finite - · simp [Set.Infinite.card_eq_zero h] - obtain ⟨hs, ht⟩ | rfl | rfl := finite_mul.1 h - · lift s to Finset α using hs - lift t to Finset α using ht - rw [← Finset.coe_mul] - simpa [-Finset.coe_mul] using Finset.card_mul_le - all_goals simp - -end IsCancelMul - -section InvolutiveInv -variable [InvolutiveInv α] {s t : Set α} - -@[to_additive (attr := simp)] -lemma card_inv (s : Set α) : Nat.card ↥(s⁻¹) = Nat.card s := by - rw [← image_inv, Nat.card_image_of_injective inv_injective] - -end InvolutiveInv - -section Group -variable [Group α] {s t : Set α} - -@[to_additive] -lemma card_div_le : Nat.card (s / t) ≤ Nat.card s * Nat.card t := by - rw [div_eq_mul_inv, ← card_inv t]; exact card_mul_le - -end Group end Set instance Nat.decidablePred_mem_vadd_set {s : Set ℕ} [DecidablePred (· ∈ s)] (a : ℕ) : diff --git a/Mathlib/Data/Finset/Pointwise/Card.lean b/Mathlib/Data/Finset/Pointwise/Card.lean new file mode 100644 index 0000000000000..5516539f11c50 --- /dev/null +++ b/Mathlib/Data/Finset/Pointwise/Card.lean @@ -0,0 +1,62 @@ +/- +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.Finset.Pointwise.Basic +import Mathlib.SetTheory.Cardinal.Finite + +/-! +# Cardinalities of pointwise operations on sets. +-/ + +namespace Set + +open Pointwise + +variable {α β : Type*} + +section MulAction +variable [Group α] [MulAction α β] + +@[to_additive (attr := simp)] +lemma card_smul_set (a : α) (s : Set β) : Nat.card ↥(a • s) = Nat.card s := + Nat.card_image_of_injective (MulAction.injective a) _ + +end MulAction + +section IsCancelMul +variable [Mul α] [IsCancelMul α] {s t : Set α} + +@[to_additive] +lemma card_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by + classical + obtain h | h := (s * t).infinite_or_finite + · simp [Set.Infinite.card_eq_zero h] + obtain ⟨hs, ht⟩ | rfl | rfl := finite_mul.1 h + · lift s to Finset α using hs + lift t to Finset α using ht + rw [← Finset.coe_mul] + simpa [-Finset.coe_mul] using Finset.card_mul_le + all_goals simp + +end IsCancelMul + +section InvolutiveInv +variable [InvolutiveInv α] {s t : Set α} + +@[to_additive (attr := simp)] +lemma card_inv (s : Set α) : Nat.card ↥(s⁻¹) = Nat.card s := by + rw [← image_inv, Nat.card_image_of_injective inv_injective] + +end InvolutiveInv + +section Group +variable [Group α] {s t : Set α} + +@[to_additive] +lemma card_div_le : Nat.card (s / t) ≤ Nat.card s * Nat.card t := by + rw [div_eq_mul_inv, ← card_inv t]; exact card_mul_le + +end Group +end Set diff --git a/Mathlib/Data/Finset/Pointwise/Interval.lean b/Mathlib/Data/Finset/Pointwise/Interval.lean index 2df7980014e25..9178e9e993251 100644 --- a/Mathlib/Data/Finset/Pointwise/Interval.lean +++ b/Mathlib/Data/Finset/Pointwise/Interval.lean @@ -3,8 +3,9 @@ Copyright (c) 2023 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.Data.Finset.Pointwise +import Mathlib.Data.Finset.Pointwise.Basic import Mathlib.Data.Set.Pointwise.Interval +import Mathlib.Order.Interval.Finset.Defs /-! # Pointwise operations on intervals diff --git a/Mathlib/Data/Multiset/Interval.lean b/Mathlib/Data/Multiset/Interval.lean index bc119c09440fe..d4db3e807d935 100644 --- a/Mathlib/Data/Multiset/Interval.lean +++ b/Mathlib/Data/Multiset/Interval.lean @@ -6,6 +6,7 @@ Authors: Eric Wieser import Mathlib.Data.DFinsupp.Interval import Mathlib.Data.DFinsupp.Multiset import Mathlib.Order.Interval.Finset.Nat +import Mathlib.Data.Nat.Lattice /-! # Finite intervals of multisets diff --git a/Mathlib/RingTheory/HahnSeries/Multiplication.lean b/Mathlib/RingTheory/HahnSeries/Multiplication.lean index 0ddaf67084bec..148770477311f 100644 --- a/Mathlib/RingTheory/HahnSeries/Multiplication.lean +++ b/Mathlib/RingTheory/HahnSeries/Multiplication.lean @@ -32,7 +32,6 @@ Hahn series. - [J. van der Hoeven, *Operators on Generalized Power Series*][van_der_hoeven] -/ - open Finset Function Pointwise noncomputable section diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index 84c7d158bcbab..9b6d457b4a950 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -1,3 +1,4 @@ +Mathlib/SetTheory/Ordinal/Segfault.lean : line 6 : ERR_MOD : Module docstring missing, or too late Mathlib/Algebra/BigOperators/Group/Finset.lean : line 1 : ERR_NUM_LIN : 2400 file contains 2245 lines, try to split it up Mathlib/Algebra/Group/Subgroup/Basic.lean : line 1 : ERR_NUM_LIN : 3000 file contains 2893 lines, try to split it up Mathlib/Algebra/MonoidAlgebra/Basic.lean : line 1 : ERR_NUM_LIN : 2100 file contains 1965 lines, try to split it up @@ -15,9 +16,9 @@ Mathlib/Data/DFinsupp/Basic.lean : line 1 : ERR_NUM_LIN : 2200 file contains 207 Mathlib/Data/Fin/Basic.lean : line 1 : ERR_NUM_LIN : 1700 file contains 1524 lines, try to split it up Mathlib/Data/Finset/Basic.lean : line 1 : ERR_NUM_LIN : 3100 file contains 2999 lines, try to split it up Mathlib/Data/Finset/Lattice.lean : line 1 : ERR_NUM_LIN : 2100 file contains 1968 lines, try to split it up -Mathlib/Data/Finset/Pointwise.lean : line 1 : ERR_NUM_LIN : 2200 file contains 2032 lines, try to split it up +Mathlib/Data/Finset/Pointwise/Basic.lean : line 1 : ERR_NUM_LIN : 2100 file contains 1997 lines, try to split it up Mathlib/Data/Finsupp/Basic.lean : line 1 : ERR_NUM_LIN : 1900 file contains 1703 lines, try to split it up -Mathlib/Data/List/Basic.lean : line 1 : ERR_NUM_LIN : 2900 file contains 2764 lines, try to split it up +Mathlib/Data/List/Basic.lean : line 1 : ERR_NUM_LIN : 2700 file contains 2598 lines, try to split it up Mathlib/Data/Matrix/Basic.lean : line 1 : ERR_NUM_LIN : 2700 file contains 2519 lines, try to split it up Mathlib/Data/Multiset/Basic.lean : line 1 : ERR_NUM_LIN : 2900 file contains 2725 lines, try to split it up Mathlib/Data/Num/Lemmas.lean : line 1 : ERR_NUM_LIN : 1700 file contains 1576 lines, try to split it up From 43a2b849b82673af012a27628b3e76dc46bc85ca Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Thu, 8 Aug 2024 15:42:09 +0000 Subject: [PATCH 112/359] chore: add ascription for delicate-to-elaborate expression (#15597) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The expression `2 • (Finsupp.single 0 1 : α →₀ F) = Finsupp.single 0 1` only accidentally elaborates. The binop elaborator fails to compute a "max type" for this, since the right-hand `Finsupp.single 0 1` does not have a fully specified type yet, so it falls back to a strategy where it propagates the LHS type to the RHS. However, the rest of the process then depends on the exact order that default instances are applied. If the default instances for the RHS apply first, it specializes to `ℕ →₀ ℕ` and elaboration fails, but if the LHS ones apply first, then the SMul manages to specialize to `ℕ → (α →₀ F) → (α →₀ F)`, and it succeeds. This came up while experimenting with performance improvements to the binop elaborator. --- Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean b/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean index c5225ad3646fb..f12b01482968d 100644 --- a/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean +++ b/Counterexamples/ZeroDivisorsInAddMonoidAlgebras.lean @@ -233,7 +233,8 @@ example {α} [Ring α] [Nontrivial α] : ∃ f g : AddMonoidAlgebra α F, f ≠ zero_divisors_of_periodic (1 : F) le_rfl (by simp [two_smul]) z01.ne' example {α} [Zero α] : - 2 • (Finsupp.single 0 1 : α →₀ F) = Finsupp.single 0 1 ∧ (Finsupp.single 0 1 : α →₀ F) ≠ 0 := + 2 • (Finsupp.single 0 1 : α →₀ F) = (Finsupp.single 0 1 : α →₀ F) + ∧ (Finsupp.single 0 1 : α →₀ F) ≠ 0 := ⟨smul_single _ _ _, by simp [Ne, Finsupp.single_eq_zero, z01.ne]⟩ end F From 2eea5591850ffa539e684e58396480e9766df586 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 8 Aug 2024 15:42:10 +0000 Subject: [PATCH 113/359] feat(RelSeries): reverse_reverse (#15601) from the Carleson project. --- Mathlib/Order/RelSeries.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index c3676e955e9ca..36c5009b8d87c 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -399,6 +399,9 @@ def reverse (p : RelSeries r) : RelSeries (fun (a b : α) ↦ r b a) where @[simp] lemma head_reverse (p : RelSeries r) : p.reverse.head = p.last := by simp [RelSeries.last, RelSeries.head] +@[simp] lemma reverse_reverse {r : Rel α α} (p : RelSeries r) : p.reverse.reverse = p := by + ext <;> simp + /-- Given a series `a₀ -r→ a₁ -r→ ... -r→ aₙ` and an `a` such that `a₀ -r→ a` holds, there is a series of length `n+1`: `a -r→ a₀ -r→ a₁ -r→ ... -r→ aₙ`. From 182be4fbbee5a608b22fd31e3fd0c16a584fa97b Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 8 Aug 2024 16:24:10 +0000 Subject: [PATCH 114/359] =?UTF-8?q?feat:=20star=20algebra=20homomorphisms?= =?UTF-8?q?=20between=20non-unital=20C=E2=8B=86-algebras=20are=20contracti?= =?UTF-8?q?ve=20(#14544)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This generalizes the existing results for unital algebras to non-unital algebras by passing to the unitization within the proof. We have to add a `StarModule` hypothesis to recover the unital case, but this is minor. --- .../Instances.lean | 9 ++- Mathlib/Analysis/CStarAlgebra/Spectrum.lean | 77 +++++++++++-------- 2 files changed, 52 insertions(+), 34 deletions(-) diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index 4f2c56e68c51b..b7a308f06b1b0 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -153,8 +153,13 @@ instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [NormedRing A exists_cfc_of_predicate a ha := by refine ⟨(elementalStarAlgebra ℂ a).subtype.comp <| continuousFunctionalCalculus a, ?hom_closedEmbedding, ?hom_id, ?hom_map_spectrum, ?predicate_hom⟩ - case hom_closedEmbedding => exact Isometry.closedEmbedding <| - isometry_subtype_coe.comp <| StarAlgEquiv.isometry (continuousFunctionalCalculus a) + case hom_closedEmbedding => + -- note: Lean should find these for `StarAlgEquiv.isometry`, but it doesn't and so we + -- provide them manually. + have : SMulCommClass ℂ C(σ ℂ a, ℂ) C(σ ℂ a, ℂ) := Algebra.to_smulCommClass (A := C(σ ℂ a, ℂ)) + have : IsScalarTower ℂ C(σ ℂ a, ℂ) C(σ ℂ a, ℂ) := IsScalarTower.right (A := C(σ ℂ a, ℂ)) + exact Isometry.closedEmbedding <| + isometry_subtype_coe.comp <| StarAlgEquiv.isometry (continuousFunctionalCalculus a) case hom_id => exact congr_arg Subtype.val <| continuousFunctionalCalculus_map_id a case hom_map_spectrum => intro f diff --git a/Mathlib/Analysis/CStarAlgebra/Spectrum.lean b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean index 924eaed27bea7..ceac7a25b3c1e 100644 --- a/Mathlib/Analysis/CStarAlgebra/Spectrum.lean +++ b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.CStarAlgebra.Basic import Mathlib.Analysis.Normed.Algebra.Spectrum +import Mathlib.Analysis.CStarAlgebra.Unitization import Mathlib.Analysis.SpecialFunctions.Exponential import Mathlib.Algebra.Star.StarAlgHom @@ -119,46 +119,59 @@ theorem selfAdjoint.val_re_map_spectrum [StarModule ℂ A] (a : selfAdjoint A) : end ComplexScalars -namespace StarAlgHom - -variable {F A B : Type*} [NormedRing A] [NormedAlgebra ℂ A] [CompleteSpace A] [StarRing A] - [CStarRing A] [NormedRing B] [NormedAlgebra ℂ B] [CompleteSpace B] [StarRing B] [CStarRing B] - [FunLike F A B] [AlgHomClass F ℂ A B] [StarAlgHomClass F ℂ A B] (φ : F) - -/-- A star algebra homomorphism of complex C⋆-algebras is norm contractive. -/ -theorem nnnorm_apply_le (a : A) : ‖(φ a : B)‖₊ ≤ ‖a‖₊ := by - suffices ∀ s : A, IsSelfAdjoint s → ‖φ s‖₊ ≤ ‖s‖₊ by - exact nonneg_le_nonneg_of_sq_le_sq zero_le' <| by - simpa only [nnnorm_star_mul_self, map_star, map_mul] - using this _ (IsSelfAdjoint.star_mul_self a) - intro s hs - simpa only [hs.spectralRadius_eq_nnnorm, (hs.starHom_apply φ).spectralRadius_eq_nnnorm, - coe_le_coe] using - show spectralRadius ℂ (φ s) ≤ spectralRadius ℂ s from - iSup_le_iSup_of_subset (AlgHom.spectrum_apply_subset φ s) - -/-- A star algebra homomorphism of complex C⋆-algebras is norm contractive. -/ -theorem norm_apply_le (a : A) : ‖(φ a : B)‖ ≤ ‖a‖ := - nnnorm_apply_le φ a - -/-- Star algebra homomorphisms between C⋆-algebras are continuous linear maps. +namespace NonUnitalStarAlgHom + +variable {F A B : Type*} +variable [NonUnitalNormedRing A] [CompleteSpace A] [StarRing A] [CStarRing A] +variable [NormedSpace ℂ A] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] [StarModule ℂ A] +variable [NonUnitalNormedRing B] [CompleteSpace B] [StarRing B] [CStarRing B] +variable [NormedSpace ℂ B] [IsScalarTower ℂ B B] [SMulCommClass ℂ B B] [StarModule ℂ B] +variable [FunLike F A B] [NonUnitalAlgHomClass F ℂ A B] [NonUnitalStarAlgHomClass F ℂ A B] + +open Unitization + +/-- A non-unital star algebra homomorphism of complex C⋆-algebras is norm contractive. -/ +lemma nnnorm_apply_le (φ : F) (a : A) : ‖φ a‖₊ ≤ ‖a‖₊ := by + have h (ψ : Unitization ℂ A →⋆ₐ[ℂ] Unitization ℂ B) (x : Unitization ℂ A) : + ‖ψ x‖₊ ≤ ‖x‖₊ := by + suffices ∀ {s}, IsSelfAdjoint s → ‖ψ s‖₊ ≤ ‖s‖₊ by + refine nonneg_le_nonneg_of_sq_le_sq zero_le' ?_ + simp_rw [← nnnorm_star_mul_self, ← map_star, ← map_mul] + exact this <| .star_mul_self x + intro s hs + suffices this : spectralRadius ℂ (ψ s) ≤ spectralRadius ℂ s by + -- changing the order of `rw`s below runs into https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/weird.20type.20class.20synthesis.20error/near/421224482 + rwa [(hs.starHom_apply ψ).spectralRadius_eq_nnnorm, hs.spectralRadius_eq_nnnorm, coe_le_coe] + at this + exact iSup_le_iSup_of_subset (AlgHom.spectrum_apply_subset ψ s) + simpa [nnnorm_inr] using h (starLift (inrNonUnitalStarAlgHom ℂ B |>.comp (φ : A →⋆ₙₐ[ℂ] B))) a + +/-- A non-unital star algebra homomorphism of complex C⋆-algebras is norm contractive. -/ +lemma norm_apply_le (φ : F) (a : A) : ‖φ a‖ ≤ ‖a‖ := by + exact_mod_cast nnnorm_apply_le φ a + +/-- Non-unital star algebra homomorphisms between C⋆-algebras are continuous linear maps. See note [lower instance priority] -/ -noncomputable instance (priority := 100) : ContinuousLinearMapClass F ℂ A B := - { AlgHomClass.linearMapClass with +lemma instContinuousLinearMapClassComplex : ContinuousLinearMapClass F ℂ A B := + { NonUnitalAlgHomClass.instLinearMapClass with map_continuous := fun φ => AddMonoidHomClass.continuous_of_bound φ 1 (by simpa only [one_mul] using nnnorm_apply_le φ) } -end StarAlgHom +scoped[CStarAlgebra] attribute [instance] NonUnitalStarAlgHom.instContinuousLinearMapClassComplex + +end NonUnitalStarAlgHom namespace StarAlgEquiv -variable {F A B : Type*} [NormedRing A] [NormedAlgebra ℂ A] [CompleteSpace A] [StarRing A] - [CStarRing A] [NormedRing B] [NormedAlgebra ℂ B] [CompleteSpace B] [StarRing B] [CStarRing B] - [EquivLike F A B] [NonUnitalAlgEquivClass F ℂ A B] [StarAlgEquivClass F ℂ A B] +variable {F A B : Type*} [NormedRing A] [NormedSpace ℂ A] [SMulCommClass ℂ A A] +variable [IsScalarTower ℂ A A] [CompleteSpace A] [StarRing A] [CStarRing A] [StarModule ℂ A] +variable [NormedRing B] [NormedSpace ℂ B] [SMulCommClass ℂ B B] [IsScalarTower ℂ B B] +variable [CompleteSpace B] [StarRing B] [CStarRing B] [StarModule ℂ B] [EquivLike F A B] +variable [NonUnitalAlgEquivClass F ℂ A B] [StarAlgEquivClass F ℂ A B] lemma nnnorm_map (φ : F) (a : A) : ‖φ a‖₊ = ‖a‖₊ := - le_antisymm (StarAlgHom.nnnorm_apply_le φ a) <| by - simpa using StarAlgHom.nnnorm_apply_le (symm (φ : A ≃⋆ₐ[ℂ] B)) ((φ : A ≃⋆ₐ[ℂ] B) a) + le_antisymm (NonUnitalStarAlgHom.nnnorm_apply_le φ a) <| by + simpa using NonUnitalStarAlgHom.nnnorm_apply_le (symm (φ : A ≃⋆ₐ[ℂ] B)) ((φ : A ≃⋆ₐ[ℂ] B) a) lemma norm_map (φ : F) (a : A) : ‖φ a‖ = ‖a‖ := congr_arg NNReal.toReal (nnnorm_map φ a) From 4fe4d4db825a4132ab89c9c107c72629ba7badce Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 8 Aug 2024 16:24:12 +0000 Subject: [PATCH 115/359] =?UTF-8?q?feat:=20`(a=E2=82=81=20=3D=3D=20a?= =?UTF-8?q?=E2=82=82)=20=3D=20(b=E2=82=81=20=3D=3D=20b=E2=82=82)=20?= =?UTF-8?q?=E2=86=94=20(a=E2=82=81=20=3D=20a=E2=82=82=20=E2=86=94=20b?= =?UTF-8?q?=E2=82=81=20=3D=20b=E2=82=82)`=20(#15229)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is useful to simplify `beq` away --- Mathlib/Logic/Basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 1d1460adf46e1..3c7429f258290 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -977,6 +977,9 @@ theorem beq_eq_decide {α : Type*} [BEq α] [LawfulBEq α] {a b : α} : (a == b) rw [← beq_iff_eq a b] cases a == b <;> simp +@[simp] lemma beq_eq_beq {α β : Type*} [BEq α] [LawfulBEq α] [BEq β] [LawfulBEq β] {a₁ a₂ : α} + {b₁ b₂ : β} : (a₁ == a₂) = (b₁ == b₂) ↔ (a₁ = a₂ ↔ b₁ = b₂) := by rw [Bool.eq_iff_iff]; simp + @[ext] theorem beq_ext {α : Type*} (inst1 : BEq α) (inst2 : BEq α) (h : ∀ x y, @BEq.beq _ inst1 x y = @BEq.beq _ inst2 x y) : From 340c789afdcbe26f967370ab99449ed1c87c5019 Mon Sep 17 00:00:00 2001 From: Shuhao Song Date: Thu, 8 Aug 2024 16:24:13 +0000 Subject: [PATCH 116/359] feat(Tactic/NthRewrite): Allow multiple occurrences (#15532) Allow `nth_rewrite` and `nth_rw` tactic to use multiple occurrences, such as `nth_rw 2 3 [p]`. --- Mathlib/Tactic/NthRewrite.lean | 34 +++++++++++++++++----------------- test/NthRewrite.lean | 3 +++ 2 files changed, 20 insertions(+), 17 deletions(-) diff --git a/Mathlib/Tactic/NthRewrite.lean b/Mathlib/Tactic/NthRewrite.lean index 3ba52c3db0334..eff8ce34cc201 100644 --- a/Mathlib/Tactic/NthRewrite.lean +++ b/Mathlib/Tactic/NthRewrite.lean @@ -17,19 +17,19 @@ namespace Mathlib.Tactic open Lean Elab Tactic Meta Parser.Tactic -/-- `nth_rewrite` is a variant of `rewrite` that only changes the `n`ᵗʰ _occurrence_ of the -expression to be rewritten. `nth_rewrite n [eq₁, eq₂,..., eqₘ]` will rewrite the `n`ᵗʰ _occurrence_ -of each of the `m` equalities `eqᵢ`in that order. Occurrences are counted beginning with `1` in -order of precedence. +/-- `nth_rewrite` is a variant of `rewrite` that only changes the `n₁, ..., nₖ`ᵗʰ _occurrence_ of +the expression to be rewritten. `nth_rewrite n₁ ... nₖ [eq₁, eq₂,..., eqₘ]` will rewrite the +`n₁, ..., nₖ`ᵗʰ _occurrence_ of each of the `m` equalities `eqᵢ`in that order. Occurrences are +counted beginning with `1` in order of precedence. For example, ```lean example (h : a = 1) : a + a + a + a + a = 5 := by - nth_rewrite 2 [h] + nth_rewrite 2 3 [h] /- a: ℕ h: a = 1 -⊢ a + 1 + a + a + a = 5 +⊢ a + 1 + 1 + a + a = 5 -/ ``` Notice that the second occurrence of `a` from the left has been rewritten by `nth_rewrite`. @@ -67,16 +67,16 @@ h: a = a + b This new instance of `a` also turns out to be the third _occurrence_ of `a`. Therefore, the next `nth_rewrite` with `h` rewrites this `a`. -/ -syntax (name := nthRewriteSeq) "nth_rewrite" (config)? ppSpace num rwRuleSeq (location)? : tactic +syntax (name := nthRewriteSeq) "nth_rewrite" (config)? ppSpace num+ rwRuleSeq (location)? : tactic @[inherit_doc nthRewriteSeq, tactic nthRewriteSeq] def evalNthRewriteSeq : Tactic := fun stx => do match stx with - | `(tactic| nth_rewrite $[$_cfg]? $n $_rules $[$_loc]?) => + | `(tactic| nth_rewrite $[$_cfg]? $[$n]* $_rules:rwRuleSeq $[$_loc]?) => -- [TODO] `stx` should not be used directly, but the corresponding functions do not yet exist -- in Lean 4 core let cfg ← elabRewriteConfig stx[1] let loc := expandOptLocation stx[4] - let occ := Occurrences.pos [n.getNat] + let occ := Occurrences.pos (n.map TSyntax.getNat).toList let cfg := { cfg with occs := occ } withRWRulesSeq stx[0] stx[3] fun symm term => do withLocation loc @@ -86,18 +86,18 @@ syntax (name := nthRewriteSeq) "nth_rewrite" (config)? ppSpace num rwRuleSeq (lo | _ => throwUnsupportedSyntax /-- -`nth_rw` is a variant of `rw` that only changes the `n`ᵗʰ _occurrence_ of the expression to be -rewritten. Like `rw`, and unlike `nth_rewrite`, it will try to close the goal by trying `rfl` -afterwards. `nth_rw n [eq₁, eq₂,..., eqₘ]` will rewrite the `n`ᵗʰ _occurrence_ of each of the -`m` equalities `eqᵢ`in that order. Occurrences are counted beginning with `1` in +`nth_rw` is a variant of `rw` that only changes the `n₁, ..., nₖ`ᵗʰ _occurrence_ of the expression +to be rewritten. Like `rw`, and unlike `nth_rewrite`, it will try to close the goal by trying `rfl` +afterwards. `nth_rw n₁ ... nₖ [eq₁, eq₂,..., eqₘ]` will rewrite the `n₁, ..., nₖ`ᵗʰ _occurrence_ of +each of the `m` equalities `eqᵢ`in that order. Occurrences are counted beginning with `1` in order of precedence. For example, ```lean example (h : a = 1) : a + a + a + a + a = 5 := by - nth_rw 2 [h] + nth_rw 2 3 [h] /- a: ℕ h: a = 1 -⊢ a + 1 + a + a + a = 5 +⊢ a + 1 + 1 + a + a = 5 -/ ``` Notice that the second occurrence of `a` from the left has been rewritten by `nth_rewrite`. @@ -137,11 +137,11 @@ the next `nth_rw` with `h` rewrites this `a`. Further, `nth_rw` will close the remaining goal with `rfl` if possible. -/ -macro (name := nthRwSeq) "nth_rw" c:(config)? ppSpace n:num s:rwRuleSeq l:(location)? : tactic => +macro (name := nthRwSeq) "nth_rw" c:(config)? ppSpace n:num+ s:rwRuleSeq l:(location)? : tactic => -- Note: This is a direct copy of `nth_rw` from core. match s with | `(rwRuleSeq| [$rs,*]%$rbrak) => -- We show the `rfl` state on `]` - `(tactic| (nth_rewrite $(c)? $n [$rs,*] $(l)?; with_annotate_state $rbrak + `(tactic| (nth_rewrite $(c)? $[$n]* [$rs,*] $(l)?; with_annotate_state $rbrak (try (with_reducible rfl)))) | _ => Macro.throwUnsupported diff --git a/test/NthRewrite.lean b/test/NthRewrite.lean index 5b9aaa207077c..608d18fdb6bf8 100644 --- a/test/NthRewrite.lean +++ b/test/NthRewrite.lean @@ -103,3 +103,6 @@ example (x y : ℕ) (h₁ : x = y) (h₂ : x = x + x) : x + x = x := by nth_rewrite 1 [← h₁] at h₂ nth_rewrite 3 [h₂] rfl + +example (x y : ℕ) (h : x = y) : x + x + x = x + y + y := by + nth_rw 2 3 [h] From 066f498929dd3240f275f6e607e716650781b1c4 Mon Sep 17 00:00:00 2001 From: Peter Nelson <71660771+apnelson1@users.noreply.github.com> Date: Thu, 8 Aug 2024 17:32:49 +0000 Subject: [PATCH 117/359] feat(Order/Zorn): change zorn to use Maximal API (#14781) We change statements of Zorn's lemma to use the new `Maximal` API from #14721. The new definitions eliminate the need for separate `PartialOrder` and `Preorder` versions of Zorn's lemma, and slightly simplify most of its applications. The new statements of Zorn aren't defeq to the old ones, so a good number of files are touched, but very few theorem statements outside `Order.Zorn` itself are actually changed - most changes are small modifications of existing proofs. - [x] depends on: #14721 --- Mathlib/Algebra/Module/Injective.lean | 8 +- Mathlib/Analysis/Convex/Cone/Extension.lean | 4 +- Mathlib/Analysis/Convex/KreinMilman.lean | 7 +- Mathlib/Analysis/Convex/StoneSeparation.lean | 7 +- Mathlib/Analysis/InnerProductSpace/Basic.lean | 5 +- Mathlib/Data/Matroid/IndepAxioms.lean | 34 +++---- Mathlib/FieldTheory/Extension.lean | 9 +- Mathlib/GroupTheory/Sylow.lean | 4 +- Mathlib/LinearAlgebra/LinearIndependent.lean | 10 +- Mathlib/MeasureTheory/Covering/Vitali.lean | 22 +++-- .../MeasureTheory/Measure/Haar/Unique.lean | 13 ++- Mathlib/Order/CompactlyGenerated/Basic.lean | 15 +-- Mathlib/Order/Extension/Linear.lean | 9 +- Mathlib/Order/PrimeSeparator.lean | 7 +- Mathlib/Order/Zorn.lean | 98 +++++++------------ Mathlib/Order/ZornAtoms.lean | 17 ++-- Mathlib/RingTheory/AlgebraicIndependent.lean | 29 ++---- Mathlib/RingTheory/Ideal/MinimalPrime.lean | 12 +-- Mathlib/RingTheory/Ideal/Operations.lean | 14 +-- Mathlib/RingTheory/PrincipalIdealDomain.lean | 19 ++-- .../Cardinal/SchroederBernstein.lean | 18 ++-- Mathlib/Topology/Algebra/Semigroup.lean | 10 +- Mathlib/Topology/Compactness/Compact.lean | 5 +- Mathlib/Topology/ExtremallyDisconnected.lean | 5 +- Mathlib/Topology/Irreducible.lean | 4 +- Mathlib/Topology/ShrinkingLemma.lean | 7 +- 26 files changed, 171 insertions(+), 221 deletions(-) diff --git a/Mathlib/Algebra/Module/Injective.lean b/Mathlib/Algebra/Module/Injective.lean index fd2e6fec70ebc..536a41a974e42 100644 --- a/Mathlib/Algebra/Module/Injective.lean +++ b/Mathlib/Algebra/Module/Injective.lean @@ -205,13 +205,13 @@ instance ExtensionOf.inhabited : Inhabited (ExtensionOf i f) where /-- Since every nonempty chain has a maximal element, by Zorn's lemma, there is a maximal `extension_of i f`. -/ def extensionOfMax : ExtensionOf i f := - (@zorn_nonempty_partialOrder (ExtensionOf i f) _ ⟨Inhabited.default⟩ fun _ hchain hnonempty => + (@zorn_le_nonempty (ExtensionOf i f) _ ⟨Inhabited.default⟩ fun _ hchain hnonempty => ⟨ExtensionOf.max hchain hnonempty, ExtensionOf.le_max hchain hnonempty⟩).choose theorem extensionOfMax_is_max : - ∀ a : ExtensionOf i f, extensionOfMax i f ≤ a → a = extensionOfMax i f := - (@zorn_nonempty_partialOrder (ExtensionOf i f) _ ⟨Inhabited.default⟩ fun _ hchain hnonempty => - ⟨ExtensionOf.max hchain hnonempty, ExtensionOf.le_max hchain hnonempty⟩).choose_spec + ∀ (a : ExtensionOf i f), extensionOfMax i f ≤ a → a = extensionOfMax i f := + fun _ ↦ (@zorn_le_nonempty (ExtensionOf i f) _ ⟨Inhabited.default⟩ fun _ hchain hnonempty => + ⟨ExtensionOf.max hchain hnonempty, ExtensionOf.le_max hchain hnonempty⟩).choose_spec.eq_of_ge -- Porting note: helper function. Lean looks for an instance of `Sup (Type u)` when the -- right hand side is substituted in directly diff --git a/Mathlib/Analysis/Convex/Cone/Extension.lean b/Mathlib/Analysis/Convex/Cone/Extension.lean index 256439bf7adb0..8b2c4f3992353 100644 --- a/Mathlib/Analysis/Convex/Cone/Extension.lean +++ b/Mathlib/Analysis/Convex/Cone/Extension.lean @@ -125,14 +125,14 @@ theorem exists_top (p : E →ₗ.[ℝ] ℝ) (hp_nonneg : ∀ x : p.domain, (x : have : f ≤ LinearPMap.sSup c hcd := LinearPMap.le_sSup _ hfc convert ← hcs hfc ⟨x, hfx⟩ hxs using 1 exact this.2 rfl - obtain ⟨q, hqs, hpq, hq⟩ := zorn_nonempty_partialOrder₀ S hSc p hp_nonneg + obtain ⟨q, hpq, hqs, hq⟩ := zorn_le_nonempty₀ S hSc p hp_nonneg refine ⟨q, hpq, ?_, hqs⟩ contrapose! hq have hqd : ∀ y, ∃ x : q.domain, (x : E) + y ∈ s := fun y ↦ let ⟨x, hx⟩ := hp_dense y ⟨Submodule.inclusion hpq.left x, hx⟩ rcases step s q hqs hqd hq with ⟨r, hqr, hr⟩ - exact ⟨r, hr, hqr.le, hqr.ne'⟩ + exact ⟨r, hr, hqr.le, fun hrq ↦ hqr.ne' <| hrq.antisymm hqr.le⟩ end RieszExtension diff --git a/Mathlib/Analysis/Convex/KreinMilman.lean b/Mathlib/Analysis/Convex/KreinMilman.lean index f8a590f768ad8..8b3207fdbc5e4 100644 --- a/Mathlib/Analysis/Convex/KreinMilman.lean +++ b/Mathlib/Analysis/Convex/KreinMilman.lean @@ -61,8 +61,9 @@ variable {E F : Type*} [AddCommGroup E] [Module ℝ E] [TopologicalSpace E] [T2S theorem IsCompact.extremePoints_nonempty (hscomp : IsCompact s) (hsnemp : s.Nonempty) : (s.extremePoints ℝ).Nonempty := by let S : Set (Set E) := { t | t.Nonempty ∧ IsClosed t ∧ IsExtreme ℝ s t } - rsuffices ⟨t, ⟨⟨x, hxt⟩, htclos, hst⟩, hBmin⟩ : ∃ t ∈ S, ∀ u ∈ S, u ⊆ t → u = t - · refine ⟨x, IsExtreme.mem_extremePoints ?_⟩ + rsuffices ⟨t, ht⟩ : ∃ t, Minimal (· ∈ S) t + · obtain ⟨⟨x,hxt⟩, htclos, hst⟩ := ht.prop + refine ⟨x, IsExtreme.mem_extremePoints ?_⟩ rwa [← eq_singleton_iff_unique_mem.2 ⟨hxt, fun y hyB => ?_⟩] by_contra hyx obtain ⟨l, hl⟩ := geometric_hahn_banach_point_point hyx @@ -70,7 +71,7 @@ theorem IsCompact.extremePoints_nonempty (hscomp : IsCompact s) (hsnemp : s.None (hscomp.of_isClosed_subset htclos hst.1).exists_isMaxOn ⟨x, hxt⟩ l.continuous.continuousOn have h : IsExposed ℝ t ({ z ∈ t | ∀ w ∈ t, l w ≤ l z }) := fun _ => ⟨l, rfl⟩ - rw [← hBmin ({ z ∈ t | ∀ w ∈ t, l w ≤ l z }) + rw [ ht.eq_of_ge (y := ({ z ∈ t | ∀ w ∈ t, l w ≤ l z })) ⟨⟨z, hzt, hz⟩, h.isClosed htclos, hst.trans h.isExtreme⟩ (t.sep_subset _)] at hyB exact hl.not_le (hyB.2 x hxt) refine zorn_superset _ fun F hFS hF => ?_ diff --git a/Mathlib/Analysis/Convex/StoneSeparation.lean b/Mathlib/Analysis/Convex/StoneSeparation.lean index 9b8ca97841e8b..3a0fdf9701a5f 100644 --- a/Mathlib/Analysis/Convex/StoneSeparation.lean +++ b/Mathlib/Analysis/Convex/StoneSeparation.lean @@ -78,7 +78,7 @@ theorem not_disjoint_segment_convexHull_triple {p q u v x y z : E} (hz : z ∈ s theorem exists_convex_convex_compl_subset (hs : Convex 𝕜 s) (ht : Convex 𝕜 t) (hst : Disjoint s t) : ∃ C : Set E, Convex 𝕜 C ∧ Convex 𝕜 Cᶜ ∧ s ⊆ C ∧ t ⊆ Cᶜ := by let S : Set (Set E) := { C | Convex 𝕜 C ∧ Disjoint C t } - obtain ⟨C, hC, hsC, hCmax⟩ := + obtain ⟨C, hsC, hmax⟩ := zorn_subset_nonempty S (fun c hcS hc ⟨_, _⟩ => ⟨⋃₀ c, @@ -86,6 +86,7 @@ theorem exists_convex_convex_compl_subset (hs : Convex 𝕜 s) (ht : Convex 𝕜 disjoint_sUnion_left.2 fun c hc => (hcS hc).2⟩, fun s => subset_sUnion_of_mem⟩) s ⟨hs, hst⟩ + obtain hC : _ ∧ _ := hmax.prop refine ⟨C, hC.1, convex_iff_segment_subset.2 fun x hx y hy z hz hzC => ?_, hsC, hC.2.subset_compl_left⟩ suffices h : ∀ c ∈ Cᶜ, ∃ a ∈ C, (segment 𝕜 c a ∩ t).Nonempty by @@ -98,8 +99,8 @@ theorem exists_convex_convex_compl_subset (hs : Convex 𝕜 s) (ht : Convex 𝕜 rintro c hc by_contra! h suffices h : Disjoint (convexHull 𝕜 (insert c C)) t by - rw [← - hCmax _ ⟨convex_convexHull _ _, h⟩ ((subset_insert _ _).trans <| subset_convexHull _ _)] at hc + rw [hmax.eq_of_subset ⟨convex_convexHull _ _, h⟩ <| + (subset_insert ..).trans <| subset_convexHull ..] at hc exact hc (subset_convexHull _ _ <| mem_insert _ _) rw [convexHull_insert ⟨z, hzC⟩, convexJoin_singleton_left] refine disjoint_iUnion₂_left.2 fun a ha => disjoint_iff_inter_eq_empty.2 (h a ?_) diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 8121332042b11..6b2938e160a15 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -981,9 +981,8 @@ theorem exists_maximal_orthonormal {s : Set E} (hs : Orthonormal 𝕜 (Subtype.v ∃ w ⊇ s, Orthonormal 𝕜 (Subtype.val : w → E) ∧ ∀ u ⊇ w, Orthonormal 𝕜 (Subtype.val : u → E) → u = w := by have := zorn_subset_nonempty { b | Orthonormal 𝕜 (Subtype.val : b → E) } ?_ _ hs - · obtain ⟨b, bi, sb, h⟩ := this - refine ⟨b, sb, bi, ?_⟩ - exact fun u hus hu => h u hu hus + · obtain ⟨b, hb⟩ := this + exact ⟨b, hb.1, hb.2.1, fun u hus hu => hb.2.eq_of_ge hu hus ⟩ · refine fun c hc cc _c0 => ⟨⋃₀ c, ?_, ?_⟩ · exact orthonormal_sUnion_of_directed cc.directedOn fun x xc => hc xc · exact fun _ => Set.subset_sUnion_of_mem diff --git a/Mathlib/Data/Matroid/IndepAxioms.lean b/Mathlib/Data/Matroid/IndepAxioms.lean index a71adf7dde620..3d56e013d70fa 100644 --- a/Mathlib/Data/Matroid/IndepAxioms.lean +++ b/Mathlib/Data/Matroid/IndepAxioms.lean @@ -215,27 +215,21 @@ namespace IndepMatroid rw [ncard_insert_of_not_mem heI₀ hI₀fin, ← Nat.lt_iff_add_one_le] at hcard obtain ⟨f, hfJ, hfI₀, hfi⟩ := indep_aug (indep_subset hI hI₀I) hI₀fin hJ hJfin hcard - exact hI₀ f ⟨Or.elim (hJss hfJ) (fun hfe ↦ (heJ <| hfe ▸ hfJ).elim) (by aesop), hfI₀⟩ hfi ) + exact hI₀ f ⟨Or.elim (hJss hfJ) (fun hfe ↦ (heJ <| hfe ▸ hfJ).elim) (by aesop), hfI₀⟩ hfi) (indep_maximal := by - rintro X - I hI hIX - have hzorn := zorn_subset_nonempty {Y | Indep Y ∧ I ⊆ Y ∧ Y ⊆ X} ?_ I ⟨hI, Subset.rfl, hIX⟩ - · obtain ⟨J, ⟨hJi, hIJ, hJX⟩, -, hJmax⟩ := hzorn - simp_rw [maximal_subset_iff] - exact ⟨J, hIJ, ⟨hJi, hJX⟩, fun K hK hJK ↦ (hJmax K ⟨hK.1, hIJ.trans hJK, hK.2⟩ hJK).symm⟩ - - refine fun Is hIs hchain ⟨K, hK⟩ ↦ ⟨⋃₀ Is, ⟨?_,?_,?_⟩, fun _ ↦ subset_sUnion_of_mem⟩ - · refine indep_compact _ fun J hJ hJfin ↦ ?_ - have hchoose : ∀ e, e ∈ J → ∃ I, I ∈ Is ∧ (e : α) ∈ I := fun _ he ↦ mem_sUnion.1 <| hJ he - choose! f hf using hchoose - refine J.eq_empty_or_nonempty.elim (fun hJ ↦ hJ ▸ indep_empty) (fun hne ↦ ?_) - obtain ⟨x, hxJ, hxmax⟩ := Finite.exists_maximal_wrt f _ hJfin hne - refine indep_subset (hIs (hf x hxJ).1).1 fun y hyJ ↦ ?_ - obtain (hle | hle) := hchain.total (hf _ hxJ).1 (hf _ hyJ).1 - · rw [hxmax _ hyJ hle]; exact (hf _ hyJ).2 - exact hle (hf _ hyJ).2 - - · exact subset_sUnion_of_subset _ K (hIs hK).2.1 hK - exact sUnion_subset fun X hX ↦ (hIs hX).2.2) + refine fun X _ I hI hIX ↦ zorn_subset_nonempty {Y | Indep Y ∧ Y ⊆ X} ?_ I ⟨hI, hIX⟩ + refine fun Is hIs hchain _ ↦ + ⟨⋃₀ Is, ⟨?_, sUnion_subset fun Y hY ↦ (hIs hY).2⟩, fun _ ↦ subset_sUnion_of_mem⟩ + refine indep_compact _ fun J hJ hJfin ↦ ?_ + have hchoose : ∀ e, e ∈ J → ∃ I, I ∈ Is ∧ (e : α) ∈ I := fun _ he ↦ mem_sUnion.1 <| hJ he + choose! f hf using hchoose + refine J.eq_empty_or_nonempty.elim (fun hJ ↦ hJ ▸ indep_empty) (fun hne ↦ ?_) + obtain ⟨x, hxJ, hxmax⟩ := Finite.exists_maximal_wrt f _ hJfin hne + refine indep_subset (hIs (hf x hxJ).1).1 fun y hyJ ↦ ?_ + obtain (hle | hle) := hchain.total (hf _ hxJ).1 (hf _ hyJ).1 + · rw [hxmax _ hyJ hle]; exact (hf _ hyJ).2 + exact hle (hf _ hyJ).2) + (subset_ground := subset_ground) @[simp] theorem ofFinitary_indep (E : Set α) (Indep : Set α → Prop) diff --git a/Mathlib/FieldTheory/Extension.lean b/Mathlib/FieldTheory/Extension.lean index 024192b91c90d..2ded21368a221 100644 --- a/Mathlib/FieldTheory/Extension.lean +++ b/Mathlib/FieldTheory/Extension.lean @@ -97,7 +97,7 @@ section private theorem exists_algHom_adjoin_of_splits'' {L : IntermediateField F E} (f : L →ₐ[F] K) (hK : ∀ s ∈ S, IsIntegral L s ∧ (minpoly L s).Splits f.toRingHom) : ∃ φ : adjoin L S →ₐ[F] K, φ.comp (IsScalarTower.toAlgHom F L _) = f := by - obtain ⟨φ, hfφ, hφ⟩ := zorn_nonempty_Ici₀ _ + obtain ⟨φ, hfφ, hφ⟩ := zorn_le_nonempty_Ici₀ _ (fun c _ hc _ _ ↦ Lifts.exists_upper_bound c hc) ⟨L, f⟩ le_rfl refine ⟨φ.emb.comp (inclusion <| (le_extendScalars_iff hfφ.1 <| adjoin L S).mp <| adjoin_le_iff.mpr fun s h ↦ ?_), AlgHom.ext hfφ.2⟩ @@ -105,7 +105,8 @@ private theorem exists_algHom_adjoin_of_splits'' {L : IntermediateField F E} letI : SMul L φ.carrier := Algebra.toSMul have : IsScalarTower L φ.carrier E := ⟨(smul_assoc · (· : E))⟩ have := φ.exists_lift_of_splits' (hK s h).1.tower_top ((hK s h).1.minpoly_splits_tower_top' ?_) - · obtain ⟨y, h1, h2⟩ := this; exact (hφ y h1).1 h2 + · obtain ⟨y, h1, h2⟩ := this + exact (hφ h1).1 h2 · convert (hK s h).2; ext; apply hfφ.2 variable {L : Type*} [Field L] [Algebra F L] [Algebra L E] [IsScalarTower F L E] @@ -146,11 +147,11 @@ variable (hK : ∀ s ∈ S, IsIntegral F s ∧ (minpoly F s).Splits (algebraMap -- The following uses the hypothesis `hK`. theorem exists_algHom_adjoin_of_splits : ∃ φ : adjoin F S →ₐ[F] K, φ.comp (inclusion hL) = f := by - obtain ⟨φ, hfφ, hφ⟩ := zorn_nonempty_Ici₀ _ + obtain ⟨φ, hfφ, hφ⟩ := zorn_le_nonempty_Ici₀ _ (fun c _ hc _ _ ↦ Lifts.exists_upper_bound c hc) ⟨L, f⟩ le_rfl refine ⟨φ.emb.comp (inclusion <| adjoin_le_iff.mpr fun s hs ↦ ?_), ?_⟩ · rcases φ.exists_lift_of_splits (hK s hs).1 (hK s hs).2 with ⟨y, h1, h2⟩ - exact (hφ y h1).1 h2 + exact (hφ h1).1 h2 · ext; apply hfφ.2 theorem nonempty_algHom_adjoin_of_splits : Nonempty (adjoin F S →ₐ[F] K) := diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index 578db97ec19ca..45a76037e7714 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -126,7 +126,7 @@ end Sylow Every `p`-subgroup is contained in a Sylow `p`-subgroup. -/ theorem IsPGroup.exists_le_sylow {P : Subgroup G} (hP : IsPGroup p P) : ∃ Q : Sylow p G, P ≤ Q := Exists.elim - (zorn_nonempty_partialOrder₀ { Q : Subgroup G | IsPGroup p Q } + (zorn_le_nonempty₀ { Q : Subgroup G | IsPGroup p Q } (fun c hc1 hc2 Q hQ => ⟨{ carrier := ⋃ R : c, R one_mem' := ⟨Q, ⟨⟨Q, hQ⟩, rfl⟩, Q.one_mem⟩ @@ -138,7 +138,7 @@ theorem IsPGroup.exists_le_sylow {P : Subgroup G} (hP : IsPGroup p P) : ∃ Q : refine Exists.imp (fun k hk => ?_) (hc1 S.2 ⟨g, hg⟩) rwa [Subtype.ext_iff, coe_pow] at hk ⊢, fun M hM g hg => ⟨M, ⟨⟨M, hM⟩, rfl⟩, hg⟩⟩) P hP) - fun {Q} ⟨hQ1, hQ2, hQ3⟩ => ⟨⟨Q, hQ1, hQ3 _⟩, hQ2⟩ + fun {Q} h => ⟨⟨Q, h.2.prop, h.2.eq_of_ge⟩, h.1⟩ instance Sylow.nonempty : Nonempty (Sylow p G) := nonempty_of_exists IsPGroup.of_bot.exists_le_sylow diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 9420ed64bedfb..89caba69587a6 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -1339,13 +1339,9 @@ theorem exists_linearIndependent_extension (hs : LinearIndependent K ((↑) : s · exact sUnion_subset fun x xc => (hc xc).1 · exact linearIndependent_sUnion_of_directed cc.directedOn fun x xc => (hc xc).2 · exact subset_sUnion_of_mem - rcases this with - ⟨b, ⟨bt, bi⟩, sb, h⟩ - refine ⟨b, bt, sb, fun x xt => ?_, bi⟩ - by_contra hn - apply hn - rw [← h _ ⟨insert_subset_iff.2 ⟨xt, bt⟩, bi.insert hn⟩ (subset_insert _ _)] - exact subset_span (mem_insert _ _) + obtain ⟨b, sb, h⟩ := this + refine ⟨b, h.prop.1, sb, fun x xt => by_contra fun hn ↦ hn ?_, h.prop.2⟩ + exact subset_span <| h.mem_of_prop_insert ⟨insert_subset xt h.prop.1, h.prop.2.insert hn⟩ variable (K t) diff --git a/Mathlib/MeasureTheory/Covering/Vitali.lean b/Mathlib/MeasureTheory/Covering/Vitali.lean index 94ab1f0331870..5928bb689a9c9 100644 --- a/Mathlib/MeasureTheory/Covering/Vitali.lean +++ b/Mathlib/MeasureTheory/Covering/Vitali.lean @@ -74,7 +74,7 @@ theorem exists_disjoint_subfamily_covering_enlargment (B : ι → Set α) (t : S let T : Set (Set ι) := { u | u ⊆ t ∧ u.PairwiseDisjoint B ∧ ∀ a ∈ t, ∀ b ∈ u, (B a ∩ B b).Nonempty → ∃ c ∈ u, (B a ∩ B c).Nonempty ∧ δ a ≤ τ * δ c } -- By Zorn, choose a maximal family in the good set `T` of disjoint families. - obtain ⟨u, uT, hu⟩ : ∃ u ∈ T, ∀ v ∈ T, u ⊆ v → v = u := by + have hzorn : ∃ m, Maximal (fun x ↦ x ∈ T) m := by refine zorn_subset _ fun U UT hU => ?_ refine ⟨⋃₀ U, ?_, fun s hs => subset_sUnion_of_mem hs⟩ simp only [T, Set.sUnion_subset_iff, and_imp, exists_prop, forall_exists_index, mem_sUnion, @@ -85,17 +85,18 @@ theorem exists_disjoint_subfamily_covering_enlargment (B : ι → Set α) (t : S obtain ⟨c, cu, ac, hc⟩ : ∃ c, c ∈ u ∧ (B a ∩ B c).Nonempty ∧ δ a ≤ τ * δ c := (UT uU).2.2 a hat b hbu hab exact ⟨c, ⟨u, uU, cu⟩, ac, hc⟩ + obtain ⟨u, hu⟩ := hzorn -- The only nontrivial bit is to check that every `a ∈ t` intersects an element `b ∈ u` with -- comparatively large `δ b`. Assume this is not the case, then we will contradict the maximality. - refine ⟨u, uT.1, uT.2.1, fun a hat => ?_⟩ - contrapose! hu + refine ⟨u, hu.prop.1, hu.prop.2.1, fun a hat => ?_⟩ + by_contra! hcon have a_disj : ∀ c ∈ u, Disjoint (B a) (B c) := by intro c hc by_contra h rw [not_disjoint_iff_nonempty_inter] at h obtain ⟨d, du, ad, hd⟩ : ∃ d, d ∈ u ∧ (B a ∩ B d).Nonempty ∧ δ a ≤ τ * δ d := - uT.2.2 a hat c hc h - exact lt_irrefl _ ((hu d du ad).trans_le hd) + hu.prop.2.2 a hat c hc h + exact lt_irrefl _ ((hcon d du ad).trans_le hd) -- Let `A` be all the elements of `t` which do not intersect the family `u`. It is nonempty as it -- contains `a`. We will pick an element `a'` of `A` with `δ a'` almost as large as possible. let A := { a' | a' ∈ t ∧ ∀ c ∈ u, Disjoint (B a') (B c) } @@ -117,16 +118,17 @@ theorem exists_disjoint_subfamily_covering_enlargment (B : ι → Set α) (t : S rcases exists_lt_of_lt_csSup (Anonempty.image _) I with ⟨x, xA, hx⟩ rcases (mem_image _ _ _).1 xA with ⟨a', ha', rfl⟩ exact ⟨a', ha', hx.le⟩ - clear hat hu a_disj a + clear hat hcon a_disj a have a'_ne_u : a' ∉ u := fun H => (hne _ a'A.1).ne_empty (disjoint_self.1 (a'A.2 _ H)) -- we claim that `u ∪ {a'}` still belongs to `T`, contradicting the maximality of `u`. - refine ⟨insert a' u, ⟨?_, ?_, ?_⟩, subset_insert _ _, (ne_insert_of_not_mem _ a'_ne_u).symm⟩ + refine a'_ne_u (hu.mem_of_prop_insert ⟨?_, ?_, ?_⟩) + -- refine ⟨insert a' u, ⟨?_, ?_, ?_⟩, subset_insert _ _, (ne_insert_of_not_mem _ a'_ne_u).symm⟩ · -- check that `u ∪ {a'}` is made of elements of `t`. rw [insert_subset_iff] - exact ⟨a'A.1, uT.1⟩ + exact ⟨a'A.1, hu.prop.1⟩ · -- Check that `u ∪ {a'}` is a disjoint family. This follows from the fact that `a'` does not -- intersect `u`. - exact uT.2.1.insert fun b bu _ => a'A.2 b bu + exact hu.prop.2.1.insert fun b bu _ => a'A.2 b bu · -- check that every element `c` of `t` intersecting `u ∪ {a'}` intersects an element of this -- family with large `δ`. intro c ct b ba'u hcb @@ -134,7 +136,7 @@ theorem exists_disjoint_subfamily_covering_enlargment (B : ι → Set α) (t : S -- large `δ` by the assumption on `u`, and there is nothing left to do. by_cases H : ∃ d ∈ u, (B c ∩ B d).Nonempty · rcases H with ⟨d, du, hd⟩ - rcases uT.2.2 c ct d du hd with ⟨d', d'u, hd'⟩ + rcases hu.prop.2.2 c ct d du hd with ⟨d', d'u, hd'⟩ exact ⟨d', mem_insert_of_mem _ d'u, hd'⟩ · -- Otherwise, `c` belongs to `A`. The element of `u ∪ {a'}` that it intersects has to be `a'`. -- Moreover, `δ c` is smaller than the maximum `m` of `δ` over `A`, which is `≤ δ a' / τ` diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index 91e4b80b8b62d..aa8ebde841979 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -709,7 +709,7 @@ theorem measure_isHaarMeasure_eq_smul_of_isEverywherePos [LocallyCompactSpace G] exact ⟨closure k, hk.closure, isClosed_closure, mem_of_superset hmem subset_closure⟩ have one_k : 1 ∈ k := mem_of_mem_nhds k_mem let A : Set (Set G) := {t | t ⊆ s ∧ PairwiseDisjoint t (fun x ↦ x • k)} - obtain ⟨m, mA, m_max⟩ : ∃ m ∈ A, ∀ a ∈ A, m ⊆ a → a = m := by + obtain ⟨m, m_max⟩ : ∃ m, Maximal (· ∈ A) m := by apply zorn_subset intro c cA hc refine ⟨⋃ a ∈ c, a, ⟨?_, ?_⟩, ?_⟩ @@ -725,17 +725,16 @@ theorem measure_isHaarMeasure_eq_smul_of_isEverywherePos [LocallyCompactSpace G] exact (cA mc).2 (am xa) (bm yb) hxy · intro a ac exact subset_biUnion_of_mem (u := id) ac - change m ⊆ s ∧ PairwiseDisjoint m (fun x ↦ x • k) at mA + obtain ⟨hms : m ⊆ s, hdj : PairwiseDisjoint m (fun x ↦ x • k)⟩ := m_max.prop have sm : s ⊆ ⋃ x ∈ m, x • (k * k⁻¹) := by intro y hy by_cases h'y : m ∪ {y} ∈ A - · have : m ∪ {y} = m := m_max _ h'y subset_union_left - have ym : y ∈ m := by simpa using subset_union_right.trans this.subset + · have ym : y ∈ m := m_max.mem_of_prop_insert (by simpa using h'y) have : y ∈ y • (k * k⁻¹) := by simpa using mem_leftCoset y (Set.mul_mem_mul one_k (Set.inv_mem_inv.mpr one_k)) exact mem_biUnion ym this · obtain ⟨x, xm, -, z, zy, zx⟩ : ∃ x ∈ m, y ≠ x ∧ ∃ z, z ∈ y • k ∧ z ∈ x • k := by - simpa [A, mA.1, hy, insert_subset_iff, pairwiseDisjoint_insert, mA.2, not_disjoint_iff] + simpa [A, hms, hy, insert_subset_iff, pairwiseDisjoint_insert, hdj, not_disjoint_iff] using h'y have : y ∈ x • (k * k⁻¹) := by rw [show y = x * ((x⁻¹ * z) * (y⁻¹ * z)⁻¹) by group] @@ -771,14 +770,14 @@ theorem measure_isHaarMeasure_eq_smul_of_isEverywherePos [LocallyCompactSpace G] have : ∑' (x : m), ρ (s ∩ ((x : G) • k)) < ∞ := by apply lt_of_le_of_lt (MeasureTheory.tsum_meas_le_meas_iUnion_of_disjoint _ M _) _ · have I : PairwiseDisjoint m fun x ↦ s ∩ x • k := - mA.2.mono (fun x ↦ inter_subset_right) + hdj.mono (fun x ↦ inter_subset_right) exact I.on_injective Subtype.val_injective (fun x ↦ x.2) · exact lt_of_le_of_lt (measure_mono (by simp [inter_subset_left])) h'm.lt_top have C : Set.Countable (support fun (i : m) ↦ ρ (s ∩ (i : G) • k)) := Summable.countable_support_ennreal this.ne have : support (fun (i : m) ↦ ρ (s ∩ (i : G) • k)) = univ := by apply eq_univ_iff_forall.2 (fun i ↦ ?_) - apply ne_of_gt (hρ (i : G) (mA.1 i.2) _ _) + apply ne_of_gt (hρ (i : G) (hms i.2) _ _) exact inter_mem_nhdsWithin s (by simpa using smul_mem_nhds (i : G) k_mem) rw [this] at C have : Countable m := countable_univ_iff.mp C diff --git a/Mathlib/Order/CompactlyGenerated/Basic.lean b/Mathlib/Order/CompactlyGenerated/Basic.lean index 8623c948b430d..5a2890bd05342 100644 --- a/Mathlib/Order/CompactlyGenerated/Basic.lean +++ b/Mathlib/Order/CompactlyGenerated/Basic.lean @@ -471,9 +471,9 @@ theorem Iic_coatomic_of_compact_element {k : α} (h : IsCompactElement k) : obtain rfl | H := eq_or_ne b k · left; ext; simp only [Set.Iic.coe_top, Subtype.coe_mk] right - have ⟨a, a₀, ba, h⟩ := zorn_nonempty_partialOrder₀ (Set.Iio k) ?_ b (lt_of_le_of_ne hbk H) - · refine ⟨⟨a, le_of_lt a₀⟩, ⟨ne_of_lt a₀, fun c hck => by_contradiction fun c₀ => ?_⟩, ba⟩ - cases h c.1 (lt_of_le_of_ne c.2 fun con => c₀ (Subtype.ext con)) hck.le + have ⟨a, ba, h⟩ := zorn_le_nonempty₀ (Set.Iio k) ?_ b (lt_of_le_of_ne hbk H) + · refine ⟨⟨a, le_of_lt h.prop⟩, ⟨ne_of_lt h.prop, fun c hck => by_contradiction fun c₀ => ?_⟩, ba⟩ + cases h.eq_of_le (y := c.1) (lt_of_le_of_ne c.2 fun con ↦ c₀ (Subtype.ext con)) hck.le exact lt_irrefl _ hck · intro S SC cC I _ by_cases hS : S.Nonempty @@ -542,8 +542,8 @@ theorem exists_setIndependent_isCompl_sSup_atoms (h : sSup { a : α | IsAtom a } 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_subset - {s : Set α | CompleteLattice.SetIndependent s ∧ Disjoint b (sSup s) ∧ ∀ a ∈ s, IsAtom a} + have zorn := zorn_subset + (S := {s : Set α | CompleteLattice.SetIndependent 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, ?_, @@ -555,7 +555,8 @@ theorem exists_setIndependent_isCompl_sSup_atoms (h : sSup { a : α | IsAtom a } exact (hc1 hs).2.1 · rw [directedOn_image] exact hc2.directedOn.mono @fun s t => sSup_le_sSup - obtain ⟨s, ⟨s_ind, b_inf_Sup_s, s_atoms⟩, s_max⟩ := this + simp_rw [maximal_subset_iff] at zorn + obtain ⟨s, ⟨s_ind, b_inf_Sup_s, s_atoms⟩, s_max⟩ := zorn refine ⟨s, s_ind, ⟨b_inf_Sup_s, ?_⟩, s_atoms⟩ rw [codisjoint_iff_le_sup, ← h, sSup_le_iff] intro a ha @@ -566,7 +567,7 @@ theorem exists_setIndependent_isCompl_sSup_atoms (h : sSup { a : α | IsAtom a } rw [← disjoint_iff] at con have a_dis_Sup_s : Disjoint a (sSup s) := con.mono_right le_sup_right -- Porting note: The two following `fun x hx => _` are no-op - rw [← s_max (s ∪ {a}) ⟨fun x hx => _, _, fun x hx => _⟩ Set.subset_union_left] + rw [@s_max (s ∪ {a}) ⟨fun x hx => _, _, fun x hx => _⟩ Set.subset_union_left] · exact Set.mem_union_right _ (Set.mem_singleton _) · intro x hx rw [Set.mem_union, Set.mem_singleton_iff] at hx diff --git a/Mathlib/Order/Extension/Linear.lean b/Mathlib/Order/Extension/Linear.lean index fe6aaffb272d7..92f3dc8345584 100644 --- a/Mathlib/Order/Extension/Linear.lean +++ b/Mathlib/Order/Extension/Linear.lean @@ -45,13 +45,14 @@ theorem extend_partialOrder {α : Type u} (r : α → α → Prop) [IsPartialOrd cases' hc₂.total h₁s₁ h₁s₂ with h h · exact antisymm (h _ _ h₂s₁) h₂s₂ · apply antisymm h₂s₁ (h _ _ h₂s₂) - obtain ⟨s, hs₁ : IsPartialOrder _ _, rs, hs₂⟩ := zorn_nonempty_partialOrder₀ S hS r ‹_› - haveI : IsPartialOrder α s := hs₁ - refine ⟨s, { total := ?_, refl := hs₁.refl, trans := hs₁.trans, antisymm := hs₁.antisymm } , rs⟩ + obtain ⟨s, hrs, hs⟩ := zorn_le_nonempty₀ S hS r ‹_› + haveI : IsPartialOrder α s := hs.prop + refine ⟨s, + { total := ?_, refl := hs.1.refl, trans := hs.1.trans, antisymm := hs.1.antisymm }, hrs⟩ intro x y by_contra! h let s' x' y' := s x' y' ∨ s x' x ∧ s y y' - rw [← hs₂ s' _ fun _ _ ↦ Or.inl] at h + rw [hs.eq_of_le (y := s') ?_ fun _ _ ↦ Or.inl] at h · apply h.1 (Or.inr ⟨refl _, refl _⟩) · refine { refl := fun x ↦ Or.inl (refl _) diff --git a/Mathlib/Order/PrimeSeparator.lean b/Mathlib/Order/PrimeSeparator.lean index 6e37a106316ce..25826ad97ce55 100644 --- a/Mathlib/Order/PrimeSeparator.lean +++ b/Mathlib/Order/PrimeSeparator.lean @@ -63,7 +63,8 @@ theorem prime_ideal_of_disjoint_filter_ideal (hFI : Disjoint (F : Set α) (I : S ⟨le_trans (hcS hJ).2.1 (le_sSup hJ), fun J hJ ↦ (hcS hJ).2.2⟩⟩ -- Thus, by Zorn's lemma, we can pick a maximal ideal J in S. - obtain ⟨Jset, ⟨Jidl, IJ, JF⟩, ⟨_, Jmax⟩⟩ := zorn_subset_nonempty S chainub I IinS + obtain ⟨Jset, _, hmax⟩ := zorn_subset_nonempty S chainub I IinS + obtain ⟨Jidl, IJ, JF⟩ := hmax.prop set J := IsIdeal.toIdeal Jidl use J have IJ' : I ≤ J := IJ @@ -94,8 +95,8 @@ theorem prime_ideal_of_disjoint_filter_ideal (hFI : Disjoint (F : Set α) (I : S have J₂J : ↑J₂ ≠ Jset := ne_of_mem_of_not_mem' a₂J₂ ha₂ -- Therefore, since J is maximal, we must have Jᵢ ∉ S. - have J₁S : ↑J₁ ∉ S := fun h => J₁J (Jmax J₁ h (le_sup_left : J ≤ J₁)) - have J₂S : ↑J₂ ∉ S := fun h => J₂J (Jmax J₂ h (le_sup_left : J ≤ J₂)) + have J₁S : ↑J₁ ∉ S := fun h => J₁J (hmax.eq_of_le h (le_sup_left : J ≤ J₁)).symm + have J₂S : ↑J₂ ∉ S := fun h => J₂J (hmax.eq_of_le h (le_sup_left : J ≤ J₂)).symm -- Since Jᵢ is an ideal that contains I, we have that Jᵢ is not disjoint from F. have J₁F : ¬ (Disjoint (F : Set α) J₁) := by diff --git a/Mathlib/Order/Zorn.lean b/Mathlib/Order/Zorn.lean index c42459be022d5..e31a5347f9d81 100644 --- a/Mathlib/Order/Zorn.lean +++ b/Mathlib/Order/Zorn.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Order.Chain +import Mathlib.Order.Minimal /-! # Zorn's lemmas @@ -14,7 +15,7 @@ This file proves several formulations of Zorn's Lemma. The primary statement of Zorn's lemma is `exists_maximal_of_chains_bounded`. Then it is specialized to particular relations: -* `(≤)` with `zorn_partialOrder` +* `(≤)` with `zorn_le` * `(⊆)` with `zorn_subset` * `(⊇)` with `zorn_superset` @@ -58,6 +59,8 @@ Originally ported from Isabelle/HOL. The Fleuriot, Tobias Nipkow, Christian Sternagel. -/ + +open scoped Classical open Set variable {α β : Type*} {r : α → α → Prop} {c : Set α} @@ -95,98 +98,67 @@ section Preorder variable [Preorder α] -theorem zorn_preorder (h : ∀ c : Set α, IsChain (· ≤ ·) c → BddAbove c) : - ∃ m : α, ∀ a, m ≤ a → a ≤ m := +theorem zorn_le (h : ∀ c : Set α, IsChain (· ≤ ·) c → BddAbove c) : ∃ m : α, IsMax m := exists_maximal_of_chains_bounded h le_trans -theorem zorn_nonempty_preorder [Nonempty α] - (h : ∀ c : Set α, IsChain (· ≤ ·) c → c.Nonempty → BddAbove c) : ∃ m : α, ∀ a, m ≤ a → a ≤ m := +theorem zorn_le_nonempty [Nonempty α] + (h : ∀ c : Set α, IsChain (· ≤ ·) c → c.Nonempty → BddAbove c) : ∃ m : α, IsMax m := exists_maximal_of_nonempty_chains_bounded h le_trans -theorem zorn_preorder₀ (s : Set α) - (ih : ∀ c ⊆ s, IsChain (· ≤ ·) c → ∃ ub ∈ s, ∀ z ∈ c, z ≤ ub) : - ∃ m ∈ s, ∀ z ∈ s, m ≤ z → z ≤ m := +theorem zorn_le₀ (s : Set α) (ih : ∀ c ⊆ s, IsChain (· ≤ ·) c → ∃ ub ∈ s, ∀ z ∈ c, z ≤ ub) : + ∃ m, Maximal (· ∈ s) m := let ⟨⟨m, hms⟩, h⟩ := - @zorn_preorder s _ fun c hc => + @zorn_le s _ fun c hc => let ⟨ub, hubs, hub⟩ := ih (Subtype.val '' c) (fun _ ⟨⟨_, hx⟩, _, h⟩ => h ▸ hx) (by rintro _ ⟨p, hpc, rfl⟩ _ ⟨q, hqc, rfl⟩ hpq exact hc hpc hqc fun t => hpq (Subtype.ext_iff.1 t)) ⟨⟨ub, hubs⟩, fun ⟨y, hy⟩ hc => hub _ ⟨_, hc, rfl⟩⟩ - ⟨m, hms, fun z hzs hmz => h ⟨z, hzs⟩ hmz⟩ + ⟨m, hms, fun z hzs hmz => @h ⟨z, hzs⟩ hmz⟩ -theorem zorn_nonempty_preorder₀ (s : Set α) - (ih : ∀ c ⊆ s, IsChain (· ≤ ·) c → ∀ y ∈ c, ∃ ub ∈ s, ∀ z ∈ c, z ≤ ub) (x : α) - (hxs : x ∈ s) : ∃ m ∈ s, x ≤ m ∧ ∀ z ∈ s, m ≤ z → z ≤ m := by +theorem zorn_le_nonempty₀ (s : Set α) + (ih : ∀ c ⊆ s, IsChain (· ≤ ·) c → ∀ y ∈ c, ∃ ub ∈ s, ∀ z ∈ c, z ≤ ub) (x : α) (hxs : x ∈ s) : + ∃ m, x ≤ m ∧ Maximal (· ∈ s) m := by -- Porting note: the first three lines replace the following two lines in mathlib3. -- The mathlib3 `rcases` supports holes for proof obligations, this is not yet implemented in 4. -- rcases zorn_preorder₀ ({ y ∈ s | x ≤ y }) fun c hcs hc => ?_ with ⟨m, ⟨hms, hxm⟩, hm⟩ -- · exact ⟨m, hms, hxm, fun z hzs hmz => hm _ ⟨hzs, hxm.trans hmz⟩ hmz⟩ - have H := zorn_preorder₀ ({ y ∈ s | x ≤ y }) fun c hcs hc => ?_ + have H := zorn_le₀ ({ y ∈ s | x ≤ y }) fun c hcs hc => ?_ · rcases H with ⟨m, ⟨hms, hxm⟩, hm⟩ - exact ⟨m, hms, hxm, fun z hzs hmz => hm _ ⟨hzs, hxm.trans hmz⟩ hmz⟩ + exact ⟨m, hxm, hms, fun z hzs hmz => @hm _ ⟨hzs, hxm.trans hmz⟩ hmz⟩ · rcases c.eq_empty_or_nonempty with (rfl | ⟨y, hy⟩) · exact ⟨x, ⟨hxs, le_rfl⟩, fun z => False.elim⟩ · rcases ih c (fun z hz => (hcs hz).1) hc y hy with ⟨z, hzs, hz⟩ exact ⟨z, ⟨hzs, (hcs hy).2.trans <| hz _ hy⟩, hz⟩ -theorem zorn_nonempty_Ici₀ (a : α) - (ih : ∀ c ⊆ Ici a, IsChain (· ≤ ·) c → ∀ y ∈ c, ∃ ub, ∀ z ∈ c, z ≤ ub) - (x : α) (hax : a ≤ x) : ∃ m, x ≤ m ∧ ∀ z, m ≤ z → z ≤ m := by - let ⟨m, _, hxm, hm⟩ := zorn_nonempty_preorder₀ (Ici a) (fun c hca hc y hy ↦ ?_) x hax - · exact ⟨m, hxm, fun z hmz => hm _ (hax.trans <| hxm.trans hmz) hmz⟩ - · have ⟨ub, hub⟩ := ih c hca hc y hy; exact ⟨ub, (hca hy).trans (hub y hy), hub⟩ +theorem zorn_le_nonempty_Ici₀ (a : α) + (ih : ∀ c ⊆ Ici a, IsChain (· ≤ ·) c → ∀ y ∈ c, ∃ ub, ∀ z ∈ c, z ≤ ub) (x : α) (hax : a ≤ x) : + ∃ m, x ≤ m ∧ IsMax m := by + let ⟨m, hxm, ham, hm⟩ := zorn_le_nonempty₀ (Ici a) (fun c hca hc y hy ↦ ?_) x hax + · exact ⟨m, hxm, fun z hmz => hm (ham.trans hmz) hmz⟩ + · have ⟨ub, hub⟩ := ih c hca hc y hy + exact ⟨ub, (hca hy).trans (hub y hy), hub⟩ end Preorder -section PartialOrder - -variable [PartialOrder α] - -theorem zorn_partialOrder (h : ∀ c : Set α, IsChain (· ≤ ·) c → BddAbove c) : - ∃ m : α, ∀ a, m ≤ a → a = m := - let ⟨m, hm⟩ := zorn_preorder h - ⟨m, fun a ha => le_antisymm (hm a ha) ha⟩ - -theorem zorn_nonempty_partialOrder [Nonempty α] - (h : ∀ c : Set α, IsChain (· ≤ ·) c → c.Nonempty → BddAbove c) : ∃ m : α, ∀ a, m ≤ a → a = m := - let ⟨m, hm⟩ := zorn_nonempty_preorder h - ⟨m, fun a ha => le_antisymm (hm a ha) ha⟩ - -theorem zorn_partialOrder₀ (s : Set α) - (ih : ∀ c ⊆ s, IsChain (· ≤ ·) c → ∃ ub ∈ s, ∀ z ∈ c, z ≤ ub) : - ∃ m ∈ s, ∀ z ∈ s, m ≤ z → z = m := - let ⟨m, hms, hm⟩ := zorn_preorder₀ s ih - ⟨m, hms, fun z hzs hmz => (hm z hzs hmz).antisymm hmz⟩ - -theorem zorn_nonempty_partialOrder₀ (s : Set α) - (ih : ∀ c ⊆ s, IsChain (· ≤ ·) c → ∀ y ∈ c, ∃ ub ∈ s, ∀ z ∈ c, z ≤ ub) (x : α) - (hxs : x ∈ s) : ∃ m ∈ s, x ≤ m ∧ ∀ z ∈ s, m ≤ z → z = m := - let ⟨m, hms, hxm, hm⟩ := zorn_nonempty_preorder₀ s ih x hxs - ⟨m, hms, hxm, fun z hzs hmz => (hm z hzs hmz).antisymm hmz⟩ - -end PartialOrder - theorem zorn_subset (S : Set (Set α)) - (h : ∀ c ⊆ S, IsChain (· ⊆ ·) c → ∃ ub ∈ S, ∀ s ∈ c, s ⊆ ub) : - ∃ m ∈ S, ∀ a ∈ S, m ⊆ a → a = m := - zorn_partialOrder₀ S h + (h : ∀ c ⊆ S, IsChain (· ⊆ ·) c → ∃ ub ∈ S, ∀ s ∈ c, s ⊆ ub) : ∃ m, Maximal (· ∈ S) m := + zorn_le₀ S h theorem zorn_subset_nonempty (S : Set (Set α)) - (H : ∀ c ⊆ S, IsChain (· ⊆ ·) c → c.Nonempty → ∃ ub ∈ S, ∀ s ∈ c, s ⊆ ub) (x) - (hx : x ∈ S) : ∃ m ∈ S, x ⊆ m ∧ ∀ a ∈ S, m ⊆ a → a = m := - zorn_nonempty_partialOrder₀ _ (fun _ cS hc y yc => H _ cS hc ⟨y, yc⟩) _ hx + (H : ∀ c ⊆ S, IsChain (· ⊆ ·) c → c.Nonempty → ∃ ub ∈ S, ∀ s ∈ c, s ⊆ ub) (x) (hx : x ∈ S) : + ∃ m, x ⊆ m ∧ Maximal (· ∈ S) m := + zorn_le_nonempty₀ _ (fun _ cS hc y yc => H _ cS hc ⟨y, yc⟩) _ hx theorem zorn_superset (S : Set (Set α)) - (h : ∀ c ⊆ S, IsChain (· ⊆ ·) c → ∃ lb ∈ S, ∀ s ∈ c, lb ⊆ s) : - ∃ m ∈ S, ∀ a ∈ S, a ⊆ m → a = m := - (@zorn_partialOrder₀ (Set α)ᵒᵈ _ S) fun c cS hc => h c cS hc.symm + (h : ∀ c ⊆ S, IsChain (· ⊆ ·) c → ∃ lb ∈ S, ∀ s ∈ c, lb ⊆ s) : ∃ m, Minimal (· ∈ S) m := + (@zorn_le₀ (Set α)ᵒᵈ _ S) fun c cS hc => h c cS hc.symm theorem zorn_superset_nonempty (S : Set (Set α)) - (H : ∀ c ⊆ S, IsChain (· ⊆ ·) c → c.Nonempty → ∃ lb ∈ S, ∀ s ∈ c, lb ⊆ s) (x) - (hx : x ∈ S) : ∃ m ∈ S, m ⊆ x ∧ ∀ a ∈ S, a ⊆ m → a = m := - @zorn_nonempty_partialOrder₀ (Set α)ᵒᵈ _ S (fun _ cS hc y yc => H _ cS hc.symm ⟨y, yc⟩) _ hx + (H : ∀ c ⊆ S, IsChain (· ⊆ ·) c → c.Nonempty → ∃ lb ∈ S, ∀ s ∈ c, lb ⊆ s) (x) (hx : x ∈ S) : + ∃ m, m ⊆ x ∧ Minimal (· ∈ S) m := + @zorn_le_nonempty₀ (Set α)ᵒᵈ _ S (fun _ cS hc y yc => H _ cS hc.symm ⟨y, yc⟩) _ hx /-- Every chain is contained in a maximal chain. This generalizes Hausdorff's maximality principle. -/ @@ -196,8 +168,8 @@ theorem IsChain.exists_maxChain (hc : IsChain r c) : ∃ M, @IsMaxChain _ r M -- obtain ⟨M, ⟨_, hM₀⟩, hM₁, hM₂⟩ := -- zorn_subset_nonempty { s | c ⊆ s ∧ IsChain r s } _ c ⟨Subset.rfl, hc⟩ have H := zorn_subset_nonempty { s | c ⊆ s ∧ IsChain r s } ?_ c ⟨Subset.rfl, hc⟩ - · obtain ⟨M, ⟨_, hM₀⟩, hM₁, hM₂⟩ := H - exact ⟨M, ⟨hM₀, fun d hd hMd => (hM₂ _ ⟨hM₁.trans hMd, hd⟩ hMd).symm⟩, hM₁⟩ + · obtain ⟨M, hcM, hM⟩ := H + exact ⟨M, ⟨hM.prop.2, fun d hd hMd ↦ hM.eq_of_subset ⟨hcM.trans hMd, hd⟩ hMd⟩, hcM⟩ rintro cs hcs₀ hcs₁ ⟨s, hs⟩ refine ⟨⋃₀cs, ⟨fun _ ha => Set.mem_sUnion_of_mem ((hcs₀ hs).left ha) hs, ?_⟩, fun _ => diff --git a/Mathlib/Order/ZornAtoms.lean b/Mathlib/Order/ZornAtoms.lean index f175ed8d92770..99e546394afc7 100644 --- a/Mathlib/Order/ZornAtoms.lean +++ b/Mathlib/Order/ZornAtoms.lean @@ -20,18 +20,15 @@ open Set /-- **Zorn's lemma**: A partial order is coatomic if every nonempty chain `c`, `⊤ ∉ c`, has an upper bound not equal to `⊤`. -/ theorem IsCoatomic.of_isChain_bounded {α : Type*} [PartialOrder α] [OrderTop α] - (h : - ∀ c : Set α, - IsChain (· ≤ ·) c → c.Nonempty → ⊤ ∉ c → ∃ x ≠ ⊤, x ∈ upperBounds c) : + (h : ∀ c : Set α, IsChain (· ≤ ·) c → c.Nonempty → ⊤ ∉ c → ∃ x ≠ ⊤, x ∈ upperBounds c) : IsCoatomic α := by refine ⟨fun x => le_top.eq_or_lt.imp_right fun hx => ?_⟩ - have : ∃ y ∈ Ico x ⊤, x ≤ y ∧ ∀ z ∈ Ico x ⊤, y ≤ z → z = y := by - refine zorn_nonempty_partialOrder₀ (Ico x ⊤) (fun c hxc hc y hy => ?_) x (left_mem_Ico.2 hx) - rcases h c hc ⟨y, hy⟩ fun h => (hxc h).2.ne rfl with ⟨z, hz, hcz⟩ - exact ⟨z, ⟨le_trans (hxc hy).1 (hcz hy), hz.lt_top⟩, hcz⟩ - rcases this with ⟨y, ⟨hxy, hy⟩, -, hy'⟩ - refine ⟨y, ⟨hy.ne, fun z hyz => le_top.eq_or_lt.resolve_right fun hz => ?_⟩, hxy⟩ - exact hyz.ne' (hy' z ⟨hxy.trans hyz.le, hz⟩ hyz.le) + have := zorn_le_nonempty₀ (Ico x ⊤) (fun c hxc hc y hy => ?_) x (left_mem_Ico.2 hx) + · obtain ⟨y, hxy, hmax⟩ := this + refine ⟨y, ⟨hmax.prop.2.ne, fun z hyz ↦ le_top.eq_or_lt.resolve_right fun hz => ?_⟩, hxy⟩ + exact hyz.ne <| hmax.eq_of_le ⟨hxy.trans hyz.le, hz⟩ hyz.le + rcases h c hc ⟨y, hy⟩ fun h => (hxc h).2.ne rfl with ⟨z, hz, hcz⟩ + exact ⟨z, ⟨le_trans (hxc hy).1 (hcz hy), hz.lt_top⟩, hcz⟩ /-- **Zorn's lemma**: A partial order is atomic if every nonempty chain `c`, `⊥ ∉ c`, has a lower bound not equal to `⊥`. -/ diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean index 331ebad059d46..8898f192deb78 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent.lean @@ -296,24 +296,12 @@ theorem algebraicIndependent_sUnion_of_directed {s : Set (Set A)} (hsn : s.Nonem exact algebraicIndependent_iUnion_of_directed hs.directed_val (by simpa using h) theorem exists_maximal_algebraicIndependent (s t : Set A) (hst : s ⊆ t) - (hs : AlgebraicIndependent R ((↑) : s → A)) : - ∃ u : Set A, AlgebraicIndependent R ((↑) : u → A) ∧ s ⊆ u ∧ u ⊆ t ∧ - ∀ x : Set A, AlgebraicIndependent R ((↑) : x → A) → u ⊆ x → x ⊆ t → x = u := by - rcases zorn_subset_nonempty { u : Set A | AlgebraicIndependent R ((↑) : u → A) ∧ s ⊆ u ∧ u ⊆ t } - (fun c hc chainc hcn => - ⟨⋃₀ c, by - refine ⟨⟨algebraicIndependent_sUnion_of_directed hcn chainc.directedOn - fun a ha => (hc ha).1, ?_, ?_⟩, ?_⟩ - · cases' hcn with x hx - exact subset_sUnion_of_subset _ x (hc hx).2.1 hx - · exact sUnion_subset fun x hx => (hc hx).2.2 - · intro s - exact subset_sUnion_of_mem⟩) - s ⟨hs, Set.Subset.refl s, hst⟩ with - ⟨u, ⟨huai, _, hut⟩, hsu, hx⟩ - use u, huai, hsu, hut - intro x hxai huv hxt - exact hx _ ⟨hxai, _root_.trans hsu huv, hxt⟩ huv + (hs : AlgebraicIndependent R ((↑) : s → A)) : ∃ u, s ⊆ u ∧ + Maximal (fun (x : Set A) ↦ AlgebraicIndependent R ((↑) : x → A) ∧ x ⊆ t) u := by + refine zorn_subset_nonempty { u : Set A | AlgebraicIndependent R ((↑) : u → A) ∧ u ⊆ t} + (fun c hc chainc hcn ↦ ⟨⋃₀ c, ⟨?_, ?_⟩, fun _ ↦ subset_sUnion_of_mem⟩) s ⟨hs, hst⟩ + · exact algebraicIndependent_sUnion_of_directed hcn chainc.directedOn (fun x hxc ↦ (hc hxc).1) + exact fun x ⟨w, hyc, hwy⟩ ↦ (hc hyc).2 hwy section repr @@ -443,10 +431,9 @@ theorem exists_isTranscendenceBasis (h : Injective (algebraMap R A)) : cases' exists_maximal_algebraicIndependent (∅ : Set A) Set.univ (Set.subset_univ _) ((algebraicIndependent_empty_iff R A).2 h) with s hs - use s, hs.1 - intro t ht hr + refine ⟨s, hs.2.1.1, fun t ht hst ↦ ?_⟩ simp only [Subtype.range_coe_subtype, setOf_mem_eq] at * - exact Eq.symm (hs.2.2.2 t ht hr (Set.subset_univ _)) + exact hs.2.eq_of_le ⟨ht, subset_univ _⟩ hst variable {R} diff --git a/Mathlib/RingTheory/Ideal/MinimalPrime.lean b/Mathlib/RingTheory/Ideal/MinimalPrime.lean index c4c6347f666fb..b3b7e15843969 100644 --- a/Mathlib/RingTheory/Ideal/MinimalPrime.lean +++ b/Mathlib/RingTheory/Ideal/MinimalPrime.lean @@ -50,13 +50,11 @@ lemma minimalPrimes_eq_minimals : minimalPrimes R = {x | Minimal Ideal.IsPrime x variable {I J} theorem Ideal.exists_minimalPrimes_le [J.IsPrime] (e : I ≤ J) : ∃ p ∈ I.minimalPrimes, p ≤ J := by - suffices - ∃ m ∈ { p : (Ideal R)ᵒᵈ | Ideal.IsPrime p ∧ I ≤ OrderDual.ofDual p }, - OrderDual.toDual J ≤ m ∧ ∀ z ∈ { p : (Ideal R)ᵒᵈ | Ideal.IsPrime p ∧ I ≤ p }, m ≤ z → z = m by - obtain ⟨p, h₁, h₂, h₃⟩ := this - simp_rw [← @eq_comm _ p] at h₃ - exact ⟨p, ⟨h₁, fun a b c => le_of_eq (h₃ a b c)⟩, h₂⟩ - apply zorn_nonempty_partialOrder₀ + set S := { p : (Ideal R)ᵒᵈ | Ideal.IsPrime p ∧ I ≤ OrderDual.ofDual p } + suffices h : ∃ m, OrderDual.toDual J ≤ m ∧ Maximal (· ∈ S) m by + obtain ⟨p, hJp, hp⟩ := h + exact ⟨p, ⟨hp.prop, fun q hq hle ↦ hp.le_of_ge hq hle⟩, hJp⟩ + apply zorn_le_nonempty₀ swap · refine ⟨show J.IsPrime by infer_instance, e⟩ rintro (c : Set (Ideal R)) hc hc' J' hJ' diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 3b6ae09d0078d..3a108417c37a7 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -898,19 +898,19 @@ theorem IsPrime.radical_le_iff (hJ : IsPrime J) : I.radical ≤ J ↔ I ≤ J := theorem radical_eq_sInf (I : Ideal R) : radical I = sInf { J : Ideal R | I ≤ J ∧ IsPrime J } := le_antisymm (le_sInf fun J hJ ↦ hJ.2.radical_le_iff.2 hJ.1) fun r hr ↦ by_contradiction fun hri ↦ - let ⟨m, (hrm : r ∉ radical m), him, hm⟩ := - zorn_nonempty_partialOrder₀ { K : Ideal R | r ∉ radical K } + let ⟨m, hIm, hm⟩ := + zorn_le_nonempty₀ { K : Ideal R | r ∉ radical K } (fun c hc hcc y hyc => ⟨sSup c, fun ⟨n, hrnc⟩ => let ⟨y, hyc, hrny⟩ := (Submodule.mem_sSup_of_directed ⟨y, hyc⟩ hcc.directedOn).1 hrnc hc hyc ⟨n, hrny⟩, fun z => le_sSup⟩) I hri + have hrm : r ∉ radical m := hm.prop have : ∀ x ∉ m, r ∈ radical (m ⊔ span {x}) := fun x hxm => - by_contradiction fun hrmx => - hxm <| - hm (m ⊔ span {x}) hrmx le_sup_left ▸ - (le_sup_right : _ ≤ m ⊔ span {x}) (subset_span <| Set.mem_singleton _) + by_contradiction fun hrmx => hxm <| by + rw [hm.eq_of_le hrmx le_sup_left] + exact Submodule.mem_sup_right <| mem_span_singleton_self x have : IsPrime m := ⟨by rintro rfl; rw [radical_top] at hrm; exact hrm trivial, fun {x y} hxym => or_iff_not_imp_left.2 fun hxm => @@ -929,7 +929,7 @@ theorem radical_eq_sInf (I : Ideal R) : radical I = sInf { J : Ideal R | I ≤ J m.add_mem (m.mul_mem_right _ hpm) (m.add_mem (m.mul_mem_left _ hfm) (m.mul_mem_left _ hxym))⟩⟩ hrm <| - this.radical.symm ▸ (sInf_le ⟨him, this⟩ : sInf { J : Ideal R | I ≤ J ∧ IsPrime J } ≤ m) hr + this.radical.symm ▸ (sInf_le ⟨hIm, this⟩ : sInf { J : Ideal R | I ≤ J ∧ IsPrime J } ≤ m) hr theorem isRadical_bot_of_noZeroDivisors {R} [CommSemiring R] [NoZeroDivisors R] : (⊥ : Ideal R).IsRadical := fun _ hx => hx.recOn fun _ hn => pow_eq_zero hn diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index e2acb507eec86..29bea641d71f3 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -497,28 +497,29 @@ theorem IsPrincipalIdealRing.of_prime (H : ∀ P : Ideal R, P.IsPrime → P.IsPr rw [← nonPrincipals_eq_empty_iff, Set.eq_empty_iff_forall_not_mem] intro J hJ -- We will show a maximal element `I ∈ nonPrincipals R` (which exists by Zorn) is prime. - obtain ⟨I, Ibad, -, Imax⟩ := zorn_nonempty_partialOrder₀ (nonPrincipals R) nonPrincipals_zorn _ hJ + obtain ⟨I, hJI, hI⟩ := zorn_le_nonempty₀ (nonPrincipals R) nonPrincipals_zorn _ hJ + have Imax' : ∀ {J}, I < J → J.IsPrincipal := by - intro J hJ - by_contra He - exact hJ.ne (Imax _ ((nonPrincipals_def R).2 He) hJ.le).symm + intro K hK + simpa [nonPrincipals] using hI.not_prop_of_gt hK + by_cases hI1 : I = ⊤ · subst hI1 - exact Ibad top_isPrincipal + exact hI.prop top_isPrincipal -- Let `x y : R` with `x * y ∈ I` and suppose WLOG `y ∉ I`. - refine Ibad (H I ⟨hI1, fun {x y} hxy => or_iff_not_imp_right.mpr fun hy => ?_⟩) + refine hI.prop (H I ⟨hI1, fun {x y} hxy => or_iff_not_imp_right.mpr fun hy => ?_⟩) obtain ⟨a, ha⟩ : (I ⊔ span {y}).IsPrincipal := Imax' (left_lt_sup.mpr (mt I.span_singleton_le_iff_mem.mp hy)) -- Then `x ∈ I.colon (span {y})`, which is equal to `I` if it's not principal. suffices He : ¬(I.colon (span {y})).IsPrincipal by - rw [← Imax _ ((nonPrincipals_def R).2 He) fun a ha => - Ideal.mem_colon_singleton.2 (mul_mem_right _ _ ha)] + rw [hI.eq_of_le ((nonPrincipals_def R).2 He) fun a ha ↦ + Ideal.mem_colon_singleton.2 (mul_mem_right _ _ ha)] exact Ideal.mem_colon_singleton.2 hxy -- So suppose for the sake of contradiction that both `I ⊔ span {y}` and `I.colon (span {y})` -- are principal. rintro ⟨b, hb⟩ -- We will show `I` is generated by `a * b`. - refine (nonPrincipals_def _).1 Ibad ⟨a * b, ?_⟩ + refine (nonPrincipals_def _).1 hI.prop ⟨a * b, ?_⟩ refine le_antisymm (α := Ideal R) (fun i hi => ?_) <| (span_singleton_mul_span_singleton a b).ge.trans ?_ diff --git a/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean b/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean index 33ccbbb9807a1..4e8640d0749be 100644 --- a/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean +++ b/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean @@ -93,13 +93,11 @@ private abbrev sets := there is an element that injects into the others. See `Cardinal.conditionallyCompleteLinearOrderBot` for (one of) the lattice instances. -/ theorem min_injective [I : Nonempty ι] : ∃ i, Nonempty (∀ j, β i ↪ β j) := - let ⟨s, hs, ms⟩ := - show ∃ s ∈ sets β, ∀ a ∈ sets β, s ⊆ a → a = s from - zorn_subset (sets β) fun c hc hcc => - ⟨⋃₀c, fun x ⟨p, hpc, hxp⟩ y ⟨q, hqc, hyq⟩ i hi => - (hcc.total hpc hqc).elim (fun h => hc hqc x (h hxp) y hyq i hi) fun h => - hc hpc x hxp y (h hyq) i hi, - fun _ => subset_sUnion_of_mem⟩ + let ⟨s, hs⟩ := show ∃ s, Maximal (· ∈ sets β) s by + refine zorn_subset _ fun c hc hcc ↦ + ⟨⋃₀ c, fun x ⟨p, hpc, hxp⟩ y ⟨q, hqc, hyq⟩ i hi ↦ ?_, fun _ ↦ subset_sUnion_of_mem⟩ + exact (hcc.total hpc hqc).elim (fun h ↦ hc hqc x (h hxp) y hyq i hi) + fun h ↦ hc hpc x hxp y (h hyq) i hi let ⟨i, e⟩ := show ∃ i, ∀ y, ∃ x ∈ s, (x : ∀ i, β i) i = y from Classical.by_contradiction fun h => @@ -113,8 +111,8 @@ theorem min_injective [I : Nonempty ι] : ∃ i, Nonempty (∀ j, β i ↪ β j) exact fun i e => (hf i y hy e.symm).elim · subst y exact fun i e => (hf i x hx e).elim - · exact hs x hx y hy - ms _ this (subset_insert f s) ▸ mem_insert _ _ + · exact hs.prop x hx y hy + hs.eq_of_subset this (subset_insert _ _) ▸ mem_insert .. let ⟨i⟩ := I hf i f this rfl let ⟨f, hf⟩ := Classical.axiom_of_choice e @@ -123,7 +121,7 @@ theorem min_injective [I : Nonempty ι] : ∃ i, Nonempty (∀ j, β i ↪ β j) ⟨fun a => f a j, fun a b e' => by let ⟨sa, ea⟩ := hf a let ⟨sb, eb⟩ := hf b - rw [← ea, ← eb, hs _ sa _ sb _ e']⟩⟩⟩ + rw [← ea, ← eb, hs.prop _ sa _ sb _ e']⟩⟩⟩ end Wo diff --git a/Mathlib/Topology/Algebra/Semigroup.lean b/Mathlib/Topology/Algebra/Semigroup.lean index 36fa20518b569..9069ed9f766ed 100644 --- a/Mathlib/Topology/Algebra/Semigroup.lean +++ b/Mathlib/Topology/Algebra/Semigroup.lean @@ -29,12 +29,13 @@ theorem exists_idempotent_of_compact_t2_of_continuous_mul_left {M} [Nonempty M] It will turn out that any minimal element is `{m}` for an idempotent `m : M`. -/ let S : Set (Set M) := { N | IsClosed N ∧ N.Nonempty ∧ ∀ (m) (_ : m ∈ N) (m') (_ : m' ∈ N), m * m' ∈ N } - rsuffices ⟨N, ⟨N_closed, ⟨m, hm⟩, N_mul⟩, N_minimal⟩ : ∃ N ∈ S, ∀ N' ∈ S, N' ⊆ N → N' = N - · use m + rsuffices ⟨N, hN⟩ : ∃ N', Minimal (· ∈ S) N' + · obtain ⟨N_closed, ⟨m, hm⟩, N_mul⟩ := hN.prop + use m /- We now have an element `m : M` of a minimal subsemigroup `N`, and want to show `m + m = m`. We first show that every element of `N` is of the form `m' + m`. -/ have scaling_eq_self : (· * m) '' N = N := by - apply N_minimal + apply hN.eq_of_subset · refine ⟨(continuous_mul_left m).isClosedMap _ N_closed, ⟨_, ⟨m, hm, rfl⟩⟩, ?_⟩ rintro _ ⟨m'', hm'', rfl⟩ _ ⟨m', hm', rfl⟩ exact ⟨m'' * m * m', N_mul _ (N_mul _ hm'' _ hm) _ hm', mul_assoc _ _ _⟩ @@ -43,14 +44,13 @@ theorem exists_idempotent_of_compact_t2_of_continuous_mul_left {M} [Nonempty M] /- In particular, this means that `m' * m = m` for some `m'`. We now use minimality again to show that this holds for all `m' ∈ N`. -/ have absorbing_eq_self : N ∩ { m' | m' * m = m } = N := by - apply N_minimal + apply hN.eq_of_subset · refine ⟨N_closed.inter ((T1Space.t1 m).preimage (continuous_mul_left m)), ?_, ?_⟩ · rwa [← scaling_eq_self] at hm · rintro m'' ⟨mem'', eq'' : _ = m⟩ m' ⟨mem', eq' : _ = m⟩ refine ⟨N_mul _ mem'' _ mem', ?_⟩ rw [Set.mem_setOf_eq, mul_assoc, eq', eq''] apply Set.inter_subset_left - -- Thus `m * m = m` as desired. rw [← absorbing_eq_self] at hm exact hm.2 refine zorn_superset _ fun c hcs hc => ?_ diff --git a/Mathlib/Topology/Compactness/Compact.lean b/Mathlib/Topology/Compactness/Compact.lean index e8f56fd2a5354..1a468a8d2b236 100644 --- a/Mathlib/Topology/Compactness/Compact.lean +++ b/Mathlib/Topology/Compactness/Compact.lean @@ -1090,7 +1090,7 @@ theorem IsClosed.exists_minimal_nonempty_closed_subset [CompactSpace X] {S : Set ∃ V : Set X, V ⊆ S ∧ V.Nonempty ∧ IsClosed V ∧ ∀ V' : Set X, V' ⊆ V → V'.Nonempty → IsClosed V' → V' = V := by let opens := { U : Set X | Sᶜ ⊆ U ∧ IsOpen U ∧ Uᶜ.Nonempty } - obtain ⟨U, ⟨Uc, Uo, Ucne⟩, h⟩ := + obtain ⟨U, h⟩ := zorn_subset opens fun c hc hz => by by_cases hcne : c.Nonempty · obtain ⟨U₀, hU₀⟩ := hcne @@ -1115,10 +1115,11 @@ theorem IsClosed.exists_minimal_nonempty_closed_subset [CompactSpace X] {S : Set refine ⟨⟨Set.Subset.refl _, isOpen_compl_iff.mpr hS, ?_⟩, fun U Uc => (hcne ⟨U, Uc⟩).elim⟩ rw [compl_compl] exact hne + obtain ⟨Uc, Uo, Ucne⟩ := h.prop refine ⟨Uᶜ, Set.compl_subset_comm.mp Uc, Ucne, Uo.isClosed_compl, ?_⟩ intro V' V'sub V'ne V'cls have : V'ᶜ = U := by - refine h V'ᶜ ⟨?_, isOpen_compl_iff.mpr V'cls, ?_⟩ (Set.subset_compl_comm.mp V'sub) + refine h.eq_of_ge ⟨?_, isOpen_compl_iff.mpr V'cls, ?_⟩ (subset_compl_comm.2 V'sub) · exact Set.Subset.trans Uc (Set.subset_compl_comm.mp V'sub) · simp only [compl_compl, V'ne] rw [← this, compl_compl] diff --git a/Mathlib/Topology/ExtremallyDisconnected.lean b/Mathlib/Topology/ExtremallyDisconnected.lean index bdba4f48c8dba..a5d3b08560a46 100644 --- a/Mathlib/Topology/ExtremallyDisconnected.lean +++ b/Mathlib/Topology/ExtremallyDisconnected.lean @@ -143,12 +143,13 @@ lemma exists_compact_surjective_zorn_subset [T1Space A] [CompactSpace D] {π : D -- suffices to apply Zorn's lemma on the subsets of $D$ that are closed and mapped onto $A$ let S : Set <| Set D := {E : Set D | IsClosed E ∧ π '' E = univ} suffices ∀ (C : Set <| Set D) (_ : C ⊆ S) (_ : IsChain (· ⊆ ·) C), ∃ s ∈ S, ∀ c ∈ C, s ⊆ c by - rcases zorn_superset S this with ⟨E, ⟨E_closed, E_surj⟩, E_min⟩ + rcases zorn_superset S this with ⟨E, E_min⟩ + obtain ⟨E_closed, E_surj⟩ := E_min.prop refine ⟨E, isCompact_iff_compactSpace.mp E_closed.isCompact, E_surj, ?_⟩ intro E₀ E₀_min E₀_closed contrapose! E₀_min exact eq_univ_of_image_val_eq <| - E_min E₀ ⟨E₀_closed.trans E_closed, image_image_val_eq_restrict_image ▸ E₀_min⟩ + E_min.eq_of_subset ⟨E₀_closed.trans E_closed, image_image_val_eq_restrict_image ▸ E₀_min⟩ image_val_subset -- suffices to prove intersection of chain is minimal intro C C_sub C_chain diff --git a/Mathlib/Topology/Irreducible.lean b/Mathlib/Topology/Irreducible.lean index cf88aaae25d57..8a868fe36faa7 100644 --- a/Mathlib/Topology/Irreducible.lean +++ b/Mathlib/Topology/Irreducible.lean @@ -74,7 +74,7 @@ protected alias ⟨_, IsIrreducible.closure⟩ := isIrreducible_iff_closure theorem exists_preirreducible (s : Set X) (H : IsPreirreducible s) : ∃ t : Set X, IsPreirreducible t ∧ s ⊆ t ∧ ∀ u, IsPreirreducible u → t ⊆ u → u = t := - let ⟨m, hm, hsm, hmm⟩ := + let ⟨m, hsm, hm⟩ := zorn_subset_nonempty { t : Set X | IsPreirreducible t } (fun c hc hcc _ => ⟨⋃₀ c, fun u v hu hv ⟨y, hy, hyu⟩ ⟨x, hx, hxv⟩ => @@ -89,7 +89,7 @@ theorem exists_preirreducible (s : Set X) (H : IsPreirreducible s) : ⟨x, mem_sUnion_of_mem hxp hpc, hxuv⟩, fun _ hxc => subset_sUnion_of_mem hxc⟩) s H - ⟨m, hm, hsm, fun _u hu hmu => hmm _ hu hmu⟩ + ⟨m, hm.prop, hsm, fun _u hu hmu => (hm.eq_of_subset hu hmu).symm⟩ /-- The set of irreducible components of a topological space. -/ def irreducibleComponents (X : Type*) [TopologicalSpace X] : Set (Set X) := diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index 144ab42061543..49f989797cbfc 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -206,13 +206,12 @@ theorem exists_subset_iUnion_closure_subset (hs : IsClosed s) (uo : ∀ i, IsOpe have : ∀ c : Set (PartialRefinement u s), IsChain (· ≤ ·) c → c.Nonempty → ∃ ub, ∀ v ∈ c, v ≤ ub := fun c hc ne => ⟨.chainSup c hc ne uf us, fun v hv => PartialRefinement.le_chainSup _ _ _ _ hv⟩ - rcases zorn_nonempty_partialOrder this with ⟨v, hv⟩ + rcases zorn_le_nonempty this with ⟨v, hv⟩ suffices ∀ i, i ∈ v.carrier from ⟨v, v.subset_iUnion, fun i => v.isOpen _, fun i => v.closure_subset (this i)⟩ - contrapose! hv - rcases hv with ⟨i, hi⟩ + refine fun i ↦ by_contra fun hi ↦ ?_ rcases v.exists_gt hs i hi with ⟨v', hlt⟩ - exact ⟨v', hlt.le, hlt.ne'⟩ + exact hv.not_lt hlt /-- **Shrinking lemma**. A point-finite open cover of a closed subset of a normal space can be "shrunk" to a new closed cover so that each new closed set is contained in the corresponding From ac927311b415408f4b81e13a67720f5a4456d77d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 8 Aug 2024 17:32:51 +0000 Subject: [PATCH 118/359] feat(MeasureTheory/../DomMulAct): action of `DomMulAct` on `Lp` is continuous (#15206) --- Mathlib.lean | 1 + .../Constructions/BorelSpace/Basic.lean | 1 + .../Function/LpSpace/DomAct/Continuous.lean | 56 +++++++++++++++++++ 3 files changed, 58 insertions(+) create mode 100644 Mathlib/MeasureTheory/Function/LpSpace/DomAct/Continuous.lean diff --git a/Mathlib.lean b/Mathlib.lean index c4fa6e74e8a44..4d95f16d8fc02 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3136,6 +3136,7 @@ import Mathlib.MeasureTheory.Function.LpSeminorm.Trim import Mathlib.MeasureTheory.Function.LpSpace import Mathlib.MeasureTheory.Function.LpSpace.ContinuousCompMeasurePreserving import Mathlib.MeasureTheory.Function.LpSpace.DomAct.Basic +import Mathlib.MeasureTheory.Function.LpSpace.DomAct.Continuous import Mathlib.MeasureTheory.Function.SimpleFunc import Mathlib.MeasureTheory.Function.SimpleFuncDense import Mathlib.MeasureTheory.Function.SimpleFuncDenseLp diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 06adebcec0517..6aecc1704584b 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -504,6 +504,7 @@ instance (priority := 100) TopologicalGroup.measurableInv [Group γ] [Topologica MeasurableInv γ := ⟨continuous_inv.measurable⟩ +@[to_additive] instance (priority := 100) ContinuousSMul.measurableSMul {M α} [TopologicalSpace M] [TopologicalSpace α] [MeasurableSpace M] [MeasurableSpace α] [OpensMeasurableSpace M] [BorelSpace α] [SMul M α] [ContinuousSMul M α] : MeasurableSMul M α := diff --git a/Mathlib/MeasureTheory/Function/LpSpace/DomAct/Continuous.lean b/Mathlib/MeasureTheory/Function/LpSpace/DomAct/Continuous.lean new file mode 100644 index 0000000000000..3a9b91acac9f1 --- /dev/null +++ b/Mathlib/MeasureTheory/Function/LpSpace/DomAct/Continuous.lean @@ -0,0 +1,56 @@ +/- +Copyright (c) 2024 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.MeasureTheory.Function.LpSpace.DomAct.Basic +import Mathlib.MeasureTheory.Function.LpSpace.ContinuousCompMeasurePreserving +import Mathlib.Topology.Algebra.Constructions.DomMulAct + +/-! +# Continuity of the action of `Mᵈᵐᵃ` on `MeasureSpace.Lp E p μ` + +In this file we prove that under certain conditions, +the action of `Mᵈᵐᵃ` on `MeasureTheory.Lp E p μ` is continuous in both variables. + +Recall that `Mᵈᵐᵃ` acts on `MeasureTheory.Lp E p μ` +by `mk c • f = MeasureTheory.Lp.compMeasurePreserving (c • ·) _ f`. +This action is defined, if `M` acts on `X` by mesaure preserving maps. + +If `M` acts on `X` by continuous maps +preserving a locally finite measure +which is inner regular for finite measure sets with respect to compact sets, +then the action of `Mᵈᵐᵃ` on `Lp E p μ` described above, `1 ≤ p < ∞`, +is continuous in both arguments. + +In particular, it applies to the case when `X = M` is a locally compact topological group, +and `μ` is the Haar measure. + +## Tags + +measure theory, group action, domain action, continuous action, Lp space +-/ + +open scoped ENNReal +open DomMulAct + +namespace MeasureTheory + +variable {X M E : Type*} + [TopologicalSpace X] [R1Space X] [MeasurableSpace X] [BorelSpace X] + [Monoid M] [TopologicalSpace M] [MeasurableSpace M] [OpensMeasurableSpace M] + [SMul M X] [ContinuousSMul M X] + [NormedAddCommGroup E] + {μ : Measure X} [IsLocallyFiniteMeasure μ] [μ.InnerRegularCompactLTTop] + [SMulInvariantMeasure M X μ] + {p : ℝ≥0∞} [Fact (1 ≤ p)] [hp : Fact (p ≠ ∞)] + +@[to_additive] +instance Lp.instContinuousSMulDomMulAct : ContinuousSMul Mᵈᵐᵃ (Lp E p μ) where + continuous_smul := + let g : C(Mᵈᵐᵃ × Lp E p μ, C(X, X)) := + (ContinuousMap.mk (fun a : M × X ↦ a.1 • a.2) continuous_smul).curry.comp <| + .comp (.mk DomMulAct.mk.symm) ContinuousMap.fst + continuous_snd.compMeasurePreservingLp g.continuous _ Fact.out + +end MeasureTheory From 71182adddbfdcb452d72959c4ddca10c54591de7 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 8 Aug 2024 17:32:52 +0000 Subject: [PATCH 119/359] chore: backports for leanprover/lean4#4814 (part 37) (#15608) --- .../Instances.lean | 20 ++++---- .../Restrict.lean | 42 ++++++++-------- .../LineDeriv/IntegrationByParts.lean | 4 +- .../Analysis/Complex/LocallyUniformLimit.lean | 6 ++- .../ContinuousFunctionalCalculus/ExpLog.lean | 48 +++++++++---------- .../MeasureTheory/Integral/PeakFunction.lean | 4 +- .../NumberTheory/LSeries/AbstractFuncEq.lean | 7 +-- 7 files changed, 72 insertions(+), 59 deletions(-) diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index b7a308f06b1b0..cd5df457cc5a8 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -42,8 +42,8 @@ local notation "σ" => spectrum section RCLike -variable {𝕜 A : Type*} [RCLike 𝕜] [NonUnitalNormedRing A] [StarRing A] [CStarRing A] -variable [CompleteSpace A] [NormedSpace 𝕜 A] [IsScalarTower 𝕜 A A] [SMulCommClass 𝕜 A A] +variable {𝕜 A : Type*} [RCLike 𝕜] [NonUnitalNormedRing A] [StarRing A] +variable [NormedSpace 𝕜 A] [IsScalarTower 𝕜 A A] [SMulCommClass 𝕜 A A] variable [StarModule 𝕜 A] {p : A → Prop} {p₁ : Unitization 𝕜 A → Prop} local postfix:max "⁺¹" => Unitization 𝕜 @@ -99,6 +99,8 @@ lemma spec_cfcₙAux (f : C(σₙ 𝕜 a, 𝕜)₀) : σ 𝕜 (cfcₙAux hp₁ a · exact ⟨⟨x, (Unitization.quasispectrum_eq_spectrum_inr' 𝕜 𝕜 a).symm ▸ x.property⟩, rfl⟩ · exact ⟨⟨x, Unitization.quasispectrum_eq_spectrum_inr' 𝕜 𝕜 a ▸ x.property⟩, rfl⟩ +variable [CompleteSpace A] + lemma cfcₙAux_mem_range_inr (f : C(σₙ 𝕜 a, 𝕜)₀) : cfcₙAux hp₁ a ha f ∈ NonUnitalStarAlgHom.range (Unitization.inrNonUnitalStarAlgHom 𝕜 A) := by have h₁ := (closedEmbedding_cfcₙAux hp₁ a ha).continuous.range_subset_closure_image_dense @@ -116,6 +118,8 @@ lemma cfcₙAux_mem_range_inr (f : C(σₙ 𝕜 a, 𝕜)₀) : convert IsClosed.preimage this (isClosed_singleton (x := 0)) aesop +variable [CStarRing A] + open Unitization NonUnitalStarAlgHom in theorem RCLike.nonUnitalContinuousFunctionalCalculus : NonUnitalContinuousFunctionalCalculus 𝕜 (p : A → Prop) where @@ -183,7 +187,7 @@ end Normal section SelfAdjointNonUnital variable {A : Type*} [TopologicalSpace A] [NonUnitalRing A] [StarRing A] [Module ℂ A] - [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] [StarModule ℂ A] + [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] [NonUnitalContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop)] /-- An element in a non-unital C⋆-algebra is selfadjoint if and only if it is normal and its @@ -224,7 +228,7 @@ section SelfAdjointUnital variable {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] [Algebra ℂ A] - [StarModule ℂ A] [ContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop)] + [ContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop)] /- TODO: REMOVE @@ -293,7 +297,7 @@ lemma CFC.exists_sqrt_of_isSelfAdjoint_of_quasispectrumRestricts {A : Type*} [No variable {A : Type*} [NonUnitalRing A] [PartialOrder A] [StarRing A] [StarOrderedRing A] variable [TopologicalSpace A] [Module ℝ A] [IsScalarTower ℝ A A] [SMulCommClass ℝ A A] variable [NonUnitalContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] -variable [NonnegSpectrumClass ℝ A] [UniqueNonUnitalContinuousFunctionalCalculus ℝ A] +variable [NonnegSpectrumClass ℝ A] lemma nonneg_iff_isSelfAdjoint_and_quasispectrumRestricts {a : A} : 0 ≤ a ↔ IsSelfAdjoint a ∧ QuasispectrumRestricts a ContinuousMap.realToNNReal := by @@ -333,7 +337,7 @@ lemma CFC.exists_sqrt_of_isSelfAdjoint_of_spectrumRestricts {A : Type*} [Ring A] variable {A : Type*} [Ring A] [PartialOrder A] [StarRing A] [StarOrderedRing A] [TopologicalSpace A] variable [Algebra ℝ A] [ContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] -variable [NonnegSpectrumClass ℝ A] [UniqueContinuousFunctionalCalculus ℝ A] +variable [NonnegSpectrumClass ℝ A] -- TODO: REMOVE (duplicate; see comment on `isSelfAdjoint_iff_isStarNormal_and_spectrumRestricts`) lemma nonneg_iff_isSelfAdjoint_and_spectrumRestricts {a : A} : @@ -589,7 +593,7 @@ end NonnegSpectrumClass section RealEqComplex variable {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] [Algebra ℂ A] - [StarModule ℂ A] [ContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop)] + [ContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop)] [ContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] [UniqueContinuousFunctionalCalculus ℝ A] [UniqueContinuousFunctionalCalculus ℂ A] @@ -614,7 +618,7 @@ section NNRealEqReal open NNReal variable {A : Type*} [TopologicalSpace A] [Ring A] [PartialOrder A] [StarRing A] - [StarOrderedRing A] [Algebra ℝ A] [TopologicalRing A] [ContinuousStar A] [StarModule ℝ A] + [StarOrderedRing A] [Algebra ℝ A] [TopologicalRing A] [ContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] [ContinuousFunctionalCalculus ℝ≥0 ((0 : A) ≤ ·)] [UniqueContinuousFunctionalCalculus ℝ A] [NonnegSpectrumClass ℝ A] diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean index 6bd6ffb31d9d7..b3720734734ab 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Restrict.lean @@ -68,8 +68,18 @@ def starAlgHom {R : Type u} {S : Type v} {A : Type w} [Semifield R] variable {R S A : Type*} {p q : A → Prop} variable [Semifield R] [StarRing R] [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] variable [Semifield S] [StarRing S] [MetricSpace S] [TopologicalSemiring S] [ContinuousStar S] -variable [TopologicalSpace A] [Ring A] [StarRing A] [Algebra S A] [ContinuousFunctionalCalculus S q] +variable [Ring A] [StarRing A] [Algebra S A] variable [Algebra R S] [Algebra R A] [IsScalarTower R S A] [StarModule R S] [ContinuousSMul R S] + +lemma starAlgHom_id {a : A} {φ : C(spectrum S a, S) →⋆ₐ[S] A} {f : C(S, R)} + (h : SpectrumRestricts a f) (h_id : φ (.restrict (spectrum S a) <| .id S) = a) : + h.starAlgHom φ (.restrict (spectrum R a) <| .id R) = a := by + simp only [SpectrumRestricts.starAlgHom_apply] + convert h_id + ext x + exact h.rightInvOn x.2 + +variable [TopologicalSpace A] [ContinuousFunctionalCalculus S q] variable [CompleteSpace R] lemma closedEmbedding_starAlgHom {a : A} {φ : C(spectrum S a, S) →⋆ₐ[S] A} @@ -81,14 +91,6 @@ lemma closedEmbedding_starAlgHom {a : A} {φ : C(spectrum S a, S) →⋆ₐ[S] A (ContinuousMap.uniformEmbedding_comp _ halg) (UniformEquiv.arrowCongr h.homeomorph.symm (.refl _) |>.uniformEmbedding) -lemma starAlgHom_id {a : A} {φ : C(spectrum S a, S) →⋆ₐ[S] A} {f : C(S, R)} - (h : SpectrumRestricts a f) (h_id : φ (.restrict (spectrum S a) <| .id S) = a) : - h.starAlgHom φ (.restrict (spectrum R a) <| .id R) = a := by - simp only [SpectrumRestricts.starAlgHom_apply] - convert h_id - ext x - exact h.rightInvOn x.2 - /-- Given a `ContinuousFunctionalCalculus S q`. If we form the predicate `p` for `a : A` characterized by: `q a` and the spectrum of `a` restricts to the scalar subring `R` via `f : C(S, R)`, then we can get a restricted functional calculus @@ -201,9 +203,19 @@ def nonUnitalStarAlgHom {R : Type u} {S : Type v} {A : Type w} [Semifield R] variable {R S A : Type*} {p q : A → Prop} variable [Semifield R] [StarRing R] [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] variable [Field S] [StarRing S] [MetricSpace S] [TopologicalRing S] [ContinuousStar S] -variable [TopologicalSpace A] [NonUnitalRing A] [StarRing A] [Module S A] [IsScalarTower S A A] -variable [SMulCommClass S A A] [NonUnitalContinuousFunctionalCalculus S q] +variable [NonUnitalRing A] [StarRing A] [Module S A] [IsScalarTower S A A] +variable [SMulCommClass S A A] variable [Algebra R S] [Module R A] [IsScalarTower R S A] [StarModule R S] [ContinuousSMul R S] + +lemma nonUnitalStarAlgHom_id {a : A} {φ : C(σₙ S a, S)₀ →⋆ₙₐ[S] A} {f : C(S, R)} + (h : QuasispectrumRestricts a f) (h_id : φ (.id rfl) = a) : + h.nonUnitalStarAlgHom φ (.id rfl) = a := by + simp only [QuasispectrumRestricts.nonUnitalStarAlgHom_apply] + convert h_id + ext x + exact h.rightInvOn x.2 + +variable [TopologicalSpace A] [NonUnitalContinuousFunctionalCalculus S q] variable [CompleteSpace R] lemma closedEmbedding_nonUnitalStarAlgHom {a : A} {φ : C(σₙ S a, S)₀ →⋆ₙₐ[S] A} @@ -216,14 +228,6 @@ lemma closedEmbedding_nonUnitalStarAlgHom {a : A} {φ : C(σₙ S a, S)₀ → (ContinuousMapZero.uniformEmbedding_comp _ halg) (UniformEquiv.arrowCongrLeft₀ h.homeomorph.symm this |>.uniformEmbedding) -lemma nonUnitalStarAlgHom_id {a : A} {φ : C(σₙ S a, S)₀ →⋆ₙₐ[S] A} {f : C(S, R)} - (h : QuasispectrumRestricts a f) (h_id : φ (.id rfl) = a) : - h.nonUnitalStarAlgHom φ (.id rfl) = a := by - simp only [QuasispectrumRestricts.nonUnitalStarAlgHom_apply] - convert h_id - ext x - exact h.rightInvOn x.2 - variable [IsScalarTower R A A] [SMulCommClass R A A] /-- Given a `NonUnitalContinuousFunctionalCalculus S q`. If we form the predicate `p` for `a : A` diff --git a/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean b/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean index ba10c92e4cc25..e46ccfaeb5601 100644 --- a/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean +++ b/Mathlib/Analysis/Calculus/LineDeriv/IntegrationByParts.lean @@ -48,7 +48,7 @@ open MeasureTheory Measure FiniteDimensional variable {E F G W : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [NormedAddCommGroup F] [NormedSpace ℝ F] [NormedAddCommGroup G] [NormedSpace ℝ G] [NormedAddCommGroup W] - [NormedSpace ℝ W] [MeasurableSpace E] [BorelSpace E] {μ : Measure E} + [NormedSpace ℝ W] [MeasurableSpace E] {μ : Measure E} lemma integral_bilinear_hasLineDerivAt_right_eq_neg_left_of_integrable_aux1 [SigmaFinite μ] {f f' : E × ℝ → F} {g g' : E × ℝ → G} {B : F →L[ℝ] G →L[ℝ] W} @@ -72,6 +72,8 @@ lemma integral_bilinear_hasLineDerivAt_right_eq_neg_left_of_integrable_aux1 [Sig <;> simp _ = - ∫ x, B (f' x) (g x) ∂(μ.prod volume) := by rw [integral_neg, integral_prod _ hf'g] +variable [BorelSpace E] + lemma integral_bilinear_hasLineDerivAt_right_eq_neg_left_of_integrable_aux2 [FiniteDimensional ℝ E] {μ : Measure (E × ℝ)} [IsAddHaarMeasure μ] {f f' : E × ℝ → F} {g g' : E × ℝ → G} {B : F →L[ℝ] G →L[ℝ] W} diff --git a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean index 8151998153a2a..09792d0f7a4bd 100644 --- a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean +++ b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean @@ -26,7 +26,7 @@ open Set Metric MeasureTheory Filter Complex intervalIntegral open scoped Real Topology -variable {E ι : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] [CompleteSpace E] {U K : Set ℂ} +variable {E ι : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] {U K : Set ℂ} {z : ℂ} {M r δ : ℝ} {φ : Filter ι} {F : ι → ℂ → E} {f g : ℂ → E} namespace Complex @@ -39,7 +39,7 @@ holomorphic, because it depends continuously on `f` for the uniform topology. -/ noncomputable def cderiv (r : ℝ) (f : ℂ → E) (z : ℂ) : E := (2 * π * I : ℂ)⁻¹ • ∮ w in C(z, r), ((w - z) ^ 2)⁻¹ • f w -theorem cderiv_eq_deriv (hU : IsOpen U) (hf : DifferentiableOn ℂ f U) (hr : 0 < r) +theorem cderiv_eq_deriv [CompleteSpace E] (hU : IsOpen U) (hf : DifferentiableOn ℂ f U) (hr : 0 < r) (hzr : closedBall z r ⊆ U) : cderiv r f z = deriv f z := two_pi_I_inv_smul_circleIntegral_sub_sq_inv_smul_of_differentiable hU hzr hf (mem_ball_self hr) @@ -103,6 +103,8 @@ theorem _root_.TendstoUniformlyOn.cderiv (hF : TendstoUniformlyOn F f φ (cthick end Cderiv +variable [CompleteSpace E] + section Weierstrass theorem tendstoUniformlyOn_deriv_of_cthickening_subset (hf : TendstoLocallyUniformlyOn F f φ U) diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean index e781a032ee7fc..4879f5d427b76 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/ExpLog.lean @@ -56,10 +56,9 @@ end general_exponential namespace CFC section RCLikeNormed -variable {𝕜 : Type*} {A : Type*} [RCLike 𝕜] {p : A → Prop} [PartialOrder A] [NormedRing A] - [StarRing A] [StarOrderedRing A] [TopologicalRing A] [NormedAlgebra 𝕜 A] [CompleteSpace A] +variable {𝕜 : Type*} {A : Type*} [RCLike 𝕜] {p : A → Prop} [NormedRing A] + [StarRing A] [TopologicalRing A] [NormedAlgebra 𝕜 A] [CompleteSpace A] [ContinuousFunctionalCalculus 𝕜 p] - [UniqueContinuousFunctionalCalculus 𝕜 A] lemma exp_eq_normedSpace_exp {a : A} (ha : p a := by cfc_tac) : cfc (exp 𝕜 : 𝕜 → 𝕜) a = exp 𝕜 a := by @@ -75,10 +74,9 @@ end RCLikeNormed section RealNormed -variable {A : Type*} {p : A → Prop} [PartialOrder A] [NormedRing A] [StarRing A] [StarOrderedRing A] +variable {A : Type*} {p : A → Prop} [NormedRing A] [StarRing A] [TopologicalRing A] [NormedAlgebra ℝ A] [CompleteSpace A] [ContinuousFunctionalCalculus ℝ p] - [UniqueContinuousFunctionalCalculus ℝ A] lemma real_exp_eq_normedSpace_exp {a : A} (ha : p a := by cfc_tac) : cfc Real.exp a = exp ℝ a := @@ -88,10 +86,9 @@ end RealNormed section ComplexNormed -variable {A : Type*} {p : A → Prop} [PartialOrder A] [NormedRing A] [StarRing A] [StarOrderedRing A] +variable {A : Type*} {p : A → Prop} [NormedRing A] [StarRing A] [TopologicalRing A] [NormedAlgebra ℂ A] [CompleteSpace A] [ContinuousFunctionalCalculus ℂ p] - [UniqueContinuousFunctionalCalculus ℂ A] lemma complex_exp_eq_normedSpace_exp {a : A} (ha : p a := by cfc_tac) : cfc Complex.exp a = exp ℂ a := @@ -104,10 +101,8 @@ section real_log open scoped ComplexOrder -variable {A : Type*} [PartialOrder A] [NormedRing A] [StarRing A] [StarOrderedRing A] - [TopologicalRing A] [NormedAlgebra ℝ A] [CompleteSpace A] +variable {A : Type*} [NormedRing A] [StarRing A] [NormedAlgebra ℝ A] [ContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] - [UniqueContinuousFunctionalCalculus ℝ A] /-- The real logarithm, defined via the continuous functional calculus. This can be used on matrices, operators on a Hilbert space, elements of a C⋆-algebra, etc. -/ @@ -116,20 +111,6 @@ noncomputable def log (a : A) : A := cfc Real.log a @[simp] protected lemma _root_.IsSelfAdjoint.log {a : A} : IsSelfAdjoint (log a) := cfc_predicate _ a -lemma log_exp (a : A) (ha : IsSelfAdjoint a := by cfc_tac) : log (NormedSpace.exp ℝ a) = a := by - have hcont : ContinuousOn Real.log (Real.exp '' spectrum ℝ a) := by fun_prop (disch := aesop) - rw [log, ← real_exp_eq_normedSpace_exp, ← cfc_comp' Real.log Real.exp a hcont] - simp [cfc_id' (R := ℝ) a] - --- TODO: Relate the hypothesis to a notion of strict positivity -lemma exp_log (a : A) (ha₂ : ∀ x ∈ spectrum ℝ a, 0 < x) (ha₁ : IsSelfAdjoint a := by cfc_tac) : - NormedSpace.exp ℝ (log a) = a := by - have ha₃ : ContinuousOn Real.log (spectrum ℝ a) := by - have : ∀ x ∈ spectrum ℝ a, x ≠ 0 := by peel ha₂ with x hx h; exact h.ne' - fun_prop (disch := assumption) - rw [← real_exp_eq_normedSpace_exp .log, log, ← cfc_comp' Real.exp Real.log a (by fun_prop) ha₃] - conv_rhs => rw [← cfc_id (R := ℝ) a ha₁] - exact cfc_congr (Real.exp_log <| ha₂ · ·) @[simp] lemma log_zero : log (0 : A) = 0 := by simp [log] @@ -139,6 +120,8 @@ lemma exp_log (a : A) (ha₂ : ∀ x ∈ spectrum ℝ a, 0 < x) (ha₁ : IsSelfA lemma log_algebraMap {r : ℝ} : log (algebraMap ℝ A r) = algebraMap ℝ A (Real.log r) := by simp [log] +variable [UniqueContinuousFunctionalCalculus ℝ A] + -- TODO: Relate the hypothesis to a notion of strict positivity lemma log_smul {r : ℝ} (a : A) (ha₂ : ∀ x ∈ spectrum ℝ a, 0 < x) (hr : 0 < r) (ha₁ : IsSelfAdjoint a := by cfc_tac) : @@ -159,5 +142,22 @@ lemma log_pow (n : ℕ) (a : A) (ha₂ : ∀ x ∈ spectrum ℝ a, 0 < x) rw [log, ← cfc_pow_id (R := ℝ) a n ha₁, ← cfc_comp' Real.log (· ^ n) a ha₂'', log] simp_rw [Real.log_pow, ← Nat.cast_smul_eq_nsmul ℝ n, cfc_const_mul (n : ℝ) Real.log a ha₂'] +variable [CompleteSpace A] + +lemma log_exp (a : A) (ha : IsSelfAdjoint a := by cfc_tac) : log (NormedSpace.exp ℝ a) = a := by + have hcont : ContinuousOn Real.log (Real.exp '' spectrum ℝ a) := by fun_prop (disch := aesop) + rw [log, ← real_exp_eq_normedSpace_exp, ← cfc_comp' Real.log Real.exp a hcont] + simp [cfc_id' (R := ℝ) a] + +-- TODO: Relate the hypothesis to a notion of strict positivity +lemma exp_log (a : A) (ha₂ : ∀ x ∈ spectrum ℝ a, 0 < x) (ha₁ : IsSelfAdjoint a := by cfc_tac) : + NormedSpace.exp ℝ (log a) = a := by + have ha₃ : ContinuousOn Real.log (spectrum ℝ a) := by + have : ∀ x ∈ spectrum ℝ a, x ≠ 0 := by peel ha₂ with x hx h; exact h.ne' + fun_prop (disch := assumption) + rw [← real_exp_eq_normedSpace_exp .log, log, ← cfc_comp' Real.exp Real.log a (by fun_prop) ha₃] + conv_rhs => rw [← cfc_id (R := ℝ) a ha₁] + exact cfc_congr (Real.exp_log <| ha₂ · ·) + end real_log end CFC diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index 3f601b21d0e48..cf188e01e74b8 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -87,8 +87,6 @@ theorem integrableOn_peak_smul_of_integrableOn_of_tendsto alias integrableOn_peak_smul_of_integrableOn_of_continuousWithinAt := integrableOn_peak_smul_of_integrableOn_of_tendsto -variable [CompleteSpace E] - /-- If a sequence of peak functions `φᵢ` converges uniformly to zero away from a point `x₀` and its integral on some finite-measure neighborhood of `x₀` converges to `1`, and `g` is integrable and has a limit `a` at `x₀`, then `∫ φᵢ • g` converges to `a`. @@ -182,6 +180,8 @@ theorem tendsto_setIntegral_peak_smul_of_integrableOn_of_tendsto_aux alias tendsto_setIntegral_peak_smul_of_integrableOn_of_continuousWithinAt_aux := tendsto_setIntegral_peak_smul_of_integrableOn_of_tendsto_aux +variable [CompleteSpace E] + /-- If a sequence of peak functions `φᵢ` converges uniformly to zero away from a point `x₀` and its integral on some finite-measure neighborhood of `x₀` converges to `1`, and `g` is integrable and has a limit `a` at `x₀`, then `∫ φᵢ • g` converges to `a`. Version localized to a subset. -/ diff --git a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean index 8dc4f268a2571..b7ad4b399bedb 100644 --- a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean +++ b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean @@ -67,7 +67,7 @@ noncomputable section open Real Complex Filter Topology Asymptotics Set MeasureTheory -variable (E : Type*) [NormedAddCommGroup E] [NormedSpace ℂ E] [CompleteSpace E] +variable (E : Type*) [NormedAddCommGroup E] [NormedSpace ℂ E] /-! ## Definitions and symmetry @@ -335,7 +335,7 @@ lemma f_modif_aux1 : EqOn (fun x ↦ P.f_modif x - P.f x + P.f₀) /-- Compute the Mellin transform of the modifying term used to kill off the constants at `0` and `∞`. -/ -lemma f_modif_aux2 {s : ℂ} (hs : P.k < re s) : +lemma f_modif_aux2 [CompleteSpace E] {s : ℂ} (hs : P.k < re s) : mellin (fun x ↦ P.f_modif x - P.f x + P.f₀) s = (1 / s) • P.f₀ + (P.ε / (P.k - s)) • P.g₀ := by have h_re1 : -1 < re (s - 1) := by simpa using P.hk.trans hs have h_re2 : -1 < re (s - P.k - 1) := by simpa using hs @@ -406,7 +406,8 @@ theorem differentiableAt_Λ {s : ℂ} (hs : s ≠ 0 ∨ P.f₀ = 0) (hs' : s ≠ · simpa only [hs', smul_zero] using differentiableAt_const (0 : E) /-- Relation between `Λ s` and the Mellin transform of `f - f₀`, where the latter is defined. -/ -theorem hasMellin {s : ℂ} (hs : P.k < s.re) : HasMellin (P.f · - P.f₀) s (P.Λ s) := by +theorem hasMellin [CompleteSpace E] + {s : ℂ} (hs : P.k < s.re) : HasMellin (P.f · - P.f₀) s (P.Λ s) := by have hc1 : MellinConvergent (P.f · - P.f₀) s := let ⟨_, ht⟩ := exists_gt s.re mellinConvergent_of_isBigO_rpow (P.hf_int.sub (locallyIntegrableOn_const _)) (P.hf_top _) ht From 4866b677f1a4b1862b74f6db599d2ed63f00d4da Mon Sep 17 00:00:00 2001 From: Ralf Stephan Date: Thu, 8 Aug 2024 18:23:33 +0000 Subject: [PATCH 120/359] feat(Algebra/GeomSum): add two lemmas (#15050) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add missing lemma `x ^ m - 1 ∣ x ^ (m * n) - 1` for `x` both in `Nat` and `Ring`. Co-authored-by: grunweg --- Mathlib/Algebra/GeomSum.lean | 23 +++++++++++++++++------ 1 file changed, 17 insertions(+), 6 deletions(-) diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 7bcba9c410122..3232d89dff8f1 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -5,6 +5,7 @@ Authors: Neil Strickland -/ import Mathlib.Algebra.BigOperators.Intervals import Mathlib.Algebra.BigOperators.Ring +import Mathlib.Algebra.Group.NatPowAssoc import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Order.Field.Basic import Mathlib.Algebra.Order.Ring.Abs @@ -174,6 +175,13 @@ theorem Commute.sub_dvd_pow_sub_pow [Ring α] {x y : α} (h : Commute x y) (n : theorem sub_dvd_pow_sub_pow [CommRing α] (x y : α) (n : ℕ) : x - y ∣ x ^ n - y ^ n := (Commute.all x y).sub_dvd_pow_sub_pow n +theorem nat_sub_dvd_pow_sub_pow (x y n : ℕ) : x - y ∣ x ^ n - y ^ n := by + rcases le_or_lt y x with h | h + · have : y ^ n ≤ x ^ n := Nat.pow_le_pow_left h _ + exact mod_cast sub_dvd_pow_sub_pow (x : ℤ) (↑y) n + · have : x ^ n ≤ y ^ n := Nat.pow_le_pow_left h.le _ + exact (Nat.sub_eq_zero_of_le this).symm ▸ dvd_zero (x - y) + theorem one_sub_dvd_one_sub_pow [Ring α] (x : α) (n : ℕ) : 1 - x ∣ 1 - x ^ n := by conv_rhs => rw [← one_pow n] @@ -184,12 +192,15 @@ theorem sub_one_dvd_pow_sub_one [Ring α] (x : α) (n : ℕ) : conv_rhs => rw [← one_pow n] exact (Commute.one_right x).sub_dvd_pow_sub_pow n -theorem nat_sub_dvd_pow_sub_pow (x y n : ℕ) : x - y ∣ x ^ n - y ^ n := by - rcases le_or_lt y x with h | h - · have : y ^ n ≤ x ^ n := Nat.pow_le_pow_left h _ - exact mod_cast sub_dvd_pow_sub_pow (x : ℤ) (↑y) n - · have : x ^ n ≤ y ^ n := Nat.pow_le_pow_left h.le _ - exact (Nat.sub_eq_zero_of_le this).symm ▸ dvd_zero (x - y) +lemma pow_one_sub_dvd_pow_mul_sub_one [Ring α] (x : α) (m n : ℕ) : + ((x ^ m) - 1 : α) ∣ (x ^ (m * n) - 1) := by + rw [npow_mul] + exact sub_one_dvd_pow_sub_one (x := x ^ m) (n := n) + +lemma nat_pow_one_sub_dvd_pow_mul_sub_one (x m n : ℕ) : x ^ m - 1 ∣ x ^ (m * n) - 1 := by + nth_rw 2 [← Nat.one_pow n] + rw [Nat.pow_mul x m n] + apply nat_sub_dvd_pow_sub_pow (x ^ m) 1 theorem Odd.add_dvd_pow_add_pow [CommRing α] (x y : α) {n : ℕ} (h : Odd n) : x + y ∣ x ^ n + y ^ n := by From 61e3806f0285fb536527b28a7fa17fadec57c17a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 8 Aug 2024 18:23:35 +0000 Subject: [PATCH 121/359] feat(Topology/../DomMulAct): add missing instances (#15199) --- .../Topology/Algebra/Constructions/DomMulAct.lean | 13 +++++++++++++ Mathlib/Topology/Compactness/LocallyCompact.lean | 13 +++++++++++-- 2 files changed, 24 insertions(+), 2 deletions(-) diff --git a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean index 78a83ef90b77e..1da11ac38b7b9 100644 --- a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean +++ b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean @@ -95,6 +95,19 @@ instance instFirstCountableTopology [FirstCountableTopology M] : FirstCountableT instance instSecondCountableTopology [SecondCountableTopology M] : SecondCountableTopology Mᵈᵐᵃ := inducing_mk_symm.secondCountableTopology +@[to_additive] +instance instCompactSpace [CompactSpace M] : CompactSpace Mᵈᵐᵃ := + mkHomeomorph.compactSpace + +@[to_additive] +instance instLocallyCompactSpace [LocallyCompactSpace M] : LocallyCompactSpace Mᵈᵐᵃ := + openEmbedding_mk_symm.locallyCompactSpace + +@[to_additive] +instance instWeaklyLocallyCompactSpace [WeaklyLocallyCompactSpace M] : + WeaklyLocallyCompactSpace Mᵈᵐᵃ := + closedEmbedding_mk_symm.weaklyLocallyCompactSpace + @[to_additive (attr := simp)] theorem map_mk_nhds (x : M) : map (mk : M → Mᵈᵐᵃ) (𝓝 x) = 𝓝 (mk x) := mkHomeomorph.map_nhds_eq x diff --git a/Mathlib/Topology/Compactness/LocallyCompact.lean b/Mathlib/Topology/Compactness/LocallyCompact.lean index 346a860454e8b..1881270210e57 100644 --- a/Mathlib/Topology/Compactness/LocallyCompact.lean +++ b/Mathlib/Topology/Compactness/LocallyCompact.lean @@ -18,7 +18,6 @@ open Set Filter Topology TopologicalSpace variable {X : Type*} {Y : Type*} {ι : Type*} variable [TopologicalSpace X] [TopologicalSpace Y] {s t : Set X} - instance [WeaklyLocallyCompactSpace X] [WeaklyLocallyCompactSpace Y] : WeaklyLocallyCompactSpace (X × Y) where exists_compact_mem_nhds x := @@ -29,13 +28,23 @@ instance [WeaklyLocallyCompactSpace X] [WeaklyLocallyCompactSpace Y] : instance {ι : Type*} [Finite ι] {X : ι → Type*} [(i : ι) → TopologicalSpace (X i)] [(i : ι) → WeaklyLocallyCompactSpace (X i)] : WeaklyLocallyCompactSpace ((i : ι) → X i) where - exists_compact_mem_nhds := fun f ↦ by + exists_compact_mem_nhds f := by choose s hsc hs using fun i ↦ exists_compact_mem_nhds (f i) exact ⟨pi univ s, isCompact_univ_pi hsc, set_pi_mem_nhds univ.toFinite fun i _ ↦ hs i⟩ instance (priority := 100) [CompactSpace X] : WeaklyLocallyCompactSpace X where exists_compact_mem_nhds _ := ⟨univ, isCompact_univ, univ_mem⟩ +protected theorem ClosedEmbedding.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace Y] + {f : X → Y} (hf : ClosedEmbedding f) : WeaklyLocallyCompactSpace X where + exists_compact_mem_nhds x := + let ⟨K, hK, hKx⟩ := exists_compact_mem_nhds (f x) + ⟨f ⁻¹' K, hf.isCompact_preimage hK, hf.continuous.continuousAt hKx⟩ + +protected theorem IsClosed.weaklyLocallyCompactSpace [WeaklyLocallyCompactSpace X] + {s : Set X} (hs : IsClosed s) : WeaklyLocallyCompactSpace s := + (closedEmbedding_subtype_val hs).weaklyLocallyCompactSpace + /-- In a weakly locally compact space, every compact set is contained in the interior of a compact set. -/ theorem exists_compact_superset [WeaklyLocallyCompactSpace X] {K : Set X} (hK : IsCompact K) : From ec7748d19ae0e9d710906dd1ee15dfaf07d487f0 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Thu, 8 Aug 2024 19:22:43 +0000 Subject: [PATCH 122/359] perf(Module.LinearMap.Defs): de-simp `LinearMapClass.map_smul` (#15613) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The keys for this are `@DFunLike.coe _ _ _ _ _ (@HSMul.hSMul _ _ _ _ _ _)` so it applies to *anything* of the form `↑f (a * b)` where `f` has a `DFunLike` instance. It builds and lints fine without this simp theorem. --- Mathlib/Algebra/Module/LinearMap/Defs.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index 4549f739077ee..4d7974c2e677b 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -126,7 +126,6 @@ abbrev LinearMapClass (F : Type*) (R : outParam Type*) (M M₂ : Type*) [FunLike F M M₂] := SemilinearMapClass F (RingHom.id R) M M₂ -@[simp high] protected lemma LinearMapClass.map_smul {R M M₂ : outParam Type*} [Semiring R] [AddCommMonoid M] [AddCommMonoid M₂] [Module R M] [Module R M₂] {F : Type*} [FunLike F M M₂] [LinearMapClass F R M M₂] (f : F) (r : R) (x : M) : From 849267ff0fb09be49756988a6e65ab362aff561f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 8 Aug 2024 20:12:25 +0000 Subject: [PATCH 123/359] feat(ContinuousCompMeasurePreserving): generalize to `NullMeasurableSet` (#15112) --- .../ContinuousCompMeasurePreserving.lean | 21 +++++++++---------- Mathlib/MeasureTheory/Measure/Regular.lean | 16 +++++++++++++- 2 files changed, 25 insertions(+), 12 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean b/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean index 8594c514eaaad..09380e0b12e50 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean @@ -45,7 +45,7 @@ More precisely, the measure of the symmetric difference of these two sets tends theorem tendsto_measure_symmDiff_preimage_nhds_zero {l : Filter α} {f : α → C(X, Y)} {g : C(X, Y)} {s : Set Y} (hfg : Tendsto f l (𝓝 g)) (hf : ∀ᶠ a in l, MeasurePreserving (f a) μ ν) (hg : MeasurePreserving g μ ν) - (hs : MeasurableSet s) (hνs : ν s ≠ ∞) : + (hs : NullMeasurableSet s ν) (hνs : ν s ≠ ∞) : Tendsto (fun a ↦ μ ((f a ⁻¹' s) ∆ (g ⁻¹' s))) l (𝓝 0) := by have : ν.InnerRegularCompactLTTop := by rw [← hg.map_eq] @@ -58,7 +58,7 @@ theorem tendsto_measure_symmDiff_preimage_nhds_zero -- Indeed, we can choose an open set `U` such that `ν (U ∆ s) < ε / 3`, -- apply the lemma to `U`, then use the triangle inequality for `μ (_ ∆ _)`. rcases hs.exists_isOpen_symmDiff_lt hνs H.ne' with ⟨U, hUo, hU, hUs⟩ - have hmU : MeasurableSet U := hUo.measurableSet + have hmU : NullMeasurableSet U ν := hUo.measurableSet.nullMeasurableSet replace hUs := hUs.le filter_upwards [hf, this hmU hU.ne _ H hUo] with a hfa ha calc @@ -70,16 +70,15 @@ theorem tendsto_measure_symmDiff_preimage_nhds_zero apply measure_symmDiff_le _ ≤ ε / 3 + ε / 3 + ε / 3 := by gcongr - · rwa [← preimage_symmDiff, hfa.measure_preimage (hs.symmDiff hmU).nullMeasurableSet, - symmDiff_comm] - · rwa [← preimage_symmDiff, hg.measure_preimage (hmU.symmDiff hs).nullMeasurableSet] + · rwa [← preimage_symmDiff, hfa.measure_preimage (hs.symmDiff hmU), symmDiff_comm] + · rwa [← preimage_symmDiff, hg.measure_preimage (hmU.symmDiff hs)] _ = ε := by simp -- Take a compact closed subset `K ⊆ g ⁻¹' s` of almost full measure, -- `μ (g ⁻¹' s \ K) < ε / 2`. - have hνs' : μ (g ⁻¹' s) ≠ ∞ := by rwa [hg.measure_preimage hs.nullMeasurableSet] + have hνs' : μ (g ⁻¹' s) ≠ ∞ := by rwa [hg.measure_preimage hs] obtain ⟨K, hKg, hKco, hKcl, hKμ⟩ : ∃ K, MapsTo g K s ∧ IsCompact K ∧ IsClosed K ∧ μ (g ⁻¹' s \ K) < ε / 2 := - (hs.preimage hg.measurable).exists_isCompact_isClosed_diff_lt hνs' <| by simp [hε.ne'] + (hg.measurable hso.measurableSet).exists_isCompact_isClosed_diff_lt hνs' <| by simp [hε.ne'] have hKm : MeasurableSet K := hKcl.measurableSet -- Take `a` such that `f a` is measure preserving and maps `K` to `s`. -- This is possible, because `K` is a compact set and `s` is an open set. @@ -91,9 +90,8 @@ theorem tendsto_measure_symmDiff_preimage_nhds_zero rw [symmDiff_of_ge ha.subset_preimage, symmDiff_of_le hKg.subset_preimage] gcongr have hK' : μ K ≠ ∞ := ne_top_of_le_ne_top hνs' <| measure_mono hKg.subset_preimage - rw [measure_diff_le_iff_le_add hKm ha.subset_preimage hK', - hfa.measure_preimage hs.nullMeasurableSet, ← hg.measure_preimage hs.nullMeasurableSet, - ← measure_diff_le_iff_le_add hKm hKg.subset_preimage hK'] + rw [measure_diff_le_iff_le_add hKm ha.subset_preimage hK', hfa.measure_preimage hs, + ← hg.measure_preimage hs, ← measure_diff_le_iff_le_add hKm hKg.subset_preimage hK'] exact hKμ.le namespace Lp @@ -122,7 +120,8 @@ theorem compMeasurePreserving_continuous (hp : p ≠ ∞) : | @h_ind c s hs hνs => dsimp only [Lp.simpleFunc.coe_indicatorConst, Lp.indicatorConstLp_compMeasurePreserving] refine continuous_indicatorConstLp_set hp fun f ↦ ?_ - apply tendsto_measure_symmDiff_preimage_nhds_zero continuousAt_subtype_val _ f.2 hs hνs.ne + apply tendsto_measure_symmDiff_preimage_nhds_zero continuousAt_subtype_val _ f.2 + hs.nullMeasurableSet hνs.ne exact eventually_of_forall Subtype.property end Lp diff --git a/Mathlib/MeasureTheory/Measure/Regular.lean b/Mathlib/MeasureTheory/Measure/Regular.lean index da222d3e2a93c..72a01a4c40abd 100644 --- a/Mathlib/MeasureTheory/Measure/Regular.lean +++ b/Mathlib/MeasureTheory/Measure/Regular.lean @@ -815,7 +815,7 @@ protected theorem _root_.IsCompact.exists_isOpen_lt_add [InnerRegularCompactLTTo /-- Let `μ` be a locally finite measure on an R₁ topological space with Borel σ-algebra. If `μ` is inner regular for finite measure sets with respect to compact sets, -then any finite measurable set can be approximated in measure by an open set. +then any measurable set of finite measure can be approximated in measure by an open set. See also `Set.exists_isOpen_lt_of_lt` and `MeasurableSet.exists_isOpen_diff_lt` for the case of an outer regular measure. -/ protected theorem _root_.MeasurableSet.exists_isOpen_symmDiff_lt [InnerRegularCompactLTTop μ] @@ -835,6 +835,20 @@ protected theorem _root_.MeasurableSet.exists_isOpen_symmDiff_lt [InnerRegularCo exact ne_top_of_le_ne_top hμs (by gcongr) · exact lt_of_le_of_lt (by gcongr) hμK +/-- Let `μ` be a locally finite measure on an R₁ topological space with Borel σ-algebra. +If `μ` is inner regular for finite measure sets with respect to compact sets, +then any null measurable set of finite measure can be approximated in measure by an open set. +See also `Set.exists_isOpen_lt_of_lt` and `MeasurableSet.exists_isOpen_diff_lt` +for the case of an outer regular measure. -/ +protected theorem _root_.MeasureTheory.NullMeasurableSet.exists_isOpen_symmDiff_lt + [InnerRegularCompactLTTop μ] [IsLocallyFiniteMeasure μ] [R1Space α] [BorelSpace α] + {s : Set α} (hs : NullMeasurableSet s μ) (hμs : μ s ≠ ∞) {ε : ℝ≥0∞} (hε : ε ≠ 0) : + ∃ U, IsOpen U ∧ μ U < ∞ ∧ μ (U ∆ s) < ε := by + rcases hs with ⟨t, htm, hst⟩ + rcases htm.exists_isOpen_symmDiff_lt (by rwa [← measure_congr hst]) hε with ⟨U, hUo, hμU, hUs⟩ + refine ⟨U, hUo, hμU, ?_⟩ + rwa [measure_congr <| (ae_eq_refl _).symmDiff hst] + instance smul [h : InnerRegularCompactLTTop μ] (c : ℝ≥0∞) : InnerRegularCompactLTTop (c • μ) := by by_cases hc : c = 0 · simp only [hc, zero_smul] From a951fed0840429482e50309ee47ff17c45fecbd8 Mon Sep 17 00:00:00 2001 From: FR Date: Thu, 8 Aug 2024 20:12:27 +0000 Subject: [PATCH 124/359] fix: make `to_additive` deal with `abbrev` correctly (#15474) --- Mathlib/Tactic/ToAdditive/Frontend.lean | 15 +++++++++++++++ test/toAdditive.lean | 14 ++++++++++++++ 2 files changed, 29 insertions(+) diff --git a/Mathlib/Tactic/ToAdditive/Frontend.lean b/Mathlib/Tactic/ToAdditive/Frontend.lean index 3fb558afa5aee..f3d84bce1ec87 100644 --- a/Mathlib/Tactic/ToAdditive/Frontend.lean +++ b/Mathlib/Tactic/ToAdditive/Frontend.lean @@ -729,6 +729,16 @@ def findAuxDecls (e : Expr) (pre : Name) : NameSet := else l +/-- It's just the same as `Lean.Meta.setInlineAttribute` but with type `CoreM Unit`. + +TODO (lean4#4965): make `Lean.Meta.setInlineAttribute` a `CoreM Unit` and remove this definition. -/ +def setInlineAttribute (declName : Name) (kind := Compiler.InlineAttributeKind.inline) : + CoreM Unit := do + let env ← getEnv + match Compiler.setInlineAttribute env declName kind with + | .ok env => setEnv env + | .error msg => throwError msg + /-- transform the declaration `src` and all declarations `pre._proof_i` occurring in `src` using the transforms dictionary. `replace_all`, `trace`, `ignore` and `reorder` are configuration options. @@ -793,6 +803,11 @@ partial def transformDeclAux setEnv <| addNoncomputable (← getEnv) tgt else addAndCompile trgDecl.toDeclaration! + if let .defnDecl { hints := .abbrev, .. } := trgDecl.toDeclaration! then + if (← getReducibilityStatus src) == .reducible then + setReducibilityStatus tgt .reducible + if Compiler.getInlineAttribute? (← getEnv) src == some .inline then + setInlineAttribute tgt -- now add declaration ranges so jump-to-definition works -- note: we currently also do this for auxiliary declarations, while they are not normally -- generated for those. We could change that. diff --git a/test/toAdditive.lean b/test/toAdditive.lean index 069467d17980a..13390631ccf6c 100644 --- a/test/toAdditive.lean +++ b/test/toAdditive.lean @@ -146,6 +146,20 @@ def foo21 {N} {A} [Pow A N] (a : A) (n : N) : A := a ^ n run_cmd liftCoreM <| MetaM.run' <| guard <| relevantArgAttr.find? (← getEnv) `Test.foo21 == some 1 +@[to_additive bar22] +abbrev foo22 {α} [Monoid α] (a : α) : ℕ → α + | 0 => 1 + | _ => a + +run_cmd liftCoreM <| MetaM.run' <| do + -- make `abbrev` definition `reducible` automatically + guard <| (← getReducibilityStatus `Test.bar22) == .reducible + -- make `abbrev` definition `inline` automatically + guard <| (Compiler.getInlineAttribute? (← getEnv) `Test.bar22) == some .inline + -- some auxiliary definitions are also `abbrev` but not `reducible` + guard <| (← getReducibilityStatus `Test.bar22.match_1) != .reducible + guard <| (Compiler.getInlineAttribute? (← getEnv) `Test.bar22.match_1) == some .inline + /- test the eta-expansion applied on `foo6`. -/ run_cmd do let c ← getConstInfo `Test.foo6 From 3c03695ca2ac47379aea8d6b9c86f2988145cc4d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 8 Aug 2024 20:12:28 +0000 Subject: [PATCH 125/359] perf(CategoryTheory/Sites/Sheaf): speed up a slow proof (#15619) --- Mathlib/CategoryTheory/Sites/Sheaf.lean | 36 +++++++++++-------------- 1 file changed, 16 insertions(+), 20 deletions(-) diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index 1eb3f3d744895..bc886bdc1fcf9 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -636,26 +636,22 @@ def isSheafForIsSheafFor' (P : Cᵒᵖ ⥤ A) (s : A ⥤ Type max v₁ u₁) [∀ J, PreservesLimitsOfShape (Discrete.{max v₁ u₁} J) s] (U : C) (R : Presieve U) : IsLimit (s.mapCone (Fork.ofι _ (w R P))) ≃ IsLimit (Fork.ofι _ (Equalizer.Presieve.w (P ⋙ s) R)) := by - apply Equiv.trans (isLimitMapConeForkEquiv _ _) _ - apply (IsLimit.postcomposeHomEquiv _ _).symm.trans (IsLimit.equivIsoLimit _) - · apply NatIso.ofComponents _ _ - · rintro (_ | _) - · apply PreservesProduct.iso s - · apply PreservesProduct.iso s - · rintro _ _ (_ | _) - · refine limit.hom_ext (fun j => ?_) - dsimp [Equalizer.Presieve.firstMap, firstMap] - simp only [limit.lift_π, map_lift_piComparison, assoc, Fan.mk_π_app, Functor.map_comp] - rw [piComparison_comp_π_assoc] - · refine limit.hom_ext (fun j => ?_) - dsimp [Equalizer.Presieve.secondMap, secondMap] - simp only [limit.lift_π, map_lift_piComparison, assoc, Fan.mk_π_app, Functor.map_comp] - rw [piComparison_comp_π_assoc] - · dsimp - simp - · refine Fork.ext (Iso.refl _) ?_ - dsimp [Equalizer.forkMap, forkMap] - simp [Fork.ι] + let e : parallelPair (s.map (firstMap R P)) (s.map (secondMap R P)) ≅ + parallelPair (Equalizer.Presieve.firstMap (P ⋙ s) R) + (Equalizer.Presieve.secondMap (P ⋙ s) R) := by + refine parallelPair.ext (PreservesProduct.iso s _) ((PreservesProduct.iso s _)) + (limit.hom_ext (fun j => ?_)) (limit.hom_ext (fun j => ?_)) + · dsimp [Equalizer.Presieve.firstMap, firstMap] + simp only [map_lift_piComparison, Functor.map_comp, limit.lift_π, Fan.mk_pt, + Fan.mk_π_app, assoc, piComparison_comp_π_assoc] + · dsimp [Equalizer.Presieve.secondMap, secondMap] + simp only [map_lift_piComparison, Functor.map_comp, limit.lift_π, Fan.mk_pt, + Fan.mk_π_app, assoc, piComparison_comp_π_assoc] + refine Equiv.trans (isLimitMapConeForkEquiv _ _) ?_ + refine (IsLimit.postcomposeHomEquiv e _).symm.trans + (IsLimit.equivIsoLimit (Fork.ext (Iso.refl _) ?_)) + dsimp [Equalizer.forkMap, forkMap, e, Fork.ι] + simp only [id_comp, map_lift_piComparison] -- Remark : this lemma uses `A'` not `A`; `A'` is `A` but with a universe -- restriction. Can it be generalised? From 78019567d35dd571716b39ad0eb8345d2e3c8af1 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Thu, 8 Aug 2024 20:12:29 +0000 Subject: [PATCH 126/359] perf(RepresentationTheory.LowDegree): scope simp theorems with weak keys (#15622) Both of these have `HSMul.hSMul _ _ _ _ _ _` as keys so apply to any occurrence of scalar multiplication. --- Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index 62c5e0264b362..bd825851e6089 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -416,7 +416,7 @@ section variable {G A : Type*} [Group G] [AddCommGroup A] [MulAction G A] -@[simp] theorem map_inv_of_isOneCocycle {f : G → A} (hf : IsOneCocycle f) (g : G) : +@[scoped simp] theorem map_inv_of_isOneCocycle {f : G → A} (hf : IsOneCocycle f) (g : G) : g • f g⁻¹ = - f g := by rw [← add_eq_zero_iff_eq_neg, ← map_one_of_isOneCocycle hf, ← mul_inv_self g, hf g g⁻¹] @@ -539,7 +539,7 @@ section variable {G M : Type*} [Group G] [CommGroup M] [MulAction G M] -@[simp] theorem map_inv_of_isMulOneCocycle {f : G → M} (hf : IsMulOneCocycle f) (g : G) : +@[scoped simp] theorem map_inv_of_isMulOneCocycle {f : G → M} (hf : IsMulOneCocycle f) (g : G) : g • f g⁻¹ = (f g)⁻¹ := by rw [← mul_eq_one_iff_eq_inv, ← map_one_of_isMulOneCocycle hf, ← mul_inv_self g, hf g g⁻¹] From 961016cd11daa0966729e114d1e61b84a789d143 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Thu, 8 Aug 2024 20:12:31 +0000 Subject: [PATCH 127/359] perf(NumberTheory.MulChar): scope simp theorem with weak keys (#15623) The keys for this `DFunLike.coe _ _ _ _ _ _` so applly to any expression that has a `DFunLike.coe` subexpression. --- Mathlib/NumberTheory/MulChar/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index 0aad41a44a431..500ac261a79da 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -75,12 +75,12 @@ class MulCharClass (F : Type*) (R R' : outParam Type*) [CommMonoid R] initialize_simps_projections MulChar (toFun → apply, -toMonoidHom) -attribute [simp] MulCharClass.map_nonunit - end Defi namespace MulChar +attribute [scoped simp] MulCharClass.map_nonunit + section Group -- The domain of our multiplicative characters From 419ac86bd755c072f503b33089c07b911988782c Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Thu, 8 Aug 2024 20:12:32 +0000 Subject: [PATCH 128/359] perf(Data.Nat.Cast.Basic): de-simp theorem with weak keys (#15626) The keys are `DFunLike.coe _ _ _ _ _ _` which means we are trying to synthesize two `NonAssocSemiring`'s each time any expression has this as a subterm. --- Archive/Wiedijk100Theorems/AbelRuffini.lean | 1 + Mathlib/Algebra/Star/CHSH.lean | 2 +- Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean | 7 ++++--- Mathlib/Analysis/InnerProductSpace/OfNorm.lean | 1 + Mathlib/Analysis/InnerProductSpace/Symmetric.lean | 1 + Mathlib/Data/Nat/Cast/Basic.lean | 2 -- 6 files changed, 8 insertions(+), 6 deletions(-) diff --git a/Archive/Wiedijk100Theorems/AbelRuffini.lean b/Archive/Wiedijk100Theorems/AbelRuffini.lean index 0d0bde8da3d54..df179eb859924 100644 --- a/Archive/Wiedijk100Theorems/AbelRuffini.lean +++ b/Archive/Wiedijk100Theorems/AbelRuffini.lean @@ -92,6 +92,7 @@ theorem irreducible_Phi (p : ℕ) (hp : p.Prime) (hpa : p ∣ a) (hpb : p ∣ b) exact mt Int.natCast_dvd_natCast.mp hp2b all_goals exact Monic.isPrimitive (monic_Phi a b) +attribute [local simp] map_ofNat in -- use `ofNat` simp theorem with bad keys theorem real_roots_Phi_le : Fintype.card ((Φ ℚ a b).rootSet ℝ) ≤ 3 := by rw [← map_Phi a b (algebraMap ℤ ℚ), Φ, ← one_mul (X ^ 5), ← C_1] apply (card_rootSet_le_derivative _).trans diff --git a/Mathlib/Algebra/Star/CHSH.lean b/Mathlib/Algebra/Star/CHSH.lean index 8247bb1bd6655..58d2771af2470 100644 --- a/Mathlib/Algebra/Star/CHSH.lean +++ b/Mathlib/Algebra/Star/CHSH.lean @@ -120,7 +120,7 @@ theorem CHSH_inequality_of_comm [OrderedCommRing R] [StarRing R] [StarOrderedRin have i₁ : 0 ≤ P := by have idem : P * P = 4 * P := CHSH_id T.A₀_inv T.A₁_inv T.B₀_inv T.B₁_inv have idem' : P = (1 / 4 : ℝ) • (P * P) := by - have h : 4 * P = (4 : ℝ) • P := by simp [Algebra.smul_def] + have h : 4 * P = (4 : ℝ) • P := by simp [map_ofNat, Algebra.smul_def] rw [idem, h, ← mul_smul] norm_num have sa : star P = P := by diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean index 38abc3bb0ed03..6042ae7f510d6 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Jacobian.lean @@ -1573,9 +1573,10 @@ protected lemma map_smul (u : R) : f ∘ (u • P) = f u • (f ∘ P) := by ext i; fin_cases i <;> simp [smul_fin3] @[simp] lemma map_addZ : addZ (f ∘ P) (f ∘ Q) = f (addZ P Q) := by simp [addZ] -@[simp] lemma map_addX : addX (W'.map f) (f ∘ P) (f ∘ Q) = f (W'.addX P Q) := by simp [addX] +@[simp] lemma map_addX : addX (W'.map f) (f ∘ P) (f ∘ Q) = f (W'.addX P Q) := by + simp [map_ofNat, addX] @[simp] lemma map_negAddY : negAddY (W'.map f) (f ∘ P) (f ∘ Q) = f (W'.negAddY P Q) := by - simp [negAddY] + simp [map_ofNat, negAddY] @[simp] lemma map_negY : negY (W'.map f) (f ∘ P) = f (W'.negY P) := by simp [negY] @[simp] protected lemma map_neg : neg (W'.map f) (f ∘ P) = f ∘ W'.neg P := by @@ -1603,7 +1604,7 @@ lemma map_polynomialZ : (W'.map f).toJacobian.polynomialZ = MvPolynomial.map f W @[simp] lemma map_dblU : dblU (W'.map f) (f ∘ P) = f (W'.dblU P) := by simp [dblU, map_polynomialX, ← eval₂_id, eval₂_comp_left] -@[simp] lemma map_dblX : dblX (W'.map f) (f ∘ P) = f (W'.dblX P) := by simp [dblX] +@[simp] lemma map_dblX : dblX (W'.map f) (f ∘ P) = f (W'.dblX P) := by simp [map_ofNat, dblX] @[simp] lemma map_negDblY : negDblY (W'.map f) (f ∘ P) = f (W'.negDblY P) := by simp [negDblY] @[simp] lemma map_dblY : dblY (W'.map f) (f ∘ P) = f (W'.dblY P) := by simp [dblY, ← comp_fin3] diff --git a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean index 7a796a9460a36..542e4fe0e33dc 100644 --- a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean +++ b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean @@ -127,6 +127,7 @@ theorem inner_.norm_sq (x : E) : ‖x‖ ^ 2 = re (inner_ 𝕜 x x) := by map_add, ofReal_re, ofNat_im, ofReal_im, mul_im, I_re, inv_im] ring +attribute [local simp] map_ofNat in -- use `ofNat` simp theorem with bad keys theorem inner_.conj_symm (x y : E) : conj (inner_ 𝕜 y x) = inner_ 𝕜 x y := by simp only [inner_] have h4 : conj (4⁻¹ : 𝕜) = 4⁻¹ := by norm_num diff --git a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean index a9ba99321c951..cc99f2f374367 100644 --- a/Mathlib/Analysis/InnerProductSpace/Symmetric.lean +++ b/Mathlib/Analysis/InnerProductSpace/Symmetric.lean @@ -124,6 +124,7 @@ section Complex variable {V : Type*} [NormedAddCommGroup V] [InnerProductSpace ℂ V] +attribute [local simp] map_ofNat in -- use `ofNat` simp theorem with bad keys /-- A linear operator on a complex inner product space is symmetric precisely when `⟪T v, v⟫_ℂ` is real for all v. -/ theorem isSymmetric_iff_inner_map_self_real (T : V →ₗ[ℂ] V) : diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index 1f78a839c5614..57acfff19ec9f 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -163,8 +163,6 @@ 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 --- See note [no_index around OfNat.ofNat] -@[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 f1c8635ed23f681f13e9875b96c18fd8d7fba655 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 8 Aug 2024 21:19:40 +0000 Subject: [PATCH 129/359] doc(Topology/../DomMulAct): improve docs (#15198) --- .../Algebra/Constructions/DomMulAct.lean | 19 +++++++++++++------ 1 file changed, 13 insertions(+), 6 deletions(-) diff --git a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean index 1da11ac38b7b9..9c399af02e789 100644 --- a/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean +++ b/Mathlib/Topology/Algebra/Constructions/DomMulAct.lean @@ -7,11 +7,17 @@ import Mathlib.Topology.Homeomorph import Mathlib.GroupTheory.GroupAction.DomAct.Basic /-! -# Topological structure on `DomMulAct _` +# Topological space structure on `Mᵈᵐᵃ` and `Mᵈᵃᵃ` -In this file we define topology on `DomMulAct _` and prove basic facts about this topology. -The topology on `Mᵈᵐᵃ` is the same as the topology on `M` -(formally, it is induced by `DomMulAct.mk.symm`, since the types aren't definitionally equal). +In this file we define `TopologicalSpace` structure on `Mᵈᵐᵃ` and `Mᵈᵃᵃ` +and prove basic theorems about these topologies. +The topologies on `Mᵈᵐᵃ` and `Mᵈᵃᵃ` are the same as the topology on `M`. +Formally, they are induced by `DomMulAct.mk.symm` and `DomAddAct.mk.symm`, +since the types aren't definitionally equal. + +## Tags + +topological space, group action, domain action -/ open Filter TopologicalSpace @@ -21,7 +27,8 @@ namespace DomMulAct variable {M : Type*} [TopologicalSpace M] -@[to_additive] +/-- Put the same topological space structure on `Mᵈᵐᵃ` as on the original space. -/ +@[to_additive "Put the same topological space structure on `Mᵈᵃᵃ` as on the original space."] instance instTopologicalSpace : TopologicalSpace Mᵈᵐᵃ := .induced mk.symm ‹_› @[to_additive (attr := continuity, fun_prop)] @@ -31,7 +38,7 @@ theorem continuous_mk : Continuous (@mk M) := continuous_induced_rng.2 continuou theorem continuous_mk_symm : Continuous (@mk M).symm := continuous_induced_dom /-- `DomMulAct.mk` as a homeomorphism. -/ -@[to_additive (attr := simps toEquiv) "`DomAddAct.mk` as a homeomorphism"] +@[to_additive (attr := simps toEquiv) "`DomAddAct.mk` as a homeomorphism."] def mkHomeomorph : M ≃ₜ Mᵈᵐᵃ where toEquiv := mk From faf7c200e4b9b8220bec3e33b01f8494b7a7e1e5 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 8 Aug 2024 21:19:42 +0000 Subject: [PATCH 130/359] golf: found by the have vs let linter (#15595) The `have` vs `let` linter had flagged these `have` as "should be `let`s" and changing them made the proof simpler. --- Mathlib/Analysis/Convolution.lean | 23 +++++++---------------- 1 file changed, 7 insertions(+), 16 deletions(-) diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 44fc258910888..82f6dd1ab1554 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1227,10 +1227,10 @@ theorem contDiffOn_convolution_right_with_param {f : G → E} {n : ℕ∞} (L : let eE' : Type max uE' uG uF uP := ULift.{max uG uF uP} E' let eF : Type max uF uG uE' uP := ULift.{max uG uE' uP} F let eP : Type max uP uG uE' uF := ULift.{max uG uE' uF} P - have isoG : eG ≃L[𝕜] G := ContinuousLinearEquiv.ulift - have isoE' : eE' ≃L[𝕜] E' := ContinuousLinearEquiv.ulift - have isoF : eF ≃L[𝕜] F := ContinuousLinearEquiv.ulift - have isoP : eP ≃L[𝕜] P := ContinuousLinearEquiv.ulift + let isoG : eG ≃L[𝕜] G := ContinuousLinearEquiv.ulift + let isoE' : eE' ≃L[𝕜] E' := ContinuousLinearEquiv.ulift + let isoF : eF ≃L[𝕜] F := ContinuousLinearEquiv.ulift + let isoP : eP ≃L[𝕜] P := ContinuousLinearEquiv.ulift let ef := f ∘ isoG let eμ : Measure eG := Measure.map isoG.symm μ let eg : eP → eG → eE' := fun ep ex => isoE'.symm (g (isoP ep) (isoG ex)) @@ -1246,11 +1246,7 @@ theorem contDiffOn_convolution_right_with_param {f : G → E} {n : ℕ∞} (L : simp only [eg, (· ∘ ·), ContinuousLinearEquiv.prod_apply, LinearIsometryEquiv.coe_coe, ContinuousLinearEquiv.map_eq_zero_iff] exact hgs _ _ hp hx - · apply (locallyIntegrable_map_homeomorph isoG.symm.toHomeomorph).2 - convert hf - ext1 x - simp only [ef, ContinuousLinearEquiv.coe_toHomeomorph, (· ∘ ·), - ContinuousLinearEquiv.apply_symm_apply] + · exact (locallyIntegrable_map_homeomorph isoG.symm.toHomeomorph).2 hf · apply isoE'.symm.contDiff.comp_contDiffOn apply hg.comp (isoP.prod isoG).contDiff.contDiffOn rintro ⟨p, x⟩ ⟨hp, -⟩ @@ -1269,13 +1265,8 @@ theorem contDiffOn_convolution_right_with_param {f : G → E} {n : ℕ∞} (L : ContinuousLinearEquiv.prod_apply] simp only [R, convolution, coe_comp', ContinuousLinearEquiv.coe_coe, (· ∘ ·)] rw [ClosedEmbedding.integral_map, ← isoF.integral_comp_comm] - swap; · exact isoG.symm.toHomeomorph.closedEmbedding - congr 1 - ext1 a - simp only [ef, eg, eL, (· ∘ ·), ContinuousLinearEquiv.apply_symm_apply, coe_comp', - ContinuousLinearEquiv.prod_apply, ContinuousLinearEquiv.map_sub, - ContinuousLinearEquiv.arrowCongr, ContinuousLinearEquiv.arrowCongrSL_symm_apply, - ContinuousLinearEquiv.coe_coe, Function.comp_apply, ContinuousLinearEquiv.apply_symm_apply] + · rfl + · exact isoG.symm.toHomeomorph.closedEmbedding simp_rw [this] at A exact A From e8110d4997635efe76820a06ee6f009ada2ce683 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Thu, 8 Aug 2024 23:56:34 +0000 Subject: [PATCH 131/359] chore: rm relaxedAutoImplicit from lakefile (#15604) The `relaxedAutoImplicit` option has no effect if `autoImplicit` is false, so this has no impact on most of mathlib. However, when enabling `autoImplicit` locally, it is annoying (and slightly non-discoverable) that one must also enable `relaxedAutoImplicit` in many cases to actually get auto-implicits. Co-authored-by: Mario Carneiro --- Mathlib/Tactic/Ring/Basic.lean | 2 -- lakefile.lean | 3 +-- 2 files changed, 1 insertion(+), 4 deletions(-) diff --git a/Mathlib/Tactic/Ring/Basic.lean b/Mathlib/Tactic/Ring/Basic.lean index 9eee27032abc3..49b7a2931f4f8 100644 --- a/Mathlib/Tactic/Ring/Basic.lean +++ b/Mathlib/Tactic/Ring/Basic.lean @@ -136,7 +136,6 @@ end -- In this file, we would like to use multi-character auto-implicits. -set_option relaxedAutoImplicit true set_option autoImplicit true mutual -- partial only to speed up compilation @@ -212,7 +211,6 @@ partial def ExSum.cast {a : Q($arg)} : ExSum sα a → Σ a, ExSum sβ a end set_option autoImplicit false -set_option relaxedAutoImplicit false variable {u : Lean.Level} diff --git a/lakefile.lean b/lakefile.lean index a4540d64f35f6..23162bc1da869 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -5,8 +5,7 @@ open Lake DSL package mathlib where leanOptions := #[ ⟨`pp.unicode.fun, true⟩, -- pretty-prints `fun a ↦ b` - ⟨`autoImplicit, false⟩, - ⟨`relaxedAutoImplicit, false⟩ + ⟨`autoImplicit, false⟩ ] -- These are additional settings which do not affect the lake hash, -- so they can be enabled in CI and disabled locally or vice versa. From 9101a5ceaaef32a09b26cbd804d64bfbfb3b0775 Mon Sep 17 00:00:00 2001 From: grunweg Date: Fri, 9 Aug 2024 00:58:17 +0000 Subject: [PATCH 132/359] refactor: polish `lint_style` and rename to lint-style (#14757) - Use upperCamelCase for all local variables, per the naming convention. - Rename the executable to lint-style, per [zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Naming.20convention/near/451072858). --- .github/workflows/nolints.yml | 2 +- Mathlib/Tactic/Linter/TextBased.lean | 58 ++++++++++---------- lakefile.lean | 4 +- scripts/{lint_style.lean => lint-style.lean} | 14 ++--- scripts/lint-style.py | 4 +- scripts/lint-style.sh | 2 +- scripts/print-style-errors.sh | 2 +- 7 files changed, 43 insertions(+), 43 deletions(-) rename scripts/{lint_style.lean => lint-style.lean} (84%) diff --git a/.github/workflows/nolints.yml b/.github/workflows/nolints.yml index 6536b44c24c78..ed9b666b0d7eb 100644 --- a/.github/workflows/nolints.yml +++ b/.github/workflows/nolints.yml @@ -57,7 +57,7 @@ jobs: shell: bash run: | env LEAN_ABORT_ON_PANIC=1 lake exe runLinter --update Mathlib - lake exe lint_style --update + lake exe lint-style --update - name: configure git setup run: | diff --git a/Mathlib/Tactic/Linter/TextBased.lean b/Mathlib/Tactic/Linter/TextBased.lean index 09a1bfbf6335c..caacb32947720 100644 --- a/Mathlib/Tactic/Linter/TextBased.lean +++ b/Mathlib/Tactic/Linter/TextBased.lean @@ -17,7 +17,7 @@ All of these have been rewritten from the `lint-style.py` script. For now, this only contains the linters for the copyright and author headers and large files: further linters will be ported in subsequent PRs. -An executable running all these linters is defined in `scripts/lint_style.lean`. +An executable running all these linters is defined in `scripts/lint-style.lean`. -/ open System @@ -51,7 +51,7 @@ inductive StyleError where as well as a size limit (slightly larger). On future runs, this linter will allow this file to grow up to this limit. For diagnostic purposes, this may also contain a previous size limit, which is now exceeded. -/ - | fileTooLong (number_lines : ℕ) (new_size_limit : ℕ) (previousLimit : Option ℕ) : StyleError + | fileTooLong (numberLines : ℕ) (newSizeLimit : ℕ) (previousLimit : Option ℕ) : StyleError deriving BEq /-- How to format style errors -/ @@ -81,16 +81,16 @@ def StyleError.errorMessage (err : StyleError) (style : ErrorFormat) : String := e.g. mathlib4#13779). Please consider carefully if this import is useful and make sure to \ benchmark it. If this is fine, feel free to allow this linter." | StyleError.lineLength n => s!"Line has {n} characters, which is more than 100" - | StyleError.fileTooLong current_size size_limit previousLimit => + | StyleError.fileTooLong currentSize sizeLimit previousLimit => match style with | ErrorFormat.github => if let some n := previousLimit then - s!"file contains {current_size} lines (at most {n} allowed), try to split it up" + s!"file contains {currentSize} lines (at most {n} allowed), try to split it up" else - s!"file contains {current_size} lines, try to split it up" + s!"file contains {currentSize} lines, try to split it up" | ErrorFormat.exceptionsFile => - s!"{size_limit} file contains {current_size} lines, try to split it up" - | ErrorFormat.humanReadable => s!"file contains {current_size} lines, try to split it up" + s!"{sizeLimit} file contains {currentSize} lines, try to split it up" + | ErrorFormat.humanReadable => s!"file contains {currentSize} lines, try to split it up" /-- 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; @@ -164,7 +164,7 @@ def ErrorContext.find?_comparable (e : ErrorContext) (exceptions : Array ErrorCo `style` specifies if the error should be formatted for humans to read, github problem matchers to consume, or for the style exceptions file. -/ def outputMessage (errctx : ErrorContext) (style : ErrorFormat) : String := - let error_message := errctx.error.errorMessage style + let errorMessage := errctx.error.errorMessage style match style with | ErrorFormat.github => -- We are outputting for github: duplicate file path, line number and error code, @@ -172,31 +172,31 @@ def outputMessage (errctx : ErrorContext) (style : ErrorFormat) : String := let path := errctx.path let nr := errctx.lineNumber let code := errctx.error.errorCode - s!"::ERR file={path},line={nr},code={code}::{path}:{nr} {code}: {error_message}" + s!"::ERR file={path},line={nr},code={code}::{path}:{nr} {code}: {errorMessage}" | ErrorFormat.exceptionsFile => -- Produce an entry in the exceptions file: with error code and "line" in front of the number. - s!"{errctx.path} : line {errctx.lineNumber} : {errctx.error.errorCode} : {error_message}" + s!"{errctx.path} : line {errctx.lineNumber} : {errctx.error.errorCode} : {errorMessage}" | ErrorFormat.humanReadable => -- Print for humans: clickable file name and omit the error code - s!"error: {errctx.path}:{errctx.lineNumber}: {error_message}" + s!"error: {errctx.path}:{errctx.lineNumber}: {errorMessage}" /-- Try parsing an `ErrorContext` from a string: return `some` if successful, `none` otherwise. -/ def parse?_errorContext (line : String) : Option ErrorContext := Id.run do let parts := line.split (· == ' ') match parts with - | filename :: ":" :: "line" :: line_number :: ":" :: error_code :: ":" :: error_message => + | filename :: ":" :: "line" :: lineNumber :: ":" :: errorCode :: ":" :: errorMessage => -- Turn the filename into a path. In general, this is ambiguous if we don't know if we're -- dealing with e.g. Windows or POSIX paths. In our setting, this is fine, since no path -- component contains any path separator. let path := mkFilePath (filename.split (FilePath.pathSeparators.contains ·)) -- Parse the error kind from the error code, ugh. -- NB: keep this in sync with `StyleError.errorCode` above! - let err : Option StyleError := match error_code with + let err : Option StyleError := match errorCode with -- Use default values for parameters which are ignored for comparing style exceptions. -- NB: keep this in sync with `compare` above! | "ERR_COP" => some (StyleError.copyright none) | "ERR_LIN" => - if let some n := error_message.get? 2 then + if let some n := errorMessage.get? 2 then match String.toNat? n with | some n => return StyleError.lineLength n | none => none @@ -205,21 +205,21 @@ def parse?_errorContext (line : String) : Option ErrorContext := Id.run do | "ERR_ADN" => some (StyleError.adaptationNote) | "ERR_IMP" => -- XXX tweak exceptions messages to ease parsing? - if (error_message.get! 0).containsSubstr "tactic" then + if (errorMessage.get! 0).containsSubstr "tactic" then some (StyleError.broadImport BroadImports.TacticFolder) else some (StyleError.broadImport BroadImports.Lake) | "ERR_NUM_LIN" => -- Parse the error message in the script. `none` indicates invalid input. - match (error_message.get? 0, error_message.get? 3) with + match (errorMessage.get? 0, errorMessage.get? 3) with | (some limit, some current) => match (String.toNat? limit, String.toNat? current) with - | (some size_limit, some current_size) => - some (StyleError.fileTooLong current_size size_limit (some size_limit)) + | (some sizeLimit, some currentSize) => + some (StyleError.fileTooLong currentSize sizeLimit (some sizeLimit)) | _ => none | _ => none | _ => none - match String.toNat? line_number with + match String.toNat? lineNumber with | some n => err.map fun e ↦ (ErrorContext.mk e n path) | _ => none -- It would be nice to print an error on any line which doesn't match the above format, @@ -336,9 +336,9 @@ def broadImportsLinter : TextbasedLinter := fun lines ↦ Id.run do We allow URLs to be longer, though. -/ def lineLengthLinter : TextbasedLinter := fun lines ↦ Id.run do - let errors := (lines.toList.enumFrom 1).filterMap (fun (line_number, line) ↦ + let errors := (lines.toList.enumFrom 1).filterMap (fun (lineNumber, line) ↦ if line.length > 101 && !line.containsSubstr "http" then - some (StyleError.lineLength line.length, line_number) + some (StyleError.lineLength line.length, lineNumber) else none) errors.toArray @@ -353,17 +353,17 @@ def isImportsOnlyFile (lines : Array String) : Bool := **and** longer than an optional previous limit. If the file is too large, return a matching `StyleError`, which includes a new size limit (which is somewhat larger than the current size). -/ -def checkFileLength (lines : Array String) (existing_limit : Option ℕ) : Option StyleError := +def checkFileLength (lines : Array String) (existingLimit : Option ℕ) : Option StyleError := Id.run do if lines.size > 1500 then - let is_larger : Bool := match existing_limit with + let isLarger : Bool := match existingLimit with | some mark => lines.size > mark | none => true - if is_larger then + if isLarger then -- We add about 200 lines of slack to the current file size: small PRs will be unaffected, -- but sufficiently large PRs will get nudged towards splitting up this file. return some (StyleError.fileTooLong lines.size - ((Nat.div lines.size 100) * 100 + 200) existing_limit) + ((Nat.div lines.size 100) * 100 + 200) existingLimit) none end @@ -444,12 +444,12 @@ def lintModules (moduleNames : Array String) (mode : OutputSetting) : IO UInt32 formatErrors allUnexpectedErrors style if numberErrorFiles > 0 && mode matches OutputSetting.print _ then IO.println s!"error: found {numberErrorFiles} new style errors\n\ - run `lake exe lint_style --update` to ignore all of them" + run `lake exe lint-style --update` to ignore all of them" | OutputSetting.update => formatErrors allUnexpectedErrors ErrorFormat.humanReadable -- Regenerate the style exceptions file, including the Python output. IO.FS.writeFile exceptionsFilePath "" - let python_output ← IO.Process.run { cmd := "./scripts/print-style-errors.sh" } + let pythonOutput ← IO.Process.run { cmd := "./scripts/print-style-errors.sh" } -- Combine style exception entries: for each new error, replace by a corresponding -- previous exception if that is preferred. let mut tweaked := allUnexpectedErrors.map fun err ↦ @@ -457,7 +457,7 @@ def lintModules (moduleNames : Array String) (mode : OutputSetting) : IO UInt32 if let ComparisonResult.Comparable (true) := _root_.compare err existing then existing else err else err - let this_output := "\n".intercalate (tweaked.map + let thisOutput := "\n".intercalate (tweaked.map (fun err ↦ outputMessage err ErrorFormat.exceptionsFile)).toList - IO.FS.writeFile exceptionsFilePath s!"{python_output}{this_output}\n" + IO.FS.writeFile exceptionsFilePath s!"{pythonOutput}{thisOutput}\n" return numberErrorFiles diff --git a/lakefile.lean b/lakefile.lean index 23162bc1da869..2d1665c48244d 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -65,8 +65,8 @@ lean_exe shake where root := `Shake.Main supportInterpreter := true -/-- `lake exe lint_style` runs text-based style linters. -/ -lean_exe lint_style where +/-- `lake exe lint-style` runs text-based style linters. -/ +lean_exe «lint-style» where srcDir := "scripts" /-- diff --git a/scripts/lint_style.lean b/scripts/lint-style.lean similarity index 84% rename from scripts/lint_style.lean rename to scripts/lint-style.lean index 5138eca39af87..93a1068c14cad 100644 --- a/scripts/lint_style.lean +++ b/scripts/lint-style.lean @@ -10,13 +10,13 @@ import Cli.Basic /-! # Text-based style linters -This files defines the `lint_style` executable which runs all text-based style linters. +This files defines the `lint-style` executable which runs all text-based style linters. The linters themselves are defined in `Mathlib.Tactic.Linter.TextBased`. -/ open Cli -/-- Implementation of the `lint_style` command line program. -/ +/-- Implementation of the `lint-style` command line program. -/ def lintStyleCli (args : Cli.Parsed) : IO UInt32 := do let mode : OutputSetting := match (args.hasFlag "update", args.hasFlag "github") with | (true, _) => OutputSetting.update @@ -31,10 +31,10 @@ def lintStyleCli (args : Cli.Parsed) : IO UInt32 := do -- in shell scripts. return min numberErrorFiles 125 -/-- Setting up command line options and help text for `lake exe lint_style`. -/ +/-- Setting up command line options and help text for `lake exe lint-style`. -/ -- so far, no help options or so: perhaps that is fine? -def lint_style : Cmd := `[Cli| - lint_style VIA lintStyleCli; ["0.0.1"] +def lintStyle : Cmd := `[Cli| + «lint-style» VIA lintStyleCli; ["0.0.1"] "Run text-based style linters on every Lean file in Mathlib/, Archive/ and Counterexamples/. Print errors about any unexpected style errors to standard output." @@ -48,5 +48,5 @@ def lint_style : Cmd := `[Cli| and run this script again with this flag." ] -/-- The entry point to the `lake exe lint_style` command. -/ -def main (args : List String) : IO UInt32 := do lint_style.validate args +/-- The entry point to the `lake exe lint-style` command. -/ +def main (args : List String) : IO UInt32 := do lintStyle.validate args diff --git a/scripts/lint-style.py b/scripts/lint-style.py index 7ecb30e7869b5..43e6f34940475 100755 --- a/scripts/lint-style.py +++ b/scripts/lint-style.py @@ -25,9 +25,9 @@ The linters in this script are gradually being rewritten in Lean. Do not add new linters here; please write them in Lean instead. -To run all style linters, run `lake exe lint_style`. +To run all style linters, run `lake exe lint-style`. To update the list of allowed/ignored style exceptions, use - $ lake exe lint_style --regenerate + $ lake exe lint-style --update """ # TODO: This is adapted from the linter for mathlib3. It should be rewritten in Lean. diff --git a/scripts/lint-style.sh b/scripts/lint-style.sh index 7a4399e456bf3..91d550ab65b55 100755 --- a/scripts/lint-style.sh +++ b/scripts/lint-style.sh @@ -43,7 +43,7 @@ git ls-files 'Archive/*.lean' | xargs ./scripts/lint-style.py "$@" git ls-files 'Counterexamples/*.lean' | xargs ./scripts/lint-style.py "$@" # Call the in-progress Lean rewrite of these Python lints. -lake exe lint_style --github +lake exe lint-style --github # 2. Global checks on the mathlib repository diff --git a/scripts/print-style-errors.sh b/scripts/print-style-errors.sh index c62d223886831..1eda2b5b93663 100755 --- a/scripts/print-style-errors.sh +++ b/scripts/print-style-errors.sh @@ -2,7 +2,7 @@ # Print all errors of the python style linter. This script is temporary and should be removed # once the Python style linters have been rewritten in Lean. -# Humans should never run this directly, but at most through `lean exe lint_style --regenerate` +# Humans should never run this directly, but at most through `lean exe lint-style --update` # use C locale so that sorting is the same on macOS and Linux # see https://unix.stackexchange.com/questions/362728/why-does-gnu-sort-sort-differently-on-my-osx-machine-and-linux-machine From 5c40331147f7b03bc1d9c7ef5dc46f8c4f0c816c Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 9 Aug 2024 00:58:18 +0000 Subject: [PATCH 133/359] fix: use `read` in the long line linter (#15180) Follow up from the [discussion on Zulip](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/reading.20the.20.22visible.22.20file.2C.20rather.20than.20the.20saved.20one). As a consequence, the lean server should flag long lines among the imports. --- Mathlib/Tactic/Linter/Lint.lean | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index b6ea9fc557177..8c049afef1034 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -248,11 +248,10 @@ def longLineLinter : Linter where run := withSetOptionIn fun stx ↦ do -- if the linter reached the end of the file, then we scan the `import` syntax instead let stx := ← do if stx.isOfKind ``Lean.Parser.Command.eoi then - let fname ← getFileName - if !(← System.FilePath.pathExists fname) then return default - let contents ← IO.FS.readFile fname + let fileMap ← getFileMap -- `impMods` is the syntax for the modules imported in the current file - let (impMods, _) ← Parser.parseHeader (Parser.mkInputContext contents fname) + let (impMods, _) ← Parser.parseHeader + { input := fileMap.source, fileName := ← getFileName, fileMap := fileMap } return impMods else return stx let sstr := stx.getSubstring? From f248f7dbe5bfde8b89f182d72258501f6a506591 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 9 Aug 2024 00:58:20 +0000 Subject: [PATCH 134/359] chore: backports for leanprover/lean4#4814 (part 24) (#15520) --- .../Algebra/Module/LinearMap/Polynomial.lean | 9 +++-- .../EllipticCurve/Affine.lean | 8 ++--- .../InverseFunctionTheorem/FDeriv.lean | 3 +- Mathlib/Analysis/Convex/Body.lean | 20 +++++------ Mathlib/Analysis/Convex/GaugeRescale.lean | 31 ++++++++++------- .../Analysis/InnerProductSpace/Rayleigh.lean | 6 ++-- .../HahnBanach/SeparatingDual.lean | 6 +++- .../RingedSpace/PresheafedSpace/Gluing.lean | 2 -- Mathlib/GroupTheory/PushoutI.lean | 10 +++--- .../BilinearForm/Orthogonal.lean | 18 +++++----- .../BilinearForm/TensorProduct.lean | 3 +- Mathlib/LinearAlgebra/Charpoly/Basic.lean | 7 ++-- .../CliffordAlgebra/BaseChange.lean | 2 +- Mathlib/LinearAlgebra/Contraction.lean | 2 +- Mathlib/LinearAlgebra/Orientation.lean | 20 +++++------ .../TensorProduct/Graded/Internal.lean | 4 ++- Mathlib/MeasureTheory/Function/AEEqFun.lean | 21 ++++++++---- .../Function/SpecialFunctions/Arctan.lean | 4 +-- Mathlib/MeasureTheory/Integral/Lebesgue.lean | 2 +- .../MeasureTheory/Measure/AEMeasurable.lean | 2 +- .../Measure/MeasureSpaceDef.lean | 3 +- Mathlib/MeasureTheory/Measure/OpenPos.lean | 2 +- .../MeasureTheory/Order/Group/Lattice.lean | 16 ++++----- Mathlib/NumberTheory/Rayleigh.lean | 7 ++-- Mathlib/NumberTheory/SiegelsLemma.lean | 21 +++++++----- .../AdicCompletion/Functoriality.lean | 12 ++++--- Mathlib/RingTheory/Adjoin/Field.lean | 24 +++++++++----- Mathlib/RingTheory/Algebraic.lean | 10 +++--- Mathlib/RingTheory/Artinian.lean | 3 +- Mathlib/RingTheory/EssentialFiniteness.lean | 7 +++- .../IntegralClosure/Algebra/Basic.lean | 2 +- .../IntegralClosure/IntegrallyClosed.lean | 33 +++++++++++-------- .../RingTheory/Kaehler/CotangentComplex.lean | 2 +- .../RingTheory/Localization/BaseChange.lean | 2 +- Mathlib/RingTheory/PowerBasis.lean | 6 ++-- .../RingTheory/Regular/RegularSequence.lean | 6 ++-- Mathlib/RingTheory/SurjectiveOnStalks.lean | 22 ++++++++----- .../TensorProduct/MvPolynomial.lean | 8 ++++- Mathlib/RingTheory/Trace/Defs.lean | 6 ++-- .../Compactness/CompactlyGeneratedSpace.lean | 27 +++++++-------- .../Topology/ContinuousFunction/Ideals.lean | 2 +- .../MetricSpace/GromovHausdorffRealized.lean | 11 ++++--- Mathlib/Topology/VectorBundle/Hom.lean | 8 +++-- 43 files changed, 244 insertions(+), 176 deletions(-) diff --git a/Mathlib/Algebra/Module/LinearMap/Polynomial.lean b/Mathlib/Algebra/Module/LinearMap/Polynomial.lean index 4f40705b3285b..6f0b08d7782c3 100644 --- a/Mathlib/Algebra/Module/LinearMap/Polynomial.lean +++ b/Mathlib/Algebra/Module/LinearMap/Polynomial.lean @@ -467,6 +467,7 @@ noncomputable def nilRank (φ : L →ₗ[R] Module.End R M) : ℕ := nilRankAux φ (Module.Free.chooseBasis R L) +section variable [Nontrivial R] lemma nilRank_eq_polyCharpoly_natTrailingDegree (b : Basis ι R L) : @@ -480,7 +481,7 @@ lemma polyCharpoly_coeff_nilRank_ne_zero : open FiniteDimensional Module.Free -lemma nilRank_le_card (b : Basis ι R M) : nilRank φ ≤ Fintype.card ι := by +lemma nilRank_le_card {ι : Type*} [Fintype ι] (b : Basis ι R M) : nilRank φ ≤ Fintype.card ι := by apply Polynomial.natTrailingDegree_le_of_ne_zero rw [← FiniteDimensional.finrank_eq_card_basis b, ← polyCharpoly_natDegree φ (chooseBasis R L), Polynomial.coeff_natDegree, (polyCharpoly_monic _ _).leadingCoeff] @@ -498,6 +499,8 @@ lemma nilRank_le_natTrailingDegree_charpoly (x : L) : apply Polynomial.trailingCoeff_nonzero_iff_nonzero.mpr _ h apply (LinearMap.charpoly_monic _).ne_zero +end + /-- Let `L` and `M` be finite free modules over `R`, and let `φ : L →ₗ[R] Module.End R M` be a linear family of endomorphisms, and denote `n := nilRank φ`. @@ -518,7 +521,7 @@ lemma isNilRegular_iff_coeff_polyCharpoly_nilRank_ne_zero : ((polyCharpoly φ b).coeff (nilRank φ)) ≠ 0 := by rw [IsNilRegular, polyCharpoly_coeff_eval] -lemma isNilRegular_iff_natTrailingDegree_charpoly_eq_nilRank : +lemma isNilRegular_iff_natTrailingDegree_charpoly_eq_nilRank [Nontrivial R] : IsNilRegular φ x ↔ (φ x).charpoly.natTrailingDegree = nilRank φ := by rw [isNilRegular_def] constructor @@ -535,7 +538,7 @@ section IsDomain variable [IsDomain R] -open Cardinal FiniteDimensional MvPolynomial in +open Cardinal FiniteDimensional MvPolynomial Module.Free in lemma exists_isNilRegular_of_finrank_le_card (h : finrank R M ≤ #R) : ∃ x : L, IsNilRegular φ x := by let b := chooseBasis R L diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index d3e6d63300354..8e25ace403b09 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -551,11 +551,11 @@ lemma nonsingular_add {x₁ x₂ y₁ y₂ : F} (h₁ : W.Nonsingular x₁ y₁) W.Nonsingular (W.addX x₁ x₂ <| W.slope x₁ x₂ y₁ y₂) (W.addY x₁ x₂ y₁ <| W.slope x₁ x₂ y₁ y₂) := nonsingular_neg <| nonsingular_negAdd h₁ h₂ hxy -variable {x₁ x₂ : F} (y₁ y₂ : F) (hx : x₁ ≠ x₂) +variable {x₁ x₂ : F} (y₁ y₂ : F) /-- The formula x(P₁ + P₂) = x(P₁ - P₂) - ψ(P₁)ψ(P₂) / (x(P₂) - x(P₁))², where ψ(x,y) = 2y + a₁x + a₃. -/ -lemma addX_eq_addX_negY_sub : +lemma addX_eq_addX_negY_sub (hx : x₁ ≠ x₂) : W.addX x₁ x₂ (W.slope x₁ x₂ y₁ y₂) = W.addX x₁ x₂ (W.slope x₁ x₂ y₁ (W.negY x₂ y₂)) - (y₁ - W.negY x₁ y₁) * (y₂ - W.negY x₂ y₂) / (x₂ - x₁) ^ 2 := by simp_rw [slope_of_X_ne hx, addX, negY, ← neg_sub x₁, neg_sq] @@ -564,7 +564,7 @@ lemma addX_eq_addX_negY_sub : /-- The formula y(P₁)(x(P₂) - x(P₃)) + y(P₂)(x(P₃) - x(P₁)) + y(P₃)(x(P₁) - x(P₂)) = 0, assuming that P₁ + P₂ + P₃ = O. -/ -lemma cyclic_sum_Y_mul_X_sub_X : +lemma cyclic_sum_Y_mul_X_sub_X (hx : x₁ ≠ x₂) : letI x₃ := W.addX x₁ x₂ (W.slope x₁ x₂ y₁ y₂) y₁ * (x₂ - x₃) + y₂ * (x₃ - x₁) + W.negAddY x₁ x₂ y₁ (W.slope x₁ x₂ y₁ y₂) * (x₁ - x₂) = 0 := by simp_rw [slope_of_X_ne hx, negAddY, addX] @@ -574,7 +574,7 @@ lemma cyclic_sum_Y_mul_X_sub_X : /-- The formula ψ(P₁ + P₂) = (ψ(P₂)(x(P₁) - x(P₃)) - ψ(P₁)(x(P₂) - x(P₃))) / (x(P₂) - x(P₁)), where ψ(x,y) = 2y + a₁x + a₃. -/ -lemma addY_sub_negY_addY : +lemma addY_sub_negY_addY (hx : x₁ ≠ x₂) : letI x₃ := W.addX x₁ x₂ (W.slope x₁ x₂ y₁ y₂) letI y₃ := W.addY x₁ x₂ y₁ (W.slope x₁ x₂ y₁ y₂) y₃ - W.negY x₃ y₃ = diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean index a76f0b66ac46e..82d6364c31eb6 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/FDeriv.lean @@ -92,7 +92,7 @@ theorem map_nhds_eq_of_surj [CompleteSpace E] [CompleteSpace F] {f : E → F} {f apply hs.map_nhds_eq f'symm s_nhds (Or.inr (NNReal.half_lt_self _)) simp [ne_of_gt f'symm_pos] -variable [CompleteSpace E] {f : E → F} {f' : E ≃L[𝕜] F} {a : E} +variable {f : E → F} {f' : E ≃L[𝕜] F} {a : E} theorem approximates_deriv_on_open_nhds (hf : HasStrictFDerivAt f (f' : E →L[𝕜] F) a) : ∃ s : Set E, a ∈ s ∧ IsOpen s ∧ @@ -104,6 +104,7 @@ theorem approximates_deriv_on_open_nhds (hf : HasStrictFDerivAt f (f' : E →L[ f'.subsingleton_or_nnnorm_symm_pos.imp id fun hf' => half_pos <| inv_pos.2 hf' variable (f) +variable [CompleteSpace E] /-- Given a function with an invertible strict derivative at `a`, returns a `PartialHomeomorph` with `to_fun = f` and `a ∈ source`. This is a part of the inverse function theorem. diff --git a/Mathlib/Analysis/Convex/Body.lean b/Mathlib/Analysis/Convex/Body.lean index e8edd08e74e9f..334349d6395d4 100644 --- a/Mathlib/Analysis/Convex/Body.lean +++ b/Mathlib/Analysis/Convex/Body.lean @@ -89,6 +89,16 @@ theorem zero_mem_of_symmetric (K : ConvexBody V) (h_symm : ∀ x ∈ K, - x ∈ section ContinuousAdd +instance : Zero (ConvexBody V) where + zero := ⟨0, convex_singleton 0, isCompact_singleton, Set.singleton_nonempty 0⟩ + +@[simp] -- Porting note: add norm_cast; we leave it out for now to reproduce mathlib3 behavior. +theorem coe_zero : (↑(0 : ConvexBody V) : Set V) = 0 := + rfl + +instance : Inhabited (ConvexBody V) := + ⟨0⟩ + variable [ContinuousAdd V] instance : Add (ConvexBody V) where @@ -96,9 +106,6 @@ instance : Add (ConvexBody V) where ⟨K + L, K.convex.add L.convex, K.isCompact.add L.isCompact, K.nonempty.add L.nonempty⟩ -instance : Zero (ConvexBody V) where - zero := ⟨0, convex_singleton 0, isCompact_singleton, Set.singleton_nonempty 0⟩ - instance : SMul ℕ (ConvexBody V) where smul := nsmulRec @@ -114,13 +121,6 @@ instance : AddMonoid (ConvexBody V) := theorem coe_add (K L : ConvexBody V) : (↑(K + L) : Set V) = (K : Set V) + L := rfl -@[simp] -- Porting note: add norm_cast; we leave it out for now to reproduce mathlib3 behavior. -theorem coe_zero : (↑(0 : ConvexBody V) : Set V) = 0 := - rfl - -instance : Inhabited (ConvexBody V) := - ⟨0⟩ - instance : AddCommMonoid (ConvexBody V) := SetLike.coe_injective.addCommMonoid (↑) rfl (fun _ _ ↦ rfl) fun _ _ ↦ coe_nsmul _ _ diff --git a/Mathlib/Analysis/Convex/GaugeRescale.lean b/Mathlib/Analysis/Convex/GaugeRescale.lean index 9b6e1e49d9956..76930a882f4c9 100644 --- a/Mathlib/Analysis/Convex/GaugeRescale.lean +++ b/Mathlib/Analysis/Convex/GaugeRescale.lean @@ -43,7 +43,21 @@ theorem gaugeRescale_smul (s t : Set E) {c : ℝ} (hc : 0 ≤ c) (x : E) : simp only [gaugeRescale, gauge_smul_of_nonneg hc, smul_smul, smul_eq_mul] rw [mul_div_mul_comm, mul_right_comm, div_self_mul_self] -variable [TopologicalSpace E] [T1Space E] +theorem gauge_gaugeRescale' (s : Set E) {t : Set E} {x : E} (hx : gauge t x ≠ 0) : + gauge t (gaugeRescale s t x) = gauge s x := by + rw [gaugeRescale, gauge_smul_of_nonneg (div_nonneg (gauge_nonneg _) (gauge_nonneg _)), + smul_eq_mul, div_mul_cancel₀ _ hx] + +theorem gauge_gaugeRescale_le (s t : Set E) (x : E) : + gauge t (gaugeRescale s t x) ≤ gauge s x := by + by_cases hx : gauge t x = 0 + · simp [gaugeRescale, hx, gauge_nonneg] + · exact (gauge_gaugeRescale' s hx).le + +variable [TopologicalSpace E] + +section +variable [T1Space E] theorem gaugeRescale_self_apply {s : Set E} (hsa : Absorbent ℝ s) (hsb : IsVonNBounded ℝ s) (x : E) : gaugeRescale s s x = x := by @@ -55,23 +69,12 @@ theorem gaugeRescale_self {s : Set E} (hsa : Absorbent ℝ s) (hsb : IsVonNBound gaugeRescale s s = id := funext <| gaugeRescale_self_apply hsa hsb -theorem gauge_gaugeRescale' (s : Set E) {t : Set E} {x : E} (hx : gauge t x ≠ 0) : - gauge t (gaugeRescale s t x) = gauge s x := by - rw [gaugeRescale, gauge_smul_of_nonneg (div_nonneg (gauge_nonneg _) (gauge_nonneg _)), - smul_eq_mul, div_mul_cancel₀ _ hx] - theorem gauge_gaugeRescale (s : Set E) {t : Set E} (hta : Absorbent ℝ t) (htb : IsVonNBounded ℝ t) (x : E) : gauge t (gaugeRescale s t x) = gauge s x := by rcases eq_or_ne x 0 with rfl | hx · simp · exact gauge_gaugeRescale' s ((gauge_pos hta htb).2 hx).ne' -theorem gauge_gaugeRescale_le (s t : Set E) (x : E) : - gauge t (gaugeRescale s t x) ≤ gauge s x := by - by_cases hx : gauge t x = 0 - · simp [gaugeRescale, hx, gauge_nonneg] - · exact (gauge_gaugeRescale' s hx).le - theorem gaugeRescale_gaugeRescale {s t u : Set E} (hta : Absorbent ℝ t) (htb : IsVonNBounded ℝ t) (x : E) : gaugeRescale t u (gaugeRescale s t x) = gaugeRescale s u x := by rcases eq_or_ne x 0 with rfl | hx; · simp @@ -87,6 +90,8 @@ def gaugeRescaleEquiv (s t : Set E) (hsa : Absorbent ℝ s) (hsb : IsVonNBounded left_inv x := by rw [gaugeRescale_gaugeRescale, gaugeRescale_self_apply] <;> assumption right_inv x := by rw [gaugeRescale_gaugeRescale, gaugeRescale_self_apply] <;> assumption +end + variable [TopologicalAddGroup E] [ContinuousSMul ℝ E] {s t : Set E} theorem mapsTo_gaugeRescale_interior (h₀ : t ∈ 𝓝 0) (hc : Convex ℝ t) : @@ -100,6 +105,8 @@ theorem mapsTo_gaugeRescale_closure {s t : Set E} (hsc : Convex ℝ s) (hs₀ : mem_closure_of_gauge_le_one htc ht₀ hta <| (gauge_gaugeRescale_le _ _ _).trans <| (gauge_le_one_iff_mem_closure hsc hs₀).2 hx +variable [T1Space E] + theorem continuous_gaugeRescale {s t : Set E} (hs : Convex ℝ s) (hs₀ : s ∈ 𝓝 0) (ht : Convex ℝ t) (ht₀ : t ∈ 𝓝 0) (htb : IsVonNBounded ℝ t) : Continuous (gaugeRescale s t) := by diff --git a/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean b/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean index b4320f445e51e..1758106c5c5f0 100644 --- a/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean +++ b/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean @@ -218,7 +218,7 @@ end IsSelfAdjoint section FiniteDimensional -variable [FiniteDimensional 𝕜 E] [_i : Nontrivial E] {T : E →ₗ[𝕜] E} +variable [FiniteDimensional 𝕜 E] {T : E →ₗ[𝕜] E} namespace LinearMap @@ -226,7 +226,7 @@ namespace IsSymmetric /-- The supremum of the Rayleigh quotient of a symmetric operator `T` on a nontrivial finite-dimensional vector space is an eigenvalue for that operator. -/ -theorem hasEigenvalue_iSup_of_finiteDimensional (hT : T.IsSymmetric) : +theorem hasEigenvalue_iSup_of_finiteDimensional [Nontrivial E] (hT : T.IsSymmetric) : HasEigenvalue T ↑(⨆ x : { x : E // x ≠ 0 }, RCLike.re ⟪T x, x⟫ / ‖(x : E)‖ ^ 2 : ℝ) := by haveI := FiniteDimensional.proper_rclike 𝕜 E let T' := hT.toSelfAdjoint @@ -245,7 +245,7 @@ theorem hasEigenvalue_iSup_of_finiteDimensional (hT : T.IsSymmetric) : /-- The infimum of the Rayleigh quotient of a symmetric operator `T` on a nontrivial finite-dimensional vector space is an eigenvalue for that operator. -/ -theorem hasEigenvalue_iInf_of_finiteDimensional (hT : T.IsSymmetric) : +theorem hasEigenvalue_iInf_of_finiteDimensional [Nontrivial E] (hT : T.IsSymmetric) : HasEigenvalue T ↑(⨅ x : { x : E // x ≠ 0 }, RCLike.re ⟪T x, x⟫ / ‖(x : E)‖ ^ 2 : ℝ) := by haveI := FiniteDimensional.proper_rclike 𝕜 E let T' := hT.toSelfAdjoint diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean index 5c7b0a970591d..740851e2901ba 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean @@ -74,7 +74,7 @@ end Ring section Field variable {R V : Type*} [Field R] [AddCommGroup V] [TopologicalSpace R] [TopologicalSpace V] - [TopologicalRing R] [TopologicalAddGroup V] [Module R V] [SeparatingDual R V] + [TopologicalRing R] [Module R V] -- TODO (@alreadydone): this could generalize to CommRing R if we were to add a section theorem _root_.separatingDual_iff_injective : SeparatingDual R V ↔ @@ -84,6 +84,8 @@ theorem _root_.separatingDual_iff_injective : SeparatingDual R V ↔ rw [not_imp_comm, LinearMap.ext_iff] push_neg; rfl +variable [SeparatingDual R V] + open Function in /-- Given a finite-dimensional subspace `W` of a space `V` with separating dual, any linear functional on `W` extends to a continuous linear functional on `V`. @@ -111,6 +113,8 @@ theorem exists_eq_one_ne_zero_of_ne_zero_pair {x y : V} (hx : x ≠ 0) (hy : y · exact ⟨(v x)⁻¹ • v, inv_mul_cancel vx, show (v x)⁻¹ * v y ≠ 0 by simp [vx, vy]⟩ · exact ⟨u + v, by simp [ux, vx], by simp [uy, vy]⟩ +variable [TopologicalAddGroup V] + /-- In a topological vector space with separating dual, the group of continuous linear equivalences acts transitively on the set of nonzero vectors: given two nonzero vectors `x` and `y`, there exists `A : V ≃L[R] V` mapping `x` to `y`. -/ diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index fdbc931fdead9..be48cb8051370 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -536,8 +536,6 @@ end PresheafedSpace namespace SheafedSpace -variable [HasProducts.{v} C] - /-- A family of gluing data consists of 1. An index type `J` 2. A sheafed space `U i` for each `i : J`. diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean index 2d9497effe2a7..b6db3f380021a 100644 --- a/Mathlib/GroupTheory/PushoutI.lean +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -563,7 +563,8 @@ noncomputable def equiv : PushoutI φ ≃ NormalWord d := simp only [prod_smul, prod_empty, mul_one] right_inv := fun w => prod_smul_empty w } -theorem prod_injective : Function.Injective (prod : NormalWord d → PushoutI φ) := by +theorem prod_injective {ι : Type*} {G : ι → Type*} [(i : ι) → Group (G i)] {φ : (i : ι) → H →* G i} + {d : Transversal φ} : Function.Injective (prod : NormalWord d → PushoutI φ) := by letI := Classical.decEq ι letI := fun i => Classical.decEq (G i) classical exact equiv.symm.injective @@ -636,8 +637,8 @@ theorem Reduced.exists_normalWord_prod_eq (d : Transversal φ) {w : Word G} (hw /-- For any word `w` in the coproduct, if `w` is reduced (i.e none its letters are in the image of the base monoid), and nonempty, then `w` itself is not in the image of the base group. -/ -theorem Reduced.eq_empty_of_mem_range (hφ : ∀ i, Injective (φ i)) - {w : Word G} (hw : Reduced φ w) +theorem Reduced.eq_empty_of_mem_range + (hφ : ∀ i, Injective (φ i)) {w : Word G} (hw : Reduced φ w) (h : ofCoprodI w.prod ∈ (base φ).range) : w = .empty := by rcases transversal_nonempty φ hφ with ⟨d⟩ rcases hw.exists_normalWord_prod_eq d with ⟨w', hw'prod, hw'map⟩ @@ -655,7 +656,8 @@ end Reduced /-- The intersection of the images of the maps from any two distinct groups in the diagram into the amalgamated product is the image of the map from the base group in the diagram. -/ -theorem inf_of_range_eq_base_range (hφ : ∀ i, Injective (φ i)) {i j : ι} (hij : i ≠ j) : +theorem inf_of_range_eq_base_range + (hφ : ∀ i, Injective (φ i)) {i j : ι} (hij : i ≠ j) : (of i).range ⊓ (of j).range = (base φ).range := le_antisymm (by diff --git a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean index 76485d7f3fe9f..fb7036771789e 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Orthogonal.lean @@ -290,6 +290,15 @@ lemma ker_restrict_eq_of_codisjoint {p q : Submodule R M} (hpq : Codisjoint p q) · ext ⟨x, hx⟩ simpa using LinearMap.congr_fun h x +lemma inf_orthogonal_self_le_ker_restrict {W : Submodule R M} (b₁ : B.IsRefl) : + W ⊓ B.orthogonal W ≤ (LinearMap.ker <| B.restrict W).map W.subtype := by + rintro v ⟨hv : v ∈ W, hv' : v ∈ B.orthogonal W⟩ + simp only [Submodule.mem_map, mem_ker, restrict_apply, Submodule.coeSubtype, Subtype.exists, + exists_and_left, exists_prop, exists_eq_right_right] + refine ⟨?_, hv⟩ + ext ⟨w, hw⟩ + exact b₁ w v <| hv' w hw + variable [FiniteDimensional K V] open FiniteDimensional Submodule @@ -381,15 +390,6 @@ lemma orthogonal_eq_bot_iff rw [h, eq_bot_iff] exact fun x hx ↦ b₃ x fun y ↦ b₁ y x <| by simpa using hx y -lemma inf_orthogonal_self_le_ker_restrict (b₁ : B.IsRefl) : - W ⊓ B.orthogonal W ≤ (LinearMap.ker <| B.restrict W).map W.subtype := by - rintro v ⟨hv : v ∈ W, hv' : v ∈ B.orthogonal W⟩ - simp only [Submodule.mem_map, mem_ker, restrict_apply, Submodule.coeSubtype, Subtype.exists, - exists_and_left, exists_prop, exists_eq_right_right] - refine ⟨?_, hv⟩ - ext ⟨w, hw⟩ - exact b₁ w v <| hv' w hw - end /-! We note that we cannot use `BilinForm.restrict_nondegenerate_iff_isCompl_orthogonal` for the diff --git a/Mathlib/LinearAlgebra/BilinearForm/TensorProduct.lean b/Mathlib/LinearAlgebra/BilinearForm/TensorProduct.lean index bea71533f32d4..eb6892bb52914 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/TensorProduct.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/TensorProduct.lean @@ -37,7 +37,7 @@ section CommSemiring variable [CommSemiring R] [CommSemiring A] variable [AddCommMonoid M₁] [AddCommMonoid M₂] variable [Algebra R A] [Module R M₁] [Module A M₁] -variable [SMulCommClass R A M₁] [SMulCommClass A R M₁] [IsScalarTower R A M₁] +variable [SMulCommClass R A M₁] [IsScalarTower R A M₁] variable [Module R M₂] variable (R A) in @@ -99,7 +99,6 @@ variable [AddCommGroup M₁] [AddCommGroup M₂] variable [Module R M₁] [Module R M₂] variable [Module.Free R M₁] [Module.Finite R M₁] variable [Module.Free R M₂] [Module.Finite R M₂] -variable [Nontrivial R] variable (R) in /-- `tensorDistrib` as an equivalence. -/ diff --git a/Mathlib/LinearAlgebra/Charpoly/Basic.lean b/Mathlib/LinearAlgebra/Charpoly/Basic.lean index 76fa9b5fafedb..eff887d7c175d 100644 --- a/Mathlib/LinearAlgebra/Charpoly/Basic.lean +++ b/Mathlib/LinearAlgebra/Charpoly/Basic.lean @@ -24,7 +24,7 @@ in any basis is in `LinearAlgebra/Charpoly/ToMatrix`. universe u v w -variable {R : Type u} {M : Type v} [CommRing R] [Nontrivial R] +variable {R : Type u} {M : Type v} [CommRing R] variable [AddCommGroup M] [Module R M] [Module.Free R M] [Module.Finite R M] (f : M →ₗ[R] M) open Matrix Polynomial @@ -52,7 +52,8 @@ theorem charpoly_monic : f.charpoly.Monic := Matrix.charpoly_monic _ open FiniteDimensional in -lemma charpoly_natDegree [StrongRankCondition R] : natDegree (charpoly f) = finrank R M := by +lemma charpoly_natDegree [Nontrivial R] [StrongRankCondition R] : + natDegree (charpoly f) = finrank R M := by rw [charpoly, Matrix.charpoly_natDegree_eq_dim, finrank_eq_card_chooseBasisIndex] end Coeff @@ -88,7 +89,7 @@ theorem pow_eq_aeval_mod_charpoly (k : ℕ) : f ^ k = aeval f (X ^ k %ₘ f.char variable {f} -theorem minpoly_coeff_zero_of_injective (hf : Function.Injective f) : +theorem minpoly_coeff_zero_of_injective [Nontrivial R] (hf : Function.Injective f) : (minpoly R f).coeff 0 ≠ 0 := by intro h obtain ⟨P, hP⟩ := X_dvd_iff.2 h diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean index 5dac845eb2db0..e83ba0a051353 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/BaseChange.lean @@ -32,7 +32,7 @@ We show the additional results: variable {R A V : Type*} variable [CommRing R] [CommRing A] [AddCommGroup V] -variable [Algebra R A] [Module R V] [Module A V] [IsScalarTower R A V] +variable [Algebra R A] [Module R V] variable [Invertible (2 : R)] open scoped TensorProduct diff --git a/Mathlib/LinearAlgebra/Contraction.lean b/Mathlib/LinearAlgebra/Contraction.lean index 03cf2a7e8ff71..8ba1276a30496 100644 --- a/Mathlib/LinearAlgebra/Contraction.lean +++ b/Mathlib/LinearAlgebra/Contraction.lean @@ -201,7 +201,7 @@ section CommRing variable [CommRing R] variable [AddCommGroup M] [AddCommGroup N] [AddCommGroup P] [AddCommGroup Q] variable [Module R M] [Module R N] [Module R P] [Module R Q] -variable [Free R M] [Finite R M] [Free R N] [Finite R N] [Nontrivial R] +variable [Free R M] [Finite R M] [Free R N] [Finite R N] /-- When `M` is a finite free module, the map `lTensorHomToHomLTensor` is an equivalence. Note that `lTensorHomEquivHomLTensor` is not defined directly in terms of diff --git a/Mathlib/LinearAlgebra/Orientation.lean b/Mathlib/LinearAlgebra/Orientation.lean index b95238a255fdf..789d425dda083 100644 --- a/Mathlib/LinearAlgebra/Orientation.lean +++ b/Mathlib/LinearAlgebra/Orientation.lean @@ -324,14 +324,14 @@ variable {ι : Type*} namespace Orientation -variable [Fintype ι] [_i : FiniteDimensional R M] +variable [Fintype ι] open FiniteDimensional /-- If the index type has cardinality equal to the finite dimension, any two orientations are equal or negations. -/ -theorem eq_or_eq_neg (x₁ x₂ : Orientation R M ι) (h : Fintype.card ι = finrank R M) : - x₁ = x₂ ∨ x₁ = -x₂ := by +theorem eq_or_eq_neg [FiniteDimensional R M] (x₁ x₂ : Orientation R M ι) + (h : Fintype.card ι = finrank R M) : x₁ = x₂ ∨ x₁ = -x₂ := by have e := (finBasis R M).reindex (Fintype.equivFinOfCardEq h).symm letI := Classical.decEq ι -- Porting note: this needs to be made explicit for the simp below @@ -345,14 +345,14 @@ theorem eq_or_eq_neg (x₁ x₂ : Orientation R M ι) (h : Fintype.card ι = fin /-- If the index type has cardinality equal to the finite dimension, an orientation equals the negation of another orientation if and only if they are not equal. -/ -theorem ne_iff_eq_neg (x₁ x₂ : Orientation R M ι) (h : Fintype.card ι = finrank R M) : - x₁ ≠ x₂ ↔ x₁ = -x₂ := +theorem ne_iff_eq_neg [FiniteDimensional R M] (x₁ x₂ : Orientation R M ι) + (h : Fintype.card ι = finrank R M) : x₁ ≠ x₂ ↔ x₁ = -x₂ := ⟨fun hn => (eq_or_eq_neg x₁ x₂ h).resolve_left hn, fun he => he.symm ▸ (Module.Ray.ne_neg_self x₂).symm⟩ /-- The value of `Orientation.map` when the index type has cardinality equal to the finite dimension, in terms of `f.det`. -/ -theorem map_eq_det_inv_smul (x : Orientation R M ι) (f : M ≃ₗ[R] M) +theorem map_eq_det_inv_smul [FiniteDimensional R M] (x : Orientation R M ι) (f : M ≃ₗ[R] M) (h : Fintype.card ι = finrank R M) : Orientation.map ι f x = (LinearEquiv.det f)⁻¹ • x := haveI e := (finBasis R M).reindex (Fintype.equivFinOfCardEq h).symm e.map_orientation_eq_det_inv_smul x f @@ -360,7 +360,7 @@ theorem map_eq_det_inv_smul (x : Orientation R M ι) (f : M ≃ₗ[R] M) /-- If the index type has cardinality equal to the finite dimension, composing an alternating map with the same linear equiv on each argument gives the same orientation if and only if the determinant is positive. -/ -theorem map_eq_iff_det_pos (x : Orientation R M ι) (f : M ≃ₗ[R] M) +theorem map_eq_iff_det_pos [FiniteDimensional R M] (x : Orientation R M ι) (f : M ≃ₗ[R] M) (h : Fintype.card ι = finrank R M) : Orientation.map ι f x = x ↔ 0 < LinearMap.det (f : M →ₗ[R] M) := by cases isEmpty_or_nonempty ι @@ -385,14 +385,14 @@ theorem map_eq_neg_iff_det_neg (x : Orientation R M ι) (f : M ≃ₗ[R] M) /-- If the index type has cardinality equal to the finite dimension, a basis with the given orientation. -/ -def someBasis [Nonempty ι] [DecidableEq ι] (x : Orientation R M ι) +def someBasis [Nonempty ι] [DecidableEq ι] [FiniteDimensional R M] (x : Orientation R M ι) (h : Fintype.card ι = finrank R M) : Basis ι R M := ((finBasis R M).reindex (Fintype.equivFinOfCardEq h).symm).adjustToOrientation x /-- `some_basis` gives a basis with the required orientation. -/ @[simp] -theorem someBasis_orientation [Nonempty ι] [DecidableEq ι] (x : Orientation R M ι) - (h : Fintype.card ι = finrank R M) : (x.someBasis h).orientation = x := +theorem someBasis_orientation [Nonempty ι] [DecidableEq ι] [FiniteDimensional R M] + (x : Orientation R M ι) (h : Fintype.card ι = finrank R M) : (x.someBasis h).orientation = x := Basis.orientation_adjustToOrientation _ _ end Orientation diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean index 4aaefe468fd08..b66de04b692fd 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean @@ -53,7 +53,7 @@ suppress_compilation open scoped TensorProduct variable {R ι A B : Type*} -variable [CommSemiring ι] [Module ι (Additive ℤˣ)] [DecidableEq ι] +variable [CommSemiring ι] [DecidableEq ι] variable [CommRing R] [Ring A] [Ring B] [Algebra R A] [Algebra R B] variable (𝒜 : ι → Submodule R A) (ℬ : ι → Submodule R B) variable [GradedAlgebra 𝒜] [GradedAlgebra ℬ] @@ -135,6 +135,8 @@ theorem auxEquiv_one : auxEquiv R 𝒜 ℬ 1 = 1 := by theorem auxEquiv_symm_one : (auxEquiv R 𝒜 ℬ).symm 1 = 1 := (LinearEquiv.symm_apply_eq _).mpr (auxEquiv_one _ _).symm +variable [Module ι (Additive ℤˣ)] + /-- Auxiliary construction used to build the `Mul` instance and get distributivity of `+` and `\smul`. -/ noncomputable def mulHom : (𝒜 ᵍ⊗[R] ℬ) →ₗ[R] (𝒜 ᵍ⊗[R] ℬ) →ₗ[R] (𝒜 ᵍ⊗[R] ℬ) := by diff --git a/Mathlib/MeasureTheory/Function/AEEqFun.lean b/Mathlib/MeasureTheory/Function/AEEqFun.lean index a95d9cfc9f057..4a391c3dc6ad6 100644 --- a/Mathlib/MeasureTheory/Function/AEEqFun.lean +++ b/Mathlib/MeasureTheory/Function/AEEqFun.lean @@ -101,9 +101,12 @@ notation:25 α " →ₘ[" μ "] " β => AEEqFun α β μ end MeasurableSpace +variable [TopologicalSpace δ] + namespace AEEqFun -variable [TopologicalSpace β] [TopologicalSpace γ] [TopologicalSpace δ] +section +variable [TopologicalSpace β] /-- Construct the equivalence class `[f]` of an almost everywhere measurable function `f`, based on the equivalence relation of being almost everywhere equal. -/ @@ -183,13 +186,15 @@ theorem induction_on₃ {α' β' : Type*} [MeasurableSpace α'] [TopologicalSpac (H : ∀ f hf f' hf' f'' hf'', p (mk f hf) (mk f' hf') (mk f'' hf'')) : p f f' f'' := induction_on f fun f hf => induction_on₂ f' f'' <| H f hf +end + /-! ### Composition of an a.e. equal function with a (quasi) measure preserving function -/ section compQuasiMeasurePreserving -variable [MeasurableSpace β] {ν : MeasureTheory.Measure β} {f : α → β} +variable [TopologicalSpace γ] [MeasurableSpace β] {ν : MeasureTheory.Measure β} {f : α → β} open MeasureTheory.Measure (QuasiMeasurePreserving) @@ -221,7 +226,8 @@ end compQuasiMeasurePreserving section compMeasurePreserving -variable [MeasurableSpace β] {ν : MeasureTheory.Measure β} {f : α → β} {g : β → γ} +variable [TopologicalSpace γ] [MeasurableSpace β] {ν : MeasureTheory.Measure β} + {f : α → β} {g : β → γ} /-- Composition of an almost everywhere equal function and a quasi measure preserving function. @@ -247,6 +253,8 @@ theorem coeFn_compMeasurePreserving (g : β →ₘ[ν] γ) (hf : MeasurePreservi end compMeasurePreserving +variable [TopologicalSpace β] [TopologicalSpace γ] + /-- Given a continuous function `g : β → γ`, and an almost everywhere equal function `[f] : α →ₘ β`, return the equivalence class of `g ∘ f`, i.e., the almost everywhere equal function `[g ∘ f] : α →ₘ γ`. -/ @@ -267,7 +275,8 @@ theorem coeFn_comp (g : β → γ) (hg : Continuous g) (f : α →ₘ[μ] β) : rw [comp_eq_mk] apply coeFn_mk -theorem comp_compQuasiMeasurePreserving [MeasurableSpace β] {ν} (g : γ → δ) (hg : Continuous g) +theorem comp_compQuasiMeasurePreserving + {β : Type*} [MeasurableSpace β] {ν} (g : γ → δ) (hg : Continuous g) (f : β →ₘ[ν] γ) {φ : α → β} (hφ : Measure.QuasiMeasurePreserving φ μ ν) : (comp g hg f).compQuasiMeasurePreserving φ hφ = comp g hg (f.compQuasiMeasurePreserving φ hφ) := by @@ -408,13 +417,13 @@ theorem toGerm_injective : Injective (toGerm : (α →ₘ[μ] β) → Germ (ae ext <| Germ.coe_eq.1 <| by rwa [← toGerm_eq, ← toGerm_eq] @[simp] -theorem compQuasiMeasurePreserving_toGerm [MeasurableSpace β] {f : α → β} {ν} +theorem compQuasiMeasurePreserving_toGerm {β : Type*} [MeasurableSpace β] {f : α → β} {ν} (g : β →ₘ[ν] γ) (hf : Measure.QuasiMeasurePreserving f μ ν) : (g.compQuasiMeasurePreserving f hf).toGerm = g.toGerm.compTendsto f hf.tendsto_ae := by rcases g; rfl @[simp] -theorem compMeasurePreserving_toGerm [MeasurableSpace β] {f : α → β} {ν} +theorem compMeasurePreserving_toGerm {β : Type*} [MeasurableSpace β] {f : α → β} {ν} (g : β →ₘ[ν] γ) (hf : MeasurePreserving f μ ν) : (g.compMeasurePreserving f hf).toGerm = g.toGerm.compTendsto f hf.quasiMeasurePreserving.tendsto_ae := diff --git a/Mathlib/MeasureTheory/Function/SpecialFunctions/Arctan.lean b/Mathlib/MeasureTheory/Function/SpecialFunctions/Arctan.lean index 517c91b2394f6..15eda1c1f1ac0 100644 --- a/Mathlib/MeasureTheory/Function/SpecialFunctions/Arctan.lean +++ b/Mathlib/MeasureTheory/Function/SpecialFunctions/Arctan.lean @@ -24,10 +24,10 @@ section RealComposition open Real -variable {α : Type*} {m : MeasurableSpace α} {f : α → ℝ} (hf : Measurable f) +variable {α : Type*} {m : MeasurableSpace α} {f : α → ℝ} @[measurability] -theorem Measurable.arctan : Measurable fun x => arctan (f x) := +theorem Measurable.arctan (hf : Measurable f) : Measurable fun x => arctan (f x) := measurable_arctan.comp hf end RealComposition diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index bd086669e7f0d..a7db0fcab3509 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -2041,7 +2041,7 @@ lemma tendsto_measure_of_ae_tendsto_indicator {μ : Measure α} (A_mble : Measur /-- If `μ` is a finite measure and the indicators of measurable sets `Aᵢ` tend pointwise almost everywhere to the indicator of a measurable set `A`, then the measures `μ Aᵢ` tend to the measure `μ A`. -/ -lemma tendsto_measure_of_ae_tendsto_indicator_of_isFiniteMeasure [IsCountablyGenerated L] +lemma tendsto_measure_of_ae_tendsto_indicator_of_isFiniteMeasure {μ : Measure α} [IsFiniteMeasure μ] (A_mble : MeasurableSet A) (As_mble : ∀ i, MeasurableSet (As i)) (h_lim : ∀ᵐ x ∂μ, ∀ᶠ i in L, x ∈ As i ↔ x ∈ A) : Tendsto (fun i ↦ μ (As i)) L (𝓝 (μ A)) := diff --git a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean index 7f07b91e54462..c1459922a78e9 100644 --- a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean @@ -39,7 +39,7 @@ theorem aemeasurable_id'' (μ : Measure α) {m : MeasurableSpace α} (hm : m ≤ @AEMeasurable α α m m0 id μ := @Measurable.aemeasurable α α m0 m id μ (measurable_id'' hm) -lemma aemeasurable_of_map_neZero {mβ : MeasurableSpace β} {μ : Measure α} +lemma aemeasurable_of_map_neZero {μ : Measure α} {f : α → β} (h : NeZero (μ.map f)) : AEMeasurable f μ := by by_contra h' diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean index 61c419b87e67f..3c9f822f568c8 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean @@ -265,7 +265,8 @@ disjoint unions, then the predicate holds for almost every `x : β` and all meas This is an AE version of `MeasurableSpace.induction_on_inter` where the condition is dependent on a measurable space `β`. -/ -theorem _root_.MeasurableSpace.ae_induction_on_inter {β} [MeasurableSpace β] {μ : Measure β} +theorem _root_.MeasurableSpace.ae_induction_on_inter + {α β : Type*} [MeasurableSpace β] {μ : Measure β} {C : β → Set α → Prop} {s : Set (Set α)} [m : MeasurableSpace α] (h_eq : m = MeasurableSpace.generateFrom s) (h_inter : IsPiSystem s) (h_empty : ∀ᵐ x ∂μ, C x ∅) (h_basic : ∀ᵐ x ∂μ, ∀ t ∈ s, C x t) diff --git a/Mathlib/MeasureTheory/Measure/OpenPos.lean b/Mathlib/MeasureTheory/Measure/OpenPos.lean index 69f77c4c06a0b..611930fcce591 100644 --- a/Mathlib/MeasureTheory/Measure/OpenPos.lean +++ b/Mathlib/MeasureTheory/Measure/OpenPos.lean @@ -238,7 +238,7 @@ However, with respect to a measure which is positive on non-empty open sets, *cl zero sets are nowhere dense and σ-compact measure zero sets in a Hausdorff space are meagre. -/ -variable {X : Type*} [TopologicalSpace X] [MeasurableSpace X] [BorelSpace X] {s : Set X} +variable {X : Type*} [TopologicalSpace X] [MeasurableSpace X] {s : Set X} {μ : Measure X} [IsOpenPosMeasure μ] /-- A *closed* measure zero subset is nowhere dense. (Closedness is required: for instance, the diff --git a/Mathlib/MeasureTheory/Order/Group/Lattice.lean b/Mathlib/MeasureTheory/Order/Group/Lattice.lean index e19df239fc3aa..da3f3746b9d09 100644 --- a/Mathlib/MeasureTheory/Order/Group/Lattice.lean +++ b/Mathlib/MeasureTheory/Order/Group/Lattice.lean @@ -16,26 +16,26 @@ measurable function, group, lattice operation -/ variable {α β : Type*} [Lattice α] [Group α] [MeasurableSpace α] - [MeasurableSpace β] {f : β → α} (hf : Measurable f) - -variable [MeasurableSup α] + [MeasurableSpace β] {f : β → α} @[to_additive (attr := measurability)] -theorem measurable_oneLePart : Measurable (oneLePart : α → α) := +theorem measurable_oneLePart [MeasurableSup α] : Measurable (oneLePart : α → α) := measurable_sup_const _ @[to_additive (attr := measurability)] -protected theorem Measurable.oneLePart : Measurable fun x ↦ oneLePart (f x) := +protected theorem Measurable.oneLePart [MeasurableSup α] (hf : Measurable f) : + Measurable fun x ↦ oneLePart (f x) := measurable_oneLePart.comp hf variable [MeasurableInv α] @[to_additive (attr := measurability)] -theorem measurable_leOnePart : Measurable (leOnePart : α → α) := +theorem measurable_leOnePart [MeasurableSup α] : Measurable (leOnePart : α → α) := (measurable_sup_const _).comp measurable_inv @[to_additive (attr := measurability)] -protected theorem Measurable.leOnePart : Measurable fun x ↦ leOnePart (f x) := +protected theorem Measurable.leOnePart [MeasurableSup α] (hf : Measurable f) : + Measurable fun x ↦ leOnePart (f x) := measurable_leOnePart.comp hf variable [MeasurableSup₂ α] @@ -45,5 +45,5 @@ theorem measurable_mabs : Measurable (mabs : α → α) := measurable_id'.sup measurable_inv @[to_additive (attr := measurability)] -protected theorem Measurable.mabs : Measurable fun x ↦ mabs (f x) := +protected theorem Measurable.mabs (hf : Measurable f) : Measurable fun x ↦ mabs (f x) := measurable_mabs.comp hf diff --git a/Mathlib/NumberTheory/Rayleigh.lean b/Mathlib/NumberTheory/Rayleigh.lean index 0aac719cef568..efa4c9a35de90 100644 --- a/Mathlib/NumberTheory/Rayleigh.lean +++ b/Mathlib/NumberTheory/Rayleigh.lean @@ -55,11 +55,12 @@ noncomputable def beattySeq' (r : ℝ) : ℤ → ℤ := namespace Beatty -variable {r s : ℝ} (hrs : r.IsConjExponent s) {j k : ℤ} +variable {r s : ℝ} {j k : ℤ} /-- Let `r > 1` and `1/r + 1/s = 1`. Then `B_r` and `B'_s` are disjoint (i.e. no collision exists). -/ -private theorem no_collision : Disjoint {beattySeq r k | k} {beattySeq' s k | k} := by +private theorem no_collision (hrs : r.IsConjExponent s) : + Disjoint {beattySeq r k | k} {beattySeq' s k | k} := by rw [Set.disjoint_left] intro j ⟨k, h₁⟩ ⟨m, h₂⟩ rw [beattySeq, Int.floor_eq_iff, ← div_le_iff hrs.pos, ← lt_div_iff hrs.pos] at h₁ @@ -74,7 +75,7 @@ private theorem no_collision : Disjoint {beattySeq r k | k} {beattySeq' s k | k} /-- Let `r > 1` and `1/r + 1/s = 1`. Suppose there is an integer `j` where `B_r` and `B'_s` both jump over `j` (i.e. an anti-collision). Then this leads to a contradiction. -/ -private theorem no_anticollision : +private theorem no_anticollision (hrs : r.IsConjExponent s) : ¬∃ j k m : ℤ, k < j / r ∧ (j + 1) / r ≤ k + 1 ∧ m ≤ j / s ∧ (j + 1) / s < m + 1 := by intro ⟨j, k, m, h₁₁, h₁₂, h₂₁, h₂₂⟩ have h₃ := add_lt_add_of_lt_of_le h₁₁ h₂₁ diff --git a/Mathlib/NumberTheory/SiegelsLemma.lean b/Mathlib/NumberTheory/SiegelsLemma.lean index 957448cbd13d6..998f8940d66b9 100644 --- a/Mathlib/NumberTheory/SiegelsLemma.lean +++ b/Mathlib/NumberTheory/SiegelsLemma.lean @@ -39,7 +39,7 @@ open Matrix Finset namespace Int.Matrix variable {α β : Type*} [Fintype α] [Fintype β] - (A : Matrix α β ℤ) (v : β → ℤ) (hn : Fintype.card α < Fintype.card β) (hm : 0 < Fintype.card α) + (A : Matrix α β ℤ) (v : β → ℤ) -- Some definitions and relative properties @@ -58,8 +58,6 @@ local notation3 "S" => Finset.Icc N P section preparation -variable [DecidableEq α] [DecidableEq β] - /- In order to apply Pigeonhole we need: # Step 1: ∀ v ∈ T, A *ᵥ v ∈ S and @@ -70,7 +68,7 @@ Their difference is the solution we are looking for -- # Step 1: ∀ v ∈ T, A *ᵥ v ∈ S -private lemma image_T_subset_S (v) (hv : v ∈ T) : A *ᵥ v ∈ S := by +private lemma image_T_subset_S [DecidableEq α] [DecidableEq β] (v) (hv : v ∈ T) : A *ᵥ v ∈ S := by rw [mem_Icc] at hv ⊢ have mulVec_def : A.mulVec v = fun i ↦ Finset.sum univ fun j : β ↦ A i j * v j := rfl @@ -93,7 +91,7 @@ private lemma image_T_subset_S (v) (hv : v ∈ T) : A *ᵥ v ∈ S := by -- # Preparation for Step 2 -private lemma card_T_eq : (T).card = (B + 1) ^ n := by +private lemma card_T_eq [DecidableEq β] : (T).card = (B + 1) ^ n := by rw [Pi.card_Icc 0 B'] simp only [Pi.zero_apply, card_Icc, sub_zero, toNat_ofNat_add_one, prod_const, card_univ, add_pos_iff, zero_lt_one, or_true] @@ -111,7 +109,7 @@ private lemma N_le_P_add_one (i : α) : N i ≤ P i + 1 := by intro j _ exact mul_nonneg (Nat.cast_nonneg B) (posPart_nonneg (A i j)) -private lemma card_S_eq : (Finset.Icc N P).card = ∏ i : α, (P i - N i + 1) := by +private lemma card_S_eq [DecidableEq α] : (Finset.Icc N P).card = ∏ i : α, (P i - N i + 1) := by rw [Pi.card_Icc N P, Nat.cast_prod] congr ext i @@ -134,7 +132,9 @@ lemma one_le_norm_A_of_ne_zero (hA : A ≠ 0) : 1 ≤ ‖A‖ := by open Real Nat -private lemma card_S_lt_card_T : (S).card < (T).card := by +private lemma card_S_lt_card_T [DecidableEq α] [DecidableEq β] + (hn : Fintype.card α < Fintype.card β) (hm : 0 < Fintype.card α) : + (S).card < (T).card := by zify -- This is necessary to use card_S_eq rw [card_T_eq A, card_S_eq] rify -- This is necessary because ‖A‖ is a real number @@ -173,7 +173,8 @@ private lemma card_S_lt_card_T : (S).card < (T).card := by end preparation -theorem exists_ne_zero_int_vec_norm_le : ∃ t : β → ℤ, t ≠ 0 ∧ +theorem exists_ne_zero_int_vec_norm_le + (hn : Fintype.card α < Fintype.card β) (hm : 0 < Fintype.card α) : ∃ t : β → ℤ, t ≠ 0 ∧ A *ᵥ t = 0 ∧ ‖t‖ ≤ (n * max 1 ‖A‖) ^ ((m : ℝ) / (n - m)) := by classical -- Pigeonhole @@ -205,7 +206,9 @@ theorem exists_ne_zero_int_vec_norm_le : ∃ t : β → ℤ, t ≠ 0 ∧ exact hyT.1 i -theorem exists_ne_zero_int_vec_norm_le' (hA : A ≠ 0) : ∃ t : β → ℤ, t ≠ 0 ∧ +theorem exists_ne_zero_int_vec_norm_le' + (hn : Fintype.card α < Fintype.card β) (hm : 0 < Fintype.card α) (hA : A ≠ 0) : + ∃ t : β → ℤ, t ≠ 0 ∧ A *ᵥ t = 0 ∧ ‖t‖ ≤ (n * ‖A‖) ^ ((m : ℝ) / (n - m)) := by have := exists_ne_zero_int_vec_norm_le A hn hm rwa [max_eq_right] at this diff --git a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean index 8455083320dad..671d719a25adf 100644 --- a/Mathlib/RingTheory/AdicCompletion/Functoriality.lean +++ b/Mathlib/RingTheory/AdicCompletion/Functoriality.lean @@ -225,7 +225,7 @@ inverse to each other. -/ -variable {ι : Type*} [DecidableEq ι] (M : ι → Type*) [∀ i, AddCommGroup (M i)] +variable {ι : Type*} (M : ι → Type*) [∀ i, AddCommGroup (M i)] [∀ i, Module R (M i)] section Pi @@ -244,19 +244,19 @@ open DirectSum /-- The canonical map from the sum of the adic completions to the adic completion of the sum. -/ -def sum : +def sum [DecidableEq ι] : (⨁ j, (AdicCompletion I (M j))) →ₗ[AdicCompletion I R] AdicCompletion I (⨁ j, M j) := toModule (AdicCompletion I R) ι (AdicCompletion I (⨁ j, M j)) (fun j ↦ map I (lof R ι M j)) @[simp] -theorem sum_lof (j : ι) (x : AdicCompletion I (M j)) : +theorem sum_lof [DecidableEq ι] (j : ι) (x : AdicCompletion I (M j)) : sum I M ((DirectSum.lof (AdicCompletion I R) ι (fun i ↦ AdicCompletion I (M i)) j) x) = map I (lof R ι M j) x := by simp [sum] @[simp] -theorem sum_of (j : ι) (x : AdicCompletion I (M j)) : +theorem sum_of [DecidableEq ι] (j : ι) (x : AdicCompletion I (M j)) : sum I M ((DirectSum.of (fun i ↦ AdicCompletion I (M i)) j) x) = map I (lof R ι M j) x := by rw [← lof_eq_of R] @@ -284,6 +284,8 @@ theorem sumInv_apply (x : AdicCompletion I (⨁ j, M j)) (j : ι) : apply induction_on I _ x (fun x ↦ ?_) rfl +variable [DecidableEq ι] + theorem sumInv_comp_sum : sumInv I M ∘ₗ sum I M = LinearMap.id := by ext j x apply DirectSum.ext (AdicCompletion I R) (fun i ↦ ?_) @@ -326,7 +328,7 @@ section Pi open DirectSum -variable [Fintype ι] +variable [DecidableEq ι] [Fintype ι] /-- If `ι` is finite, `pi` is a linear equiv. -/ def piEquivOfFintype : diff --git a/Mathlib/RingTheory/Adjoin/Field.lean b/Mathlib/RingTheory/Adjoin/Field.lean index 69adb77259598..eb204b8922d11 100644 --- a/Mathlib/RingTheory/Adjoin/Field.lean +++ b/Mathlib/RingTheory/Adjoin/Field.lean @@ -79,28 +79,36 @@ theorem Polynomial.lift_of_splits {F K L : Type*} [Field F] [Field K] [Field L] end Embeddings -variable {R K L M : Type*} [CommRing R] [Field K] [Field L] [CommRing M] [Algebra R K] [Algebra R L] - [Algebra R M] {x : L} (int : IsIntegral R x) (h : Splits (algebraMap R K) (minpoly R x)) +variable {R K L M : Type*} [CommRing R] [Field K] [Field L] [CommRing M] [Algebra R K] + [Algebra R M] {x : L} -theorem IsIntegral.mem_range_algHom_of_minpoly_splits (f : K →ₐ[R] L) : x ∈ f.range := +section +variable [Algebra R L] + +theorem IsIntegral.mem_range_algHom_of_minpoly_splits + (int : IsIntegral R x) (h : Splits (algebraMap R K) (minpoly R x))(f : K →ₐ[R] L) : + x ∈ f.range := show x ∈ Set.range f from Set.image_subset_range _ _ <| by rw [image_rootSet h f, mem_rootSet'] exact ⟨((minpoly.monic int).map _).ne_zero, minpoly.aeval R x⟩ -theorem IsIntegral.mem_range_algebraMap_of_minpoly_splits [Algebra K L] [IsScalarTower R K L] : +theorem IsIntegral.mem_range_algebraMap_of_minpoly_splits [Algebra K L] [IsScalarTower R K L] + (int : IsIntegral R x) (h : Splits (algebraMap R K) (minpoly R x)) : x ∈ (algebraMap K L).range := int.mem_range_algHom_of_minpoly_splits h (IsScalarTower.toAlgHom R K L) -variable [Algebra K M] [IsScalarTower R K M] {x : M} (int : IsIntegral R x) +end + +variable [Algebra K M] [IsScalarTower R K M] {x : M} -theorem IsIntegral.minpoly_splits_tower_top' {f : K →+* L} +theorem IsIntegral.minpoly_splits_tower_top' (int : IsIntegral R x) {f : K →+* L} (h : Splits (f.comp <| algebraMap R K) (minpoly R x)) : Splits f (minpoly K x) := splits_of_splits_of_dvd _ ((minpoly.monic int).map _).ne_zero ((splits_map_iff _ _).mpr h) (minpoly.dvd_map_of_isScalarTower R _ x) -theorem IsIntegral.minpoly_splits_tower_top [Algebra K L] [IsScalarTower R K L] - (h : Splits (algebraMap R L) (minpoly R x)) : +theorem IsIntegral.minpoly_splits_tower_top [Algebra K L] [Algebra R L] [IsScalarTower R K L] + (int : IsIntegral R x) (h : Splits (algebraMap R L) (minpoly R x)) : Splits (algebraMap K L) (minpoly K x) := by rw [IsScalarTower.algebraMap_eq R K L] at h exact int.minpoly_splits_tower_top' h diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index 6b435db2667aa..e72d6a31ef4c3 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -306,9 +306,9 @@ section NoZeroSMulDivisors namespace Algebra.IsAlgebraic variable [CommRing K] [Field L] -variable [Algebra K L] [NoZeroSMulDivisors K L] +variable [Algebra K L] -theorem algHom_bijective [Algebra.IsAlgebraic K L] (f : L →ₐ[K] L) : +theorem algHom_bijective [NoZeroSMulDivisors K L] [Algebra.IsAlgebraic K L] (f : L →ₐ[K] L) : Function.Bijective f := by refine ⟨f.injective, fun b ↦ ?_⟩ obtain ⟨p, hp, he⟩ := Algebra.IsAlgebraic.isAlgebraic (R := K) b @@ -318,12 +318,12 @@ theorem algHom_bijective [Algebra.IsAlgebraic K L] (f : L →ₐ[K] L) : obtain ⟨a, ha⟩ := this ⟨b, mem_rootSet.2 ⟨hp, he⟩⟩ exact ⟨a, Subtype.ext_iff.1 ha⟩ -theorem algHom_bijective₂ [Field R] [Algebra K R] +theorem algHom_bijective₂ [NoZeroSMulDivisors K L] [Field R] [Algebra K R] [Algebra.IsAlgebraic K L] (f : L →ₐ[K] R) (g : R →ₐ[K] L) : Function.Bijective f ∧ Function.Bijective g := (g.injective.bijective₂_of_surjective f.injective (algHom_bijective <| g.comp f).2).symm -theorem bijective_of_isScalarTower [Algebra.IsAlgebraic K L] +theorem bijective_of_isScalarTower [NoZeroSMulDivisors K L] [Algebra.IsAlgebraic K L] [Field R] [Algebra K R] [Algebra L R] [IsScalarTower K L R] (f : R →ₐ[K] L) : Function.Bijective f := (algHom_bijective₂ (IsScalarTower.toAlgHom K L R) f).2 @@ -338,7 +338,7 @@ variable (K L) /-- Bijection between algebra equivalences and algebra homomorphisms -/ @[simps] -noncomputable def algEquivEquivAlgHom [Algebra.IsAlgebraic K L] : +noncomputable def algEquivEquivAlgHom [NoZeroSMulDivisors K L] [Algebra.IsAlgebraic K L] : (L ≃ₐ[K] L) ≃* (L →ₐ[K] L) where toFun ϕ := ϕ.toAlgHom invFun ϕ := AlgEquiv.ofBijective ϕ (algHom_bijective ϕ) diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian.lean index c56c93f43f6ad..916bb9144d3c4 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian.lean @@ -531,6 +531,7 @@ instance [IsReduced R] : DecompositionMonoid (Polynomial R) := theorem isSemisimpleRing_of_isReduced [IsReduced R] : IsSemisimpleRing R := (equivPi R).symm.isSemisimpleRing -proof_wanted IsSemisimpleRing.isArtinianRing [IsSemisimpleRing R] : IsArtinianRing R +proof_wanted IsSemisimpleRing.isArtinianRing (R : Type*) [CommRing R] [IsSemisimpleRing R] : + IsArtinianRing R end IsArtinianRing diff --git a/Mathlib/RingTheory/EssentialFiniteness.lean b/Mathlib/RingTheory/EssentialFiniteness.lean index 1ccf45d8121b9..65b40d237017c 100644 --- a/Mathlib/RingTheory/EssentialFiniteness.lean +++ b/Mathlib/RingTheory/EssentialFiniteness.lean @@ -22,7 +22,7 @@ open scoped TensorProduct namespace Algebra variable (R S T : Type*) [CommRing R] [CommRing S] [CommRing T] -variable [Algebra R S] [Algebra R T] [Algebra S T] [IsScalarTower R S T] +variable [Algebra R S] [Algebra R T] /-- An `R`-algebra is essentially of finite type if @@ -102,6 +102,9 @@ lemma EssFiniteType.of_isLocalization (M : Submonoid R) [IsLocalization M S] : lemma EssFiniteType.of_id : EssFiniteType R R := inferInstance +section +variable [Algebra S T] [IsScalarTower R S T] + lemma EssFiniteType.aux (σ : Subalgebra R S) (hσ : ∀ s : S, ∃ t ∈ σ, IsUnit t ∧ s * t ∈ σ) (τ : Set T) (t : T) (ht : t ∈ Algebra.adjoin S τ) : @@ -198,6 +201,8 @@ lemma EssFiniteType.comp_iff [EssFiniteType R S] : EssFiniteType R T ↔ EssFiniteType S T := ⟨fun _ ↦ of_comp R S T, fun _ ↦ comp R S T⟩ +end + variable {R S} in lemma EssFiniteType.algHom_ext [EssFiniteType R S] (f g : S →ₐ[R] T) (H : ∀ s ∈ finset R S, f s = g s) : f = g := by diff --git a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean index 92836e01ac84a..c84cfb1553f9f 100644 --- a/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/Algebra/Basic.lean @@ -108,7 +108,7 @@ nonrec theorem IsIntegral.mul {x y : A} (hx : IsIntegral R x) (hy : IsIntegral R IsIntegral R (x * y) := hx.mul (algebraMap R A) hy -theorem IsIntegral.smul {R} [CommSemiring R] [CommRing S] [Algebra R B] [Algebra S B] [Algebra R S] +theorem IsIntegral.smul {R} [CommSemiring R] [Algebra R B] [Algebra S B] [Algebra R S] [IsScalarTower R S B] {x : B} (r : R)(hx : IsIntegral S x) : IsIntegral S (r • x) := .of_mem_of_fg _ hx.fg_adjoin_singleton _ <| by rw [← algebraMap_smul S]; apply Subalgebra.smul_mem; exact Algebra.subset_adjoin rfl diff --git a/Mathlib/RingTheory/IntegralClosure/IntegrallyClosed.lean b/Mathlib/RingTheory/IntegralClosure/IntegrallyClosed.lean index 8e6c66bbbd17b..aa5706c755a4c 100644 --- a/Mathlib/RingTheory/IntegralClosure/IntegrallyClosed.lean +++ b/Mathlib/RingTheory/IntegralClosure/IntegrallyClosed.lean @@ -126,15 +126,18 @@ end Iff namespace IsIntegrallyClosedIn -variable {R A : Type*} [CommRing R] [CommRing A] [Algebra R A] [iic : IsIntegrallyClosedIn R A] +variable {R A : Type*} [CommRing R] [CommRing A] [Algebra R A] -theorem algebraMap_eq_of_integral {x : A} : IsIntegral R x → ∃ y : R, algebraMap R A y = x := +theorem algebraMap_eq_of_integral [IsIntegrallyClosedIn R A] {x : A} : + IsIntegral R x → ∃ y : R, algebraMap R A y = x := IsIntegralClosure.isIntegral_iff.mp -theorem isIntegral_iff {x : A} : IsIntegral R x ↔ ∃ y : R, algebraMap R A y = x := +theorem isIntegral_iff [IsIntegrallyClosedIn R A] {x : A} : + IsIntegral R x ↔ ∃ y : R, algebraMap R A y = x := IsIntegralClosure.isIntegral_iff -theorem exists_algebraMap_eq_of_isIntegral_pow {x : A} {n : ℕ} (hn : 0 < n) +theorem exists_algebraMap_eq_of_isIntegral_pow [IsIntegrallyClosedIn R A] + {x : A} {n : ℕ} (hn : 0 < n) (hx : IsIntegral R <| x ^ n) : ∃ y : R, algebraMap R A y = x := isIntegral_iff.mp <| hx.of_pow hn @@ -160,7 +163,8 @@ theorem integralClosure_eq_bot_iff (hRA : Function.Injective (algebraMap R A)) : variable (R) @[simp] -theorem integralClosure_eq_bot [NoZeroSMulDivisors R A] [Nontrivial A] : integralClosure R A = ⊥ := +theorem integralClosure_eq_bot [IsIntegrallyClosedIn R A] [NoZeroSMulDivisors R A] [Nontrivial A] : + integralClosure R A = ⊥ := (integralClosure_eq_bot_iff A (NoZeroSMulDivisors.algebraMap_injective _ _)).mpr ‹_› variable {A} {B : Type*} [CommRing B] @@ -187,21 +191,23 @@ end IsIntegrallyClosedIn namespace IsIntegrallyClosed -variable {R S : Type*} [CommRing R] [CommRing S] [id : IsDomain R] [iic : IsIntegrallyClosed R] +variable {R S : Type*} [CommRing R] [CommRing S] variable {K : Type*} [CommRing K] [Algebra R K] [ifr : IsFractionRing R K] /-- Note that this is not a duplicate instance, since `IsIntegrallyClosed R` is instead defined as `IsIntegrallyClosed R R (FractionRing R)`. -/ -instance : IsIntegralClosure R R K := +instance [iic : IsIntegrallyClosed R] : IsIntegralClosure R R K := (isIntegrallyClosed_iff_isIntegralClosure K).mp iic -theorem algebraMap_eq_of_integral {x : K} : IsIntegral R x → ∃ y : R, algebraMap R K y = x := +theorem algebraMap_eq_of_integral [IsIntegrallyClosed R] {x : K} : + IsIntegral R x → ∃ y : R, algebraMap R K y = x := IsIntegralClosure.isIntegral_iff.mp -theorem isIntegral_iff {x : K} : IsIntegral R x ↔ ∃ y : R, algebraMap R K y = x := +theorem isIntegral_iff [IsIntegrallyClosed R] {x : K} : + IsIntegral R x ↔ ∃ y : R, algebraMap R K y = x := IsIntegrallyClosedIn.isIntegral_iff -theorem exists_algebraMap_eq_of_isIntegral_pow {x : K} {n : ℕ} (hn : 0 < n) +theorem exists_algebraMap_eq_of_isIntegral_pow [IsIntegrallyClosed R] {x : K} {n : ℕ} (hn : 0 < n) (hx : IsIntegral R <| x ^ n) : ∃ y : R, algebraMap R K y = x := IsIntegrallyClosedIn.exists_algebraMap_eq_of_isIntegral_pow hn hx @@ -212,7 +218,7 @@ theorem exists_algebraMap_eq_of_pow_mem_subalgebra {K : Type*} [CommRing K] [Alg variable (R S K) -instance _root_.IsIntegralClosure.of_isIntegrallyClosed +instance _root_.IsIntegralClosure.of_isIntegrallyClosed [IsIntegrallyClosed R] [Algebra S R] [Algebra S K] [IsScalarTower S R K] [Algebra.IsIntegral S R] : IsIntegralClosure R S K := IsIntegralClosure.of_isIntegrallyClosedIn @@ -224,7 +230,8 @@ theorem integralClosure_eq_bot_iff : integralClosure R K = ⊥ ↔ IsIntegrallyC (isIntegrallyClosed_iff_isIntegrallyClosedIn _).symm @[simp] -theorem pow_dvd_pow_iff {n : ℕ} (hn : n ≠ 0) {a b : R} : a ^ n ∣ b ^ n ↔ a ∣ b := by +theorem pow_dvd_pow_iff [IsDomain R] [IsIntegrallyClosed R] + {n : ℕ} (hn : n ≠ 0) {a b : R} : a ^ n ∣ b ^ n ↔ a ∣ b := by refine ⟨fun ⟨x, hx⟩ ↦ ?_, fun h ↦ pow_dvd_pow_of_dvd h n⟩ by_cases ha : a = 0 · simpa [ha, hn] using hx @@ -247,7 +254,7 @@ variable (R) /-- This is almost a duplicate of `IsIntegrallyClosedIn.integralClosure_eq_bot`, except the `NoZeroSMulDivisors` hypothesis isn't inferred automatically from `IsFractionRing`. -/ @[simp] -theorem integralClosure_eq_bot : integralClosure R K = ⊥ := +theorem integralClosure_eq_bot [IsIntegrallyClosed R] : integralClosure R K = ⊥ := (integralClosure_eq_bot_iff K).mpr ‹_› end IsIntegrallyClosed diff --git a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean index 23c177a98c8b0..3f687662317e1 100644 --- a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean +++ b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean @@ -122,7 +122,7 @@ variable [Algebra R R''] [Algebra S S''] [Algebra R S''] [IsScalarTower R R'' S''] [IsScalarTower R S S''] variable [Algebra R' R''] [Algebra S' S''] [Algebra R' S''] [IsScalarTower R' R'' S''] [IsScalarTower R' S' S''] -variable [IsScalarTower R R' R''] [IsScalarTower S S' S''] +variable [IsScalarTower S S' S''] @[simp] lemma map_id : diff --git a/Mathlib/RingTheory/Localization/BaseChange.lean b/Mathlib/RingTheory/Localization/BaseChange.lean index 31e5f767447d1..fcb23d52941ff 100644 --- a/Mathlib/RingTheory/Localization/BaseChange.lean +++ b/Mathlib/RingTheory/Localization/BaseChange.lean @@ -19,7 +19,7 @@ localize `M` by `S`. This gives us a `Localization S`-module. variable {R : Type*} [CommSemiring R] (S : Submonoid R) (A : Type*) [CommRing A] [Algebra R A] [IsLocalization S A] - {M : Type*} [AddCommMonoid M] [Module R M] [Module A M] [IsScalarTower R A M] + {M : Type*} [AddCommMonoid M] [Module R M] {M' : Type*} [AddCommMonoid M'] [Module R M'] [Module A M'] [IsScalarTower R A M'] (f : M →ₗ[R] M') diff --git a/Mathlib/RingTheory/PowerBasis.lean b/Mathlib/RingTheory/PowerBasis.lean index e8ed57c1d9fd9..ce69b35f56744 100644 --- a/Mathlib/RingTheory/PowerBasis.lean +++ b/Mathlib/RingTheory/PowerBasis.lean @@ -45,7 +45,7 @@ open Polynomial open Polynomial variable {R S T : Type*} [CommRing R] [Ring S] [Algebra R S] -variable {A B : Type*} [CommRing A] [CommRing B] [IsDomain B] [Algebra A B] +variable {A B : Type*} [CommRing A] [CommRing B] [Algebra A B] variable {K : Type*} [Field K] /-- `pb : PowerBasis R S` states that `1, pb.gen, ..., pb.gen ^ (pb.dim - 1)` @@ -316,7 +316,7 @@ noncomputable def liftEquiv (pb : PowerBasis A S) : /-- `pb.liftEquiv'` states that elements of the root set of the minimal polynomial of `pb.gen` correspond to maps sending `pb.gen` to that root. -/ @[simps! (config := .asFn)] -noncomputable def liftEquiv' (pb : PowerBasis A S) : +noncomputable def liftEquiv' [IsDomain B] (pb : PowerBasis A S) : (S →ₐ[A] B) ≃ { y : B // y ∈ (minpoly A pb.gen).aroots B } := pb.liftEquiv.trans ((Equiv.refl _).subtypeEquiv fun x => by rw [Equiv.refl_apply, mem_roots_iff_aeval_eq_zero] @@ -325,7 +325,7 @@ noncomputable def liftEquiv' (pb : PowerBasis A S) : /-- There are finitely many algebra homomorphisms `S →ₐ[A] B` if `S` is of the form `A[x]` and `B` is an integral domain. -/ -noncomputable def AlgHom.fintype (pb : PowerBasis A S) : Fintype (S →ₐ[A] B) := +noncomputable def AlgHom.fintype [IsDomain B] (pb : PowerBasis A S) : Fintype (S →ₐ[A] B) := letI := Classical.decEq B Fintype.ofEquiv _ pb.liftEquiv'.symm diff --git a/Mathlib/RingTheory/Regular/RegularSequence.lean b/Mathlib/RingTheory/Regular/RegularSequence.lean index 18e7989807eab..27b18cacdb514 100644 --- a/Mathlib/RingTheory/Regular/RegularSequence.lean +++ b/Mathlib/RingTheory/Regular/RegularSequence.lean @@ -153,7 +153,7 @@ end Definitions section Congr variable {S M} [CommRing R] [CommRing S] [AddCommGroup M] [AddCommGroup M₂] - [Module R M] [Module R M₂] [Module S M₂] + [Module R M] [Module S M₂] {σ : R →+* S} {σ' : S →+* R} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] open DistribMulAction AddSubgroup in @@ -192,7 +192,7 @@ lemma _root_.LinearEquiv.isWeaklyRegular_congr' (e : M ≃ₛₗ[σ] M₂) (rs : e.toAddEquiv.isWeaklyRegular_congr <| List.forall₂_map_right_iff.mpr <| List.forall₂_same.mpr fun r _ x => e.map_smul' r x -lemma _root_.LinearEquiv.isWeaklyRegular_congr (e : M ≃ₗ[R] M₂) (rs : List R) : +lemma _root_.LinearEquiv.isWeaklyRegular_congr [Module R M₂] (e : M ≃ₗ[R] M₂) (rs : List R) : IsWeaklyRegular M rs ↔ IsWeaklyRegular M₂ rs := Iff.trans (e.isWeaklyRegular_congr' rs) <| iff_of_eq <| congrArg _ rs.map_id @@ -210,7 +210,7 @@ lemma _root_.LinearEquiv.isRegular_congr' (e : M ≃ₛₗ[σ] M₂) (rs : List e.toAddEquiv.isRegular_congr <| List.forall₂_map_right_iff.mpr <| List.forall₂_same.mpr fun r _ x => e.map_smul' r x -lemma _root_.LinearEquiv.isRegular_congr (e : M ≃ₗ[R] M₂) (rs : List R) : +lemma _root_.LinearEquiv.isRegular_congr [Module R M₂] (e : M ≃ₗ[R] M₂) (rs : List R) : IsRegular M rs ↔ IsRegular M₂ rs := Iff.trans (e.isRegular_congr' rs) <| iff_of_eq <| congrArg _ rs.map_id diff --git a/Mathlib/RingTheory/SurjectiveOnStalks.lean b/Mathlib/RingTheory/SurjectiveOnStalks.lean index e73656d303cd0..67f0451472874 100644 --- a/Mathlib/RingTheory/SurjectiveOnStalks.lean +++ b/Mathlib/RingTheory/SurjectiveOnStalks.lean @@ -19,8 +19,8 @@ that surjections and localizations satisfy this. -/ -variable {R : Type*} [CommRing R] (M : Submonoid R) {S : Type*} [CommRing S] [Algebra R S] -variable {T : Type*} [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] +variable {R : Type*} [CommRing R] (M : Submonoid R) {S : Type*} [CommRing S] +variable {T : Type*} [CommRing T] variable {g : S →+* T} {f : R →+* S} namespace RingHom @@ -97,13 +97,6 @@ lemma surjectiveOnStalks_of_surjective (h : Function.Surjective f) : let ⟨r, hr⟩ := h s ⟨r, 1, 1, by simpa [← Ideal.eq_top_iff_one], by simpa [← Ideal.eq_top_iff_one], by simp [hr]⟩ -variable (S) in -lemma surjectiveOnStalks_of_isLocalization [IsLocalization M S] : - SurjectiveOnStalks (algebraMap R S) := by - refine surjectiveOnStalks_of_exists_div fun s ↦ ?_ - obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective M s - exact ⟨x, s, IsLocalization.map_units S s, IsLocalization.mk'_spec' S x s⟩ - lemma SurjectiveOnStalks.comp (hg : SurjectiveOnStalks g) (hf : SurjectiveOnStalks f) : SurjectiveOnStalks (g.comp f) := by intros I hI @@ -118,8 +111,10 @@ lemma SurjectiveOnStalks.of_comp (hg : SurjectiveOnStalks (g.comp f)) : RingHom.coe_comp] at this exact this.of_comp + open TensorProduct +variable [Algebra R T] [Algebra R S] in /-- If `R → T` is surjective on stalks, and `J` is some prime of `T`, then every element `x` in `S ⊗[R] T` satisfies `(1 ⊗ r • t) * x = a ⊗ t` for some @@ -152,7 +147,16 @@ lemma SurjectiveOnStalks.exists_mul_eq_tmul Algebra.TensorProduct.tmul_mul_tmul, one_mul, smul_mul_assoc, ← TensorProduct.smul_tmul, TensorProduct.add_tmul, mul_comm t₁ t₂] +variable (S) in +lemma surjectiveOnStalks_of_isLocalization + [Algebra R S] [IsLocalization M S] : + SurjectiveOnStalks (algebraMap R S) := by + refine surjectiveOnStalks_of_exists_div fun s ↦ ?_ + obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective M s + exact ⟨x, s, IsLocalization.map_units S s, IsLocalization.mk'_spec' S x s⟩ + lemma SurjectiveOnStalks.baseChange + [Algebra R T] [Algebra R S] (hf : (algebraMap R T).SurjectiveOnStalks) : (algebraMap S (S ⊗[R] T)).SurjectiveOnStalks := by let g : T →+* S ⊗[R] T := Algebra.TensorProduct.includeRight.toRingHom diff --git a/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean b/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean index a06929ffda005..10e82ca9edd87 100644 --- a/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean +++ b/Mathlib/RingTheory/TensorProduct/MvPolynomial.lean @@ -50,12 +50,13 @@ open Set LinearMap Submodule variable {R : Type u} {M : Type v} {N : Type w} [CommSemiring R] [AddCommMonoid M] [Module R M] -variable {σ : Type*} [DecidableEq σ] +variable {σ : Type*} variable {S : Type*} [CommSemiring S] [Algebra R S] section Module +variable [DecidableEq σ] variable [AddCommMonoid N] [Module R N] /-- The tensor product of a polynomial ring by a module is @@ -141,6 +142,9 @@ lemma coeff_rTensorAlgHom_tmul rw [algebraMap_eq, mul_comm, coeff_C_mul] simp [mapAlgHom, coeff_map] +section DecidableEq +variable [DecidableEq σ] + lemma coeff_rTensorAlgHom_monomial_tmul (e : σ →₀ ℕ) (s : S) (n : N) (d : σ →₀ ℕ) : coeff d (rTensorAlgHom (monomial e s ⊗ₜ[R] n)) = @@ -186,6 +190,8 @@ noncomputable def scalarRTensorAlgEquiv : MvPolynomial σ R ⊗[R] N ≃ₐ[R] MvPolynomial σ N := rTensorAlgEquiv.trans (mapAlgEquiv σ (Algebra.TensorProduct.lid R N)) +end DecidableEq + variable (R) variable (A : Type*) [CommSemiring A] [Algebra R A] diff --git a/Mathlib/RingTheory/Trace/Defs.lean b/Mathlib/RingTheory/Trace/Defs.lean index 2b04f816b9079..86218f0b3ecff 100644 --- a/Mathlib/RingTheory/Trace/Defs.lean +++ b/Mathlib/RingTheory/Trace/Defs.lean @@ -59,7 +59,6 @@ open scoped Matrix namespace Algebra -variable (b : Basis ι R S) variable (R S) /-- The trace of an element `s` of an `R`-algebra is the trace of `(s * ·)`, @@ -85,7 +84,8 @@ theorem trace_eq_matrix_trace [DecidableEq ι] (b : Basis ι R S) (s : S) : rw [trace_apply, LinearMap.trace_eq_matrix_trace _ b, ← toMatrix_lmul_eq]; rfl /-- If `x` is in the base field `K`, then the trace is `[L : K] * x`. -/ -theorem trace_algebraMap_of_basis (x : R) : trace R S (algebraMap R S x) = Fintype.card ι • x := by +theorem trace_algebraMap_of_basis (b : Basis ι R S) (x : R) : + trace R S (algebraMap R S x) = Fintype.card ι • x := by haveI := Classical.decEq ι rw [trace_apply, LinearMap.trace_eq_matrix_trace R b, Matrix.trace] convert Finset.sum_const x @@ -165,7 +165,7 @@ theorem traceForm_apply (x y : S) : traceForm R S x y = trace R S (x * y) := theorem traceForm_isSymm : (traceForm R S).IsSymm := fun _ _ => congr_arg (trace R S) (mul_comm _ _) -theorem traceForm_toMatrix [DecidableEq ι] (i j) : +theorem traceForm_toMatrix [DecidableEq ι] (b : Basis ι R S) (i j) : BilinForm.toMatrix b (traceForm R S) i j = trace R S (b i * b j) := by rw [BilinForm.toMatrix_apply, traceForm_apply] diff --git a/Mathlib/Topology/Compactness/CompactlyGeneratedSpace.lean b/Mathlib/Topology/Compactness/CompactlyGeneratedSpace.lean index 27e7dae060724..ebc17cc09ad5a 100644 --- a/Mathlib/Topology/Compactness/CompactlyGeneratedSpace.lean +++ b/Mathlib/Topology/Compactness/CompactlyGeneratedSpace.lean @@ -47,7 +47,7 @@ open TopologicalSpace Filter Topology Set section UCompactlyGeneratedSpace -variable {X : Type w} {Y : Type x} [TopologicalSpace X] [TopologicalSpace Y] +variable {X : Type w} {Y : Type x} /-- The compactly generated topology on a topological space `X`. This is the finest topology @@ -61,7 +61,7 @@ def TopologicalSpace.compactlyGenerated (X : Type w) [TopologicalSpace X] : Topo let f : (Σ (i : (S : CompHaus.{u}) × C(S, X)), i.fst) → X := fun ⟨⟨_, i⟩, s⟩ ↦ i s coinduced f inferInstance -lemma continuous_from_compactlyGenerated [t : TopologicalSpace Y] (f : X → Y) +lemma continuous_from_compactlyGenerated [TopologicalSpace X] [t : TopologicalSpace Y] (f : X → Y) (h : ∀ (S : CompHaus.{u}) (g : C(S, X)), Continuous (f ∘ g)) : Continuous[compactlyGenerated.{u} X, t] f := by rw [continuous_coinduced_dom] @@ -94,14 +94,6 @@ instance (X : Type v) [t : TopologicalSpace X] [DiscreteTopology X] : rw [DiscreteTopology.eq_bot (t := t)] exact bot_le -/-- If `X` is compactly generated, to prove that `f : X → Y` is continuous it is enough to show -that for every compact Hausdorff space `K` and every continuous map `g : K → X`, -`f ∘ g` is continuous. -/ -lemma continuous_from_uCompactlyGeneratedSpace [UCompactlyGeneratedSpace.{u} X] (f : X → Y) - (h : ∀ (S : CompHaus.{u}) (g : C(S, X)), Continuous (f ∘ g)) : Continuous f := by - apply continuous_le_dom UCompactlyGeneratedSpace.le_compactlyGenerated - exact continuous_from_compactlyGenerated f h - /-- Let `f : X → Y`. Suppose that to prove that `f` is continuous, it suffices to show that for every compact Hausdorff space `K` and every continuous map `g : K → X`, `f ∘ g` is continuous. Then `X` is compactly generated. -/ @@ -120,6 +112,16 @@ lemma uCompactlyGeneratedSpace_of_continuous_maps [t : TopologicalSpace X] rw [← @continuous_sigma_iff] apply continuous_coinduced_rng +variable [tX : TopologicalSpace X] [tY : TopologicalSpace Y] + +/-- If `X` is compactly generated, to prove that `f : X → Y` is continuous it is enough to show +that for every compact Hausdorff space `K` and every continuous map `g : K → X`, +`f ∘ g` is continuous. -/ +lemma continuous_from_uCompactlyGeneratedSpace [UCompactlyGeneratedSpace.{u} X] (f : X → Y) + (h : ∀ (S : CompHaus.{u}) (g : C(S, X)), Continuous (f ∘ g)) : Continuous f := by + apply continuous_le_dom UCompactlyGeneratedSpace.le_compactlyGenerated + exact continuous_from_compactlyGenerated f h + /-- A topological space `X` is compactly generated if a set `s` is closed when `f ⁻¹' s` is closed for every continuous map `f : K → X`, where `K` is compact Hausdorff. -/ theorem uCompactlyGeneratedSpace_of_isClosed @@ -155,7 +157,6 @@ theorem UCompactlyGeneratedSpace.isOpen [UCompactlyGeneratedSpace.{u} X] {s : Se /-- If the topology of `X` is coinduced by a continuous function whose domain is compactly generated, then so is `X`. -/ theorem uCompactlyGeneratedSpace_of_coinduced - [tX : TopologicalSpace X] [tY : TopologicalSpace Y] [UCompactlyGeneratedSpace.{u} X] {f : X → Y} (hf : Continuous f) (ht : tY = coinduced f tX) : UCompactlyGeneratedSpace.{u} Y := by refine uCompactlyGeneratedSpace_of_isClosed fun s h ↦ ?_ @@ -291,8 +292,8 @@ theorem CompactlyGeneratedSpace.isOpen [CompactlyGeneratedSpace X] {s : Set X} /-- If the topology of `X` is coinduced by a continuous function whose domain is compactly generated, then so is `X`. -/ -theorem compactlyGeneratedSpace_of_coinduced {Y : Type u} - [tX : TopologicalSpace X] [tY : TopologicalSpace Y] +theorem compactlyGeneratedSpace_of_coinduced + {X : Type u} [tX : TopologicalSpace X] {Y : Type u} [tY : TopologicalSpace Y] [CompactlyGeneratedSpace X] {f : X → Y} (hf : Continuous f) (ht : tY = coinduced f tX) : CompactlyGeneratedSpace Y := uCompactlyGeneratedSpace_of_coinduced hf ht diff --git a/Mathlib/Topology/ContinuousFunction/Ideals.lean b/Mathlib/Topology/ContinuousFunction/Ideals.lean index 93810a823e816..2bda7f44edbd2 100644 --- a/Mathlib/Topology/ContinuousFunction/Ideals.lean +++ b/Mathlib/Topology/ContinuousFunction/Ideals.lean @@ -380,7 +380,7 @@ variable (X 𝕜 : Type*) [TopologicalSpace X] section ContinuousMapEval -variable [LocallyCompactSpace X] [CommRing 𝕜] [TopologicalSpace 𝕜] [TopologicalRing 𝕜] +variable [CommRing 𝕜] [TopologicalSpace 𝕜] [TopologicalRing 𝕜] variable [Nontrivial 𝕜] [NoZeroDivisors 𝕜] /-- The natural continuous map from a locally compact topological space `X` to the diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean index e4fb9f1ae0d5b..73e63a878282e 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorffRealized.lean @@ -51,8 +51,8 @@ distance, and show that they form a compact family by applying Arzela-Ascoli theorem. The existence of a minimizer follows. -/ section Definitions -variable (X : Type u) (Y : Type v) [MetricSpace X] [CompactSpace X] [Nonempty X] [MetricSpace Y] - [CompactSpace Y] [Nonempty Y] +variable (X : Type u) (Y : Type v) [MetricSpace X] [MetricSpace Y] + private abbrev ProdSpaceFun : Type _ := (X ⊕ Y) × (X ⊕ Y) → ℝ @@ -86,12 +86,13 @@ end Definitions section Constructions -variable {X : Type u} {Y : Type v} [MetricSpace X] [CompactSpace X] [Nonempty X] [MetricSpace Y] - [CompactSpace Y] [Nonempty Y] {f : ProdSpaceFun X Y} {x y z t : X ⊕ Y} +variable {X : Type u} {Y : Type v} [MetricSpace X] [MetricSpace Y] + {f : ProdSpaceFun X Y} {x y z t : X ⊕ Y} attribute [local instance 10] Classical.inhabited_of_nonempty' -private theorem maxVar_bound : dist x y ≤ maxVar X Y := +private theorem maxVar_bound [CompactSpace X] [Nonempty X] [CompactSpace Y] [Nonempty Y] : + dist x y ≤ maxVar X Y := calc dist x y ≤ diam (univ : Set (X ⊕ Y)) := dist_le_diam_of_mem isBounded_of_compactSpace (mem_univ _) (mem_univ _) diff --git a/Mathlib/Topology/VectorBundle/Hom.lean b/Mathlib/Topology/VectorBundle/Hom.lean index e1f049db2c704..db2f9b8915afc 100644 --- a/Mathlib/Topology/VectorBundle/Hom.lean +++ b/Mathlib/Topology/VectorBundle/Hom.lean @@ -41,7 +41,7 @@ open scoped Bundle open Bundle Set ContinuousLinearMap variable {𝕜₁ : Type*} [NontriviallyNormedField 𝕜₁] {𝕜₂ : Type*} [NontriviallyNormedField 𝕜₂] - (σ : 𝕜₁ →+* 𝕜₂) [iσ : RingHomIsometric σ] + (σ : 𝕜₁ →+* 𝕜₂) variable {B : Type*} variable {F₁ : Type*} [NormedAddCommGroup F₁] [NormedSpace 𝕜₁ F₁] (E₁ : B → Type*) @@ -82,9 +82,10 @@ def continuousLinearMapCoordChange [e₁.IsLinear 𝕜₁] [e₁'.IsLinear 𝕜 variable {σ e₁ e₁' e₂ e₂'} variable [∀ x, TopologicalSpace (E₁ x)] [FiberBundle F₁ E₁] -variable [∀ x, TopologicalSpace (E₂ x)] [ita : ∀ x, TopologicalAddGroup (E₂ x)] [FiberBundle F₂ E₂] +variable [∀ x, TopologicalSpace (E₂ x)] [FiberBundle F₂ E₂] -theorem continuousOn_continuousLinearMapCoordChange [VectorBundle 𝕜₁ F₁ E₁] [VectorBundle 𝕜₂ F₂ E₂] +theorem continuousOn_continuousLinearMapCoordChange [RingHomIsometric σ] + [VectorBundle 𝕜₁ F₁ E₁] [VectorBundle 𝕜₂ F₂ E₂] [MemTrivializationAtlas e₁] [MemTrivializationAtlas e₁'] [MemTrivializationAtlas e₂] [MemTrivializationAtlas e₂'] : ContinuousOn (continuousLinearMapCoordChange σ e₁ e₁' e₂ e₂') @@ -194,6 +195,7 @@ variable (F₁ E₁ F₂ E₂) variable [∀ x : B, TopologicalSpace (E₁ x)] [FiberBundle F₁ E₁] [VectorBundle 𝕜₁ F₁ E₁] variable [∀ x : B, TopologicalSpace (E₂ x)] [FiberBundle F₂ E₂] [VectorBundle 𝕜₂ F₂ E₂] variable [∀ x, TopologicalAddGroup (E₂ x)] [∀ x, ContinuousSMul 𝕜₂ (E₂ x)] +variable [RingHomIsometric σ] /-- The continuous `σ`-semilinear maps between two topological vector bundles form a `VectorPrebundle` (this is an auxiliary construction for the From 0d9639489a83f0d6a743b55a2aee06f59ac9fe9a Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Fri, 9 Aug 2024 00:58:21 +0000 Subject: [PATCH 135/359] perf: eagerly elaborate leaf nodes in `linear_combination` (#15599) Since the arithmetic type is known from the equality goal, we can eagerly elaborate all leaf nodes in `expandLinearCombo` while also avoiding re-elaborating them when deciding whether they are proofs or not. This brings the elaboration of a complicated example (see tests) from 3.91s to 0.29s. The main issue is that large expressions can be slow to elaborate. In particular, exponentiation relies on the default instance mechanism, which appears can take quadratic time to resolve. Forcing the resolution of default instances within each coefficient of the linear combination bounds the default instance problems. This is similar to PR #15570, but it sticks with the syntax transformation paradigm. Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> --- Mathlib/Tactic/LinearCombination.lean | 104 ++++++++++++++------------ test/linear_combination.lean | 43 +++++++++++ 2 files changed, 101 insertions(+), 46 deletions(-) diff --git a/Mathlib/Tactic/LinearCombination.lean b/Mathlib/Tactic/LinearCombination.lean index 9b6a3f8d3b1c3..78472a591957f 100644 --- a/Mathlib/Tactic/LinearCombination.lean +++ b/Mathlib/Tactic/LinearCombination.lean @@ -51,60 +51,70 @@ theorem pf_div_c [Div α] (p : a = b) (c : α) : a / c = b / c := p ▸ rfl theorem c_div_pf [Div α] (p : b = c) (a : α) : a / b = a / c := p ▸ rfl theorem div_pf [Div α] (p₁ : (a₁:α) = b₁) (p₂ : a₂ = b₂) : a₁ / a₂ = b₁ / b₂ := p₁ ▸ p₂ ▸ rfl +/-- Result of `expandLinearCombo`, either an equality proof or a value. -/ +inductive Expanded + /-- A proof of `a = b`. -/ + | proof (pf : Syntax.Term) + /-- A value, equivalently a proof of `c = c`. -/ + | const (c : Syntax.Term) + /-- Performs macro expansion of a linear combination expression, using `+`/`-`/`*`/`/` on equations and values. -* `some p` means that `p` is a syntax corresponding to a proof of an equation. - For example, if `h : a = b` then `expandLinearCombo (2 * h)` returns `some (c_add_pf 2 h)` +* `.proof p` means that `p` is a syntax corresponding to a proof of an equation. + For example, if `h : a = b` then `expandLinearCombo (2 * h)` returns `.proof (c_add_pf 2 h)` which is a proof of `2 * a = 2 * b`. -* `none` means that the input expression is not an equation but a value; - the input syntax itself is used in this case. +* `.const c` means that the input expression is not an equation but a value. -/ -partial def expandLinearCombo (stx : Syntax.Term) : TermElabM (Option Syntax.Term) := do - let mut result ← match stx with - | `(($e)) => expandLinearCombo e +partial def expandLinearCombo (ty : Expr) (stx : Syntax.Term) : TermElabM Expanded := withRef stx do + match stx with + | `(($e)) => expandLinearCombo ty e | `($e₁ + $e₂) => do - match ← expandLinearCombo e₁, ← expandLinearCombo e₂ with - | none, none => pure none - | some p₁, none => ``(pf_add_c $p₁ $e₂) - | none, some p₂ => ``(c_add_pf $p₂ $e₁) - | some p₁, some p₂ => ``(add_pf $p₁ $p₂) + match ← expandLinearCombo ty e₁, ← expandLinearCombo ty e₂ with + | .const c₁, .const c₂ => .const <$> ``($c₁ + $c₂) + | .proof p₁, .const c₂ => .proof <$> ``(pf_add_c $p₁ $c₂) + | .const c₁, .proof p₂ => .proof <$> ``(c_add_pf $p₂ $c₁) + | .proof p₁, .proof p₂ => .proof <$> ``(add_pf $p₁ $p₂) | `($e₁ - $e₂) => do - match ← expandLinearCombo e₁, ← expandLinearCombo e₂ with - | none, none => pure none - | some p₁, none => ``(pf_sub_c $p₁ $e₂) - | none, some p₂ => ``(c_sub_pf $p₂ $e₁) - | some p₁, some p₂ => ``(sub_pf $p₁ $p₂) + match ← expandLinearCombo ty e₁, ← expandLinearCombo ty e₂ with + | .const c₁, .const c₂ => .const <$> ``($c₁ - $c₂) + | .proof p₁, .const c₂ => .proof <$> ``(pf_sub_c $p₁ $c₂) + | .const c₁, .proof p₂ => .proof <$> ``(c_sub_pf $p₂ $c₁) + | .proof p₁, .proof p₂ => .proof <$> ``(sub_pf $p₁ $p₂) | `(-$e) => do - match ← expandLinearCombo e with - | none => pure none - | some p => ``(neg_pf $p) + match ← expandLinearCombo ty e with + | .const c => .const <$> `(-$c) + | .proof p => .proof <$> ``(neg_pf $p) | `(← $e) => do - match ← expandLinearCombo e with - | none => pure none - | some p => ``(Eq.symm $p) + match ← expandLinearCombo ty e with + | .const c => return .const c + | .proof p => .proof <$> ``(Eq.symm $p) | `($e₁ * $e₂) => do - match ← expandLinearCombo e₁, ← expandLinearCombo e₂ with - | none, none => pure none - | some p₁, none => ``(pf_mul_c $p₁ $e₂) - | none, some p₂ => ``(c_mul_pf $p₂ $e₁) - | some p₁, some p₂ => ``(mul_pf $p₁ $p₂) + match ← expandLinearCombo ty e₁, ← expandLinearCombo ty e₂ with + | .const c₁, .const c₂ => .const <$> ``($c₁ * $c₂) + | .proof p₁, .const c₂ => .proof <$> ``(pf_mul_c $p₁ $c₂) + | .const c₁, .proof p₂ => .proof <$> ``(c_mul_pf $p₂ $c₁) + | .proof p₁, .proof p₂ => .proof <$> ``(mul_pf $p₁ $p₂) | `($e⁻¹) => do - match ← expandLinearCombo e with - | none => pure none - | some p => ``(inv_pf $p) + match ← expandLinearCombo ty e with + | .const c => .const <$> `($c⁻¹) + | .proof p => .proof <$> ``(inv_pf $p) | `($e₁ / $e₂) => do - match ← expandLinearCombo e₁, ← expandLinearCombo e₂ with - | none, none => pure none - | some p₁, none => ``(pf_div_c $p₁ $e₂) - | none, some p₂ => ``(c_div_pf $p₂ $e₁) - | some p₁, some p₂ => ``(div_pf $p₁ $p₂) - | e => do - let e ← elabTerm e none - let eType ← inferType e - let .true := (← withReducible do whnf eType).isEq | pure none - some <$> e.toSyntax - return result.map fun r => ⟨r.raw.setInfo (SourceInfo.fromRef stx true)⟩ + match ← expandLinearCombo ty e₁, ← expandLinearCombo ty e₂ with + | .const c₁, .const c₂ => .const <$> ``($c₁ / $c₂) + | .proof p₁, .const c₂ => .proof <$> ``(pf_div_c $p₁ $c₂) + | .const c₁, .proof p₂ => .proof <$> ``(c_div_pf $p₂ $c₁) + | .proof p₁, .proof p₂ => .proof <$> ``(div_pf $p₁ $p₂) + | e => + -- We have the expected type from the goal, so we can fully synthesize this leaf node. + withSynthesize do + -- It is OK to use `ty` as the expected type even if `e` is a proof. + -- The expected type is just a hint. + let c ← withSynthesizeLight <| Term.elabTerm e ty + if (← whnfR (← inferType c)).isEq then + .proof <$> c.toSyntax + else + .const <$> c.toSyntax theorem eq_trans₃ (p : (a:α) = b) (p₁ : a = a') (p₂ : b = b') : a' = b' := p₁ ▸ p₂ ▸ p @@ -119,12 +129,14 @@ theorem eq_of_add_pow [Ring α] [NoZeroDivisors α] (n : ℕ) (p : (a:α) = b) def elabLinearCombination (norm? : Option Syntax.Tactic) (exp? : Option Syntax.NumLit) (input : Option Syntax.Term) (twoGoals := false) : Tactic.TacticM Unit := Tactic.withMainContext do + let some (ty, _) := (← (← Tactic.getMainGoal).getType').eq? | + throwError "'linear_combination' only proves equalities" let p ← match input with | none => `(Eq.refl 0) - | some e => withSynthesize do - match ← expandLinearCombo e with - | none => `(Eq.refl $e) - | some p => pure p + | some e => + match ← expandLinearCombo ty e with + | .const c => `(Eq.refl $c) + | .proof p => pure p let norm := norm?.getD (Unhygienic.run `(tactic| ring1)) Tactic.evalTactic <| ← withFreshMacroScope <| if twoGoals then diff --git a/test/linear_combination.lean b/test/linear_combination.lean index 05d7b32149fd0..efb18bd1c26d2 100644 --- a/test/linear_combination.lean +++ b/test/linear_combination.lean @@ -4,6 +4,8 @@ import Mathlib.Tactic.Linarith set_option autoImplicit true +private axiom test_sorry : ∀ {α}, α + -- We deliberately mock R here so that we don't have to import the deps axiom Real : Type notation "ℝ" => Real @@ -118,6 +120,10 @@ example {α} [h : CommRing α] {a b c d e f : α} (h1 : a * d = b * c) (h2 : c * example (x y z w : ℚ) (hzw : z = w) : x * z + 2 * y * z = x * w + 2 * y * w := by linear_combination (x + 2 * y) * hzw +example (x : ℤ) : x ^ 2 = x ^ 2 := by linear_combination x ^ 2 + +example (x y : ℤ) (h : x = 0) : y ^ 2 * x = 0 := by linear_combination y ^ 2 * h + /-! ### Cases that explicitly use a config -/ example (x y : ℚ) (h1 : 3 * x + 2 * y = 10) (h2 : 2 * x + 5 * y = 3) : -11 * y + 1 = 11 + 1 := by @@ -230,3 +236,40 @@ example (h : g a = g b) : a ^ 4 = b ^ 4 := by example {r s a b : ℕ} (h₁ : (r : ℤ) = a + 1) (h₂ : (s : ℤ) = b + 1) : r * s = (a + 1 : ℤ) * (b + 1) := by linear_combination (↑b + 1) * h₁ + ↑r * h₂ + +-- Implementation at the time of the port (Nov 2022) was 110,000 heartbeats. +-- Eagerly elaborating leaf nodes brings this to 7,540 heartbeats. +set_option maxHeartbeats 8000 in +example (K : Type*) [Field K] [CharZero K] {x y z p q : K} + (h₀ : 3 * x ^ 2 + z ^ 2 * p = 0) + (h₁ : z * (2 * y) = 0) + (h₂ : -y ^ 2 + p * x * (2 * z) + q * (3 * z ^ 2) = 0) : + ((27 * q ^ 2 + 4 * p ^ 3) * x) ^ 4 = 0 := by + linear_combination (norm := skip) + (256 / 3 * p ^ 12 * x ^ 2 + 128 * q * p ^ 11 * x * z + 2304 * q ^ 2 * p ^ 9 * x ^ 2 + + 2592 * q ^ 3 * p ^ 8 * x * z - + 64 * q * p ^ 10 * y ^ 2 + + 23328 * q ^ 4 * p ^ 6 * x ^ 2 + + 17496 * q ^ 5 * p ^ 5 * x * z - + 1296 * q ^ 3 * p ^ 7 * y ^ 2 + + 104976 * q ^ 6 * p ^ 3 * x ^ 2 + + 39366 * q ^ 7 * p ^ 2 * x * z - + 8748 * q ^ 5 * p ^ 4 * y ^ 2 + + 177147 * q ^ 8 * x ^ 2 - + 19683 * q ^ 7 * p * y ^ 2) * + h₀ + + (-(64 / 3 * p ^ 12 * x * y) + 32 * q * p ^ 11 * z * y - 432 * q ^ 2 * p ^ 9 * x * y + + 648 * q ^ 3 * p ^ 8 * z * y - + 2916 * q ^ 4 * p ^ 6 * x * y + + 4374 * q ^ 5 * p ^ 5 * z * y - + 6561 * q ^ 6 * p ^ 3 * x * y + + 19683 / 2 * q ^ 7 * p ^ 2 * z * y) * + h₁ + + (-(128 / 3 * p ^ 12 * x * z) - 192 * q * p ^ 10 * x ^ 2 - 864 * q ^ 2 * p ^ 9 * x * z - + 3888 * q ^ 3 * p ^ 7 * x ^ 2 - + 5832 * q ^ 4 * p ^ 6 * x * z - + 26244 * q ^ 5 * p ^ 4 * x ^ 2 - + 13122 * q ^ 6 * p ^ 3 * x * z - + 59049 * q ^ 7 * p * x ^ 2) * + h₂ + exact test_sorry From 0f7bb9cc912092fce417c98d0fbaeea6dcfe470d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 9 Aug 2024 00:58:22 +0000 Subject: [PATCH 136/359] chore: backports for leanprover/lean4#4814 (part 33) (#15605) --- .../Algebra/Homology/DerivedCategory/Ext.lean | 23 +++-- Mathlib/Algebra/Lie/CartanExists.lean | 6 +- Mathlib/Algebra/Lie/Killing.lean | 9 +- Mathlib/Algebra/Lie/Weights/Linear.lean | 9 +- .../ContinuousFunctionalCalculus/Unitary.lean | 4 +- .../Analysis/Calculus/FDeriv/Measurable.lean | 5 +- Mathlib/Analysis/Fourier/AddCircle.lean | 2 +- .../NormedSpace/HahnBanach/Separation.lean | 13 ++- .../Combinatorics/SimpleGraph/LapMatrix.lean | 14 +-- .../Manifold/ContMDiff/NormedSpace.lean | 92 ++++++++++--------- .../Geometry/Manifold/ContMDiff/Product.lean | 16 +--- .../LinearAlgebra/Matrix/SchurComplement.lean | 4 +- .../Constructions/BorelSpace/Real.lean | 3 +- .../Constructions/Polish/Basic.lean | 10 +- .../Function/ContinuousMapDense.lean | 19 ++-- Mathlib/MeasureTheory/Group/Integral.lean | 2 +- Mathlib/MeasureTheory/Integral/Bochner.lean | 5 +- .../Integral/BoundedContinuousFunction.lean | 18 ++-- .../Measure/Haar/InnerProductSpace.lean | 44 +++++---- Mathlib/NumberTheory/Cyclotomic/Basic.lean | 15 +-- .../DedekindDomain/IntegralClosure.lean | 7 +- Mathlib/RingTheory/Discriminant.lean | 2 +- Mathlib/RingTheory/Kaehler/Basic.lean | 5 + Mathlib/RingTheory/WittVector/MulCoeff.lean | 3 +- Mathlib/Topology/Algebra/Algebra.lean | 3 + 25 files changed, 182 insertions(+), 151 deletions(-) diff --git a/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean b/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean index 6245ca808b6f0..15e28dcd178de 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean @@ -238,18 +238,7 @@ lemma mk₀_zero : mk₀ (0 : X ⟶ Y) = 0 := by section -variable [HasDerivedCategory.{w'} C] - -variable (X Y n) in -@[simp] -lemma zero_hom : (0 : Ext X Y n).hom = 0 := by - let β : Ext 0 Y n := 0 - have hβ : β.hom = 0 := by apply (Functor.map_isZero _ (isZero_zero C)).eq_of_src - have : (0 : Ext X Y n) = (0 : Ext X 0 0).comp β (zero_add n) := by simp [β] - rw [this, comp_hom, hβ, ShiftedHom.comp_zero] - -attribute [local instance] preservesBinaryBiproductsOfPreservesBiproducts - +attribute [local instance] preservesBinaryBiproductsOfPreservesBiproducts in lemma biprod_ext {X₁ X₂ : C} {α β : Ext (X₁ ⊞ X₂) Y n} (h₁ : (mk₀ biprod.inl).comp α (zero_add n) = (mk₀ biprod.inl).comp β (zero_add n)) (h₂ : (mk₀ biprod.inr).comp α (zero_add n) = (mk₀ biprod.inr).comp β (zero_add n)) : @@ -262,6 +251,16 @@ lemma biprod_ext {X₁ X₂ : C} {α β : Ext (X₁ ⊞ X₂) Y n} (BinaryBiproduct.isBilimit X₁ X₂)).isColimit all_goals assumption +variable [HasDerivedCategory.{w'} C] + +variable (X Y n) in +@[simp] +lemma zero_hom : (0 : Ext X Y n).hom = 0 := by + let β : Ext 0 Y n := 0 + have hβ : β.hom = 0 := by apply (Functor.map_isZero _ (isZero_zero C)).eq_of_src + have : (0 : Ext X Y n) = (0 : Ext X 0 0).comp β (zero_add n) := by simp [β] + rw [this, comp_hom, hβ, ShiftedHom.comp_zero] + @[simp] lemma add_hom (α β : Ext X Y n) : (α + β).hom = α.hom + β.hom := by let α' : Ext (X ⊞ X) Y n := (mk₀ biprod.fst).comp α (zero_add n) diff --git a/Mathlib/Algebra/Lie/CartanExists.lean b/Mathlib/Algebra/Lie/CartanExists.lean index 77a85c041d18b..25735e2642f2f 100644 --- a/Mathlib/Algebra/Lie/CartanExists.lean +++ b/Mathlib/Algebra/Lie/CartanExists.lean @@ -32,7 +32,7 @@ namespace LieAlgebra section CommRing variable {K R L M : Type*} -variable [Field K] [CommRing R] [Nontrivial R] +variable [Field K] [CommRing R] variable [LieRing L] [LieAlgebra K L] [LieAlgebra R L] variable [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] variable [Module.Finite K L] @@ -83,7 +83,7 @@ def lieCharpoly : Polynomial R[X] := lemma lieCharpoly_monic : (lieCharpoly R M x y).Monic := (polyCharpoly_monic _ _).map _ -lemma lieCharpoly_natDegree : (lieCharpoly R M x y).natDegree = finrank R M := by +lemma lieCharpoly_natDegree [Nontrivial R] : (lieCharpoly R M x y).natDegree = finrank R M := by rw [lieCharpoly, (polyCharpoly_monic _ _).natDegree_map, polyCharpoly_natDegree] variable {R} in @@ -97,7 +97,7 @@ lemma lieCharpoly_map_eval (r : R) : map_add, map_mul, aeval_C, Algebra.id.map_eq_id, RingHom.id_apply, aeval_X, aux, MvPolynomial.coe_aeval_eq_eval, polyCharpoly_map_eq_charpoly, LieHom.coe_toLinearMap] -lemma lieCharpoly_coeff_natDegree (i j : ℕ) (hij : i + j = finrank R M) : +lemma lieCharpoly_coeff_natDegree [Nontrivial R] (i j : ℕ) (hij : i + j = finrank R M) : ((lieCharpoly R M x y).coeff i).natDegree ≤ j := by classical rw [← mul_one j, lieCharpoly, coeff_map] diff --git a/Mathlib/Algebra/Lie/Killing.lean b/Mathlib/Algebra/Lie/Killing.lean index 72cb8a63122ad..45f1f3365045c 100644 --- a/Mathlib/Algebra/Lie/Killing.lean +++ b/Mathlib/Algebra/Lie/Killing.lean @@ -54,7 +54,7 @@ attribute [simp] IsKilling.killingCompl_top_eq_bot namespace IsKilling -variable [Module.Free R L] [Module.Finite R L] [IsKilling R L] +variable [IsKilling R L] @[simp] lemma ker_killingForm_eq_bot : LinearMap.ker (killingForm R L) = ⊥ := by @@ -65,7 +65,8 @@ lemma killingForm_nondegenerate : simp [LinearMap.BilinForm.nondegenerate_iff_ker_eq_bot] variable {R L} in -lemma ideal_eq_bot_of_isLieAbelian [IsDomain R] [IsPrincipalIdealRing R] +lemma ideal_eq_bot_of_isLieAbelian + [Module.Free R L] [Module.Finite R L] [IsDomain R] [IsPrincipalIdealRing R] (I : LieIdeal R L) [IsLieAbelian I] : I = ⊥ := by rw [eq_bot_iff, ← killingCompl_top_eq_bot] exact I.le_killingCompl_top_of_isLieAbelian @@ -83,7 +84,9 @@ over fields with positive characteristic. Note that when the coefficients are a field this instance is redundant since we have `LieAlgebra.IsKilling.instSemisimple` and `LieAlgebra.IsSemisimple.instHasTrivialRadical`. -/ -instance instHasTrivialRadical [IsDomain R] [IsPrincipalIdealRing R] : HasTrivialRadical R L := +instance instHasTrivialRadical + [Module.Free R L] [Module.Finite R L] [IsDomain R] [IsPrincipalIdealRing R] : + HasTrivialRadical R L := (hasTrivialRadical_iff_no_abelian_ideals R L).mpr IsKilling.ideal_eq_bot_of_isLieAbelian end IsKilling diff --git a/Mathlib/Algebra/Lie/Weights/Linear.lean b/Mathlib/Algebra/Lie/Weights/Linear.lean index bcd07bf9df39c..39d847bba73c0 100644 --- a/Mathlib/Algebra/Lie/Weights/Linear.lean +++ b/Mathlib/Algebra/Lie/Weights/Linear.lean @@ -55,8 +55,7 @@ class LinearWeights [LieAlgebra.IsNilpotent R L] : Prop := namespace Weight -variable [LieAlgebra.IsNilpotent R L] [LinearWeights R L M] - [NoZeroSMulDivisors R M] [IsNoetherian R M] (χ : Weight R L M) +variable [LieAlgebra.IsNilpotent R L] [LinearWeights R L M] (χ : Weight R L M) /-- A weight of a Lie module, bundled as a linear map. -/ @[simps] @@ -155,7 +154,7 @@ instance instLinearWeightsOfCharZero [CharZero R] : end FiniteDimensional -variable [LieAlgebra.IsNilpotent R L] [LinearWeights R L M] (χ : L → R) +variable [LieAlgebra.IsNilpotent R L] (χ : L → R) /-- A type synonym for the `χ`-weight space but with the action of `x : L` on `m : weightSpace M χ`, shifted to act as `⁅x, m⁆ - χ x • m`. -/ @@ -166,6 +165,8 @@ namespace shiftedWeightSpace private lemma aux [h : Nontrivial (shiftedWeightSpace R L M χ)] : weightSpace M χ ≠ ⊥ := (LieSubmodule.nontrivial_iff_ne_bot _ _ _).mp h +variable [LinearWeights R L M] + instance : LieRingModule L (shiftedWeightSpace R L M χ) where bracket x m := ⁅x, m⁆ - χ x • m add_lie x y m := by @@ -217,7 +218,7 @@ end shiftedWeightSpace /-- Given a Lie module `M` of a Lie algebra `L` with coefficients in `R`, if a function `χ : L → R` has a simultaneous generalized eigenvector for the action of `L` then it has a simultaneous true eigenvector, provided `M` is Noetherian and has linear weights. -/ -lemma exists_forall_lie_eq_smul [IsNoetherian R M] (χ : Weight R L M) : +lemma exists_forall_lie_eq_smul [LinearWeights R L M] [IsNoetherian R M] (χ : Weight R L M) : ∃ m : M, m ≠ 0 ∧ ∀ x : L, ⁅x, m⁆ = χ x • m := by replace hχ : Nontrivial (shiftedWeightSpace R L M χ) := (LieSubmodule.nontrivial_iff_ne_bot R L M).mpr χ.weightSpace_ne_bot diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unitary.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unitary.lean index bda712cdab3b1..d39bee3d59145 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unitary.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unitary.lean @@ -20,7 +20,7 @@ section Generic variable {R A : Type*} {p : A → Prop} [CommRing R] [StarRing R] [MetricSpace R] variable [TopologicalRing R] [ContinuousStar R] [TopologicalSpace A] [Ring A] [StarRing A] -variable [Algebra R A] [StarModule R A] [ContinuousFunctionalCalculus R p] +variable [Algebra R A] [ContinuousFunctionalCalculus R p] lemma cfc_unitary_iff (f : R → R) (a : A) (ha : p a := by cfc_tac) (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) : @@ -35,7 +35,7 @@ end Generic section Complex variable {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] [Algebra ℂ A] - [StarModule ℂ A] [ContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop)] + [ContinuousFunctionalCalculus ℂ (IsStarNormal : A → Prop)] lemma unitary_iff_isStarNormal_and_spectrum_subset_unitary {u : A} : u ∈ unitary A ↔ IsStarNormal u ∧ spectrum ℂ u ⊆ unitary ℂ := by diff --git a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean index b02054cbd951d..d201d59948dd7 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Measurable.lean @@ -787,8 +787,7 @@ that the differentiability set `D` is measurable. -/ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] [LocallyCompactSpace E] {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] - {α : Type*} [TopologicalSpace α] [MeasurableSpace α] [MeasurableSpace E] - [OpensMeasurableSpace α] [OpensMeasurableSpace E] + {α : Type*} [TopologicalSpace α] {f : α → E → F} (K : Set (E →L[𝕜] F)) namespace FDerivMeasurableAux @@ -873,6 +872,8 @@ end FDerivMeasurableAux open FDerivMeasurableAux +variable [MeasurableSpace α] [OpensMeasurableSpace α] [MeasurableSpace E] [OpensMeasurableSpace E] + theorem measurableSet_of_differentiableAt_of_isComplete_with_param (hf : Continuous f.uncurry) {K : Set (E →L[𝕜] F)} (hK : IsComplete K) : MeasurableSet {p : α × E | DifferentiableAt 𝕜 (f p.1) p.2 ∧ fderiv 𝕜 (f p.1) p.2 ∈ K} := by diff --git a/Mathlib/Analysis/Fourier/AddCircle.lean b/Mathlib/Analysis/Fourier/AddCircle.lean index 732c088472bb7..7ddc90e99180e 100644 --- a/Mathlib/Analysis/Fourier/AddCircle.lean +++ b/Mathlib/Analysis/Fourier/AddCircle.lean @@ -243,7 +243,7 @@ theorem coeFn_fourierLp (p : ℝ≥0∞) [Fact (1 ≤ p)] (n : ℤ) : theorem span_fourierLp_closure_eq_top {p : ℝ≥0∞} [Fact (1 ≤ p)] (hp : p ≠ ∞) : (span ℂ (range (@fourierLp T _ p _))).topologicalClosure = ⊤ := by convert - (ContinuousMap.toLp_denseRange ℂ (@haarAddCircle T hT) hp ℂ).topologicalClosure_map_submodule + (ContinuousMap.toLp_denseRange ℂ (@haarAddCircle T hT) ℂ hp).topologicalClosure_map_submodule span_fourier_closure_eq_top erw [map_span, range_comp] simp only [ContinuousLinearMap.coe_coe] diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean index e35d4c0f03deb..a054142703272 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/Separation.lean @@ -75,8 +75,11 @@ theorem separate_convex_open_set [TopologicalSpace E] [AddCommGroup E] [Topologi one_le_gauge_of_not_mem (hs₁.starConvex hs₀) (absorbent_nhds_zero <| hs₂.mem_nhds hs₀).absorbs hx₀ -variable [TopologicalSpace E] [AddCommGroup E] [TopologicalAddGroup E] [Module ℝ E] - [ContinuousSMul ℝ E] {s t : Set E} {x y : E} +variable [TopologicalSpace E] [AddCommGroup E] [Module ℝ E] + {s t : Set E} {x y : E} +section + +variable [TopologicalAddGroup E] [ContinuousSMul ℝ E] /-- A version of the **Hahn-Banach theorem**: given disjoint convex sets `s`, `t` where `s` is open, there is a continuous linear functional which separates them. -/ @@ -203,6 +206,8 @@ theorem iInter_halfspaces_eq (hs₁ : Convex ℝ s) (hs₂ : IsClosed s) : obtain ⟨y, hy, hxy⟩ := hx l exact ((hxy.trans_lt (hlA y hy)).trans hl).not_le le_rfl +end + namespace RCLike variable [RCLike 𝕜] [Module 𝕜 E] [ContinuousSMul 𝕜 E] [IsScalarTower ℝ 𝕜 E] @@ -218,12 +223,12 @@ noncomputable def extendTo𝕜'ₗ : (E →L[ℝ] ℝ) →ₗ[ℝ] (E →L[𝕜] map_smul' := by intros; ext; simp [h, real_smul_eq_coe_mul]; ring } @[simp] -lemma re_extendTo𝕜'ₗ (g : E →L[ℝ] ℝ) (x : E) : re ((extendTo𝕜'ₗ g) x : 𝕜) = g x := by +lemma re_extendTo𝕜'ₗ (g : E →L[ℝ] ℝ) (x : E) : re ((extendTo𝕜'ₗ g) x : 𝕜) = g x := by have h g (x : E) : extendTo𝕜'ₗ g x = ((g x : 𝕜) - (I : 𝕜) * (g ((I : 𝕜) • x) : 𝕜)) := rfl simp only [h , map_sub, ofReal_re, mul_re, I_re, zero_mul, ofReal_im, mul_zero, sub_self, sub_zero] -variable [ContinuousSMul ℝ E] +variable [TopologicalAddGroup E] [ContinuousSMul ℝ E] theorem separate_convex_open_set {s : Set E} (hs₀ : (0 : E) ∈ s) (hs₁ : Convex ℝ s) (hs₂ : IsOpen s) {x₀ : E} (hx₀ : x₀ ∉ s) : diff --git a/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean b/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean index d8515484fe143..570e89aa491e2 100644 --- a/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean +++ b/Mathlib/Combinatorics/SimpleGraph/LapMatrix.lean @@ -28,7 +28,14 @@ open Finset Matrix namespace SimpleGraph variable {V : Type*} (R : Type*) -variable [Fintype V] [DecidableEq V] (G : SimpleGraph V) [DecidableRel G.Adj] +variable [Fintype V] (G : SimpleGraph V) [DecidableRel G.Adj] + +theorem degree_eq_sum_if_adj {R : Type*} [AddCommMonoidWithOne R] (i : V) : + (G.degree i : R) = ∑ j : V, if G.Adj i j then 1 else 0 := by + unfold degree neighborFinset neighborSet + rw [sum_boole, Set.toFinset_setOf] + +variable [DecidableEq V] /-- The diagonal matrix consisting of the degrees of the vertices in the graph. -/ def degMatrix [AddMonoidWithOne R] : Matrix V V R := Matrix.diagonal (G.degree ·) @@ -64,11 +71,6 @@ theorem dotProduct_mulVec_degMatrix [CommRing R] (x : V → R) : variable (R) -theorem degree_eq_sum_if_adj [AddCommMonoidWithOne R] (i : V) : - (G.degree i : R) = ∑ j : V, if G.Adj i j then 1 else 0 := by - unfold degree neighborFinset neighborSet - rw [sum_boole, Set.toFinset_setOf] - /-- Let $L$ be the graph Laplacian and let $x \in \mathbb{R}$, then $$x^{\top} L x = \sum_{i \sim j} (x_{i}-x_{j})^{2}$$, where $\sim$ denotes the adjacency relation -/ diff --git a/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean b/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean index 8fb7c52d6bcbf..97804c5a82769 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean @@ -19,11 +19,10 @@ open Set ChartedSpace SmoothManifoldWithCorners open scoped Topology Manifold variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] - -- declare a smooth manifold `M` over the pair `(E, H)`. + -- declare a charted space `M` over the pair `(E, H)`. {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] - [SmoothManifoldWithCorners I M] -- declare a smooth manifold `M'` over the pair `(E', H')`. {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] @@ -120,49 +119,6 @@ theorem ContinuousLinearMap.contMDiffOn (L : E →L[𝕜] F) {s} : ContMDiffOn theorem ContinuousLinearMap.smooth (L : E →L[𝕜] F) : Smooth 𝓘(𝕜, E) 𝓘(𝕜, F) L := L.contMDiff -theorem ContMDiffWithinAt.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} {s : Set M} {x : M} - (hg : ContMDiffWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₃) n g s x) - (hf : ContMDiffWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₁) n f s x) : - ContMDiffWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) n (fun x => (g x).comp (f x)) s x := - ContDiff.comp_contMDiffWithinAt (g := fun x : (F₁ →L[𝕜] F₃) × (F₂ →L[𝕜] F₁) => x.1.comp x.2) - (f := fun x => (g x, f x)) (contDiff_fst.clm_comp contDiff_snd) (hg.prod_mk_space hf) - -theorem ContMDiffAt.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} {x : M} - (hg : ContMDiffAt I 𝓘(𝕜, F₁ →L[𝕜] F₃) n g x) (hf : ContMDiffAt I 𝓘(𝕜, F₂ →L[𝕜] F₁) n f x) : - ContMDiffAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) n (fun x => (g x).comp (f x)) x := - (hg.contMDiffWithinAt.clm_comp hf.contMDiffWithinAt).contMDiffAt Filter.univ_mem - -theorem ContMDiffOn.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} {s : Set M} - (hg : ContMDiffOn I 𝓘(𝕜, F₁ →L[𝕜] F₃) n g s) (hf : ContMDiffOn I 𝓘(𝕜, F₂ →L[𝕜] F₁) n f s) : - ContMDiffOn I 𝓘(𝕜, F₂ →L[𝕜] F₃) n (fun x => (g x).comp (f x)) s := fun x hx => - (hg x hx).clm_comp (hf x hx) - -theorem ContMDiff.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} - (hg : ContMDiff I 𝓘(𝕜, F₁ →L[𝕜] F₃) n g) (hf : ContMDiff I 𝓘(𝕜, F₂ →L[𝕜] F₁) n f) : - ContMDiff I 𝓘(𝕜, F₂ →L[𝕜] F₃) n fun x => (g x).comp (f x) := fun x => (hg x).clm_comp (hf x) - -theorem ContMDiffWithinAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {s : Set M} {x : M} - (hg : ContMDiffWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) n g s x) - (hf : ContMDiffWithinAt I 𝓘(𝕜, F₁) n f s x) : - ContMDiffWithinAt I 𝓘(𝕜, F₂) n (fun x => g x (f x)) s x := - @ContDiffWithinAt.comp_contMDiffWithinAt _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ - (fun x : (F₁ →L[𝕜] F₂) × F₁ => x.1 x.2) (fun x => (g x, f x)) s _ x - (by apply ContDiff.contDiffAt; exact contDiff_fst.clm_apply contDiff_snd) (hg.prod_mk_space hf) - (by simp_rw [preimage_univ, subset_univ]) - -nonrec theorem ContMDiffAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {x : M} - (hg : ContMDiffAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) n g x) (hf : ContMDiffAt I 𝓘(𝕜, F₁) n f x) : - ContMDiffAt I 𝓘(𝕜, F₂) n (fun x => g x (f x)) x := - hg.clm_apply hf - -theorem ContMDiffOn.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {s : Set M} - (hg : ContMDiffOn I 𝓘(𝕜, F₁ →L[𝕜] F₂) n g s) (hf : ContMDiffOn I 𝓘(𝕜, F₁) n f s) : - ContMDiffOn I 𝓘(𝕜, F₂) n (fun x => g x (f x)) s := fun x hx => (hg x hx).clm_apply (hf x hx) - -theorem ContMDiff.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} - (hg : ContMDiff I 𝓘(𝕜, F₁ →L[𝕜] F₂) n g) (hf : ContMDiff I 𝓘(𝕜, F₁) n f) : - ContMDiff I 𝓘(𝕜, F₂) n fun x => g x (f x) := fun x => (hg x).clm_apply (hf x) - theorem ContMDiffWithinAt.clm_precomp {f : M → F₁ →L[𝕜] F₂} {s : Set M} {x : M} (hf : ContMDiffWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) n f s x) : ContMDiffWithinAt I 𝓘(𝕜, (F₂ →L[𝕜] F₃) →L[𝕜] (F₁ →L[𝕜] F₃)) n @@ -211,6 +167,52 @@ theorem ContMDiff.clm_postcomp {f : M → F₂ →L[𝕜] F₃} (hf : ContMDiff (fun y ↦ (f y).postcomp F₁ : M → (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) := fun x ↦ (hf x).clm_postcomp +-- Now make `M` a smooth manifold. +variable [SmoothManifoldWithCorners I M] + +theorem ContMDiffWithinAt.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} {s : Set M} {x : M} + (hg : ContMDiffWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₃) n g s x) + (hf : ContMDiffWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₁) n f s x) : + ContMDiffWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) n (fun x => (g x).comp (f x)) s x := + ContDiff.comp_contMDiffWithinAt (g := fun x : (F₁ →L[𝕜] F₃) × (F₂ →L[𝕜] F₁) => x.1.comp x.2) + (f := fun x => (g x, f x)) (contDiff_fst.clm_comp contDiff_snd) (hg.prod_mk_space hf) + +theorem ContMDiffAt.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} {x : M} + (hg : ContMDiffAt I 𝓘(𝕜, F₁ →L[𝕜] F₃) n g x) (hf : ContMDiffAt I 𝓘(𝕜, F₂ →L[𝕜] F₁) n f x) : + ContMDiffAt I 𝓘(𝕜, F₂ →L[𝕜] F₃) n (fun x => (g x).comp (f x)) x := + (hg.contMDiffWithinAt.clm_comp hf.contMDiffWithinAt).contMDiffAt Filter.univ_mem + +theorem ContMDiffOn.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} {s : Set M} + (hg : ContMDiffOn I 𝓘(𝕜, F₁ →L[𝕜] F₃) n g s) (hf : ContMDiffOn I 𝓘(𝕜, F₂ →L[𝕜] F₁) n f s) : + ContMDiffOn I 𝓘(𝕜, F₂ →L[𝕜] F₃) n (fun x => (g x).comp (f x)) s := fun x hx => + (hg x hx).clm_comp (hf x hx) + +theorem ContMDiff.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} + (hg : ContMDiff I 𝓘(𝕜, F₁ →L[𝕜] F₃) n g) (hf : ContMDiff I 𝓘(𝕜, F₂ →L[𝕜] F₁) n f) : + ContMDiff I 𝓘(𝕜, F₂ →L[𝕜] F₃) n fun x => (g x).comp (f x) := fun x => (hg x).clm_comp (hf x) + +theorem ContMDiffWithinAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {s : Set M} {x : M} + (hg : ContMDiffWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) n g s x) + (hf : ContMDiffWithinAt I 𝓘(𝕜, F₁) n f s x) : + ContMDiffWithinAt I 𝓘(𝕜, F₂) n (fun x => g x (f x)) s x := + ContDiffWithinAt.comp_contMDiffWithinAt + (g := fun x : (F₁ →L[𝕜] F₂) × F₁ => x.1 x.2) + (by apply ContDiff.contDiffAt; exact contDiff_fst.clm_apply contDiff_snd) (hg.prod_mk_space hf) + (by simp_rw [preimage_univ, subset_univ]) + +nonrec theorem ContMDiffAt.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {x : M} + (hg : ContMDiffAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) n g x) (hf : ContMDiffAt I 𝓘(𝕜, F₁) n f x) : + ContMDiffAt I 𝓘(𝕜, F₂) n (fun x => g x (f x)) x := + hg.clm_apply hf + +theorem ContMDiffOn.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} {s : Set M} + (hg : ContMDiffOn I 𝓘(𝕜, F₁ →L[𝕜] F₂) n g s) (hf : ContMDiffOn I 𝓘(𝕜, F₁) n f s) : + ContMDiffOn I 𝓘(𝕜, F₂) n (fun x => g x (f x)) s := fun x hx => (hg x hx).clm_apply (hf x hx) + +theorem ContMDiff.clm_apply {g : M → F₁ →L[𝕜] F₂} {f : M → F₁} + (hg : ContMDiff I 𝓘(𝕜, F₁ →L[𝕜] F₂) n g) (hf : ContMDiff I 𝓘(𝕜, F₁) n f) : + ContMDiff I 𝓘(𝕜, F₂) n fun x => g x (f x) := fun x => (hg x).clm_apply (hf x) + theorem ContMDiffWithinAt.cle_arrowCongr {f : M → F₁ ≃L[𝕜] F₂} {g : M → F₃ ≃L[𝕜] F₄} {s : Set M} {x : M} (hf : ContMDiffWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₁) n (fun x ↦ ((f x).symm : F₂ →L[𝕜] F₁)) s x) diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean index 6ef4a2c6de4d0..6cd5c2919d09b 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean @@ -20,30 +20,22 @@ open Set Function Filter ChartedSpace SmoothManifoldWithCorners open scoped Topology Manifold variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] - -- declare a smooth manifold `M` over the pair `(E, H)`. + -- declare a charted space `M` over the pair `(E, H)`. {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] - [SmoothManifoldWithCorners I M] - -- declare a smooth manifold `M'` over the pair `(E', H')`. + -- declare a charted space `M'` over the pair `(E', H')`. {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] (I' : ModelWithCorners 𝕜 E' H') {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] - [SmoothManifoldWithCorners I' M'] - -- declare a manifold `M''` over the pair `(E'', H'')`. - {E'' : Type*} - [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] - {I'' : ModelWithCorners 𝕜 E'' H''} {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] - -- declare a smooth manifold `N` over the pair `(F, G)`. + -- declare a charted space `N` over the pair `(F, G)`. {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} [TopologicalSpace G] {J : ModelWithCorners 𝕜 F G} {N : Type*} [TopologicalSpace N] [ChartedSpace G N] - [SmoothManifoldWithCorners J N] - -- declare a smooth manifold `N'` over the pair `(F', G')`. + -- declare a charted space `N'` over the pair `(F', G')`. {F' : Type*} [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] {G' : Type*} [TopologicalSpace G'] {J' : ModelWithCorners 𝕜 F' G'} {N' : Type*} [TopologicalSpace N'] [ChartedSpace G' N'] - [SmoothManifoldWithCorners J' N'] -- F₁, F₂, F₃, F₄ are normed spaces {F₁ : Type*} [NormedAddCommGroup F₁] [NormedSpace 𝕜 F₁] {F₂ : Type*} [NormedAddCommGroup F₂] diff --git a/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean b/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean index 3dc06c9196321..e233de0d891e6 100644 --- a/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean +++ b/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean @@ -443,7 +443,7 @@ end CommRing section StarOrderedRing -variable {𝕜 : Type*} [CommRing 𝕜] [PartialOrder 𝕜] [StarRing 𝕜] [StarOrderedRing 𝕜] +variable {𝕜 : Type*} [CommRing 𝕜] [StarRing 𝕜] scoped infixl:65 " ⊕ᵥ " => Sum.elim @@ -491,6 +491,8 @@ theorem IsHermitian.fromBlocks₂₂ [Fintype n] [DecidableEq n] (A : Matrix m m fromBlocks_submatrix_sum_swap_sum_swap] convert IsHermitian.fromBlocks₁₁ _ _ hD <;> simp +variable [PartialOrder 𝕜] [StarOrderedRing 𝕜] + theorem PosSemidef.fromBlocks₁₁ [Fintype m] [DecidableEq m] [Fintype n] {A : Matrix m m 𝕜} (B : Matrix m n 𝕜) (D : Matrix n n 𝕜) (hA : A.PosDef) [Invertible A] : (fromBlocks A B Bᴴ D).PosSemidef ↔ (D - Bᴴ * A⁻¹ * B).PosSemidef := by diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index e47b7dd06d224..9e03eae8ea344 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -470,7 +470,8 @@ end NNReal spanning measurable sets with finite measure on which `f` is bounded. See also `StronglyMeasurable.exists_spanning_measurableSet_norm_le` for functions into normed groups. -/ -theorem exists_spanning_measurableSet_le {m : MeasurableSpace α} {f : α → ℝ≥0} +-- We redeclare `α` to temporarily avoid the `[MeasurableSpace α]` instance. +theorem exists_spanning_measurableSet_le {α : Type*} {m : MeasurableSpace α} {f : α → ℝ≥0} (hf : Measurable f) (μ : Measure α) [SigmaFinite μ] : ∃ s : ℕ → Set α, (∀ n, MeasurableSet (s n) ∧ μ (s n) < ∞ ∧ ∀ x ∈ s n, f x ≤ n) ∧ diff --git a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean index 75e4a6fbdd0a2..be0467e7351a6 100644 --- a/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Polish/Basic.lean @@ -793,9 +793,9 @@ theorem _root_.IsClosed.measurableSet_image_of_continuousOn_injOn · rwa [continuousOn_iff_continuous_restrict] at f_cont · rwa [injOn_iff_injective] at f_inj -variable {α β : Type*} [tβ : TopologicalSpace β] [T2Space β] [MeasurableSpace β] - [MeasurableSpace α] - {s : Set γ} {f : γ → β} +variable {α β : Type*} [MeasurableSpace β] +section +variable [tβ : TopologicalSpace β] [T2Space β] [MeasurableSpace α] {s : Set γ} {f : γ → β} /-- The Lusin-Souslin theorem: if `s` is Borel-measurable in a Polish space, then its image under a continuous injective map is also Borel-measurable. -/ @@ -879,7 +879,7 @@ theorem borel_eq_borel_of_le {t t' : TopologicalSpace γ} refine le_antisymm ?_ (borel_anti hle) intro s hs have e := @Continuous.measurableEmbedding - _ _ t' _ (@borel _ t') _ (@BorelSpace.mk _ _ (borel γ) rfl) + _ _ (@borel _ t') t' _ _ (@BorelSpace.mk _ _ (borel γ) rfl) t _ (@borel _ t) (@BorelSpace.mk _ t (@borel _ t) rfl) (continuous_id_of_le hle) injective_id convert e.measurableSet_image.2 hs simp only [id_eq, image_id'] @@ -898,6 +898,8 @@ theorem isClopenable_iff_measurableSet · exact MeasurableSpace.measurableSet_generateFrom s_open infer_instance +end + /-- The set of points for which a sequence of measurable functions converges to a given function is measurable. -/ @[measurability] diff --git a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean index bfe6a02c9efa5..d16ec22981a70 100644 --- a/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean +++ b/Mathlib/MeasureTheory/Function/ContinuousMapDense.lean @@ -61,7 +61,7 @@ open scoped ENNReal NNReal Topology BoundedContinuousFunction open MeasureTheory TopologicalSpace ContinuousMap Set Bornology -variable {α : Type*} [TopologicalSpace α] [NormalSpace α] [R1Space α] +variable {α : Type*} [TopologicalSpace α] [NormalSpace α] [MeasurableSpace α] [BorelSpace α] variable {E : Type*} [NormedAddCommGroup E] {μ : Measure α} {p : ℝ≥0∞} @@ -137,7 +137,8 @@ alias exists_continuous_snorm_sub_le_of_closed := exists_continuous_eLpNorm_sub_ /-- In a locally compact space, any function in `ℒp` can be approximated by compactly supported continuous functions when `p < ∞`, version in terms of `eLpNorm`. -/ -theorem Memℒp.exists_hasCompactSupport_eLpNorm_sub_le [WeaklyLocallyCompactSpace α] [μ.Regular] +theorem Memℒp.exists_hasCompactSupport_eLpNorm_sub_le + [R1Space α] [WeaklyLocallyCompactSpace α] [μ.Regular] (hp : p ≠ ∞) {f : α → E} (hf : Memℒp f p μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : ∃ g : α → E, HasCompactSupport g ∧ eLpNorm (f - g) p μ ≤ ε ∧ Continuous g ∧ Memℒp g p μ := by suffices H : @@ -194,7 +195,7 @@ alias Memℒp.exists_hasCompactSupport_snorm_sub_le := Memℒp.exists_hasCompact /-- In a locally compact space, any function in `ℒp` can be approximated by compactly supported continuous functions when `0 < p < ∞`, version in terms of `∫`. -/ theorem Memℒp.exists_hasCompactSupport_integral_rpow_sub_le - [WeaklyLocallyCompactSpace α] [μ.Regular] + [R1Space α] [WeaklyLocallyCompactSpace α] [μ.Regular] {p : ℝ} (hp : 0 < p) {f : α → E} (hf : Memℒp f (ENNReal.ofReal p) μ) {ε : ℝ} (hε : 0 < ε) : ∃ g : α → E, HasCompactSupport g ∧ @@ -215,7 +216,7 @@ theorem Memℒp.exists_hasCompactSupport_integral_rpow_sub_le /-- In a locally compact space, any integrable function can be approximated by compactly supported continuous functions, version in terms of `∫⁻`. -/ theorem Integrable.exists_hasCompactSupport_lintegral_sub_le - [WeaklyLocallyCompactSpace α] [μ.Regular] + [R1Space α] [WeaklyLocallyCompactSpace α] [μ.Regular] {f : α → E} (hf : Integrable f μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : ∃ g : α → E, HasCompactSupport g ∧ (∫⁻ x, ‖f x - g x‖₊ ∂μ) ≤ ε ∧ Continuous g ∧ Integrable g μ := by @@ -225,7 +226,7 @@ theorem Integrable.exists_hasCompactSupport_lintegral_sub_le /-- In a locally compact space, any integrable function can be approximated by compactly supported continuous functions, version in terms of `∫`. -/ theorem Integrable.exists_hasCompactSupport_integral_sub_le - [WeaklyLocallyCompactSpace α] [μ.Regular] + [R1Space α] [WeaklyLocallyCompactSpace α] [μ.Regular] {f : α → E} (hf : Integrable f μ) {ε : ℝ} (hε : 0 < ε) : ∃ g : α → E, HasCompactSupport g ∧ (∫ x, ‖f x - g x‖ ∂μ) ≤ ε ∧ Continuous g ∧ Integrable g μ := by @@ -347,13 +348,13 @@ end Lp end MeasureTheory -variable [SecondCountableTopologyEither α E] [_i : Fact (1 ≤ p)] (hp : p ≠ ∞) +variable [SecondCountableTopologyEither α E] [_i : Fact (1 ≤ p)] variable (𝕜 : Type*) [NormedField 𝕜] [NormedAlgebra ℝ 𝕜] [NormedSpace 𝕜 E] variable (E) (μ) namespace BoundedContinuousFunction -theorem toLp_denseRange [μ.WeaklyRegular] [IsFiniteMeasure μ] : +theorem toLp_denseRange [μ.WeaklyRegular] [IsFiniteMeasure μ] (hp : p ≠ ∞) : DenseRange (toLp p μ 𝕜 : (α →ᵇ E) →L[𝕜] Lp E p μ) := by haveI : NormedSpace ℝ E := RestrictScalars.normedSpace ℝ 𝕜 E simpa only [← range_toLp p μ (𝕜 := 𝕜)] @@ -366,9 +367,9 @@ namespace ContinuousMap /-- Continuous functions are dense in `MeasureTheory.Lp`, `1 ≤ p < ∞`. This theorem assumes that the domain is a compact space because otherwise `ContinuousMap.toLp` is undefined. Use `BoundedContinuousFunction.toLp_denseRange` if the domain is not a compact space. -/ -theorem toLp_denseRange [CompactSpace α] [μ.WeaklyRegular] [IsFiniteMeasure μ] : +theorem toLp_denseRange [CompactSpace α] [μ.WeaklyRegular] [IsFiniteMeasure μ] (hp : p ≠ ∞) : DenseRange (toLp p μ 𝕜 : C(α, E) →L[𝕜] Lp E p μ) := by - refine (BoundedContinuousFunction.toLp_denseRange _ _ hp 𝕜).mono ?_ + refine (BoundedContinuousFunction.toLp_denseRange _ _ 𝕜 hp).mono ?_ refine range_subset_iff.2 fun f ↦ ?_ exact ⟨f.toContinuousMap, rfl⟩ diff --git a/Mathlib/MeasureTheory/Group/Integral.lean b/Mathlib/MeasureTheory/Group/Integral.lean index aed3991da1182..665790505f3cd 100644 --- a/Mathlib/MeasureTheory/Group/Integral.lean +++ b/Mathlib/MeasureTheory/Group/Integral.lean @@ -20,7 +20,7 @@ open Measure TopologicalSpace open scoped ENNReal variable {𝕜 M α G E F : Type*} [MeasurableSpace G] -variable [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] [NormedAddCommGroup F] +variable [NormedAddCommGroup E] [NormedSpace ℝ E] [NormedAddCommGroup F] variable {μ : Measure G} {f : G → E} {g : G} section MeasurableInv diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 0e02937ac68d2..7bc444ca0b8c6 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -699,8 +699,8 @@ Define the Bochner integral on functions generally to be the `L1` Bochner integr functions, and 0 otherwise; prove its basic properties. -/ -variable [NormedAddCommGroup E] [NormedSpace ℝ E] [hE : CompleteSpace E] [NontriviallyNormedField 𝕜] - [NormedSpace 𝕜 E] [SMulCommClass ℝ 𝕜 E] [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpace F] +variable [NormedAddCommGroup E] [hE : CompleteSpace E] [NontriviallyNormedField 𝕜] + [NormedSpace 𝕜 E] [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpace F] {G : Type*} [NormedAddCommGroup G] [NormedSpace ℝ G] open Classical in @@ -730,6 +730,7 @@ section Properties open ContinuousLinearMap MeasureTheory.SimpleFunc +variable [NormedSpace ℝ E] [SMulCommClass ℝ 𝕜 E] variable {f g : α → E} {m : MeasurableSpace α} {μ : Measure α} theorem integral_eq (f : α → E) (hf : Integrable f μ) : ∫ a, f a ∂μ = L1.integral (hf.toL1 f) := by diff --git a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean index a07c2b8aa6b73..1a10dec86151f 100644 --- a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean +++ b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean @@ -26,13 +26,13 @@ lemma apply_le_nndist_zero {X : Type*} [TopologicalSpace X] (f : X →ᵇ ℝ≥ convert nndist_coe_le_nndist x simp only [coe_zero, Pi.zero_apply, NNReal.nndist_zero_eq_val] -variable {X : Type*} [MeasurableSpace X] [TopologicalSpace X] [OpensMeasurableSpace X] +variable {X : Type*} [MeasurableSpace X] [TopologicalSpace X] lemma lintegral_le_edist_mul (f : X →ᵇ ℝ≥0) (μ : Measure X) : (∫⁻ x, f x ∂μ) ≤ edist 0 f * (μ Set.univ) := le_trans (lintegral_mono (fun x ↦ ENNReal.coe_le_coe.mpr (f.apply_le_nndist_zero x))) (by simp) -theorem measurable_coe_ennreal_comp (f : X →ᵇ ℝ≥0) : +theorem measurable_coe_ennreal_comp [OpensMeasurableSpace X] (f : X →ᵇ ℝ≥0) : Measurable fun x ↦ (f x : ℝ≥0∞) := measurable_coe_nnreal_ennreal.comp f.continuous.measurable @@ -44,12 +44,13 @@ theorem lintegral_lt_top_of_nnreal (f : X →ᵇ ℝ≥0) : ∫⁻ x, f x ∂μ have key := BoundedContinuousFunction.NNReal.upper_bound f x rwa [ENNReal.coe_le_coe] -theorem integrable_of_nnreal (f : X →ᵇ ℝ≥0) : Integrable (((↑) : ℝ≥0 → ℝ) ∘ ⇑f) μ := by +theorem integrable_of_nnreal [OpensMeasurableSpace X] (f : X →ᵇ ℝ≥0) : + Integrable (((↑) : ℝ≥0 → ℝ) ∘ ⇑f) μ := by refine ⟨(NNReal.continuous_coe.comp f.continuous).measurable.aestronglyMeasurable, ?_⟩ simp only [HasFiniteIntegral, Function.comp_apply, NNReal.nnnorm_eq] exact lintegral_lt_top_of_nnreal _ f -theorem integral_eq_integral_nnrealPart_sub (f : X →ᵇ ℝ) : +theorem integral_eq_integral_nnrealPart_sub [OpensMeasurableSpace X] (f : X →ᵇ ℝ) : ∫ x, f x ∂μ = (∫ x, (f.nnrealPart x : ℝ) ∂μ) - ∫ x, ((-f).nnrealPart x : ℝ) ∂μ := by simp only [f.self_eq_nnrealPart_sub_nnrealPart_neg, Pi.sub_apply, integral_sub, integrable_of_nnreal] @@ -58,7 +59,7 @@ theorem integral_eq_integral_nnrealPart_sub (f : X →ᵇ ℝ) : theorem lintegral_of_real_lt_top (f : X →ᵇ ℝ) : ∫⁻ x, ENNReal.ofReal (f x) ∂μ < ∞ := lintegral_lt_top_of_nnreal _ f.nnrealPart -theorem toReal_lintegral_coe_eq_integral (f : X →ᵇ ℝ≥0) (μ : Measure X) : +theorem toReal_lintegral_coe_eq_integral [OpensMeasurableSpace X] (f : X →ᵇ ℝ≥0) (μ : Measure X) : (∫⁻ x, (f x : ℝ≥0∞) ∂μ).toReal = ∫ x, (f x : ℝ) ∂μ := by rw [integral_eq_lintegral_of_nonneg_ae _ (by simpa [Function.comp_apply] using (NNReal.continuous_coe.comp f.continuous).measurable.aestronglyMeasurable)] @@ -69,10 +70,9 @@ end NNRealValued section BochnerIntegral -variable {X : Type*} [MeasurableSpace X] [TopologicalSpace X] [OpensMeasurableSpace X] +variable {X : Type*} [MeasurableSpace X] [TopologicalSpace X] variable (μ : Measure X) -variable {E : Type*} [NormedAddCommGroup E] [SecondCountableTopology E] -variable [MeasurableSpace E] [BorelSpace E] +variable {E : Type*} [NormedAddCommGroup E] lemma lintegral_nnnorm_le (f : X →ᵇ E) : ∫⁻ x, ‖f x‖₊ ∂μ ≤ ‖f‖₊ * (μ Set.univ) := by @@ -80,6 +80,8 @@ lemma lintegral_nnnorm_le (f : X →ᵇ E) : _ ≤ ∫⁻ _, ‖f‖₊ ∂μ := by gcongr; apply nnnorm_coe_le_nnnorm _ = ‖f‖₊ * (μ Set.univ) := by rw [lintegral_const] +variable [OpensMeasurableSpace X] [SecondCountableTopology E] [MeasurableSpace E] [BorelSpace E] + lemma integrable [IsFiniteMeasure μ] (f : X →ᵇ E) : Integrable f μ := by refine ⟨f.continuous.measurable.aestronglyMeasurable, (hasFiniteIntegral_def _ _).mp ?_⟩ diff --git a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean index cf60664ba60af..cc3daa531971b 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean @@ -19,11 +19,31 @@ the canonical `volume` from the `MeasureSpace` instance. open FiniteDimensional MeasureTheory MeasureTheory.Measure Set variable {ι E F : Type*} -variable [Fintype ι] [NormedAddCommGroup F] [InnerProductSpace ℝ F] [FiniteDimensional ℝ F] - [MeasurableSpace F] [BorelSpace F] -section +variable [NormedAddCommGroup F] [InnerProductSpace ℝ F] + [NormedAddCommGroup E] [InnerProductSpace ℝ E] + [MeasurableSpace E] [BorelSpace E] [MeasurableSpace F] [BorelSpace F] + +namespace LinearIsometryEquiv + +variable (f : E ≃ₗᵢ[ℝ] F) + +/-- Every linear isometry equivalence is a measurable equivalence. -/ +def toMeasureEquiv : E ≃ᵐ F where + toEquiv := f + measurable_toFun := f.continuous.measurable + measurable_invFun := f.symm.continuous.measurable + +@[simp] theorem coe_toMeasureEquiv : (f.toMeasureEquiv : E → F) = f := rfl + +theorem toMeasureEquiv_symm : f.toMeasureEquiv.symm = f.symm.toMeasureEquiv := rfl + +end LinearIsometryEquiv + +variable [Fintype ι] +variable [FiniteDimensional ℝ E] [FiniteDimensional ℝ F] +section variable {m n : ℕ} [_i : Fact (finrank ℝ F = n)] /-- The volume form coming from an orientation in an inner product space gives measure `1` to the @@ -122,27 +142,13 @@ end PiLp namespace LinearIsometryEquiv -variable [NormedAddCommGroup E] [InnerProductSpace ℝ E] [FiniteDimensional ℝ E] - [MeasurableSpace E] [BorelSpace E] - -variable (f : E ≃ₗᵢ[ℝ] F) - /-- Every linear isometry on a real finite dimensional Hilbert space is measure-preserving. -/ -theorem measurePreserving : MeasurePreserving f := by +theorem measurePreserving (f : E ≃ₗᵢ[ℝ] F) : + MeasurePreserving f := by refine ⟨f.continuous.measurable, ?_⟩ rcases exists_orthonormalBasis ℝ E with ⟨w, b, _hw⟩ erw [← OrthonormalBasis.addHaar_eq_volume b, ← OrthonormalBasis.addHaar_eq_volume (b.map f), Basis.map_addHaar _ f.toContinuousLinearEquiv] congr -/-- Every linear isometry equivalence is a measurable equivalence. -/ -def toMeasureEquiv : E ≃ᵐ F where - toEquiv := f - measurable_toFun := f.continuous.measurable - measurable_invFun := f.symm.continuous.measurable - -@[simp] theorem coe_toMeasureEquiv : (f.toMeasureEquiv : E → F) = f := rfl - -theorem toMeasureEquiv_symm : f.toMeasureEquiv.symm = f.symm.toMeasureEquiv := rfl - end LinearIsometryEquiv diff --git a/Mathlib/NumberTheory/Cyclotomic/Basic.lean b/Mathlib/NumberTheory/Cyclotomic/Basic.lean index 4027bbcc9a884..e066e17c51f07 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Basic.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Basic.lean @@ -521,7 +521,7 @@ end CyclotomicField section IsDomain -variable [Algebra A K] [IsFractionRing A K] +variable [Algebra A K] section CyclotomicRing @@ -541,7 +541,8 @@ instance CyclotomicField.algebra' {R : Type*} [CommRing R] [Algebra R K] : instance {R : Type*} [CommRing R] [Algebra R K] : IsScalarTower R K (CyclotomicField n K) := SplittingField.isScalarTower _ -instance CyclotomicField.noZeroSMulDivisors : NoZeroSMulDivisors A (CyclotomicField n K) := by +instance CyclotomicField.noZeroSMulDivisors [IsFractionRing A K] : + NoZeroSMulDivisors A (CyclotomicField n K) := by refine NoZeroSMulDivisors.of_algebraMap_injective ?_ rw [IsScalarTower.algebraMap_eq A K (CyclotomicField n K)] exact @@ -578,10 +579,12 @@ instance algebraBase : Algebra A (CyclotomicRing n A K) := -- but there is at `reducible_and_instances` #10906 example {n : ℕ+} : CyclotomicRing.algebraBase n ℤ ℚ = Ring.toIntAlgebra _ := rfl -instance : NoZeroSMulDivisors A (CyclotomicRing n A K) := +instance [IsFractionRing A K] : + NoZeroSMulDivisors A (CyclotomicRing n A K) := (adjoin A _).noZeroSMulDivisors_bot -theorem algebraBase_injective : Function.Injective <| algebraMap A (CyclotomicRing n A K) := +theorem algebraBase_injective [IsFractionRing A K] : + Function.Injective <| algebraMap A (CyclotomicRing n A K) := NoZeroSMulDivisors.algebraMap_injective _ _ instance : Algebra (CyclotomicRing n A K) (CyclotomicField n K) := @@ -597,7 +600,7 @@ instance : NoZeroSMulDivisors (CyclotomicRing n A K) (CyclotomicField n K) := instance : IsScalarTower A (CyclotomicRing n A K) (CyclotomicField n K) := IsScalarTower.subalgebra' _ _ _ _ -instance isCyclotomicExtension [NeZero ((n : ℕ) : A)] : +instance isCyclotomicExtension [IsFractionRing A K] [NeZero ((n : ℕ) : A)] : IsCyclotomicExtension {n} A (CyclotomicRing n A K) where exists_prim_root := @fun a han => by rw [mem_singleton_iff] at han @@ -620,7 +623,7 @@ instance isCyclotomicExtension [NeZero ((n : ℕ) : A)] : · exact Subalgebra.add_mem _ hy hz · exact Subalgebra.mul_mem _ hy hz -instance [IsDomain A] [NeZero ((n : ℕ) : A)] : +instance [IsFractionRing A K] [IsDomain A] [NeZero ((n : ℕ) : A)] : IsFractionRing (CyclotomicRing n A K) (CyclotomicField n K) where map_units' := fun ⟨x, hx⟩ => by rw [isUnit_iff_ne_zero] diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index 9fe35d4f6f364..a570a392aa683 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -40,8 +40,6 @@ variable (R A K : Type*) [CommRing R] [CommRing A] [Field K] open scoped nonZeroDivisors Polynomial -variable [IsDomain A] - section IsIntegralClosure /-! ### `IsIntegralClosure` section @@ -60,7 +58,7 @@ variable [Algebra C L] [IsIntegralClosure C A L] [Algebra A C] [IsScalarTower A /-- If `L` is an algebraic extension of `K = Frac(A)` and `L` has no zero smul divisors by `A`, then `L` is the localization of the integral closure `C` of `A` in `L` at `A⁰`. -/ -theorem IsIntegralClosure.isLocalization [Algebra.IsAlgebraic K L] : +theorem IsIntegralClosure.isLocalization [IsDomain A] [Algebra.IsAlgebraic K L] : IsLocalization (Algebra.algebraMapSubmonoid C A⁰) L := by haveI : IsDomain C := (IsIntegralClosure.equiv A C L (integralClosure A L)).toMulEquiv.isDomain (integralClosure A L) @@ -80,7 +78,7 @@ theorem IsIntegralClosure.isLocalization [Algebra.IsAlgebraic K L] : smul_def] · simp only [IsIntegralClosure.algebraMap_injective C A L h] -theorem IsIntegralClosure.isLocalization_of_isSeparable [Algebra.IsSeparable K L] : +theorem IsIntegralClosure.isLocalization_of_isSeparable [IsDomain A] [Algebra.IsSeparable K L] : IsLocalization (Algebra.algebraMapSubmonoid C A⁰) L := IsIntegralClosure.isLocalization A K L C @@ -107,6 +105,7 @@ theorem integralClosure_le_span_dualBasis [Algebra.IsSeparable K L] {ι : Type*} intro x hx exact ⟨⟨x, hx⟩, rfl⟩ +variable [IsDomain A] variable (A K) /-- Send a set of `x`s in a finite extension `L` of the fraction field of `R` diff --git a/Mathlib/RingTheory/Discriminant.lean b/Mathlib/RingTheory/Discriminant.lean index 7aa87440e0881..f6285f4b7aa27 100644 --- a/Mathlib/RingTheory/Discriminant.lean +++ b/Mathlib/RingTheory/Discriminant.lean @@ -301,7 +301,7 @@ end Field section Int /-- Two (finite) ℤ-bases have the same discriminant. -/ -theorem discr_eq_discr [Fintype ι] (b : Basis ι ℤ A) (b' : Basis ι ℤ A) : +theorem discr_eq_discr (b : Basis ι ℤ A) (b' : Basis ι ℤ A) : Algebra.discr ℤ b = Algebra.discr ℤ b' := by convert Algebra.discr_of_matrix_vecMul b' (b'.toMatrix b) · rw [Basis.toMatrix_map_vecMul] diff --git a/Mathlib/RingTheory/Kaehler/Basic.lean b/Mathlib/RingTheory/Kaehler/Basic.lean index d5a828b146a26..c57b401f8224a 100644 --- a/Mathlib/RingTheory/Kaehler/Basic.lean +++ b/Mathlib/RingTheory/Kaehler/Basic.lean @@ -675,6 +675,7 @@ theorem KaehlerDifferential.kerTotal_map' [Algebra R B] refine congr_arg Set.range ?_ ext; simp [IsScalarTower.algebraMap_eq R A B] +section variable [Algebra R B] [IsScalarTower R A B] [IsScalarTower R S B] [SMulCommClass S A B] /-- The map `Ω[A⁄R] →ₗ[A] Ω[B⁄S]` given a square @@ -786,6 +787,8 @@ lemma KaehlerDifferential.exact_mapBaseChange_map : Function.Exact (mapBaseChange R A B) (map R A B B) := SetLike.ext_iff.mp (range_mapBaseChange R A B).symm +end + /-- The map `I → B ⊗[A] B ⊗[A] Ω[A⁄R]` where `I = ker(A → B)`. -/ @[simps] noncomputable @@ -813,6 +816,8 @@ def KaehlerDifferential.kerCotangentToTensor : lemma KaehlerDifferential.kerCotangentToTensor_toCotangent (x) : kerCotangentToTensor R A B (Ideal.toCotangent _ x) = 1 ⊗ₜ D _ _ x.1 := rfl +variable [Algebra R B] [IsScalarTower R A B] + theorem KaehlerDifferential.range_kerCotangentToTensor (h : Function.Surjective (algebraMap A B)) : LinearMap.range (kerCotangentToTensor R A B) = diff --git a/Mathlib/RingTheory/WittVector/MulCoeff.lean b/Mathlib/RingTheory/WittVector/MulCoeff.lean index f5b7c0ecd8979..edb4720376f65 100644 --- a/Mathlib/RingTheory/WittVector/MulCoeff.lean +++ b/Mathlib/RingTheory/WittVector/MulCoeff.lean @@ -131,7 +131,8 @@ theorem mul_polyOfInterest_aux2 (n : ℕ) : rw [sum_range_succ, add_comm, Nat.sub_self, pow_zero, pow_one] rfl -theorem mul_polyOfInterest_aux3 (n : ℕ) : wittPolyProd p (n + 1) = +-- We redeclare `p` here to locally discard the unneeded `p.Prime` hypothesis. +theorem mul_polyOfInterest_aux3 (p n : ℕ) : wittPolyProd p (n + 1) = -((p : 𝕄) ^ (n + 1) * X (0, n + 1)) * ((p : 𝕄) ^ (n + 1) * X (1, n + 1)) + (p : 𝕄) ^ (n + 1) * X (0, n + 1) * rename (Prod.mk (1 : Fin 2)) (wittPolynomial p ℤ (n + 1)) + (p : 𝕄) ^ (n + 1) * X (1, n + 1) * rename (Prod.mk (0 : Fin 2)) (wittPolynomial p ℤ (n + 1)) + diff --git a/Mathlib/Topology/Algebra/Algebra.lean b/Mathlib/Topology/Algebra/Algebra.lean index 5eff70c7f4cc8..c229209194995 100644 --- a/Mathlib/Topology/Algebra/Algebra.lean +++ b/Mathlib/Topology/Algebra/Algebra.lean @@ -52,6 +52,7 @@ theorem continuousSMul_of_algebraMap [TopologicalSemiring A] (h : Continuous (al ContinuousSMul R A := ⟨(continuous_algebraMap_iff_smul R A).1 h⟩ +section variable [ContinuousSMul R A] /-- The inclusion of the base ring in a topological algebra as a continuous linear map. -/ @@ -67,6 +68,8 @@ theorem algebraMapCLM_coe : ⇑(algebraMapCLM R A) = algebraMap R A := theorem algebraMapCLM_toLinearMap : (algebraMapCLM R A).toLinearMap = Algebra.linearMap R A := rfl +end + /-- If `R` is a discrete topological ring, then any topological ring `S` which is an `R`-algebra is also a topological `R`-algebra. From 86d8be687742783af9bb08321e7d040c89bdc36d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 9 Aug 2024 00:58:24 +0000 Subject: [PATCH 137/359] chore: backports for leanprover/lean4#4814 (part 34) (#15607) --- Mathlib/Algebra/Lie/Derivation/Killing.lean | 7 +- Mathlib/Algebra/Lie/Weights/Killing.lean | 78 ++++++++++--------- Mathlib/Algebra/Lie/Weights/RootSystem.lean | 33 ++++---- .../BoxIntegral/Partition/Filter.lean | 11 ++- .../Calculus/BumpFunction/Normed.lean | 7 +- .../Calculus/ParametricIntervalIntegral.lean | 2 +- Mathlib/Analysis/Convolution.lean | 13 ++-- .../Distribution/AEEqOfIntegralContDiff.lean | 7 +- .../Analysis/Fourier/FourierTransform.lean | 21 +++-- .../Fourier/RiemannLebesgueLemma.lean | 3 - .../PiTensorProduct/InjectiveSeminorm.lean | 10 ++- Mathlib/Geometry/Manifold/BumpFunction.lean | 2 +- .../Geometry/Manifold/ContMDiffMFDeriv.lean | 19 +++-- .../Geometry/Manifold/LocalDiffeomorph.lean | 7 +- .../Manifold/MFDeriv/UniqueDifferential.lean | 16 +++- .../Geometry/Manifold/PartitionOfUnity.lean | 46 ++++++----- .../Geometry/Manifold/PoincareConjecture.lean | 8 +- .../Geometry/Manifold/VectorBundle/Basic.lean | 10 ++- .../Geometry/Manifold/VectorBundle/Hom.lean | 6 +- .../Manifold/VectorBundle/SmoothSection.lean | 70 ++++++++--------- .../Covering/Differentiation.lean | 12 +-- .../Decomposition/RadonNikodym.lean | 2 +- .../ConditionalExpectation/Basic.lean | 3 - Mathlib/MeasureTheory/Integral/Average.lean | 8 +- Mathlib/MeasureTheory/Integral/Bochner.lean | 5 +- .../Integral/FundThmCalculus.lean | 14 ++-- .../Integral/IntervalAverage.lean | 2 +- Mathlib/MeasureTheory/Measure/DiracProba.lean | 12 +-- .../Measure/Haar/NormedSpace.lean | 4 +- .../Measure/LevyProkhorovMetric.lean | 11 +-- .../Measure/WithDensityVectorMeasure.lean | 7 +- Mathlib/Probability/Density.lean | 5 +- Mathlib/Probability/Independence/ZeroOne.lean | 34 +++++--- Mathlib/Probability/Kernel/CondDistrib.lean | 4 +- Mathlib/Probability/Kernel/Condexp.lean | 6 +- .../Kernel/Disintegration/Basic.lean | 4 +- .../Probability/Kernel/IntegralCompProd.lean | 3 +- 37 files changed, 294 insertions(+), 218 deletions(-) diff --git a/Mathlib/Algebra/Lie/Derivation/Killing.lean b/Mathlib/Algebra/Lie/Derivation/Killing.lean index d45e4e1f60842..e6c9325939168 100644 --- a/Mathlib/Algebra/Lie/Derivation/Killing.lean +++ b/Mathlib/Algebra/Lie/Derivation/Killing.lean @@ -27,7 +27,7 @@ namespace LieDerivation.IsKilling section -variable (R L : Type*) [Field R] [LieRing L] [LieAlgebra R L] [Module.Finite R L] +variable (R L : Type*) [Field R] [LieRing L] [LieAlgebra R L] /-- A local notation for the set of (Lie) derivations on `L`. -/ local notation "𝔻" => (LieDerivation R L L) @@ -38,7 +38,8 @@ local notation "𝕀" => (LieHom.range (ad R L)) /-- A local notation for the Killing complement of the ideal range of `ad`. -/ local notation "𝕀ᗮ" => LinearMap.BilinForm.orthogonal (killingForm R 𝔻) 𝕀 -lemma killingForm_restrict_range_ad : (killingForm R 𝔻).restrict 𝕀 = killingForm R 𝕀 := by +lemma killingForm_restrict_range_ad [Module.Finite R L] : + (killingForm R 𝔻).restrict 𝕀 = killingForm R 𝕀 := by rw [← (ad_isIdealMorphism R L).eq, ← LieIdeal.killingForm_eq] rfl @@ -63,6 +64,8 @@ lemma ad_mem_orthogonal_of_mem_orthogonal {D : LieDerivation R L L} (hD : D ∈ simp only [ad_apply_lieDerivation, LieHom.range_coeSubmodule, neg_mem_iff] exact (rangeAdOrthogonal R L).lie_mem hD +variable [Module.Finite R L] + lemma ad_mem_ker_killingForm_ad_range_of_mem_orthogonal {D : LieDerivation R L L} (hD : D ∈ 𝕀ᗮ) (x : L) : ad R L (D x) ∈ (LinearMap.ker (killingForm R 𝕀)).map (LieHom.range (ad R L)).subtype := by diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean index f9fb2a4d49c1e..76f58fc396815 100644 --- a/Mathlib/Algebra/Lie/Weights/Killing.lean +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -38,9 +38,13 @@ variable (R K L : Type*) [CommRing R] [LieRing L] [LieAlgebra R L] [Field K] [Li namespace LieAlgebra +lemma restrict_killingForm (H : LieSubalgebra R L) : + (killingForm R L).restrict H = LieModule.traceForm R H L := + rfl + namespace IsKilling -variable [IsKilling R L] [Module.Finite R L] [Module.Free R L] +variable [IsKilling R L] /-- If the Killing form of a Lie algebra is non-singular, it remains non-singular when restricted to a Cartan subalgebra. -/ @@ -58,10 +62,6 @@ lemma ker_restrict_eq_bot_of_isCartanSubalgebra intro m₀ h₀ m₁ h₁ exact killingForm_eq_zero_of_mem_zeroRoot_mem_posFitting R L H (le_zeroRootSubalgebra R L H h₀) h₁ -lemma restrict_killingForm (H : LieSubalgebra R L) : - (killingForm R L).restrict H = LieModule.traceForm R H L := - rfl - @[simp] lemma ker_traceForm_eq_bot_of_isCartanSubalgebra [IsNoetherian R L] [IsArtinian R L] (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : LinearMap.ker (LieModule.traceForm R H L) = ⊥ := @@ -72,6 +72,8 @@ lemma traceForm_cartan_nondegenerate (LieModule.traceForm R H L).Nondegenerate := by simp [LinearMap.BilinForm.nondegenerate_iff_ker_eq_bot] +variable [Module.Free R L] [Module.Finite R L] + instance instIsLieAbelianOfIsCartanSubalgebra [IsDomain R] [IsPrincipalIdealRing R] [IsArtinian R L] (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : @@ -86,8 +88,10 @@ section Field open FiniteDimensional LieModule Set open Submodule (span subset_span) -variable [FiniteDimensional K L] - (H : LieSubalgebra K L) [H.IsCartanSubalgebra] [IsTriangularizable K H L] +variable [FiniteDimensional K L] (H : LieSubalgebra K L) [H.IsCartanSubalgebra] + +section +variable [IsTriangularizable K H L] /-- For any `α` and `β`, the corresponding root spaces are orthogonal with respect to the Killing form, provided `α + β ≠ 0`. -/ @@ -129,6 +133,7 @@ lemma mem_ker_killingForm_of_mem_rootSpace_of_forall_rootSpace_neg · exact killingForm_apply_eq_zero_of_mem_rootSpace_of_add_ne_zero K L H hx hy hαβ | h0 => simp | hadd => simp_all +end namespace IsKilling @@ -149,6 +154,17 @@ lemma eq_zero_of_isNilpotent_ad_of_mem_isCartanSubalgebra {x : L} (hx : x ∈ H) rw [traceForm_apply_apply, ← LinearMap.mul_eq_comp, LinearMap.zero_apply] exact (LinearMap.isNilpotent_trace_of_isNilpotent (comm.isNilpotent_mul_left hx')).eq_zero +@[simp] +lemma corootSpace_zero_eq_bot : + corootSpace (0 : H → K) = ⊥ := by + refine eq_bot_iff.mpr fun x hx ↦ ?_ + suffices {x | ∃ y ∈ H, ∃ z ∈ H, ⁅y, z⁆ = x} = {0} by simpa [mem_corootSpace, this] using hx + refine eq_singleton_iff_unique_mem.mpr ⟨⟨0, H.zero_mem, 0, H.zero_mem, zero_lie 0⟩, ?_⟩ + rintro - ⟨y, hy, z, hz, rfl⟩ + suffices ⁅(⟨y, hy⟩ : H), (⟨z, hz⟩ : H)⁆ = 0 by + simpa only [Subtype.ext_iff, LieSubalgebra.coe_bracket, ZeroMemClass.coe_zero] using this + simp + variable {K L} in /-- The restriction of the Killing form to a Cartan subalgebra, as a linear equivalence to the dual. -/ @@ -157,7 +173,21 @@ noncomputable def cartanEquivDual : H ≃ₗ[K] Module.Dual K H := (traceForm K H L).toDual <| traceForm_cartan_nondegenerate K L H -variable {K L H} in +variable {K L H} + +/-- The coroot corresponding to a root. -/ +noncomputable def coroot (α : Weight K H L) : H := + 2 • (α <| (cartanEquivDual H).symm α)⁻¹ • (cartanEquivDual H).symm α + +lemma traceForm_coroot (α : Weight K H L) (x : H) : + traceForm K H L (coroot α) x = 2 • (α <| (cartanEquivDual H).symm α)⁻¹ • α x := by + have : cartanEquivDual H ((cartanEquivDual H).symm α) x = α x := by + rw [LinearEquiv.apply_symm_apply, Weight.toLinear_apply] + rw [coroot, map_nsmul, map_smul, LinearMap.smul_apply, LinearMap.smul_apply] + congr 2 + +variable [IsTriangularizable K H L] + lemma lie_eq_killingForm_smul_of_mem_rootSpace_of_mem_rootSpace_neg_aux {α : Weight K H L} {e f : L} (heα : e ∈ rootSpace H α) (hfα : f ∈ rootSpace H (-α)) (aux : ∀ (h : H), ⁅h, e⁆ = α h • e) : @@ -199,10 +229,11 @@ lemma span_weight_eq_top : rw [← traceForm_flip K H L, ← LinearMap.dualAnnihilator_ker_eq_range_flip, ker_traceForm_eq_bot_of_isCartanSubalgebra, Submodule.dualAnnihilator_bot] +variable (K L H) in @[simp] lemma span_weight_isNonZero_eq_top : span K ({α : Weight K H L | α.IsNonZero}.image (Weight.toLinear K H L)) = ⊤ := by - rw [← span_weight_eq_top K L H] + rw [← span_weight_eq_top] refine le_antisymm (Submodule.span_mono <| by simp) ?_ suffices range (Weight.toLinear K H L) ⊆ insert 0 ({α : Weight K H L | α.IsNonZero}.image (Weight.toLinear K H L)) by @@ -218,21 +249,8 @@ lemma iInf_ker_weight_eq_bot : Submodule.dualAnnihilator_bot] simp [← LinearMap.range_dualMap_eq_dualAnnihilator_ker, ← Submodule.span_range_eq_iSup] -@[simp] -lemma corootSpace_zero_eq_bot : - corootSpace (0 : H → K) = ⊥ := by - refine eq_bot_iff.mpr fun x hx ↦ ?_ - suffices {x | ∃ y ∈ H, ∃ z ∈ H, ⁅y, z⁆ = x} = {0} by simpa [mem_corootSpace, this] using hx - refine eq_singleton_iff_unique_mem.mpr ⟨⟨0, H.zero_mem, 0, H.zero_mem, zero_lie 0⟩, ?_⟩ - rintro - ⟨y, hy, z, hz, rfl⟩ - suffices ⁅(⟨y, hy⟩ : H), (⟨z, hz⟩ : H)⁆ = 0 by - simpa only [Subtype.ext_iff, LieSubalgebra.coe_bracket, ZeroMemClass.coe_zero] using this - simp - section PerfectField -variable {K L H} - variable [PerfectField K] open Module.End in @@ -324,13 +342,12 @@ lemma coe_corootSpace_eq_span_singleton' (α : Weight K H L) : use t simp [Subtype.ext_iff] · simp only [Submodule.span_singleton_le_iff_mem, LieSubmodule.mem_coeSubmodule] - exact cartanEquivDual_symm_apply_mem_corootSpace K L H α + exact cartanEquivDual_symm_apply_mem_corootSpace α end PerfectField section CharZero -variable {K H L} variable [CharZero K] /-- The contrapositive of this result is very useful, taking `x` to be the element of `H` @@ -354,30 +371,19 @@ lemma disjoint_ker_weight_corootSpace (α : Weight K H L) : replace hαx : α x = 0 := by simpa using hαx exact eq_zero_of_apply_eq_zero_of_mem_corootSpace x α hαx hx -/-- The coroot corresponding to a root. -/ -noncomputable def coroot (α : Weight K H L) : H := - 2 • (α <| (cartanEquivDual H).symm α)⁻¹ • (cartanEquivDual H).symm α - lemma root_apply_cartanEquivDual_symm_ne_zero {α : Weight K H L} (hα : α.IsNonZero) : α ((cartanEquivDual H).symm α) ≠ 0 := by contrapose! hα suffices (cartanEquivDual H).symm α ∈ α.ker ⊓ corootSpace α by rw [(disjoint_ker_weight_corootSpace α).eq_bot] at this simpa using this - exact Submodule.mem_inf.mp ⟨hα, cartanEquivDual_symm_apply_mem_corootSpace K L H α⟩ + exact Submodule.mem_inf.mp ⟨hα, cartanEquivDual_symm_apply_mem_corootSpace α⟩ lemma root_apply_coroot {α : Weight K H L} (hα : α.IsNonZero) : α (coroot α) = 2 := by rw [← Weight.coe_coe] simpa [coroot] using inv_mul_cancel (root_apply_cartanEquivDual_symm_ne_zero hα) -lemma traceForm_coroot (α : Weight K H L) (x : H) : - traceForm K H L (coroot α) x = 2 • (α <| (cartanEquivDual H).symm α)⁻¹ • α x := by - have : cartanEquivDual H ((cartanEquivDual H).symm α) x = α x := by - rw [LinearEquiv.apply_symm_apply, Weight.toLinear_apply] - rw [coroot, map_nsmul, map_smul, LinearMap.smul_apply, LinearMap.smul_apply] - congr 2 - @[simp] lemma coroot_eq_zero_iff {α : Weight K H L} : coroot α = 0 ↔ α.IsZero := by refine ⟨fun hα ↦ ?_, fun hα ↦ ?_⟩ diff --git a/Mathlib/Algebra/Lie/Weights/RootSystem.lean b/Mathlib/Algebra/Lie/Weights/RootSystem.lean index 1054e9d82da97..b38901b4a096c 100644 --- a/Mathlib/Algebra/Lie/Weights/RootSystem.lean +++ b/Mathlib/Algebra/Lie/Weights/RootSystem.lean @@ -41,9 +41,9 @@ variable {K L : Type*} [Field K] [CharZero K] [LieRing L] [LieAlgebra K L] [IsKilling K L] [FiniteDimensional K L] {H : LieSubalgebra K L} [H.IsCartanSubalgebra] [IsTriangularizable K H L] -variable (α β : Weight K H L) (hα : α.IsNonZero) +variable (α β : Weight K H L) -private lemma chainLength_aux {x} (hx : x ∈ rootSpace H (chainTop α β)) : +private lemma chainLength_aux (hα : α.IsNonZero) {x} (hx : x ∈ rootSpace H (chainTop α β)) : ∃ n : ℕ, n • x = ⁅coroot α, x⁆ := by by_cases hx' : x = 0 · exact ⟨0, by simp [hx']⟩ @@ -106,7 +106,7 @@ lemma rootSpace_neg_nsmul_add_chainTop_of_le {n : ℕ} (hn : n ≤ chainLength simp only [← smul_neg, ne_eq, LieSubmodule.eq_bot_iff, not_forall] exact ⟨_, toEnd_pow_apply_mem hf hx n, prim.pow_toEnd_f_ne_zero_of_eq_nat rfl hn⟩ -lemma rootSpace_neg_nsmul_add_chainTop_of_lt {n : ℕ} (hn : chainLength α β < n) : +lemma rootSpace_neg_nsmul_add_chainTop_of_lt (hα : α.IsNonZero) {n : ℕ} (hn : chainLength α β < n) : rootSpace H (- (n • α) + chainTop α β) = ⊥ := by by_contra e let W : Weight K H L := ⟨_, e⟩ @@ -185,7 +185,8 @@ lemma apply_coroot_eq_cast : β (coroot α) = (chainBotCoeff α β - chainTopCoeff α β : ℤ) := by rw [apply_coroot_eq_cast', ← chainTopCoeff_add_chainBotCoeff]; congr 1; omega -lemma le_chainBotCoeff_of_rootSpace_ne_top (n : ℤ) (hn : rootSpace H (-n • α + β) ≠ ⊥) : +lemma le_chainBotCoeff_of_rootSpace_ne_top + (hα : α.IsNonZero) (n : ℤ) (hn : rootSpace H (-n • α + β) ≠ ⊥) : n ≤ chainBotCoeff α β := by contrapose! hn lift n to ℕ using (Nat.cast_nonneg _).trans hn.le @@ -196,7 +197,7 @@ lemma le_chainBotCoeff_of_rootSpace_ne_top (n : ℤ) (hn : rootSpace H (-n • ← add_smul, Nat.cast_add, neg_add, add_assoc, neg_add_self, add_zero] at this /-- Members of the `α`-chain through `β` are the only roots of the form `β - kα`. -/ -lemma rootSpace_zsmul_add_ne_bot_iff (n : ℤ) : +lemma rootSpace_zsmul_add_ne_bot_iff (hα : α.IsNonZero) (n : ℤ) : rootSpace H (n • α + β) ≠ ⊥ ↔ n ≤ chainTopCoeff α β ∧ -n ≤ chainBotCoeff α β := by constructor · refine (fun hn ↦ ⟨?_, le_chainBotCoeff_of_rootSpace_ne_top α β hα _ (by rwa [neg_neg])⟩) @@ -214,11 +215,12 @@ lemma rootSpace_zsmul_add_ne_bot_iff (n : ℤ) : rwa [coe_chainTop, ← Nat.cast_smul_eq_nsmul ℤ, ← neg_smul, ← add_assoc, ← add_smul, ← sub_eq_neg_add] at this -lemma rootSpace_zsmul_add_ne_bot_iff_mem (n : ℤ) : +lemma rootSpace_zsmul_add_ne_bot_iff_mem (hα : α.IsNonZero) (n : ℤ) : rootSpace H (n • α + β) ≠ ⊥ ↔ n ∈ Finset.Icc (-chainBotCoeff α β : ℤ) (chainTopCoeff α β) := by rw [rootSpace_zsmul_add_ne_bot_iff α β hα n, Finset.mem_Icc, and_comm, neg_le] -lemma chainTopCoeff_of_eq_zsmul_add (β' : Weight K H L) (n : ℤ) (hβ' : (β' : H → K) = n • α + β) : +lemma chainTopCoeff_of_eq_zsmul_add + (hα : α.IsNonZero) (β' : Weight K H L) (n : ℤ) (hβ' : (β' : H → K) = n • α + β) : chainTopCoeff α β' = chainTopCoeff α β - n := by apply le_antisymm · refine le_sub_iff_add_le.mpr ((rootSpace_zsmul_add_ne_bot_iff α β hα _).mp ?_).1 @@ -228,7 +230,8 @@ lemma chainTopCoeff_of_eq_zsmul_add (β' : Weight K H L) (n : ℤ) (hβ' : (β' rw [hβ', ← add_assoc, ← add_smul, sub_add_cancel, ← coe_chainTop] exact (chainTop α β).2 -lemma chainBotCoeff_of_eq_zsmul_add (β' : Weight K H L) (n : ℤ) (hβ' : (β' : H → K) = n • α + β) : +lemma chainBotCoeff_of_eq_zsmul_add + (hα : α.IsNonZero) (β' : Weight K H L) (n : ℤ) (hβ' : (β' : H → K) = n • α + β) : chainBotCoeff α β' = chainBotCoeff α β + n := by have : (β' : H → K) = -n • (-α) + β := by rwa [neg_smul, smul_neg, neg_neg] rw [chainBotCoeff, chainBotCoeff, ← Weight.coe_neg, @@ -244,7 +247,7 @@ lemma chainLength_of_eq_zsmul_add (β' : Weight K H L) (n : ℤ) (hβ' : (β' : chainBotCoeff_of_eq_zsmul_add α β hα β' n hβ', sub_eq_add_neg, add_add_add_comm, neg_add_self, add_zero] -lemma chainTopCoeff_zero_right [Nontrivial L] : +lemma chainTopCoeff_zero_right [Nontrivial L] (hα : α.IsNonZero) : chainTopCoeff α (0 : Weight K H L) = 1 := by symm apply eq_of_le_of_not_lt @@ -275,20 +278,21 @@ lemma chainTopCoeff_zero_right [Nontrivial L] : Nat.cast_inj] at this rwa [this, Nat.succ_le, two_mul, add_lt_add_iff_left] -lemma chainBotCoeff_zero_right [Nontrivial L] : chainBotCoeff α (0 : Weight K H L) = 1 := +lemma chainBotCoeff_zero_right [Nontrivial L] (hα : α.IsNonZero) : + chainBotCoeff α (0 : Weight K H L) = 1 := chainTopCoeff_zero_right (-α) hα.neg -lemma chainLength_zero_right [Nontrivial L] : chainLength α 0 = 2 := by +lemma chainLength_zero_right [Nontrivial L] (hα : α.IsNonZero) : chainLength α 0 = 2 := by rw [← chainBotCoeff_add_chainTopCoeff, chainTopCoeff_zero_right α hα, chainBotCoeff_zero_right α hα] -lemma rootSpace_two_smul : rootSpace H (2 • α) = ⊥ := by +lemma rootSpace_two_smul (hα : α.IsNonZero) : rootSpace H (2 • α) = ⊥ := by cases subsingleton_or_nontrivial L · exact IsEmpty.elim inferInstance α simpa [chainTopCoeff_zero_right α hα] using weightSpace_chainTopCoeff_add_one_nsmul_add α (0 : Weight K H L) hα -lemma rootSpace_one_div_two_smul : rootSpace H ((2⁻¹ : K) • α) = ⊥ := by +lemma rootSpace_one_div_two_smul (hα : α.IsNonZero) : rootSpace H ((2⁻¹ : K) • α) = ⊥ := by by_contra h let W : Weight K H L := ⟨_, h⟩ have hW : 2 • (W : H → K) = α := by @@ -299,7 +303,8 @@ lemma rootSpace_one_div_two_smul : rootSpace H ((2⁻¹ : K) • α) = ⊥ := by apply_fun (2 • ·) at e; simpa [hW] using e) rwa [hW] at this -lemma eq_neg_one_or_eq_zero_or_eq_one_of_eq_smul (k : K) (h : (β : H → K) = k • α) : +lemma eq_neg_one_or_eq_zero_or_eq_one_of_eq_smul + (hα : α.IsNonZero) (k : K) (h : (β : H → K) = k • α) : k = -1 ∨ k = 0 ∨ k = 1 := by cases subsingleton_or_nontrivial L · exact IsEmpty.elim inferInstance α diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean index b0558a5db054c..bfbdfec1ce738 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean @@ -326,7 +326,7 @@ theorem toFilter_inf_iUnion_eq (l : IntegrationParams) (I : Box ι) (π₀ : Pre l.toFilter I ⊓ 𝓟 { π | π.iUnion = π₀.iUnion } = l.toFilteriUnion I π₀ := (iSup_inf_principal _ _).symm -variable {r r₁ r₂ : (ι → ℝ) → Ioi (0 : ℝ)} {π π₁ π₂ : TaggedPrepartition I} +variable {r₁ r₂ : (ι → ℝ) → Ioi (0 : ℝ)} {π π₁ π₂ : TaggedPrepartition I} variable (I) in theorem MemBaseSet.mono' (h : l₁ ≤ l₂) (hc : c₁ ≤ c₂) @@ -342,13 +342,16 @@ theorem MemBaseSet.mono (h : l₁ ≤ l₂) (hc : c₁ ≤ c₂) (hr : ∀ x ∈ Box.Icc I, r₁ x ≤ r₂ x) (hπ : l₁.MemBaseSet I c₁ r₁ π) : l₂.MemBaseSet I c₂ r₂ π := hπ.mono' I h hc fun J _ => hr _ <| π.tag_mem_Icc J -theorem MemBaseSet.exists_common_compl (h₁ : l.MemBaseSet I c₁ r₁ π₁) (h₂ : l.MemBaseSet I c₂ r₂ π₂) +theorem MemBaseSet.exists_common_compl + {l : IntegrationParams} + (h₁ : l.MemBaseSet I c₁ r₁ π₁) (h₂ : l.MemBaseSet I c₂ r₂ π₂) (hU : π₁.iUnion = π₂.iUnion) : ∃ π : Prepartition I, π.iUnion = ↑I \ π₁.iUnion ∧ (l.bDistortion → π.distortion ≤ c₁) ∧ (l.bDistortion → π.distortion ≤ c₂) := by + clear l₁ l₂ wlog hc : c₁ ≤ c₂ with H · simpa [hU, _root_.and_comm] using - @H _ _ I J c c₂ c₁ _ l₂ l₁ r r₂ r₁ π π₂ π₁ h₂ h₁ hU.symm (le_of_not_le hc) + @H _ _ I J c c₂ c₁ l r₂ r₁ π π₂ π₁ _ h₂ h₁ hU.symm (le_of_not_le hc) by_cases hD : (l.bDistortion : Prop) · rcases h₁.4 hD with ⟨π, hπU, hπc⟩ exact ⟨π, hπU, fun _ => hπc, fun _ => hπc.trans hc⟩ @@ -364,6 +367,8 @@ protected theorem MemBaseSet.unionComplToSubordinate (hπ₁ : l.MemBaseSet I c fun h => (distortion_unionComplToSubordinate _ _ _ _).trans_le (max_le (hπ₁.3 h) (hc h)), fun _ => ⟨⊥, by simp⟩⟩ +variable {r : (ι → ℝ) → Ioi (0 : ℝ)} + protected theorem MemBaseSet.filter (hπ : l.MemBaseSet I c r π) (p : Box ι → Prop) : l.MemBaseSet I c r (π.filter p) := by refine ⟨fun J hJ => hπ.1 J (π.mem_filter.1 hJ).1, fun hH J hJ => hπ.2 hH J (π.mem_filter.1 hJ).1, diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean index 299acd1d93d0a..207a397d2c912 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean @@ -53,6 +53,7 @@ protected theorem integrable : Integrable f μ := protected theorem integrable_normed : Integrable (f.normed μ) μ := f.integrable.div_const _ +section variable [μ.IsOpenPosMeasure] theorem integral_pos : 0 < ∫ x, f x ∂μ := by @@ -90,13 +91,17 @@ theorem integral_normed_smul {X} [NormedAddCommGroup X] [NormedSpace ℝ X] [CompleteSpace X] (z : X) : ∫ x, f.normed μ x • z ∂μ = z := by simp_rw [integral_smul_const, f.integral_normed (μ := μ), one_smul] +end + +variable (μ) + theorem measure_closedBall_le_integral : (μ (closedBall c f.rIn)).toReal ≤ ∫ x, f x ∂μ := by calc (μ (closedBall c f.rIn)).toReal = ∫ x in closedBall c f.rIn, 1 ∂μ := by simp _ = ∫ x in closedBall c f.rIn, f x ∂μ := setIntegral_congr measurableSet_closedBall (fun x hx ↦ (one_of_mem_closedBall f hx).symm) _ ≤ ∫ x, f x ∂μ := setIntegral_le_integral f.integrable (eventually_of_forall (fun x ↦ f.nonneg)) -theorem normed_le_div_measure_closedBall_rIn (x : E) : +theorem normed_le_div_measure_closedBall_rIn [μ.IsOpenPosMeasure] (x : E) : f.normed μ x ≤ 1 / (μ (closedBall c f.rIn)).toReal := by rw [normed_def] gcongr diff --git a/Mathlib/Analysis/Calculus/ParametricIntervalIntegral.lean b/Mathlib/Analysis/Calculus/ParametricIntervalIntegral.lean index db35182f98a19..a6d43de3221c4 100644 --- a/Mathlib/Analysis/Calculus/ParametricIntervalIntegral.lean +++ b/Mathlib/Analysis/Calculus/ParametricIntervalIntegral.lean @@ -18,7 +18,7 @@ open TopologicalSpace MeasureTheory Filter Metric open scoped Topology Filter Interval variable {𝕜 : Type*} [RCLike 𝕜] {μ : Measure ℝ} {E : Type*} [NormedAddCommGroup E] - [NormedSpace ℝ E] [NormedSpace 𝕜 E] [CompleteSpace E] {H : Type*} [NormedAddCommGroup H] + [NormedSpace ℝ E] [NormedSpace 𝕜 E] {H : Type*} [NormedAddCommGroup H] [NormedSpace 𝕜 H] {a b ε : ℝ} {bound : ℝ → ℝ} namespace intervalIntegral diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 82f6dd1ab1554..a1f0062306e40 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -829,12 +829,11 @@ variable [NormedSpace 𝕜 E'] variable [NormedSpace 𝕜 E''] variable [NormedSpace ℝ F] [NormedSpace 𝕜 F] variable {n : ℕ∞} -variable [CompleteSpace F] variable [MeasurableSpace G] {μ ν : Measure G} variable (L : E →L[𝕜] E' →L[𝕜] F) section Assoc - +variable [CompleteSpace F] variable [NormedAddCommGroup F'] [NormedSpace ℝ F'] [NormedSpace 𝕜 F'] [CompleteSpace F'] variable [NormedAddCommGroup F''] [NormedSpace ℝ F''] [NormedSpace 𝕜 F''] [CompleteSpace F''] variable {k : G → E''} @@ -994,7 +993,6 @@ variable [NormedSpace ℝ F] [NormedSpace 𝕜 F] variable {f₀ : 𝕜 → E} {g₀ : 𝕜 → E'} variable {n : ℕ∞} variable (L : E →L[𝕜] E' →L[𝕜] F) -variable [CompleteSpace F] variable {μ : Measure 𝕜} variable [IsAddLeftInvariant μ] [SFinite μ] @@ -1016,7 +1014,7 @@ end Real section WithParam variable [RCLike 𝕜] [NormedSpace 𝕜 E] [NormedSpace 𝕜 E'] [NormedSpace 𝕜 E''] [NormedSpace ℝ F] - [NormedSpace 𝕜 F] [CompleteSpace F] [MeasurableSpace G] [NormedAddCommGroup G] [BorelSpace G] + [NormedSpace 𝕜 F] [MeasurableSpace G] [NormedAddCommGroup G] [BorelSpace G] [NormedSpace 𝕜 G] [NormedAddCommGroup P] [NormedSpace 𝕜 P] {μ : Measure G} (L : E →L[𝕜] E' →L[𝕜] F) @@ -1168,7 +1166,7 @@ In this version, all the types belong to the same universe (to get an induction proof). Use instead `contDiffOn_convolution_right_with_param`, which removes this restriction. -/ theorem contDiffOn_convolution_right_with_param_aux {G : Type uP} {E' : Type uP} {F : Type uP} {P : Type uP} [NormedAddCommGroup E'] [NormedAddCommGroup F] [NormedSpace 𝕜 E'] - [NormedSpace ℝ F] [NormedSpace 𝕜 F] [CompleteSpace F] [MeasurableSpace G] + [NormedSpace ℝ F] [NormedSpace 𝕜 F] [MeasurableSpace G] {μ : Measure G} [NormedAddCommGroup G] [BorelSpace G] [NormedSpace 𝕜 G] [NormedAddCommGroup P] [NormedSpace 𝕜 P] {f : G → E} {n : ℕ∞} (L : E →L[𝕜] E' →L[𝕜] F) {g : P → G → E'} {s : Set P} {k : Set G} @@ -1322,7 +1320,7 @@ end WithParam section Nonneg -variable [NormedSpace ℝ E] [NormedSpace ℝ E'] [NormedSpace ℝ F] [CompleteSpace F] +variable [NormedSpace ℝ E] [NormedSpace ℝ E'] [NormedSpace ℝ F] /-- The forward convolution of two functions `f` and `g` on `ℝ`, with respect to a continuous bilinear map `L` and measure `ν`. It is defined to be the function mapping `x` to @@ -1388,7 +1386,8 @@ theorem integrable_posConvolution {f : ℝ → E} {g : ℝ → E'} {μ ν : Meas /-- The integral over `Ioi 0` of a forward convolution of two functions is equal to the product of their integrals over this set. (Compare `integral_convolution` for the two-sided convolution.) -/ -theorem integral_posConvolution [CompleteSpace E] [CompleteSpace E'] {μ ν : Measure ℝ} +theorem integral_posConvolution [CompleteSpace E] [CompleteSpace E'] [CompleteSpace F] + {μ ν : Measure ℝ} [SFinite μ] [SFinite ν] [IsAddRightInvariant μ] [NoAtoms ν] {f : ℝ → E} {g : ℝ → E'} (hf : IntegrableOn f (Ioi 0) ν) (hg : IntegrableOn g (Ioi 0) μ) (L : E →L[ℝ] E' →L[ℝ] F) : ∫ x : ℝ in Ioi 0, ∫ t : ℝ in (0)..x, L (f t) (g (x - t)) ∂ν ∂μ = diff --git a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean index 80726dad07447..8919c573b7372 100644 --- a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean +++ b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean @@ -33,12 +33,13 @@ section Manifold variable {H : Type*} [TopologicalSpace H] (I : ModelWithCorners ℝ E H) {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] - [MeasurableSpace M] [BorelSpace M] [SigmaCompactSpace M] [T2Space M] + [MeasurableSpace M] [BorelSpace M] [T2Space M] {f f' : M → F} {μ : Measure M} /-- If a locally integrable function `f` on a finite-dimensional real manifold has zero integral when multiplied by any smooth compactly supported function, then `f` vanishes almost everywhere. -/ -theorem ae_eq_zero_of_integral_smooth_smul_eq_zero (hf : LocallyIntegrable f μ) +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) : ∀ᵐ x ∂μ, f x = 0 := by -- record topological properties of `M` @@ -138,6 +139,8 @@ nonrec theorem IsOpen.ae_eq_zero_of_integral_smooth_smul_eq_zero' {U : Set M} (h · apply zero_smul · rintro ⟨x, rfl⟩; exact x.2 +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 → ℝ, diff --git a/Mathlib/Analysis/Fourier/FourierTransform.lean b/Mathlib/Analysis/Fourier/FourierTransform.lean index 7e87d4df21d1c..7580cc618b9f9 100644 --- a/Mathlib/Analysis/Fourier/FourierTransform.lean +++ b/Mathlib/Analysis/Fourier/FourierTransform.lean @@ -143,8 +143,6 @@ theorem fourierIntegral_convergent_iff (he : Continuous e) @[deprecated (since := "2024-03-29")] alias fourier_integral_convergent_iff := VectorFourier.fourierIntegral_convergent_iff -variable [CompleteSpace E] - theorem fourierIntegral_add (he : Continuous e) (hL : Continuous fun p : V × W ↦ L p.1 p.2) {f g : V → E} (hf : Integrable f μ) (hg : Integrable g μ) : fourierIntegral e μ L f + fourierIntegral e μ L g = fourierIntegral e μ L (f + g) := by @@ -236,7 +234,7 @@ namespace VectorFourier variable {𝕜 ι E F V W : Type*} [Fintype ι] [NontriviallyNormedField 𝕜] [NormedAddCommGroup V] [NormedSpace 𝕜 V] [MeasurableSpace V] [BorelSpace V] - [NormedAddCommGroup W] [NormedSpace 𝕜 W] [MeasurableSpace W] [BorelSpace W] + [NormedAddCommGroup W] [NormedSpace 𝕜 W] {e : AddChar 𝕜 𝕊} {μ : Measure V} {L : V →L[𝕜] W →L[𝕜] 𝕜} [NormedAddCommGroup F] [NormedSpace ℝ F] [NormedAddCommGroup E] [NormedSpace ℂ E] @@ -274,8 +272,6 @@ variable {𝕜 : Type*} [CommRing 𝕜] [MeasurableSpace 𝕜] {E : Type*} [Norm section Defs -variable [CompleteSpace E] - /-- The Fourier transform integral for `f : 𝕜 → E`, with respect to the measure `μ` and additive character `e`. -/ def fourierIntegral (e : AddChar 𝕜 𝕊) (μ : Measure 𝕜) (f : 𝕜 → E) (w : 𝕜) : E := @@ -351,10 +347,9 @@ section Apply variable {ι F V W : Type*} [Fintype ι] [NormedAddCommGroup V] [NormedSpace ℝ V] [MeasurableSpace V] [BorelSpace V] - [NormedAddCommGroup W] [NormedSpace ℝ W] [MeasurableSpace W] [BorelSpace W] + [NormedAddCommGroup W] [NormedSpace ℝ W] {μ : Measure V} {L : V →L[ℝ] W →L[ℝ] ℝ} [NormedAddCommGroup F] [NormedSpace ℝ F] - [NormedAddCommGroup E] [NormedSpace ℂ E] {M : ι → Type*} [∀ i, NormedAddCommGroup (M i)] [∀ i, NormedSpace ℝ (M i)] theorem fourierIntegral_continuousLinearMap_apply' @@ -372,12 +367,18 @@ theorem fourierIntegral_continuousMultilinearMap_apply' end Apply variable {V : Type*} [NormedAddCommGroup V] - [InnerProductSpace ℝ V] [MeasurableSpace V] [BorelSpace V] [FiniteDimensional ℝ V] + [InnerProductSpace ℝ V] [MeasurableSpace V] [BorelSpace V] {W : Type*} [NormedAddCommGroup W] [InnerProductSpace ℝ W] [MeasurableSpace W] [BorelSpace W] [FiniteDimensional ℝ W] open scoped RealInnerProductSpace +@[simp] theorem fourierIntegral_convergent_iff {μ : Measure V} {f : V → E} (w : V) : + Integrable (fun v : V ↦ 𝐞 (- ⟪v, w⟫) • f v) μ ↔ Integrable f μ := + fourierIntegral_convergent_iff' (innerSL ℝ) w + +variable [FiniteDimensional ℝ V] + /-- The Fourier transform of a function on an inner product space, with respect to the standard additive character `ω ↦ exp (2 i π ω)`. -/ def fourierIntegral (f : V → E) (w : V) : E := @@ -445,10 +446,6 @@ theorem fourierIntegral_real_eq_integral_exp_smul (f : ℝ → E) (w : ℝ) : @[deprecated (since := "2024-02-21")] alias fourierIntegral_eq_integral_exp_smul := fourierIntegral_real_eq_integral_exp_smul -@[simp] theorem fourierIntegral_convergent_iff {μ : Measure V} {f : V → E} (w : V) : - Integrable (fun v : V ↦ 𝐞 (- ⟪v, w⟫) • f v) μ ↔ Integrable f μ := - fourierIntegral_convergent_iff' (innerSL ℝ) w - theorem fourierIntegral_continuousLinearMap_apply {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] {f : V → (F →L[ℝ] E)} {a : F} {v : V} (hf : Integrable f) : diff --git a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean index 6b92667a9c412..c8a6fff1cb363 100644 --- a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean +++ b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean @@ -56,8 +56,6 @@ section InnerProductSpace variable [NormedAddCommGroup V] [MeasurableSpace V] [BorelSpace V] [InnerProductSpace ℝ V] [FiniteDimensional ℝ V] -variable [CompleteSpace E] - local notation3 "i" => fun (w : V) => (1 / (2 * ‖w‖ ^ 2) : ℝ) • w /-- Shifting `f` by `(1 / (2 * ‖w‖ ^ 2)) • w` negates the integral in the Riemann-Lebesgue lemma. -/ @@ -251,7 +249,6 @@ section NoInnerProduct variable (f) [AddCommGroup V] [TopologicalSpace V] [TopologicalAddGroup V] [T2Space V] [MeasurableSpace V] [BorelSpace V] [Module ℝ V] [ContinuousSMul ℝ V] [FiniteDimensional ℝ V] - [CompleteSpace E] /-- Riemann-Lebesgue lemma for functions on a finite-dimensional real vector space, formulated via dual space. -/ diff --git a/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean b/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean index c0f66c299a7b6..a31afd3501b4f 100644 --- a/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean +++ b/Mathlib/Analysis/NormedSpace/PiTensorProduct/InjectiveSeminorm.lean @@ -397,10 +397,16 @@ noncomputable def mapLMonoidHom : (Π i, E i →L[𝕜] E i) →* ((⨂[𝕜] i, protected theorem mapL_pow (f : Π i, E i →L[𝕜] E i) (n : ℕ) : mapL (f ^ n) = mapL f ^ n := MonoidHom.map_pow mapLMonoidHom _ _ +-- We redeclare `ι` here, and later dependent arguments, +-- to avoid the `[Fintype ι]` assumption present throughout the rest of the file. open Function in -private theorem mapL_add_smul_aux [DecidableEq ι] (i : ι) (u : E i →L[𝕜] E' i) : +private theorem mapL_add_smul_aux {ι : Type uι} + {E : ι → Type uE} [(i : ι) → SeminormedAddCommGroup (E i)] [(i : ι) → NormedSpace 𝕜 (E i)] + {E' : ι → Type u_1} [(i : ι) → SeminormedAddCommGroup (E' i)] [(i : ι) → NormedSpace 𝕜 (E' i)] + (f : (i : ι) → E i →L[𝕜] E' i) + [DecidableEq ι] (i : ι) (u : E i →L[𝕜] E' i) : (fun j ↦ (update f i u j).toLinearMap) = - update (fun j ↦ (f j).toLinearMap) i u.toLinearMap := by + update (fun j ↦ (f j).toLinearMap) i u.toLinearMap := by symm rw [update_eq_iff] constructor diff --git a/Mathlib/Geometry/Manifold/BumpFunction.lean b/Mathlib/Geometry/Manifold/BumpFunction.lean index 455a203e54316..d32ef29050ef2 100644 --- a/Mathlib/Geometry/Manifold/BumpFunction.lean +++ b/Mathlib/Geometry/Manifold/BumpFunction.lean @@ -32,7 +32,7 @@ universe uE uF uH uM variable {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] {H : Type uH} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} {M : Type uM} [TopologicalSpace M] - [ChartedSpace H M] [SmoothManifoldWithCorners I M] + [ChartedSpace H M] open Function Filter FiniteDimensional Set Metric diff --git a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean index d6f70cde8ef10..2daa9b0a934c0 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean @@ -28,26 +28,23 @@ open scoped Topology Manifold Bundle variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] - -- declare a smooth manifold `M` over the pair `(E, H)`. + -- declare a charted space `M` over the pair `(E, H)`. {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] - [Is : SmoothManifoldWithCorners I M] - -- declare a smooth manifold `M'` over the pair `(E', H')`. + -- declare a charted space `M'` over the pair `(E', H')`. {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] - [I's : SmoothManifoldWithCorners I' M'] -- declare a smooth manifold `N` over the pair `(F, G)`. {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type*} [TopologicalSpace G] {J : ModelWithCorners 𝕜 F G} {N : Type*} [TopologicalSpace N] [ChartedSpace G N] [Js : SmoothManifoldWithCorners J N] - -- declare a smooth manifold `N'` over the pair `(F', G')`. + -- declare a charted space `N'` over the pair `(F', G')`. {F' : Type*} [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] {G' : Type*} [TopologicalSpace G'] {J' : ModelWithCorners 𝕜 F' G'} {N' : Type*} [TopologicalSpace N'] [ChartedSpace G' N'] - [J's : SmoothManifoldWithCorners J' N'] -- declare some additional normed spaces, used for fibers of vector bundles {F₁ : Type*} [NormedAddCommGroup F₁] [NormedSpace 𝕜 F₁] {F₂ : Type*} [NormedAddCommGroup F₂] @@ -62,6 +59,8 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] /-! ### The derivative of a smooth function is smooth -/ section mfderiv +variable [Is : SmoothManifoldWithCorners I M] [I's : SmoothManifoldWithCorners I' M'] + /-- The function that sends `x` to the `y`-derivative of `f(x,y)` at `g(x)` is `C^m` at `x₀`, where the derivative is taken as a continuous linear map. @@ -334,7 +333,9 @@ theorem ContMDiffOn.contMDiffOn_tangentMapWithin_aux {f : H → H'} {s : Set H} /-- If a function is `C^n` on a domain with unique derivatives, then its bundled derivative is `C^m` when `m+1 ≤ n`. -/ -theorem ContMDiffOn.contMDiffOn_tangentMapWithin (hf : ContMDiffOn I I' n f s) (hmn : m + 1 ≤ n) +theorem ContMDiffOn.contMDiffOn_tangentMapWithin + [Is : SmoothManifoldWithCorners I M] [I's : SmoothManifoldWithCorners I' M'] + (hf : ContMDiffOn I I' n f s) (hmn : m + 1 ≤ n) (hs : UniqueMDiffOn I s) : ContMDiffOn I.tangent I'.tangent m (tangentMapWithin I I' f s) (π E (TangentSpace I) ⁻¹' s) := by @@ -510,6 +511,8 @@ theorem ContMDiffOn.contMDiffOn_tangentMapWithin (hf : ContMDiffOn I I' n f s) ( simp only [A, B, C, D, ← E] exact diff_DrirrflilDl.congr eq_comp +variable [Is : SmoothManifoldWithCorners I M] [I's : SmoothManifoldWithCorners I' M'] + /-- If a function is `C^n` on a domain with unique derivatives, with `1 ≤ n`, then its bundled derivative is continuous there. -/ theorem ContMDiffOn.continuousOn_tangentMapWithin (hf : ContMDiffOn I I' n f s) (hmn : 1 ≤ n) @@ -557,7 +560,7 @@ coordinates on the tangent bundle in our definitions. So this statement is not a may seem. TODO define splittings of vector bundles; state this result invariantly. -/ -theorem tangentMap_tangentBundle_pure (p : TangentBundle I M) : +theorem tangentMap_tangentBundle_pure [Is : SmoothManifoldWithCorners I M] (p : TangentBundle I M) : tangentMap I I.tangent (zeroSection E (TangentSpace I)) p = ⟨⟨p.proj, 0⟩, ⟨p.2, 0⟩⟩ := by rcases p with ⟨x, v⟩ have N : I.symm ⁻¹' (chartAt H x).target ∈ 𝓝 (I ((chartAt H x) x)) := by diff --git a/Mathlib/Geometry/Manifold/LocalDiffeomorph.lean b/Mathlib/Geometry/Manifold/LocalDiffeomorph.lean index 663a416669e9b..e41ff8cd8caa3 100644 --- a/Mathlib/Geometry/Manifold/LocalDiffeomorph.lean +++ b/Mathlib/Geometry/Manifold/LocalDiffeomorph.lean @@ -95,7 +95,7 @@ def Diffeomorph.toPartialDiffeomorph (h : Diffeomorph I J M N n) : -- Add the very basic API we need. namespace PartialDiffeomorph -variable (Φ : PartialDiffeomorph I J M N n) (hn : 1 ≤ n) +variable (Φ : PartialDiffeomorph I J M N n) /-- A partial diffeomorphism is also a local homeomorphism. -/ def toPartialHomeomorph : PartialHomeomorph M N where @@ -116,10 +116,11 @@ protected def symm : PartialDiffeomorph J I N M n where protected theorem contMDiffOn : ContMDiffOn I J n Φ Φ.source := Φ.contMDiffOn_toFun -protected theorem mdifferentiableOn : MDifferentiableOn I J Φ Φ.source := +protected theorem mdifferentiableOn (hn : 1 ≤ n) : MDifferentiableOn I J Φ Φ.source := (Φ.contMDiffOn).mdifferentiableOn hn -protected theorem mdifferentiableAt {x : M} (hx : x ∈ Φ.source) : MDifferentiableAt I J Φ x := +protected theorem mdifferentiableAt (hn : 1 ≤ n) {x : M} (hx : x ∈ Φ.source) : + MDifferentiableAt I J Φ x := (Φ.mdifferentiableOn hn x hx).mdifferentiableAt (Φ.open_source.mem_nhds hx) /- We could add lots of additional API (following `Diffeomorph` and `PartialHomeomorph`), such as diff --git a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean index 74c73f56ab6fb..d73131e3999fb 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean @@ -29,10 +29,14 @@ section UniqueMDiff variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} - [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] {E' : Type*} + [TopologicalSpace M] [ChartedSpace H M] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] - [SmoothManifoldWithCorners I' M'] {s : Set M} {x : M} + [SmoothManifoldWithCorners I' M'] {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H' M''] + {s : Set M} {x : M} + +section +variable [SmoothManifoldWithCorners I M] /-- If `s` has the unique differential property at `x`, `f` is differentiable within `s` at x` and its derivative has dense range, then `f '' s` has the unique differential property at `f x`. -/ @@ -93,8 +97,8 @@ theorem UniqueMDiffOn.uniqueDiffOn_target_inter (hs : UniqueMDiffOn I s) (x : M) /-- When considering functions between manifolds, this statement shows up often. It entails the unique differential of the pullback in extended charts of the set where the function can be read in the charts. -/ -theorem UniqueMDiffOn.uniqueDiffOn_inter_preimage (hs : UniqueMDiffOn I s) (x : M) (y : M') - {f : M → M'} (hf : ContinuousOn f s) : +theorem UniqueMDiffOn.uniqueDiffOn_inter_preimage (hs : UniqueMDiffOn I s) (x : M) (y : M'') + {f : M → M''} (hf : ContinuousOn f s) : UniqueDiffOn 𝕜 ((extChartAt I x).target ∩ (extChartAt I x).symm ⁻¹' (s ∩ f ⁻¹' (extChartAt I' y).source)) := haveI : UniqueMDiffOn I (s ∩ f ⁻¹' (extChartAt I' y).source) := by @@ -104,6 +108,8 @@ theorem UniqueMDiffOn.uniqueDiffOn_inter_preimage (hs : UniqueMDiffOn I s) (x : exact (isOpen_extChartAt_source I' y).mem_nhds hz.2 this.uniqueDiffOn_target_inter _ +end + open Bundle variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {Z : M → Type*} @@ -114,6 +120,8 @@ theorem Trivialization.mdifferentiable (e : Trivialization F (π F Z)) [MemTrivi e.toPartialHomeomorph.MDifferentiable (I.prod 𝓘(𝕜, F)) (I.prod 𝓘(𝕜, F)) := ⟨(e.smoothOn I).mdifferentiableOn, (e.smoothOn_symm I).mdifferentiableOn⟩ +variable [SmoothManifoldWithCorners I M] + theorem UniqueMDiffWithinAt.smooth_bundle_preimage {p : TotalSpace F Z} (hs : UniqueMDiffWithinAt I s p.proj) : UniqueMDiffWithinAt (I.prod 𝓘(𝕜, F)) (π F Z ⁻¹' s) p := by diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index 3e1374a8c5049..70c9c57e64c52 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -65,9 +65,9 @@ open scoped Topology Manifold Classical Filter noncomputable section variable {ι : Type uι} {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] - [FiniteDimensional ℝ E] {F : Type uF} [NormedAddCommGroup F] [NormedSpace ℝ F] {H : Type uH} + {F : Type uF} [NormedAddCommGroup F] [NormedSpace ℝ F] {H : Type uH} [TopologicalSpace H] (I : ModelWithCorners ℝ E H) {M : Type uM} [TopologicalSpace M] - [ChartedSpace H M] [SmoothManifoldWithCorners I M] + [ChartedSpace H M] /-! ### Covering by supports of smooth bump functions @@ -98,7 +98,7 @@ subordinate to `U`, see `SmoothBumpCovering.exists_isSubordinate`. This covering can be used, e.g., to construct a partition of unity and to prove the weak Whitney embedding theorem. -/ -- Porting note(#5171): was @[nolint has_nonempty_instance] -structure SmoothBumpCovering (s : Set M := univ) where +structure SmoothBumpCovering [FiniteDimensional ℝ E] (s : Set M := univ) where /-- The center point of each bump in the smooth covering. -/ c : ι → M /-- A smooth bump function around `c i`. -/ @@ -171,6 +171,13 @@ theorem le_one (i : ι) (x : M) : f i x ≤ 1 := theorem sum_nonneg (x : M) : 0 ≤ ∑ᶠ i, f i x := f.toPartitionOfUnity.sum_nonneg x +theorem finsum_smul_mem_convex {g : ι → M → F} {t : Set F} {x : M} (hx : x ∈ s) + (hg : ∀ i, f i x ≠ 0 → g i x ∈ t) (ht : Convex ℝ t) : ∑ᶠ i, f i x • g i x ∈ t := + ht.finsum_mem (fun _ => f.nonneg _ _) (f.sum_eq_one hx) hg + +section SmoothManifoldWithCorners +variable [SmoothManifoldWithCorners I M] + theorem contMDiff_smul {g : M → F} {i} (hg : ∀ x ∈ tsupport (f i), ContMDiffAt I 𝓘(ℝ, F) n g x) : ContMDiff I 𝓘(ℝ, F) n fun x => f i x • g x := contMDiff_of_tsupport fun x hx => @@ -212,9 +219,7 @@ theorem contDiffAt_finsum {s : Set E} (f : SmoothPartitionOfUnity ι 𝓘(ℝ, E simp only [← contMDiffAt_iff_contDiffAt] at * exact f.contMDiffAt_finsum hφ -theorem finsum_smul_mem_convex {g : ι → M → F} {t : Set F} {x : M} (hx : x ∈ s) - (hg : ∀ i, f i x ≠ 0 → g i x ∈ t) (ht : Convex ℝ t) : ∑ᶠ i, f i x • g i x ∈ t := - ht.finsum_mem (fun _ => f.nonneg _ _) (f.sum_eq_one hx) hg +end SmoothManifoldWithCorners section finsupport @@ -288,6 +293,8 @@ theorem isSubordinate_toPartitionOfUnity : alias ⟨_, IsSubordinate.toPartitionOfUnity⟩ := isSubordinate_toPartitionOfUnity +variable [SmoothManifoldWithCorners I M] + /-- 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 $C^n$ smooth on `U i`, then the sum `fun x ↦ ∑ᶠ i, f i x • g i x` is $C^n$ smooth on the whole manifold. -/ @@ -353,6 +360,7 @@ end BumpCovering namespace SmoothBumpCovering +variable [FiniteDimensional ℝ E] variable {s : Set M} {U : M → Set M} (fs : SmoothBumpCovering ι I M s) instance : CoeFun (SmoothBumpCovering ι I M s) fun x => ∀ i : ι, SmoothBumpFunction I (x.c i) := @@ -371,8 +379,8 @@ theorem IsSubordinate.support_subset {fs : SmoothBumpCovering ι I M s} {U : M (h : fs.IsSubordinate U) (i : ι) : support (fs i) ⊆ U (fs.c i) := Subset.trans subset_closure (h i) -variable (I) - +variable (I) in +variable [SmoothManifoldWithCorners I M] in /-- Let `M` be a smooth manifold with corners modelled on a finite dimensional real vector space. Suppose also that `M` is a Hausdorff `σ`-compact topological space. Let `s` be a closed set in `M` and `U : M → Set M` be a collection of sets such that `U x ∈ 𝓝 x` for every `x ∈ s`. @@ -399,22 +407,12 @@ theorem exists_isSubordinate [T2Space M] [SigmaCompactSpace M] (hs : IsClosed s) ((f i).support_subset_source <| hVf _ hi) (hr i hi).2 · simpa only [SmoothBumpFunction.support_updateRIn, tsupport] using hfU i -variable {I} - protected theorem locallyFinite : LocallyFinite fun i => support (fs i) := fs.locallyFinite' protected theorem point_finite (x : M) : {i | fs i x ≠ 0}.Finite := fs.locallyFinite.point_finite x -theorem mem_chartAt_source_of_eq_one {i : ι} {x : M} (h : fs i x = 1) : - x ∈ (chartAt H (fs.c i)).source := - (fs i).support_subset_source <| by simp [h] - -theorem mem_extChartAt_source_of_eq_one {i : ι} {x : M} (h : fs i x = 1) : - x ∈ (extChartAt I (fs.c i)).source := by - rw [extChartAt_source]; exact fs.mem_chartAt_source_of_eq_one h - /-- Index of a bump function such that `fs i =ᶠ[𝓝 x] 1`. -/ def ind (x : M) (hx : x ∈ s) : ι := (fs.eventuallyEq_one' x hx).choose @@ -428,6 +426,16 @@ theorem apply_ind (x : M) (hx : x ∈ s) : fs (fs.ind x hx) x = 1 := theorem mem_support_ind (x : M) (hx : x ∈ s) : x ∈ support (fs <| fs.ind x hx) := by simp [fs.apply_ind x hx] +variable [SmoothManifoldWithCorners I M] + +theorem mem_chartAt_source_of_eq_one {i : ι} {x : M} (h : fs i x = 1) : + x ∈ (chartAt H (fs.c i)).source := + (fs i).support_subset_source <| by simp [h] + +theorem mem_extChartAt_source_of_eq_one {i : ι} {x : M} (h : fs i x = 1) : + x ∈ (extChartAt I (fs.c i)).source := by + rw [extChartAt_source]; exact fs.mem_chartAt_source_of_eq_one h + theorem mem_chartAt_ind_source (x : M) (hx : x ∈ s) : x ∈ (chartAt H (fs.c (fs.ind x hx))).source := fs.mem_chartAt_source_of_eq_one (fs.apply_ind x hx) @@ -498,6 +506,8 @@ theorem sum_toSmoothPartitionOfUnity_eq (x : M) : end SmoothBumpCovering variable (I) +variable [FiniteDimensional ℝ E] +variable [SmoothManifoldWithCorners I M] /-- Given two disjoint closed sets `s, t` in a Hausdorff σ-compact finite dimensional manifold, there exists an infinitely smooth function that is equal to `0` on `s` and to `1` on `t`. diff --git a/Mathlib/Geometry/Manifold/PoincareConjecture.lean b/Mathlib/Geometry/Manifold/PoincareConjecture.lean index c71d01ca2e3fc..97601179316c1 100644 --- a/Mathlib/Geometry/Manifold/PoincareConjecture.lean +++ b/Mathlib/Geometry/Manifold/PoincareConjecture.lean @@ -26,7 +26,7 @@ local macro:max "ℝ"n:superscript(term) : term => `(EuclideanSpace ℝ (Fin $( local macro:max "𝕊"n:superscript(term) : term => `(sphere (0 : EuclideanSpace ℝ (Fin ($(⟨n.raw[0]⟩) + 1))) 1) -variable (M : Type*) [TopologicalSpace M] [T2Space M] +variable (M : Type*) [TopologicalSpace M] open ContinuousMap @@ -36,17 +36,17 @@ open ContinuousMap Newman (1966) and Connell (1967) proved it without the condition. - For n = 4 it was proven by Michael Freedman in 1982. - For n = 3 it was proven by Grigori Perelman in 2003. -/ -proof_wanted ContinuousMap.HomotopyEquiv.nonempty_homeomorph_sphere +proof_wanted ContinuousMap.HomotopyEquiv.nonempty_homeomorph_sphere [T2Space M] (n : ℕ) [ChartedSpace ℝⁿ M] : M ≃ₕ 𝕊ⁿ → Nonempty (M ≃ₜ 𝕊ⁿ) /-- The 3-dimensional topological Poincaré conjecture (proven by Perelman) -/ proof_wanted SimplyConnectedSpace.nonempty_homeomorph_sphere_three - [ChartedSpace ℝ³ M] [SimplyConnectedSpace M] [CompactSpace M] : + [T2Space M] [ChartedSpace ℝ³ M] [SimplyConnectedSpace M] [CompactSpace M] : Nonempty (M ≃ₜ 𝕊³) /-- The 3-dimensional smooth Poincaré conjecture (proven by Perelman) -/ proof_wanted SimplyConnectedSpace.nonempty_diffeomorph_sphere_three - [ChartedSpace ℝ³ M] [SmoothManifoldWithCorners (𝓡 3) M] + [T2Space M] [ChartedSpace ℝ³ M] [SmoothManifoldWithCorners (𝓡 3) M] [SimplyConnectedSpace M] [CompactSpace M] : Nonempty (M ≃ₘ⟮𝓡 3, 𝓡 3⟯ 𝕊³) diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean index 91703740e9643..b4db2a8aa2149 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean @@ -125,7 +125,7 @@ variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup F] [NormedSpace 𝕜 (IB : ModelWithCorners 𝕜 EB HB) (E' : B → Type*) [∀ x, Zero (E' x)] {EM : Type*} [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type*} [TopologicalSpace HM] {IM : ModelWithCorners 𝕜 EM HM} [TopologicalSpace M] [ChartedSpace HM M] - [Is : SmoothManifoldWithCorners IM M] {n : ℕ∞} + {n : ℕ∞} variable [TopologicalSpace B] [ChartedSpace HB B] [FiberBundle F E] @@ -252,9 +252,9 @@ end variable [NontriviallyNormedField 𝕜] {EB : Type*} [NormedAddCommGroup EB] [NormedSpace 𝕜 EB] {HB : Type*} [TopologicalSpace HB] (IB : ModelWithCorners 𝕜 EB HB) [TopologicalSpace B] - [ChartedSpace HB B] [SmoothManifoldWithCorners IB B] {EM : Type*} [NormedAddCommGroup EM] + [ChartedSpace HB B] {EM : Type*} [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type*} [TopologicalSpace HM] {IM : ModelWithCorners 𝕜 EM HM} - [TopologicalSpace M] [ChartedSpace HM M] [Is : SmoothManifoldWithCorners IM M] {n : ℕ∞} + [TopologicalSpace M] [ChartedSpace HM M] {n : ℕ∞} [∀ x, AddCommMonoid (E x)] [∀ x, Module 𝕜 (E x)] [NormedAddCommGroup F] [NormedSpace 𝕜 F] section WithTopology @@ -443,6 +443,7 @@ theorem Trivialization.contMDiffWithinAt_snd_comp_iff₂ {f : M → TotalSpace F end SmoothCoordChange +variable [SmoothManifoldWithCorners IB B] in /-- For a smooth vector bundle `E` over `B` with fiber modelled on `F`, the change-of-co-ordinates between two trivializations `e`, `e'` for `E`, considered as charts to `B × F`, is smooth and fiberwise linear. -/ @@ -461,6 +462,7 @@ instance SmoothFiberwiseLinear.hasGroupoid : · rintro ⟨b, v⟩ hb exact (e.apply_symm_apply_eq_coordChangeL e' hb.1 v).symm +variable [SmoothManifoldWithCorners IB B] in /-- A smooth vector bundle `E` is naturally a smooth manifold. -/ instance Bundle.TotalSpace.smoothManifoldWithCorners [SmoothManifoldWithCorners IB B] : SmoothManifoldWithCorners (IB.prod 𝓘(𝕜, F)) (TotalSpace F E) := by @@ -603,6 +605,8 @@ variable [∀ x : B, TopologicalSpace (E₁ x)] [∀ x : B, TopologicalSpace (E [FiberBundle F₂ E₂] [VectorBundle 𝕜 F₁ E₁] [VectorBundle 𝕜 F₂ E₂] [SmoothVectorBundle F₁ E₁ IB] [SmoothVectorBundle F₂ E₂ IB] +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 diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean index e6e83dedb3116..33808cd26cd8e 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean @@ -27,12 +27,12 @@ variable {𝕜 B F₁ F₂ M : Type*} {E₁ : B → Type*} {E₂ : B → Type*} [TopologicalSpace (TotalSpace F₁ E₁)] [∀ x, TopologicalSpace (E₁ x)] [∀ x, AddCommGroup (E₂ x)] [∀ x, Module 𝕜 (E₂ x)] [NormedAddCommGroup F₂] [NormedSpace 𝕜 F₂] [TopologicalSpace (TotalSpace F₂ E₂)] [∀ x, TopologicalSpace (E₂ x)] - [∀ x, TopologicalAddGroup (E₂ x)] [∀ x, ContinuousSMul 𝕜 (E₂ x)] {EB : Type*} + {EB : Type*} [NormedAddCommGroup EB] [NormedSpace 𝕜 EB] {HB : Type*} [TopologicalSpace HB] (IB : ModelWithCorners 𝕜 EB HB) [TopologicalSpace B] [ChartedSpace HB B] {EM : Type*} [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type*} [TopologicalSpace HM] {IM : ModelWithCorners 𝕜 EM HM} [TopologicalSpace M] [ChartedSpace HM M] - [SmoothManifoldWithCorners IM M] {n : ℕ∞} [FiberBundle F₁ E₁] [VectorBundle 𝕜 F₁ E₁] + {n : ℕ∞} [FiberBundle F₁ E₁] [VectorBundle 𝕜 F₁ E₁] [FiberBundle F₂ E₂] [VectorBundle 𝕜 F₂ E₂] {e₁ e₁' : Trivialization F₁ (π F₁ E₁)} {e₂ e₂' : Trivialization F₂ (π F₂ E₂)} @@ -49,6 +49,8 @@ theorem smoothOn_continuousLinearMapCoordChange have h₂ := smoothOn_coordChangeL IB e₂ e₂' refine (h₁.mono ?_).cle_arrowCongr (h₂.mono ?_) <;> mfld_set_tac +variable [∀ x, TopologicalAddGroup (E₂ x)] [∀ x, ContinuousSMul 𝕜 (E₂ x)] + theorem hom_chart (y₀ y : LE₁E₂) : chartAt (ModelProd HB (F₁ →L[𝕜] F₂)) y₀ y = (chartAt HB y₀.1 y.1, inCoordinates F₁ E₁ F₂ E₂ y₀.1 y.1 y₀.1 y.1 y.2) := by diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index 284123838fea1..2f4709813f8fd 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -25,18 +25,16 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCom (I' : ModelWithCorners 𝕜 E' H') {M : Type*} [TopologicalSpace M] [ChartedSpace H M] {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] {E'' : Type*} [NormedAddCommGroup E''] [NormedSpace 𝕜 E''] {H'' : Type*} [TopologicalSpace H''] {I'' : ModelWithCorners 𝕜 E'' H''} - {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] [SmoothManifoldWithCorners I M] + {M'' : Type*} [TopologicalSpace M''] [ChartedSpace H'' M''] variable (F : Type*) [NormedAddCommGroup F] [NormedSpace 𝕜 F] -- `F` model fiber (n : ℕ∞) (V : M → Type*) [TopologicalSpace (TotalSpace F V)] -- `V` vector bundle - [∀ x, AddCommGroup (V x)] - [∀ x, Module 𝕜 (V x)] -variable [∀ x : M, TopologicalSpace (V x)] [FiberBundle F V] [VectorBundle 𝕜 F V] - [SmoothVectorBundle F V I] + [∀ x : M, TopologicalSpace (V x)] [FiberBundle F V] + /-- Bundled `n` times continuously differentiable sections of a vector bundle. -/ structure ContMDiffSection where @@ -76,18 +74,6 @@ 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 -protected theorem mdifferentiable' (s : Cₛ^n⟮I; F, V⟯) (hn : 1 ≤ n) : - MDifferentiable I (I.prod 𝓘(𝕜, F)) fun x => TotalSpace.mk' F x (s x : V x) := - s.contMDiff.mdifferentiable hn - -protected theorem mdifferentiable (s : Cₛ^∞⟮I; F, V⟯) : - MDifferentiable I (I.prod 𝓘(𝕜, F)) fun x => TotalSpace.mk' F x (s x : V x) := - s.contMDiff.mdifferentiable le_top - -protected theorem mdifferentiableAt (s : Cₛ^∞⟮I; F, V⟯) {x} : - MDifferentiableAt I (I.prod 𝓘(𝕜, F)) (fun x => TotalSpace.mk' F x (s x : V x)) x := - s.mdifferentiable x - theorem coe_inj ⦃s t : Cₛ^n⟮I; F, V⟯⦄ (h : (s : ∀ x, V x) = t) : s = t := DFunLike.ext' h @@ -97,6 +83,9 @@ theorem coe_injective : Injective ((↑) : Cₛ^n⟮I; F, V⟯ → ∀ x, V x) : @[ext] theorem ext (h : ∀ x, s x = t x) : s = t := DFunLike.ext _ _ h +section +variable [∀ x, AddCommGroup (V x)] [∀ x, Module 𝕜 (V x)] [VectorBundle 𝕜 F V] + instance instAdd : Add Cₛ^n⟮I; F, V⟯ := by refine ⟨fun s t => ⟨s + t, ?_⟩⟩ intro x₀ @@ -139,21 +128,6 @@ instance inhabited : Inhabited Cₛ^n⟮I; F, V⟯ := theorem coe_zero : ⇑(0 : Cₛ^n⟮I; F, V⟯) = 0 := rfl -instance instSMul : SMul 𝕜 Cₛ^n⟮I; F, V⟯ := by - refine ⟨fun c s => ⟨c • ⇑s, ?_⟩⟩ - intro x₀ - have hs := s.contMDiff x₀ - rw [contMDiffAt_section] at hs ⊢ - set e := trivializationAt F V x₀ - refine ((contMDiffAt_const (c := c)).smul hs).congr_of_eventuallyEq ?_ - refine eventually_of_mem (e.open_baseSet.mem_nhds <| mem_baseSet_trivializationAt F V x₀) ?_ - intro x hx - apply (e.linear 𝕜 hx).2 - -@[simp] -theorem coe_smul (r : 𝕜) (s : Cₛ^n⟮I; F, V⟯) : ⇑(r • s : Cₛ^n⟮I; F, V⟯) = r • ⇑s := - rfl - instance instNeg : Neg Cₛ^n⟮I; F, V⟯ := by refine ⟨fun s => ⟨-s, ?_⟩⟩ intro x₀ @@ -192,17 +166,43 @@ theorem coe_zsmul (s : Cₛ^n⟮I; F, V⟯) (z : ℤ) : ⇑(z • s : Cₛ^n⟮I instance instAddCommGroup : AddCommGroup Cₛ^n⟮I; F, V⟯ := coe_injective.addCommGroup _ coe_zero coe_add coe_neg coe_sub coe_nsmul coe_zsmul -variable (I F V n) +instance instSMul : SMul 𝕜 Cₛ^n⟮I; F, V⟯ := by + refine ⟨fun c s => ⟨c • ⇑s, ?_⟩⟩ + intro x₀ + have hs := s.contMDiff x₀ + rw [contMDiffAt_section] at hs ⊢ + set e := trivializationAt F V x₀ + refine ((contMDiffAt_const (c := c)).smul hs).congr_of_eventuallyEq ?_ + refine eventually_of_mem (e.open_baseSet.mem_nhds <| mem_baseSet_trivializationAt F V x₀) ?_ + intro x hx + apply (e.linear 𝕜 hx).2 + +@[simp] +theorem coe_smul (r : 𝕜) (s : Cₛ^n⟮I; F, V⟯) : ⇑(r • s : Cₛ^n⟮I; F, V⟯) = r • ⇑s := + rfl +variable (I F V n) in /-- The additive morphism from smooth sections to dependent maps. -/ def coeAddHom : Cₛ^n⟮I; F, V⟯ →+ ∀ x, V x where toFun := (↑) map_zero' := coe_zero map_add' := coe_add -variable {I F V n} - instance instModule : Module 𝕜 Cₛ^n⟮I; F, V⟯ := coe_injective.module 𝕜 (coeAddHom I F n V) coe_smul +end + +protected theorem mdifferentiable' (s : Cₛ^n⟮I; F, V⟯) (hn : 1 ≤ n) : + MDifferentiable I (I.prod 𝓘(𝕜, F)) fun x => TotalSpace.mk' F x (s x : V x) := + s.contMDiff.mdifferentiable hn + +protected theorem mdifferentiable (s : Cₛ^∞⟮I; F, V⟯) : + MDifferentiable I (I.prod 𝓘(𝕜, F)) fun x => TotalSpace.mk' F x (s x : V x) := + s.contMDiff.mdifferentiable le_top + +protected theorem mdifferentiableAt (s : Cₛ^∞⟮I; F, V⟯) {x} : + MDifferentiableAt I (I.prod 𝓘(𝕜, F)) (fun x => TotalSpace.mk' F x (s x : V x)) x := + s.mdifferentiable x + end ContMDiffSection diff --git a/Mathlib/MeasureTheory/Covering/Differentiation.lean b/Mathlib/MeasureTheory/Covering/Differentiation.lean index 6dfb81371c2f3..9889a9e2e7aa6 100644 --- a/Mathlib/MeasureTheory/Covering/Differentiation.lean +++ b/Mathlib/MeasureTheory/Covering/Differentiation.lean @@ -143,6 +143,12 @@ theorem measure_le_of_frequently_le [SecondCountableTopology α] [BorelSpace α] _ ≤ ν U := (measure_mono (iUnion_subset fun i => (h.covering_mem i.2).2)) _ ≤ ν s + ε := νU +theorem eventually_filterAt_integrableOn (x : α) {f : α → E} (hf : LocallyIntegrable f μ) : + ∀ᶠ a in v.filterAt x, IntegrableOn f a μ := by + rcases hf x with ⟨w, w_nhds, hw⟩ + filter_upwards [v.eventually_filterAt_subset_of_nhds w_nhds] with a ha + exact hw.mono_set ha + section variable [SecondCountableTopology α] [BorelSpace α] [IsLocallyFiniteMeasure μ] {ρ : Measure α} @@ -871,12 +877,6 @@ theorem ae_tendsto_lintegral_nnnorm_sub_div {f : α → E} (hf : LocallyIntegrab refine setLIntegral_congr_fun h'a (eventually_of_forall (fun y hy ↦ ?_)) rw [indicator_of_mem (ha hy) f, indicator_of_mem hn f] -theorem eventually_filterAt_integrableOn (x : α) {f : α → E} (hf : LocallyIntegrable f μ) : - ∀ᶠ a in v.filterAt x, IntegrableOn f a μ := by - rcases hf x with ⟨w, w_nhds, hw⟩ - filter_upwards [v.eventually_filterAt_subset_of_nhds w_nhds] with a ha - exact hw.mono_set ha - /-- *Lebesgue differentiation theorem*: for almost every point `x`, the average of `‖f y - f x‖` on `a` tends to `0` as `a` shrinks to `x` along a Vitali family. -/ theorem ae_tendsto_average_norm_sub {f : α → E} (hf : LocallyIntegrable f μ) : diff --git a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean index f2206edcca2dd..da28661366bca 100644 --- a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean +++ b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean @@ -517,7 +517,7 @@ lemma setLIntegral_rnDeriv_mul [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ @[deprecated (since := "2024-06-29")] alias set_lintegral_rnDeriv_mul := setLIntegral_rnDeriv_mul -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] theorem integrable_rnDeriv_smul_iff [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ ν) [SigmaFinite μ] {f : α → E} : diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean index ee64cc2e488fa..cdaa8675a1517 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean @@ -71,9 +71,6 @@ namespace MeasureTheory variable {α F F' 𝕜 : Type*} {p : ℝ≥0∞} [RCLike 𝕜] -- 𝕜 for ℝ or ℂ - -- F for a Lp submodule - [NormedAddCommGroup F] - [NormedSpace 𝕜 F] -- F' for integrals on a Lp submodule [NormedAddCommGroup F'] [NormedSpace 𝕜 F'] [NormedSpace ℝ F'] [CompleteSpace F'] diff --git a/Mathlib/MeasureTheory/Integral/Average.lean b/Mathlib/MeasureTheory/Integral/Average.lean index 99145e76869a1..a779b9ccf468c 100644 --- a/Mathlib/MeasureTheory/Integral/Average.lean +++ b/Mathlib/MeasureTheory/Integral/Average.lean @@ -40,7 +40,7 @@ open ENNReal MeasureTheory MeasureTheory.Measure Metric Set Filter TopologicalSp open scoped Topology ENNReal Convex variable {α E F : Type*} {m0 : MeasurableSpace α} [NormedAddCommGroup E] [NormedSpace ℝ E] - [CompleteSpace E] [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpace F] {μ ν : Measure α} + [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpace F] {μ ν : Measure α} {s t : Set α} /-! @@ -335,7 +335,8 @@ theorem average_add_measure [IsFiniteMeasure μ] {ν : Measure α} [IsFiniteMeas ← integral_add_measure hμ hν, ← ENNReal.toReal_add (measure_ne_top μ _) (measure_ne_top ν _)] rw [average_eq, Measure.add_apply] -theorem average_pair {f : α → E} {g : α → F} (hfi : Integrable f μ) (hgi : Integrable g μ) : +theorem average_pair [CompleteSpace E] + {f : α → E} {g : α → F} (hfi : Integrable f μ) (hgi : Integrable g μ) : ⨍ x, (f x, g x) ∂μ = (⨍ x, f x ∂μ, ⨍ x, g x ∂μ) := integral_pair hfi.to_average hgi.to_average @@ -384,6 +385,8 @@ theorem average_mem_openSegment_compl_self [IsFiniteMeasure μ] {f : α → E} { average_union_mem_openSegment aedisjoint_compl_right hs.compl hs₀ hsc₀ (measure_ne_top _ _) (measure_ne_top _ _) hfi.integrableOn hfi.integrableOn +variable [CompleteSpace E] + @[simp] theorem average_const (μ : Measure α) [IsFiniteMeasure μ] [h : NeZero μ] (c : E) : ⨍ _x, c ∂μ = c := by @@ -725,6 +728,7 @@ we require that `⨍ y in a i, ‖f y - c‖ ∂μ` tends to `0`), then the inte to `c` if `gₙ` is supported in `aₙ`, has integral converging to one and supremum at most `K / μ aₙ`. -/ theorem tendsto_integral_smul_of_tendsto_average_norm_sub + [CompleteSpace E] {ι : Type*} {a : ι → Set α} {l : Filter ι} {f : α → E} {c : E} {g : ι → α → ℝ} (K : ℝ) (hf : Tendsto (fun i ↦ ⨍ y in a i, ‖f y - c‖ ∂μ) l (𝓝 0)) (f_int : ∀ᶠ i in l, IntegrableOn f (a i) μ) diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 7bc444ca0b8c6..3490d60b63ba6 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1416,7 +1416,10 @@ theorem tendsto_integral_approxOn_of_measurable_of_range_subset [MeasurableSpace apply tendsto_integral_approxOn_of_measurable hf fmeas _ _ (integrable_zero _ _ _) exact eventually_of_forall fun x => subset_closure (hs (Set.mem_union_left _ (mem_range_self _))) -theorem tendsto_integral_norm_approxOn_sub [MeasurableSpace E] [BorelSpace E] {f : α → E} +-- We redeclare `E` here to temporarily avoid +-- the `[CompleteSpace E]` and `[NormedSpace ℝ E]` instances. +theorem tendsto_integral_norm_approxOn_sub + {E : Type*} [NormedAddCommGroup E] [MeasurableSpace E] [BorelSpace E] {f : α → E} (fmeas : Measurable f) (hf : Integrable f μ) [SeparableSpace (range f ∪ {0} : Set E)] : Tendsto (fun n ↦ ∫ x, ‖SimpleFunc.approxOn f fmeas (range f ∪ {0}) 0 (by simp) n x - f x‖ ∂μ) atTop (𝓝 0) := by diff --git a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean index 1a765f962c2b4..8de4e345f0f17 100644 --- a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean +++ b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean @@ -271,8 +271,6 @@ theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae' [IsMeasurablyGenera · simp_rw [intervalIntegral] abel -variable [CompleteSpace E] - /-- **Fundamental theorem of calculus-1**, local version for any measure. Let filters `l` and `l'` be related by `TendstoIxxClass Ioc`. If `f` has a finite limit `c` at `l ⊓ ae μ`, where `μ` is a measure @@ -283,7 +281,8 @@ See also `measure_integral_sub_linear_isLittleO_of_tendsto_ae_of_le` for a versi `[intervalIntegral.FTCFilter a l l']` and `[MeasureTheory.IsLocallyFiniteMeasure μ]`. If `l` is one of `𝓝[≥] a`, `𝓝[≤] a`, `𝓝 a`, then it's easier to apply the non-primed version. The primed version also works, e.g., for `l = l' = Filter.atTop`. -/ -theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae_of_le' [IsMeasurablyGenerated l'] +theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae_of_le' + [CompleteSpace E] [IsMeasurablyGenerated l'] [TendstoIxxClass Ioc l l'] (hfm : StronglyMeasurableAtFilter f l' μ) (hf : Tendsto f (l' ⊓ ae μ) (𝓝 c)) (hl : μ.FiniteAtFilter l') (hu : Tendsto u lt l) (hv : Tendsto v lt l) (huv : u ≤ᶠ[lt] v) : @@ -303,7 +302,8 @@ See also `measure_integral_sub_linear_is_o_of_tendsto_ae_of_ge` for a version as `[intervalIntegral.FTCFilter a l l']` and `[MeasureTheory.IsLocallyFiniteMeasure μ]`. If `l` is one of `𝓝[≥] a`, `𝓝[≤] a`, `𝓝 a`, then it's easier to apply the non-primed version. The primed version also works, e.g., for `l = l' = Filter.atTop`. -/ -theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae_of_ge' [IsMeasurablyGenerated l'] +theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae_of_ge' + [CompleteSpace E] [IsMeasurablyGenerated l'] [TendstoIxxClass Ioc l l'] (hfm : StronglyMeasurableAtFilter f l' μ) (hf : Tendsto f (l' ⊓ ae μ) (𝓝 c)) (hl : μ.FiniteAtFilter l') (hu : Tendsto u lt l) (hv : Tendsto v lt l) (huv : v ≤ᶠ[lt] u) : @@ -315,7 +315,7 @@ theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae_of_ge' [IsMeasurably section -variable [IsLocallyFiniteMeasure μ] [FTCFilter a l l'] +variable [IsLocallyFiniteMeasure μ] /-- **Fundamental theorem of calculus-1**, local version for any measure. @@ -328,7 +328,7 @@ for `l = l' = Filter.atTop`. We use integrals of constants instead of measures because this way it is easier to formulate a statement that works in both cases `u ≤ v` and `v ≤ u`. -/ -theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae +theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae [FTCFilter a l l'] (hfm : StronglyMeasurableAtFilter f l' μ) (hf : Tendsto f (l' ⊓ ae μ) (𝓝 c)) (hu : Tendsto u lt l) (hv : Tendsto v lt l) : (fun t => (∫ x in u t..v t, f x ∂μ) - ∫ _ in u t..v t, c ∂μ) =o[lt] fun t => @@ -345,6 +345,7 @@ finite measure. If `f` has a finite limit `c` at `l' ⊓ ae μ`, then See also `measure_integral_sub_linear_isLittleO_of_tendsto_ae_of_le'` for a version that also works, e.g., for `l = l' = Filter.atTop`. -/ theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae_of_le + [CompleteSpace E] [FTCFilter a l l'] (hfm : StronglyMeasurableAtFilter f l' μ) (hf : Tendsto f (l' ⊓ ae μ) (𝓝 c)) (hu : Tendsto u lt l) (hv : Tendsto v lt l) (huv : u ≤ᶠ[lt] v) : (fun t => (∫ x in u t..v t, f x ∂μ) - (μ (Ioc (u t) (v t))).toReal • c) =o[lt] fun t => @@ -362,6 +363,7 @@ finite measure. If `f` has a finite limit `c` at `l' ⊓ ae μ`, then See also `measure_integral_sub_linear_isLittleO_of_tendsto_ae_of_ge'` for a version that also works, e.g., for `l = l' = Filter.atTop`. -/ theorem measure_integral_sub_linear_isLittleO_of_tendsto_ae_of_ge + [CompleteSpace E] [FTCFilter a l l'] (hfm : StronglyMeasurableAtFilter f l' μ) (hf : Tendsto f (l' ⊓ ae μ) (𝓝 c)) (hu : Tendsto u lt l) (hv : Tendsto v lt l) (huv : v ≤ᶠ[lt] u) : (fun t => (∫ x in u t..v t, f x ∂μ) + (μ (Ioc (v t) (u t))).toReal • c) =o[lt] fun t => diff --git a/Mathlib/MeasureTheory/Integral/IntervalAverage.lean b/Mathlib/MeasureTheory/Integral/IntervalAverage.lean index 55de62c64aa22..0ecbef306508c 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalAverage.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalAverage.lean @@ -29,7 +29,7 @@ open MeasureTheory Set TopologicalSpace open scoped Interval -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] notation3 "⨍ "(...)" in "a".."b", "r:60:(scoped f => average (Measure.restrict volume (uIoc a b)) f) => r diff --git a/Mathlib/MeasureTheory/Measure/DiracProba.lean b/Mathlib/MeasureTheory/Measure/DiracProba.lean index 4e2bbef462da9..4dd9303f1e13b 100644 --- a/Mathlib/MeasureTheory/Measure/DiracProba.lean +++ b/Mathlib/MeasureTheory/Measure/DiracProba.lean @@ -122,7 +122,9 @@ guarantee injectivity of `diracProba`). -/ noncomputable def diracProbaInverse : range (diracProba (X := X)) → X := fun μ' ↦ (mem_range.mp μ'.prop).choose -@[simp] lemma diracProba_diracProbaInverse (μ : range (diracProba (X := X))) : +-- We redeclare `X` here to temporarily avoid the `[TopologicalSpace X]` instance. +@[simp] lemma diracProba_diracProbaInverse {X : Type*} [MeasurableSpace X] + (μ : range (diracProba (X := X))) : diracProba (diracProbaInverse μ) = μ := (mem_range.mp μ.prop).choose_spec lemma diracProbaInverse_eq [T0Space X] {x : X} {μ : range (diracProba (X := X))} @@ -171,15 +173,15 @@ lemma continuous_diracProbaEquivSymm [T0Space X] [CompletelyRegularSpace X] : /-- In a completely regular T0 topological space `X`, `diracProbaEquiv` is a homeomorphism to its image in `ProbabilityMeasure X`. -/ -noncomputable def diracProbaHomeomorph [T0Space X] [CompletelyRegularSpace X] - [MeasurableSpace X] [OpensMeasurableSpace X] : X ≃ₜ range (diracProba (X := X)) := +noncomputable def diracProbaHomeomorph [T0Space X] [CompletelyRegularSpace X] : + X ≃ₜ range (diracProba (X := X)) := @Homeomorph.mk X _ _ _ diracProbaEquiv continuous_diracProbaEquiv continuous_diracProbaEquivSymm /-- If `X` is a completely regular T0 space with its Borel sigma algebra, then the mapping that takes a point `x : X` to the delta-measure `diracProba x` is an embedding `X → ProbabilityMeasure X`. -/ -theorem embedding_diracProba [T0Space X] [CompletelyRegularSpace X] - [MeasurableSpace X] [OpensMeasurableSpace X] : Embedding (fun (x : X) ↦ diracProba x) := +theorem embedding_diracProba [T0Space X] [CompletelyRegularSpace X] : + Embedding (fun (x : X) ↦ diracProba x) := embedding_subtype_val.comp diracProbaHomeomorph.embedding end embed_to_probabilityMeasure diff --git a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean index f6983e238a808..006c73306931e 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean @@ -201,12 +201,12 @@ variable [NormedAddCommGroup F'] [InnerProductSpace ℝ F'] [FiniteDimensional [MeasurableSpace F'] [BorelSpace F'] variable (f : E' ≃ₗᵢ[ℝ] F') -variable [NormedAddCommGroup A] [NormedSpace ℝ A] +variable [NormedAddCommGroup A] theorem integrable_comp (g : F' → A) : Integrable (g ∘ f) ↔ Integrable g := f.measurePreserving.integrable_comp_emb f.toMeasureEquiv.measurableEmbedding -theorem integral_comp (g : F' → A) : ∫ (x : E'), g (f x) = ∫ (y : F'), g y := +theorem integral_comp [NormedSpace ℝ A] (g : F' → A) : ∫ (x : E'), g (f x) = ∫ (y : F'), g y := f.measurePreserving.integral_comp' (f := f.toMeasureEquiv) g end InnerProductSpace diff --git a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean index 4e0f6d1bdedd7..fd452eb706c2f 100644 --- a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean +++ b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean @@ -456,12 +456,13 @@ section Levy_Prokhorov_metrizes_convergence_in_distribution open BoundedContinuousFunction TopologicalSpace -variable {ι : Type*} (Ω : Type*) [PseudoMetricSpace Ω] [SeparableSpace Ω] +variable {ι : Type*} (Ω : Type*) [PseudoMetricSpace Ω] variable [MeasurableSpace Ω] [OpensMeasurableSpace Ω] /-- In a separable pseudometric space, for any ε > 0 there exists a countable collection of disjoint Borel measurable subsets of diameter at most ε that cover the whole space. -/ -lemma SeparableSpace.exists_measurable_partition_diam_le {ε : ℝ} (ε_pos : 0 < ε) : +lemma SeparableSpace.exists_measurable_partition_diam_le [SeparableSpace Ω] + {ε : ℝ} (ε_pos : 0 < ε) : ∃ (As : ℕ → Set Ω), (∀ n, MeasurableSet (As n)) ∧ (∀ n, Bornology.IsBounded (As n)) ∧ (∀ n, diam (As n) ≤ ε) ∧ (⋃ n, As n = univ) ∧ (Pairwise (fun (n m : ℕ) ↦ Disjoint (As n) (As m))) := by @@ -512,7 +513,7 @@ lemma ProbabilityMeasure.toMeasure_add_pos_gt_mem_nhds (P : ProbabilityMeasure convert ENNReal.add_lt_add_right ε_top hQ exact (tsub_add_cancel_of_le easy).symm -lemma ProbabilityMeasure.continuous_toLevyProkhorov : +lemma ProbabilityMeasure.continuous_toLevyProkhorov [SeparableSpace Ω] : Continuous (ProbabilityMeasure.toLevyProkhorov (Ω := Ω)) := by -- We check continuity of `id : ProbabilityMeasure Ω → LevyProkhorov (ProbabilityMeasure Ω)` at -- each point `P : ProbabilityMeasure Ω`. @@ -611,7 +612,7 @@ lemma ProbabilityMeasure.continuous_toLevyProkhorov : /-- The topology of the Lévy-Prokhorov metric on probability measures on a separable space coincides with the topology of convergence in distribution. -/ -theorem levyProkhorov_eq_convergenceInDistribution : +theorem levyProkhorov_eq_convergenceInDistribution [SeparableSpace Ω] : (inferInstance : TopologicalSpace (ProbabilityMeasure Ω)) = TopologicalSpace.coinduced (LevyProkhorov.toProbabilityMeasure (Ω := Ω)) inferInstance := le_antisymm (ProbabilityMeasure.continuous_toLevyProkhorov (Ω := Ω)).coinduced_le @@ -619,7 +620,7 @@ theorem levyProkhorov_eq_convergenceInDistribution : /-- The identity map is a homeomorphism from `ProbabilityMeasure Ω` with the topology of convergence in distribution to `ProbabilityMeasure Ω` with the Lévy-Prokhorov (pseudo)metric. -/ -def homeomorph_probabilityMeasure_levyProkhorov : +def homeomorph_probabilityMeasure_levyProkhorov [SeparableSpace Ω] : ProbabilityMeasure Ω ≃ₜ LevyProkhorov (ProbabilityMeasure Ω) where toFun := ProbabilityMeasure.toLevyProkhorov (Ω := Ω) invFun := LevyProkhorov.toProbabilityMeasure (Ω := Ω) diff --git a/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean b/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean index 297f126f93f2f..91e480a646863 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensityVectorMeasure.lean @@ -33,7 +33,7 @@ namespace MeasureTheory open TopologicalSpace variable {μ ν : Measure α} -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] open Classical in /-- Given a measure `μ` and an integrable function `f`, `μ.withDensityᵥ f` is @@ -143,7 +143,7 @@ theorem Measure.withDensityᵥ_absolutelyContinuous (μ : Measure α) (f : α exact VectorMeasure.AbsolutelyContinuous.zero _ /-- Having the same density implies the underlying functions are equal almost everywhere. -/ -theorem Integrable.ae_eq_of_withDensityᵥ_eq {f g : α → E} (hf : Integrable f μ) +theorem Integrable.ae_eq_of_withDensityᵥ_eq [CompleteSpace E] {f g : α → E} (hf : Integrable f μ) (hg : Integrable g μ) (hfg : μ.withDensityᵥ f = μ.withDensityᵥ g) : f =ᵐ[μ] g := by refine hf.ae_eq_of_forall_setIntegral_eq f g hg fun i hi _ => ?_ rw [← withDensityᵥ_apply hf hi, hfg, withDensityᵥ_apply hg hi] @@ -157,7 +157,8 @@ theorem WithDensityᵥEq.congr_ae {f g : α → E} (h : f =ᵐ[μ] g) : · have hg : ¬Integrable g μ := by intro hg; exact hf (hg.congr h.symm) rw [withDensityᵥ, withDensityᵥ, dif_neg hf, dif_neg hg] -theorem Integrable.withDensityᵥ_eq_iff {f g : α → E} (hf : Integrable f μ) (hg : Integrable g μ) : +theorem Integrable.withDensityᵥ_eq_iff [CompleteSpace E] + {f g : α → E} (hf : Integrable f μ) (hg : Integrable g μ) : μ.withDensityᵥ f = μ.withDensityᵥ g ↔ f =ᵐ[μ] g := ⟨fun hfg => hf.ae_eq_of_withDensityᵥ_eq hg hfg, fun h => WithDensityᵥEq.congr_ae h⟩ diff --git a/Mathlib/Probability/Density.lean b/Mathlib/Probability/Density.lean index 07eab99eea6d8..9efc5692f7cbf 100644 --- a/Mathlib/Probability/Density.lean +++ b/Mathlib/Probability/Density.lean @@ -223,7 +223,7 @@ theorem lintegral_pdf_mul {X : Ω → E} [HasPDF X ℙ μ] {f : E → ℝ≥0∞ ← lintegral_map' (hf.mono_ac HasPDF.absolutelyContinuous) (HasPDF.aemeasurable X ℙ μ), lintegral_rnDeriv_mul HasPDF.absolutelyContinuous hf] -variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpace F] +variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] theorem integrable_pdf_smul_iff [IsFiniteMeasure ℙ] {X : Ω → E} [HasPDF X ℙ μ] {f : E → F} (hf : AEStronglyMeasurable f μ) : @@ -238,7 +238,8 @@ theorem integrable_pdf_smul_iff [IsFiniteMeasure ℙ] {X : Ω → E} [HasPDF X /-- **The Law of the Unconscious Statistician**: Given a random variable `X` and a measurable function `f`, `f ∘ X` is a random variable with expectation `∫ x, pdf X x • f x ∂μ` where `μ` is a measure on the codomain of `X`. -/ -theorem integral_pdf_smul [IsFiniteMeasure ℙ] {X : Ω → E} [HasPDF X ℙ μ] {f : E → F} +theorem integral_pdf_smul [IsFiniteMeasure ℙ] + {X : Ω → E} [HasPDF X ℙ μ] {f : E → F} (hf : AEStronglyMeasurable f μ) : ∫ x, (pdf X ℙ μ x).toReal • f x ∂μ = ∫ x, f (X x) ∂ℙ := by rw [← integral_map (HasPDF.aemeasurable X ℙ μ) (hf.mono_ac HasPDF.absolutelyContinuous), map_eq_withDensity_pdf X ℙ μ, pdf_def, integral_rnDeriv_smul HasPDF.absolutelyContinuous, diff --git a/Mathlib/Probability/Independence/ZeroOne.lean b/Mathlib/Probability/Independence/ZeroOne.lean index 2c5d4c925a63d..8af5f5af8b3e6 100644 --- a/Mathlib/Probability/Independence/ZeroOne.lean +++ b/Mathlib/Probability/Independence/ZeroOne.lean @@ -71,7 +71,7 @@ theorem condexp_eq_zero_or_one_of_condIndepSet_self | inl h => exact Or.inl (Or.inl h) | inr h => exact Or.inr h -variable [IsMarkovKernel κ] [IsProbabilityMeasure μ] +variable [IsMarkovKernel κ] open Filter @@ -79,7 +79,8 @@ theorem Kernel.indep_biSup_compl (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s Indep (⨆ n ∈ t, s n) (⨆ n ∈ tᶜ, s n) κ μα := indep_iSup_of_disjoint h_le h_indep disjoint_compl_right -theorem indep_biSup_compl (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (t : Set ι) : +theorem indep_biSup_compl [IsProbabilityMeasure μ] + (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (t : Set ι) : Indep (⨆ n ∈ t, s n) (⨆ n ∈ tᶜ, s n) μ := Kernel.indep_biSup_compl h_le h_indep t @@ -111,7 +112,8 @@ theorem Kernel.indep_biSup_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s simp only [Set.mem_compl_iff, eventually_map] exact eventually_of_mem (hf t ht) le_iSup₂ -theorem indep_biSup_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (hf : ∀ t, p t → tᶜ ∈ f) +theorem indep_biSup_limsup [IsProbabilityMeasure μ] + (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (hf : ∀ t, p t → tᶜ ∈ f) {t : Set ι} (ht : p t) : Indep (⨆ n ∈ t, s n) (limsup s f) μ := Kernel.indep_biSup_limsup h_le h_indep hf ht @@ -136,7 +138,8 @@ theorem Kernel.indep_iSup_directed_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : · exact hc.1 hn · exact hc.2 hn -theorem indep_iSup_directed_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) +theorem indep_iSup_directed_limsup [IsProbabilityMeasure μ] + (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (hf : ∀ t, p t → tᶜ ∈ f) (hns : Directed (· ≤ ·) ns) (hnsp : ∀ a, p (ns a)) : Indep (⨆ a, ⨆ n ∈ ns a, s n) (limsup s f) μ := Kernel.indep_iSup_directed_limsup h_le h_indep hf hns hnsp @@ -161,7 +164,8 @@ theorem Kernel.indep_iSup_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s haveI : Nonempty (∃ i : α, n ∈ ns i) := ⟨hns_univ n⟩ rw [h, iSup_const] -theorem indep_iSup_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (hf : ∀ t, p t → tᶜ ∈ f) +theorem indep_iSup_limsup [IsProbabilityMeasure μ] + (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (hf : ∀ t, p t → tᶜ ∈ f) (hns : Directed (· ≤ ·) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) : Indep (⨆ n, s n) (limsup s f) μ := Kernel.indep_iSup_limsup h_le h_indep hf hns hnsp hns_univ @@ -179,7 +183,8 @@ theorem Kernel.indep_limsup_self (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s Indep (limsup s f) (limsup s f) κ μα := indep_of_indep_of_le_left (indep_iSup_limsup h_le h_indep hf hns hnsp hns_univ) limsup_le_iSup -theorem indep_limsup_self (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (hf : ∀ t, p t → tᶜ ∈ f) +theorem indep_limsup_self [IsProbabilityMeasure μ] + (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (hf : ∀ t, p t → tᶜ ∈ f) (hns : Directed (· ≤ ·) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) : Indep (limsup s f) (limsup s f) μ := Kernel.indep_limsup_self h_le h_indep hf hns hnsp hns_univ @@ -200,7 +205,8 @@ theorem Kernel.measure_zero_or_one_of_measurableSet_limsup (h_le : ∀ n, s n ((indep_limsup_self h_le h_indep hf hns hnsp hns_univ).indepSet_of_measurableSet ht_tail ht_tail) -theorem measure_zero_or_one_of_measurableSet_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) +theorem measure_zero_or_one_of_measurableSet_limsup [IsProbabilityMeasure μ] + (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) (hf : ∀ t, p t → tᶜ ∈ f) (hns : Directed (· ≤ ·) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) {t : Set Ω} (ht_tail : MeasurableSet[limsup s f] t) : μ t = 0 ∨ μ t = 1 := by @@ -243,7 +249,8 @@ theorem Kernel.indep_limsup_atTop_self (h_le : ∀ n, s n ≤ m0) (h_indep : iIn · exact Monotone.directed_le fun i j hij k hki => le_trans hki hij · exact fun n => ⟨n, le_rfl⟩ -theorem indep_limsup_atTop_self (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) : +theorem indep_limsup_atTop_self [IsProbabilityMeasure μ] + (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) : Indep (limsup s atTop) (limsup s atTop) μ := Kernel.indep_limsup_atTop_self h_le h_indep @@ -262,7 +269,8 @@ theorem Kernel.measure_zero_or_one_of_measurableSet_limsup_atTop (h_le : ∀ n, /-- **Kolmogorov's 0-1 law** : any event in the tail σ-algebra of an independent sequence of sub-σ-algebras has probability 0 or 1. The tail σ-algebra `limsup s atTop` is the same as `⋂ n, ⋃ i ≥ n, s i`. -/ -theorem measure_zero_or_one_of_measurableSet_limsup_atTop (h_le : ∀ n, s n ≤ m0) +theorem measure_zero_or_one_of_measurableSet_limsup_atTop [IsProbabilityMeasure μ] + (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) {t : Set Ω} (ht_tail : MeasurableSet[limsup s atTop] t) : μ t = 0 ∨ μ t = 1 := by simpa only [ae_dirac_eq, Filter.eventually_pure] @@ -295,7 +303,8 @@ theorem Kernel.indep_limsup_atBot_self (h_le : ∀ n, s n ≤ m0) (h_indep : iIn · exact Antitone.directed_le fun _ _ ↦ Set.Ici_subset_Ici.2 · exact fun n => ⟨n, le_rfl⟩ -theorem indep_limsup_atBot_self (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) : +theorem indep_limsup_atBot_self [IsProbabilityMeasure μ] + (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) : Indep (limsup s atBot) (limsup s atBot) μ := Kernel.indep_limsup_atBot_self h_le h_indep @@ -315,8 +324,9 @@ theorem Kernel.measure_zero_or_one_of_measurableSet_limsup_atBot (h_le : ∀ n, /-- **Kolmogorov's 0-1 law** : any event in the tail σ-algebra of an independent sequence of sub-σ-algebras has probability 0 or 1. -/ -theorem measure_zero_or_one_of_measurableSet_limsup_atBot (h_le : ∀ n, s n ≤ m0) - (h_indep : iIndep s μ) {t : Set Ω} (ht_tail : MeasurableSet[limsup s atBot] t) : +theorem measure_zero_or_one_of_measurableSet_limsup_atBot [IsProbabilityMeasure μ] + (h_le : ∀ n, s n ≤ m0) (h_indep : iIndep s μ) {t : Set Ω} + (ht_tail : MeasurableSet[limsup s atBot] t) : μ t = 0 ∨ μ t = 1 := by simpa only [ae_dirac_eq, Filter.eventually_pure] using Kernel.measure_zero_or_one_of_measurableSet_limsup_atBot h_le h_indep ht_tail diff --git a/Mathlib/Probability/Kernel/CondDistrib.lean b/Mathlib/Probability/Kernel/CondDistrib.lean index 9d9b8715ed169..b91377a2e288d 100644 --- a/Mathlib/Probability/Kernel/CondDistrib.lean +++ b/Mathlib/Probability/Kernel/CondDistrib.lean @@ -88,7 +88,7 @@ theorem _root_.MeasureTheory.AEStronglyMeasurable.ae_integrable_condDistrib_map_ Integrable f (μ.map fun a => (X a, Y a)) := by rw [condDistrib, ← hf.ae_integrable_condKernel_iff, Measure.fst_map_prod_mk₀ hY] -variable [NormedSpace ℝ F] [CompleteSpace F] +variable [NormedSpace ℝ F] theorem _root_.MeasureTheory.AEStronglyMeasurable.integral_condDistrib_map (hY : AEMeasurable Y μ) (hf : AEStronglyMeasurable f (μ.map fun a => (X a, Y a))) : @@ -153,7 +153,7 @@ theorem _root_.MeasureTheory.Integrable.integral_norm_condDistrib (hX : AEMeasur Integrable (fun a => ∫ y, ‖f (X a, y)‖ ∂condDistrib Y X μ (X a)) μ := (hf_int.integral_norm_condDistrib_map hY).comp_aemeasurable hX -variable [NormedSpace ℝ F] [CompleteSpace F] +variable [NormedSpace ℝ F] theorem _root_.MeasureTheory.Integrable.norm_integral_condDistrib_map (hY : AEMeasurable Y μ) (hf_int : Integrable f (μ.map fun a => (X a, Y a))) : diff --git a/Mathlib/Probability/Kernel/Condexp.lean b/Mathlib/Probability/Kernel/Condexp.lean index d152a5d0ad3f1..59fb17ba03c8b 100644 --- a/Mathlib/Probability/Kernel/Condexp.lean +++ b/Mathlib/Probability/Kernel/Condexp.lean @@ -151,7 +151,7 @@ theorem integrable_toReal_condexpKernel {s : Set Ω} (hs : MeasurableSet s) : end Integrability -lemma condexpKernel_ae_eq_condexp' [IsFiniteMeasure μ] {s : Set Ω} (hs : MeasurableSet s) : +lemma condexpKernel_ae_eq_condexp' {s : Set Ω} (hs : MeasurableSet s) : (fun ω ↦ (condexpKernel μ m ω s).toReal) =ᵐ[μ] μ⟦s | m ⊓ mΩ⟧ := by have h := condDistrib_ae_eq_condexp (μ := μ) (measurable_id'' (inf_le_right : m ⊓ mΩ ≤ mΩ)) measurable_id hs @@ -159,12 +159,12 @@ lemma condexpKernel_ae_eq_condexp' [IsFiniteMeasure μ] {s : Set Ω} (hs : Measu simp_rw [condexpKernel_apply_eq_condDistrib] exact h -lemma condexpKernel_ae_eq_condexp [IsFiniteMeasure μ] +lemma condexpKernel_ae_eq_condexp (hm : m ≤ mΩ) {s : Set Ω} (hs : MeasurableSet s) : (fun ω ↦ (condexpKernel μ m ω s).toReal) =ᵐ[μ] μ⟦s | m⟧ := (condexpKernel_ae_eq_condexp' hs).trans (by rw [inf_of_le_left hm]) -lemma condexpKernel_ae_eq_trim_condexp [IsFiniteMeasure μ] +lemma condexpKernel_ae_eq_trim_condexp (hm : m ≤ mΩ) {s : Set Ω} (hs : MeasurableSet s) : (fun ω ↦ (condexpKernel μ m ω s).toReal) =ᵐ[μ.trim hm] μ⟦s | m⟧ := by rw [ae_eq_trim_iff hm _ stronglyMeasurable_condexp] diff --git a/Mathlib/Probability/Kernel/Disintegration/Basic.lean b/Mathlib/Probability/Kernel/Disintegration/Basic.lean index 13f6e4c87bf7b..0b385aef78547 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Basic.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Basic.lean @@ -46,7 +46,7 @@ This section provides a predicate for a kernel to disintegrate a measure. -/ namespace MeasureTheory.Measure -variable (ρ : Measure (α × Ω)) [IsFiniteMeasure ρ] (ρCond : Kernel α Ω) +variable (ρ : Measure (α × Ω)) (ρCond : Kernel α Ω) /-- A kernel `ρCond` is a conditional kernel for a measure `ρ` if it disintegrates it in the sense that `ρ.fst ⊗ₘ ρCond = ρ`. -/ @@ -60,6 +60,8 @@ lemma disintegrate : ρ.fst ⊗ₘ ρCond = ρ := IsCondKernel.disintegrate lemma IsCondKernel.isSFiniteKernel (hρ : ρ ≠ 0) : IsSFiniteKernel ρCond := by contrapose! hρ; rwa [← ρ.disintegrate ρCond, Measure.compProd_of_not_isSFiniteKernel] +variable [IsFiniteMeasure ρ] + /-- Auxiliary lemma for `IsCondKernel.apply_of_ne_zero`. -/ private lemma IsCondKernel.apply_of_ne_zero_of_measurableSet [MeasurableSingletonClass α] {x : α} (hx : ρ.fst {x} ≠ 0) {s : Set Ω} (hs : MeasurableSet s) : diff --git a/Mathlib/Probability/Kernel/IntegralCompProd.lean b/Mathlib/Probability/Kernel/IntegralCompProd.lean index 10ea95e8e8bcb..25fa15ae04f0a 100644 --- a/Mathlib/Probability/Kernel/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/IntegralCompProd.lean @@ -139,8 +139,7 @@ theorem _root_.MeasureTheory.Integrable.integral_compProd [NormedSpace ℝ E] /-! ### Bochner integral -/ -variable [NormedSpace ℝ E] {E' : Type*} [NormedAddCommGroup E'] - [CompleteSpace E'] [NormedSpace ℝ E'] +variable [NormedSpace ℝ E] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace ℝ E'] theorem Kernel.integral_fn_integral_add ⦃f g : β × γ → E⦄ (F : E → E') (hf : Integrable f ((κ ⊗ₖ η) a)) (hg : Integrable g ((κ ⊗ₖ η) a)) : From 8a8bd683a2eabe67226980ba84d700859eedf159 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 9 Aug 2024 00:58:25 +0000 Subject: [PATCH 138/359] chore: backports for leanprover/lean4#4814 (part 38) (#15618) --- .../NonUnital.lean | 74 ++++++++-------- .../ContinuousFunctionalCalculus/Unital.lean | 84 ++++++++++--------- 2 files changed, 86 insertions(+), 72 deletions(-) diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index 7e7644559f0e2..cfbfc80013afc 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -99,7 +99,7 @@ section Main variable {R A : Type*} {p : A → Prop} [CommSemiring R] [Nontrivial R] [StarRing R] [MetricSpace R] variable [TopologicalSemiring R] [ContinuousStar R] [NonUnitalRing A] [StarRing A] variable [TopologicalSpace A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] -variable [NonUnitalContinuousFunctionalCalculus R p] +variable [instCFCₙ : NonUnitalContinuousFunctionalCalculus R p] lemma NonUnitalStarAlgHom.ext_continuousMap [UniqueNonUnitalContinuousFunctionalCalculus R A] (a : A) (φ ψ : C(σₙ R a, R)₀ →⋆ₙₐ[R] A) (hφ : Continuous φ) (hψ : Continuous ψ) @@ -383,6 +383,8 @@ lemma cfcₙ_const_mul_id (r : R) (a : A) (ha : p a := by cfc_tac) : cfcₙ (r * lemma cfcₙ_star_id : cfcₙ (star · : R → R) a = star a := by rw [cfcₙ_star _ a, cfcₙ_id' R a] +section Comp + variable [UniqueNonUnitalContinuousFunctionalCalculus R A] lemma cfcₙ_comp (g f : R → R) (a : A) @@ -429,6 +431,8 @@ lemma cfcₙ_comp_star (hf : ContinuousOn f (star '' (σₙ R a)) := by cfc_cont cfcₙ (f <| star ·) a = cfcₙ f (star a) := by rw [cfcₙ_comp' f star a, cfcₙ_star_id a] +end Comp + lemma CFC.eq_zero_of_quasispectrum_eq_zero (h_spec : σₙ R a ⊆ {0}) (ha : p a := by cfc_tac) : a = 0 := by simpa [cfcₙ_id R a] using cfcₙ_congr (a := a) (f := id) (g := fun _ : R ↦ 0) fun x ↦ by simp_all @@ -505,17 +509,17 @@ section Order section Semiring variable {R A : Type*} {p : A → Prop} [OrderedCommSemiring R] [Nontrivial R] -variable [StarRing R] [StarOrderedRing R] [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] +variable [StarRing R] [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] variable [∀ (α) [Zero α] [TopologicalSpace α], StarOrderedRing C(α, R)₀] variable [TopologicalSpace A] [NonUnitalRing A] [StarRing A] [PartialOrder A] [StarOrderedRing A] -variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] -variable [NonUnitalContinuousFunctionalCalculus R p] [NonnegSpectrumClass R A] +variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] +variable [NonUnitalContinuousFunctionalCalculus R p] lemma cfcₙHom_mono {a : A} (ha : p a) {f g : C(σₙ R a, R)₀} (hfg : f ≤ g) : cfcₙHom ha f ≤ cfcₙHom ha g := OrderHomClass.mono (cfcₙHom ha) hfg -lemma cfcₙHom_nonneg_iff {a : A} (ha : p a) {f : C(σₙ R a, R)₀} : +lemma cfcₙHom_nonneg_iff [NonnegSpectrumClass R A] {a : A} (ha : p a) {f : C(σₙ R a, R)₀} : 0 ≤ cfcₙHom ha f ↔ 0 ≤ f := by constructor · exact fun hf x ↦ @@ -533,15 +537,16 @@ lemma cfcₙ_mono {f g : R → R} {a : A} (h : ∀ x ∈ σₙ R a, f x ≤ g x) exact cfcₙHom_mono ha fun x ↦ h x.1 x.2 · simp only [cfcₙ_apply_of_not_predicate _ ha, le_rfl] -lemma cfcₙ_nonneg_iff (f : R → R) (a : A) (hf : ContinuousOn f (σₙ R a) := by cfc_cont_tac) +lemma cfcₙ_nonneg_iff [NonnegSpectrumClass R A] (f : R → R) (a : A) + (hf : ContinuousOn f (σₙ R a) := by cfc_cont_tac) (h0 : f 0 = 0 := by cfc_zero_tac) (ha : p a := by cfc_tac) : 0 ≤ cfcₙ f a ↔ ∀ x ∈ σₙ R a, 0 ≤ f x := by rw [cfcₙ_apply .., cfcₙHom_nonneg_iff, ContinuousMapZero.le_def] simp only [ContinuousMapZero.coe_mk, ContinuousMap.coe_mk, Set.restrict_apply, Subtype.forall] congr! -lemma StarOrderedRing.nonneg_iff_quasispectrum_nonneg (a : A) (ha : p a := by cfc_tac) : - 0 ≤ a ↔ ∀ x ∈ quasispectrum R a, 0 ≤ x := by +lemma StarOrderedRing.nonneg_iff_quasispectrum_nonneg [NonnegSpectrumClass R A] (a : A) + (ha : p a := by cfc_tac) : 0 ≤ a ↔ ∀ x ∈ quasispectrum R a, 0 ≤ x := by have := cfcₙ_nonneg_iff (id : R → R) a (by fun_prop) simpa [cfcₙ_id _ a ha] using this @@ -570,10 +575,10 @@ end Semiring section Ring variable {R A : Type*} {p : A → Prop} [OrderedCommRing R] [Nontrivial R] -variable [StarRing R] [StarOrderedRing R] [MetricSpace R] [TopologicalRing R] [ContinuousStar R] +variable [StarRing R] [MetricSpace R] [TopologicalRing R] [ContinuousStar R] variable [∀ (α) [Zero α] [TopologicalSpace α], StarOrderedRing C(α, R)₀] variable [TopologicalSpace A] [NonUnitalRing A] [StarRing A] [PartialOrder A] [StarOrderedRing A] -variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] +variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] variable [NonUnitalContinuousFunctionalCalculus R p] [NonnegSpectrumClass R A] lemma cfcₙHom_le_iff {a : A} (ha : p a) {f g : C(σₙ R a, R)₀} : @@ -603,10 +608,9 @@ section UnitalToNonUnital open ContinuousMapZero Set Uniformity ContinuousMap -variable {R A : Type*} {p : A → Prop} [Semifield R] [StarRing R] [MetricSpace R] [CompleteSpace R] +variable {R A : Type*} {p : A → Prop} [Semifield R] [StarRing R] [MetricSpace R] variable [TopologicalSemiring R] [ContinuousStar R] [Ring A] [StarRing A] [TopologicalSpace A] variable [Algebra R A] [ContinuousFunctionalCalculus R p] -variable [h_cpct : ∀ a : A, CompactSpace (spectrum R a)] variable (R) in /-- The non-unital continuous functional calculus obtained by restricting a unital calculus @@ -621,6 +625,30 @@ noncomputable def cfcₙHom_of_cfcHom {a : A} (ha : p a) : C(σₙ R a, R)₀ (cfcHom ha (R := R) : C(spectrum R a, R) →⋆ₙₐ[R] A).comp <| (ψ : C(σₙ R a, R) →⋆ₙₐ[R] C(spectrum R a, R)).comp e +lemma cfcₙHom_of_cfcHom_map_quasispectrum {a : A} (ha : p a) : + ∀ f : C(σₙ R a, R)₀, σₙ R (cfcₙHom_of_cfcHom R ha f) = range f := by + intro f + simp only [cfcₙHom_of_cfcHom] + rw [quasispectrum_eq_spectrum_union_zero] + simp only [NonUnitalStarAlgHom.comp_assoc, NonUnitalStarAlgHom.comp_apply, + NonUnitalStarAlgHom.coe_coe] + rw [cfcHom_map_spectrum ha] + ext x + constructor + · rintro (⟨x, rfl⟩ | rfl) + · exact ⟨⟨x.1, spectrum_subset_quasispectrum R a x.2⟩, rfl⟩ + · exact ⟨0, map_zero f⟩ + · rintro ⟨x, rfl⟩ + have hx := x.2 + simp_rw [quasispectrum_eq_spectrum_union_zero R a] at hx + obtain (hx | hx) := hx + · exact Or.inl ⟨⟨x.1, hx⟩, rfl⟩ + · apply Or.inr + simp only [Set.mem_singleton_iff] at hx ⊢ + rw [show x = 0 from Subtype.val_injective hx, map_zero] + +variable [CompleteSpace R] [h_cpct : ∀ a : A, CompactSpace (spectrum R a)] + lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : ClosedEmbedding (cfcₙHom_of_cfcHom R ha) := by let f : C(spectrum R a, σₙ R a) := @@ -645,28 +673,6 @@ lemma closedEmbedding_cfcₙHom_of_cfcHom {a : A} (ha : p a) : convert Filter.comap_const_of_mem this with ⟨u, v⟩ <;> ext ⟨x, rfl⟩ <;> [exact map_zero u; exact map_zero v] -lemma cfcₙHom_of_cfcHom_map_quasispectrum {a : A} (ha : p a) : - ∀ f : C(σₙ R a, R)₀, σₙ R (cfcₙHom_of_cfcHom R ha f) = range f := by - intro f - simp only [cfcₙHom_of_cfcHom] - rw [quasispectrum_eq_spectrum_union_zero] - simp only [NonUnitalStarAlgHom.comp_assoc, NonUnitalStarAlgHom.comp_apply, - NonUnitalStarAlgHom.coe_coe] - rw [cfcHom_map_spectrum ha] - ext x - constructor - · rintro (⟨x, rfl⟩ | rfl) - · exact ⟨⟨x.1, spectrum_subset_quasispectrum R a x.2⟩, rfl⟩ - · exact ⟨0, map_zero f⟩ - · rintro ⟨x, rfl⟩ - have hx := x.2 - simp_rw [quasispectrum_eq_spectrum_union_zero R a] at hx - obtain (hx | hx) := hx - · exact Or.inl ⟨⟨x.1, hx⟩, rfl⟩ - · apply Or.inr - simp only [Set.mem_singleton_iff] at hx ⊢ - rw [show x = 0 from Subtype.val_injective hx, map_zero] - instance ContinuousFunctionalCalculus.toNonUnital : NonUnitalContinuousFunctionalCalculus R p where predicate_zero := cfc_predicate_zero R exists_cfc_of_predicate _ ha := diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index 7defe194cdd22..ab9b6a6c87cf1 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -191,7 +191,7 @@ class UniqueContinuousFunctionalCalculus (R A : Type*) [CommSemiring R] [StarRin variable {R A : Type*} {p : A → Prop} [CommSemiring R] [StarRing R] [MetricSpace R] variable [TopologicalSemiring R] [ContinuousStar R] [TopologicalSpace A] [Ring A] [StarRing A] -variable [Algebra R A] [ContinuousFunctionalCalculus R p] +variable [Algebra R A] [instCFC : ContinuousFunctionalCalculus R p] lemma StarAlgHom.ext_continuousMap [UniqueContinuousFunctionalCalculus R A] (a : A) (φ ψ : C(spectrum R a, R) →⋆ₐ[R] A) (hφ : Continuous φ) (hψ : Continuous ψ) @@ -374,18 +374,15 @@ lemma eqOn_of_cfc_eq_cfc {f g : R → R} {a : A} (h : cfc f a = cfc g a) congrm($(this) ⟨x, hx⟩) variable {a f g} in -lemma cfc_eq_cfc_iff_eqOn (ha : p a := by cfc_tac) - (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) - (hg : ContinuousOn g (spectrum R a) := by cfc_cont_tac) : - cfc f a = cfc g a ↔ (spectrum R a).EqOn f g := +lemma cfc_eq_cfc_iff_eqOn : cfc f a = cfc g a ↔ (spectrum R a).EqOn f g := ⟨eqOn_of_cfc_eq_cfc, cfc_congr⟩ variable (R) -lemma cfc_one (ha : p a := by cfc_tac) : cfc (1 : R → R) a = 1 := +lemma cfc_one : cfc (1 : R → R) a = 1 := cfc_apply (1 : R → R) a ▸ map_one (cfcHom (show p a from ha)) -lemma cfc_const_one (ha : p a := by cfc_tac) : cfc (fun _ : R ↦ 1) a = 1 := cfc_one R a +lemma cfc_const_one : cfc (fun _ : R ↦ 1) a = 1 := cfc_one R a @[simp] lemma cfc_zero : cfc (0 : R → R) a = 0 := by @@ -493,7 +490,7 @@ lemma cfc_smul_id {S : Type*} [SMul S R] [ContinuousConstSMul S R] lemma cfc_const_mul_id (r : R) (a : A) (ha : p a := by cfc_tac) : cfc (r * ·) a = r • a := cfc_smul_id r a -lemma cfc_star_id (ha : p a := by cfc_tac) : cfc (star · : R → R) a = star a := by +lemma cfc_star_id : cfc (star · : R → R) a = star a := by rw [cfc_star .., cfc_id' ..] section Polynomial @@ -523,18 +520,7 @@ lemma cfc_polynomial (q : R[X]) (a : A) (ha : p a := by cfc_tac) : end Polynomial -lemma CFC.eq_algebraMap_of_spectrum_subset_singleton (r : R) (h_spec : spectrum R a ⊆ {r}) - (ha : p a := by cfc_tac) : a = algebraMap R A r := by - simpa [cfc_id R a, cfc_const r a] using - cfc_congr (f := id) (g := fun _ : R ↦ r) (a := a) fun x hx ↦ by simpa using h_spec hx - -lemma CFC.eq_zero_of_spectrum_subset_zero (h_spec : spectrum R a ⊆ {0}) (ha : p a := by cfc_tac) : - a = 0 := by - simpa using eq_algebraMap_of_spectrum_subset_singleton a 0 h_spec - -lemma CFC.eq_one_of_spectrum_subset_one (h_spec : spectrum R a ⊆ {1}) (ha : p a := by cfc_tac) : - a = 1 := by - simpa using eq_algebraMap_of_spectrum_subset_singleton a 1 h_spec +section Comp variable [UniqueContinuousFunctionalCalculus R A] @@ -587,6 +573,21 @@ lemma cfc_comp_polynomial (q : R[X]) (f : R → R) (a : A) cfc (f <| q.eval ·) a = cfc f (aeval a q) := by rw [cfc_comp' .., cfc_polynomial ..] +end Comp + +lemma CFC.eq_algebraMap_of_spectrum_subset_singleton (r : R) (h_spec : spectrum R a ⊆ {r}) + (ha : p a := by cfc_tac) : a = algebraMap R A r := by + simpa [cfc_id R a, cfc_const r a] using + cfc_congr (f := id) (g := fun _ : R ↦ r) (a := a) fun x hx ↦ by simpa using h_spec hx + +lemma CFC.eq_zero_of_spectrum_subset_zero (h_spec : spectrum R a ⊆ {0}) (ha : p a := by cfc_tac) : + a = 0 := by + simpa using eq_algebraMap_of_spectrum_subset_singleton a 0 h_spec + +lemma CFC.eq_one_of_spectrum_subset_one (h_spec : spectrum R a ⊆ {1}) (ha : p a := by cfc_tac) : + a = 1 := by + simpa using eq_algebraMap_of_spectrum_subset_singleton a 1 h_spec + lemma CFC.spectrum_algebraMap_subset (r : R) : spectrum R (algebraMap R A r) ⊆ {r} := by rw [← cfc_const r 0 (cfc_predicate_zero R), cfc_map_spectrum (fun _ ↦ r) 0 (cfc_predicate_zero R)] @@ -644,16 +645,15 @@ section Inv variable {R A : Type*} {p : A → Prop} [Semifield R] [StarRing R] [MetricSpace R] variable [TopologicalSemiring R] [ContinuousStar R] [TopologicalSpace A] variable [Ring A] [StarRing A] [Algebra R A] [ContinuousFunctionalCalculus R p] -variable (f : R → R) (a : A) -lemma isUnit_cfc_iff (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) +lemma isUnit_cfc_iff (f : R → R) (a : A) (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : IsUnit (cfc f a) ↔ ∀ x ∈ spectrum R a, f x ≠ 0 := by rw [← spectrum.zero_not_mem_iff R, cfc_map_spectrum ..] aesop alias ⟨_, isUnit_cfc⟩ := isUnit_cfc_iff -variable [HasContinuousInv₀ R] +variable [HasContinuousInv₀ R] (f : R → R) (a : A) /-- Bundle `cfc f a` into a unit given a proof that `f` is nonzero on the spectrum of `a`. -/ @[simps] @@ -697,7 +697,11 @@ lemma cfc_map_div (f g : R → R) (a : A) (hg' : ∀ x ∈ spectrum R a, g x ≠ simp only [div_eq_mul_inv] rw [cfc_mul .., cfc_inv g a hg'] -variable [UniqueContinuousFunctionalCalculus R A] +section ContinuousOnInvSpectrum +-- TODO: this section should probably be moved to another file altogether + +variable {R A : Type*} [Semifield R] [Ring A] [TopologicalSpace R] [HasContinuousInv₀ R] +variable [Algebra R A] @[fun_prop] lemma Units.continuousOn_inv₀_spectrum (a : Aˣ) : ContinuousOn (· ⁻¹) (spectrum R (a : A)) := @@ -705,16 +709,12 @@ lemma Units.continuousOn_inv₀_spectrum (a : Aˣ) : ContinuousOn (· ⁻¹) (sp simpa only [Set.subset_compl_singleton_iff] using spectrum.zero_not_mem R a.isUnit @[fun_prop] -lemma Units.continuousOn_zpow₀_spectrum (a : Aˣ) (n : ℤ) : +lemma Units.continuousOn_zpow₀_spectrum [ContinuousMul R] (a : Aˣ) (n : ℤ) : ContinuousOn (· ^ n) (spectrum R (a : A)) := (continuousOn_zpow₀ n).mono <| by simpa only [Set.subset_compl_singleton_iff] using spectrum.zero_not_mem R a.isUnit -lemma cfc_comp_inv (f : R → R) (a : Aˣ) - (hf : ContinuousOn f ((· ⁻¹) '' (spectrum R (a : A))) := by cfc_cont_tac) - (ha : p a := by cfc_tac) : - cfc (fun x ↦ f x⁻¹) (a : A) = cfc f (↑a⁻¹ : A) := by - rw [cfc_comp' .., cfc_inv_id _] +end ContinuousOnInvSpectrum lemma cfcUnits_zpow (hf' : ∀ x ∈ spectrum R a, f x ≠ 0) (n : ℤ) (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : @@ -737,6 +737,14 @@ lemma cfc_zpow (a : Aˣ) (n : ℤ) (ha : p a := by cfc_tac) : have := cfc_pow (fun x ↦ x⁻¹ : R → R) (n + 1) (a : A) exact this.trans <| congr($(cfc_inv_id a) ^ (n + 1)) +variable [UniqueContinuousFunctionalCalculus R A] + +lemma cfc_comp_inv (f : R → R) (a : Aˣ) + (hf : ContinuousOn f ((· ⁻¹) '' (spectrum R (a : A))) := by cfc_cont_tac) + (ha : p a := by cfc_tac) : + cfc (fun x ↦ f x⁻¹) (a : A) = cfc f (↑a⁻¹ : A) := by + rw [cfc_comp' .., cfc_inv_id _] + lemma cfc_comp_zpow (f : R → R) (n : ℤ) (a : Aˣ) (hf : ContinuousOn f ((· ^ n) '' (spectrum R (a : A))) := by cfc_cont_tac) (ha : p a := by cfc_tac) : @@ -785,17 +793,16 @@ section Order section Semiring variable {R A : Type*} {p : A → Prop} [OrderedCommSemiring R] [StarRing R] -variable [StarOrderedRing R] [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] +variable [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] variable [∀ (α) [TopologicalSpace α], StarOrderedRing C(α, R)] variable [TopologicalSpace A] [Ring A] [StarRing A] [PartialOrder A] [StarOrderedRing A] -variable [Algebra R A] [StarModule R A] [ContinuousFunctionalCalculus R p] -variable [NonnegSpectrumClass R A] +variable [Algebra R A] [ContinuousFunctionalCalculus R p] lemma cfcHom_mono {a : A} (ha : p a) {f g : C(spectrum R a, R)} (hfg : f ≤ g) : cfcHom ha f ≤ cfcHom ha g := OrderHomClass.mono (cfcHom ha) hfg -lemma cfcHom_nonneg_iff {a : A} (ha : p a) {f : C(spectrum R a, R)} : +lemma cfcHom_nonneg_iff [NonnegSpectrumClass R A] {a : A} (ha : p a) {f : C(spectrum R a, R)} : 0 ≤ cfcHom ha f ↔ 0 ≤ f := by constructor · exact fun hf x ↦ (cfcHom_map_spectrum ha (R := R) _ ▸ spectrum_nonneg_of_nonneg hf) ⟨x, rfl⟩ @@ -810,13 +817,14 @@ lemma cfc_mono {f g : R → R} {a : A} (h : ∀ x ∈ spectrum R a, f x ≤ g x) exact cfcHom_mono ha fun x ↦ h x.1 x.2 · simp only [cfc_apply_of_not_predicate _ ha, le_rfl] -lemma cfc_nonneg_iff (f : R → R) (a : A) (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) +lemma cfc_nonneg_iff [NonnegSpectrumClass R A] (f : R → R) (a : A) + (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : 0 ≤ cfc f a ↔ ∀ x ∈ spectrum R a, 0 ≤ f x := by rw [cfc_apply .., cfcHom_nonneg_iff, ContinuousMap.le_def] simp -lemma StarOrderedRing.nonneg_iff_spectrum_nonneg (a : A) (ha : p a := by cfc_tac) : - 0 ≤ a ↔ ∀ x ∈ spectrum R a, 0 ≤ x := by +lemma StarOrderedRing.nonneg_iff_spectrum_nonneg [NonnegSpectrumClass R A] (a : A) + (ha : p a := by cfc_tac) : 0 ≤ a ↔ ∀ x ∈ spectrum R a, 0 ≤ x := by have := cfc_nonneg_iff (id : R → R) a (by fun_prop) ha simpa [cfc_id _ a ha] using this @@ -871,7 +879,7 @@ variable {R A : Type*} {p : A → Prop} [OrderedCommRing R] [StarRing R] variable [MetricSpace R] [TopologicalRing R] [ContinuousStar R] variable [∀ (α) [TopologicalSpace α], StarOrderedRing C(α, R)] variable [TopologicalSpace A] [Ring A] [StarRing A] [PartialOrder A] [StarOrderedRing A] -variable [Algebra R A] [StarModule R A] [ContinuousFunctionalCalculus R p] +variable [Algebra R A] [ContinuousFunctionalCalculus R p] variable [NonnegSpectrumClass R A] lemma cfcHom_le_iff {a : A} (ha : p a) {f g : C(spectrum R a, R)} : From 1339db17c9abb2508774e59b96fe67b54fa87965 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 9 Aug 2024 02:03:51 +0000 Subject: [PATCH 139/359] feat(Topology/Algebra): add `smul_mem_nhds_smul` (#15563) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit - add `smul_mem_nhds_smul_iff` and `smul_mem_nhds_smul`, deprecate `smul_mem_nhds`; - move 3 lemmas up in the import graph; - add `smul_mem_nhds_self`; - rename `set_smul_mem_nhds_smul_iff` to `smul_mem_nhds_smul_iff₀`, deprecate the old lemma; - similarly with `set_smul_mem_nhds_smul` and `smul_mem_nhds_smul₀`; --- .../LocallyConvex/ContinuousOfBounded.lean | 3 +- .../MeasureTheory/Measure/Haar/Unique.lean | 6 +- Mathlib/Topology/Algebra/ConstMulAction.lean | 77 ++++++++++++------- Mathlib/Topology/Algebra/Group/Basic.lean | 25 ++---- .../Topology/Algebra/Group/OpenMapping.lean | 6 +- Mathlib/Topology/Maps/Basic.lean | 4 + 6 files changed, 64 insertions(+), 57 deletions(-) diff --git a/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean b/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean index de2f64500efae..9e29639ef61d0 100644 --- a/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean +++ b/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean @@ -63,8 +63,7 @@ def LinearMap.clmOfExistsBoundedImage (f : E →ₗ[𝕜] F) _ ⊆ f ⁻¹' U := by rw [inv_smul_smul₀ x_ne _] -- Using this inclusion, it suffices to show that `x⁻¹ • V` is in `𝓝 0`, which is trivial. refine mem_of_superset ?_ this - convert set_smul_mem_nhds_smul hV (inv_ne_zero x_ne) - exact (smul_zero _).symm⟩ + rwa [set_smul_mem_nhds_zero_iff (inv_ne_zero x_ne)]⟩ theorem LinearMap.clmOfExistsBoundedImage_coe {f : E →ₗ[𝕜] F} {h : ∃ V ∈ 𝓝 (0 : E), Bornology.IsVonNBounded 𝕜 (f '' V)} : diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index aa8ebde841979..fdad18683daec 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -776,9 +776,9 @@ theorem measure_isHaarMeasure_eq_smul_of_isEverywherePos [LocallyCompactSpace G] have C : Set.Countable (support fun (i : m) ↦ ρ (s ∩ (i : G) • k)) := Summable.countable_support_ennreal this.ne have : support (fun (i : m) ↦ ρ (s ∩ (i : G) • k)) = univ := by - apply eq_univ_iff_forall.2 (fun i ↦ ?_) - apply ne_of_gt (hρ (i : G) (hms i.2) _ _) - exact inter_mem_nhdsWithin s (by simpa using smul_mem_nhds (i : G) k_mem) + refine eq_univ_iff_forall.2 fun i ↦ ?_ + refine ne_of_gt (hρ (i : G) (hms i.2) _ ?_) + exact inter_mem_nhdsWithin s (by simpa) rw [this] at C have : Countable m := countable_univ_iff.mp C exact to_countable m diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index f9621b990d249..53ea74c2fe338 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -202,7 +202,7 @@ theorem continuous_const_smul_iff (c : G) : (Continuous fun x => c • f x) ↔ /-- The homeomorphism given by scalar multiplication by a given element of a group `Γ` acting on `T` is a homeomorphism from `T` to itself. -/ -@[to_additive] +@[to_additive (attr := simps!)] def Homeomorph.smul (γ : G) : α ≃ₜ α where toEquiv := MulAction.toPerm γ continuous_toFun := continuous_const_smul γ @@ -240,6 +240,29 @@ theorem Dense.smul (c : G) {s : Set α} (hs : Dense s) : Dense (c • s) := by theorem interior_smul (c : G) (s : Set α) : interior (c • s) = c • interior s := ((Homeomorph.smul c).image_interior s).symm +@[to_additive] +theorem IsOpen.smul_left {s : Set G} {t : Set α} (ht : IsOpen t) : IsOpen (s • t) := by + rw [← iUnion_smul_set] + exact isOpen_biUnion fun a _ => ht.smul _ + +@[to_additive] +theorem subset_interior_smul_right {s : Set G} {t : Set α} : s • interior t ⊆ interior (s • t) := + interior_maximal (Set.smul_subset_smul_left interior_subset) isOpen_interior.smul_left + +@[to_additive (attr := simp)] +theorem smul_mem_nhds_smul_iff {t : Set α} (g : G) {a : α} : g • t ∈ 𝓝 (g • a) ↔ t ∈ 𝓝 a := + (Homeomorph.smul g).openEmbedding.image_mem_nhds + +@[to_additive] alias ⟨_, smul_mem_nhds_smul⟩ := smul_mem_nhds_smul_iff + +@[to_additive (attr := deprecated (since := "2024-08-06"))] +alias smul_mem_nhds := smul_mem_nhds_smul + +@[to_additive (attr := simp)] +theorem smul_mem_nhds_self [TopologicalSpace G] [ContinuousConstSMul G G] {g : G} {s : Set G} : + g • s ∈ 𝓝 g ↔ s ∈ 𝓝 1 := by + rw [← smul_mem_nhds_smul_iff g⁻¹]; simp + end Group section GroupWithZero @@ -346,38 +369,35 @@ variable [Monoid M] [TopologicalSpace α] [MulAction M α] [ContinuousConstSMul nonrec theorem tendsto_const_smul_iff {f : β → α} {l : Filter β} {a : α} {c : M} (hc : IsUnit c) : Tendsto (fun x => c • f x) l (𝓝 <| c • a) ↔ Tendsto f l (𝓝 a) := - let ⟨u, hu⟩ := hc - hu ▸ tendsto_const_smul_iff u + tendsto_const_smul_iff hc.unit variable [TopologicalSpace β] {f : β → α} {b : β} {c : M} {s : Set β} nonrec theorem continuousWithinAt_const_smul_iff (hc : IsUnit c) : ContinuousWithinAt (fun x => c • f x) s b ↔ ContinuousWithinAt f s b := - let ⟨u, hu⟩ := hc - hu ▸ continuousWithinAt_const_smul_iff u + continuousWithinAt_const_smul_iff hc.unit nonrec theorem continuousOn_const_smul_iff (hc : IsUnit c) : ContinuousOn (fun x => c • f x) s ↔ ContinuousOn f s := - let ⟨u, hu⟩ := hc - hu ▸ continuousOn_const_smul_iff u + continuousOn_const_smul_iff hc.unit nonrec theorem continuousAt_const_smul_iff (hc : IsUnit c) : ContinuousAt (fun x => c • f x) b ↔ ContinuousAt f b := - let ⟨u, hu⟩ := hc - hu ▸ continuousAt_const_smul_iff u + continuousAt_const_smul_iff hc.unit nonrec theorem continuous_const_smul_iff (hc : IsUnit c) : (Continuous fun x => c • f x) ↔ Continuous f := - let ⟨u, hu⟩ := hc - hu ▸ continuous_const_smul_iff u + continuous_const_smul_iff hc.unit nonrec theorem isOpenMap_smul (hc : IsUnit c) : IsOpenMap fun x : α => c • x := - let ⟨u, hu⟩ := hc - hu ▸ isOpenMap_smul u + isOpenMap_smul hc.unit nonrec theorem isClosedMap_smul (hc : IsUnit c) : IsClosedMap fun x : α => c • x := - let ⟨u, hu⟩ := hc - hu ▸ isClosedMap_smul u + isClosedMap_smul hc.unit + +nonrec theorem smul_mem_nhds_smul_iff (hc : IsUnit c) {s : Set α} {a : α} : + c • s ∈ 𝓝 (c • a) ↔ s ∈ 𝓝 a := + smul_mem_nhds_smul_iff hc.unit end IsUnit @@ -474,19 +494,20 @@ section MulAction variable {G₀ : Type*} [GroupWithZero G₀] [MulAction G₀ α] [TopologicalSpace α] [ContinuousConstSMul G₀ α] --- Porting note: generalize to a group action + `IsUnit` -/-- Scalar multiplication preserves neighborhoods. -/ -theorem set_smul_mem_nhds_smul {c : G₀} {s : Set α} {x : α} (hs : s ∈ 𝓝 x) (hc : c ≠ 0) : - c • s ∈ 𝓝 (c • x : α) := by - rw [mem_nhds_iff] at hs ⊢ - obtain ⟨U, hs', hU, hU'⟩ := hs - exact ⟨c • U, Set.smul_set_mono hs', hU.smul₀ hc, Set.smul_mem_smul_set hU'⟩ +/-- Scalar multiplication by a nonzero scalar preserves neighborhoods. -/ +theorem smul_mem_nhds_smul_iff₀ {c : G₀} {s : Set α} {x : α} (hc : c ≠ 0) : + c • s ∈ 𝓝 (c • x : α) ↔ s ∈ 𝓝 x := + smul_mem_nhds_smul_iff (Units.mk0 c hc) -theorem set_smul_mem_nhds_smul_iff {c : G₀} {s : Set α} {x : α} (hc : c ≠ 0) : - c • s ∈ 𝓝 (c • x : α) ↔ s ∈ 𝓝 x := by - refine ⟨fun h => ?_, fun h => set_smul_mem_nhds_smul h hc⟩ - rw [← inv_smul_smul₀ hc x, ← inv_smul_smul₀ hc s] - exact set_smul_mem_nhds_smul h (inv_ne_zero hc) +@[deprecated (since := "2024-08-06")] +alias set_smul_mem_nhds_smul_iff := smul_mem_nhds_smul_iff₀ + +alias ⟨_, smul_mem_nhds_smul₀⟩ := smul_mem_nhds_smul_iff₀ + +@[deprecated smul_mem_nhds_smul₀ (since := "2024-08-06")] +theorem set_smul_mem_nhds_smul {c : G₀} {s : Set α} {x : α} (hs : s ∈ 𝓝 x) (hc : c ≠ 0) : + c • s ∈ 𝓝 (c • x : α) := + smul_mem_nhds_smul₀ hc hs end MulAction @@ -497,7 +518,7 @@ variable {G₀ : Type*} [GroupWithZero G₀] [AddMonoid α] [DistribMulAction G theorem set_smul_mem_nhds_zero_iff {s : Set α} {c : G₀} (hc : c ≠ 0) : c • s ∈ 𝓝 (0 : α) ↔ s ∈ 𝓝 (0 : α) := by - refine Iff.trans ?_ (set_smul_mem_nhds_smul_iff hc) + refine Iff.trans ?_ (smul_mem_nhds_smul_iff₀ hc) rw [smul_zero] end DistribMulAction diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index fda640ec9b168..99692392d933b 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -1047,20 +1047,6 @@ section ContinuousConstSMul variable [TopologicalSpace β] [Group α] [MulAction α β] [ContinuousConstSMul α β] {s : Set α} {t : Set β} -@[to_additive] -theorem IsOpen.smul_left (ht : IsOpen t) : IsOpen (s • t) := by - rw [← iUnion_smul_set] - exact isOpen_biUnion fun a _ => ht.smul _ - -@[to_additive] -theorem subset_interior_smul_right : s • interior t ⊆ interior (s • t) := - interior_maximal (Set.smul_subset_smul_left interior_subset) isOpen_interior.smul_left - -@[to_additive] -theorem smul_mem_nhds (a : α) {x : β} (ht : t ∈ 𝓝 x) : a • t ∈ 𝓝 (a • x) := by - rcases mem_nhds_iff.1 ht with ⟨u, ut, u_open, hu⟩ - exact mem_nhds_iff.2 ⟨a • u, smul_set_mono ut, u_open.smul a, smul_mem_smul_set hu⟩ - variable [TopologicalSpace α] @[to_additive] @@ -1138,8 +1124,7 @@ theorem subset_interior_mul : interior s * interior t ⊆ interior (s * t) := @[to_additive] theorem singleton_mul_mem_nhds (a : α) {b : α} (h : s ∈ 𝓝 b) : {a} * s ∈ 𝓝 (a * b) := by - have := smul_mem_nhds a h - rwa [← singleton_smul] at this + rwa [← smul_eq_mul, ← smul_eq_mul, singleton_smul, smul_mem_nhds_smul_iff] @[to_additive] theorem singleton_mul_mem_nhds_of_nhds_one (a : α) (h : s ∈ 𝓝 (1 : α)) : {a} * s ∈ 𝓝 a := by @@ -1153,8 +1138,8 @@ variable [TopologicalSpace α] [Group α] [ContinuousConstSMul αᵐᵒᵖ α] { @[to_additive] theorem IsOpen.mul_right (hs : IsOpen s) : IsOpen (s * t) := by - rw [← iUnion_op_smul_set] - exact isOpen_biUnion fun a _ => hs.smul _ + rw [← image_op_smul] + exact hs.smul_left @[to_additive] theorem subset_interior_mul_left : interior s * t ⊆ interior (s * t) := @@ -1166,8 +1151,8 @@ theorem subset_interior_mul' : interior s * interior t ⊆ interior (s * t) := @[to_additive] theorem mul_singleton_mem_nhds (a : α) {b : α} (h : s ∈ 𝓝 b) : s * {a} ∈ 𝓝 (b * a) := by - simp only [← iUnion_op_smul_set, mem_singleton_iff, iUnion_iUnion_eq_left] - exact smul_mem_nhds _ h + rw [mul_singleton] + exact smul_mem_nhds_smul (op a) h @[to_additive] theorem mul_singleton_mem_nhds_of_nhds_one (a : α) (h : s ∈ 𝓝 (1 : α)) : s * {a} ∈ 𝓝 a := by diff --git a/Mathlib/Topology/Algebra/Group/OpenMapping.lean b/Mathlib/Topology/Algebra/Group/OpenMapping.lean index de49dc949e475..7dbba50fb9e92 100644 --- a/Mathlib/Topology/Algebra/Group/OpenMapping.lean +++ b/Mathlib/Topology/Algebra/Group/OpenMapping.lean @@ -45,10 +45,8 @@ theorem smul_singleton_mem_nhds_of_sigmaCompact if `V` is small enough. -/ obtain ⟨V, V_mem, V_closed, V_symm, VU⟩ : ∃ V ∈ 𝓝 (1 : G), IsClosed V ∧ V⁻¹ = V ∧ V * V ⊆ U := exists_closed_nhds_one_inv_eq_mul_subset hU - obtain ⟨s, s_count, hs⟩ : ∃ (s : Set G), s.Countable ∧ ⋃ g ∈ s, g • V = univ := by - apply countable_cover_nhds_of_sigma_compact (fun g ↦ ?_) - convert smul_mem_nhds g V_mem - simp only [smul_eq_mul, mul_one] + obtain ⟨s, s_count, hs⟩ : ∃ (s : Set G), s.Countable ∧ ⋃ g ∈ s, g • V = univ := + countable_cover_nhds_of_sigma_compact fun _ ↦ by simpa let K : ℕ → Set G := compactCovering G let F : ℕ × s → Set X := fun p ↦ (K p.1 ∩ (p.2 : G) • V) • ({x} : Set X) obtain ⟨⟨n, ⟨g, hg⟩⟩, hi⟩ : ∃ i, (interior (F i)).Nonempty := by diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index 568d2b9c373a5..3f1b53b0f9042 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -559,6 +559,10 @@ theorem of_comp (f : X → Y) (hg : OpenEmbedding g) theorem of_isEmpty [IsEmpty X] (f : X → Y) : OpenEmbedding f := openEmbedding_of_embedding_open (.of_subsingleton f) (IsOpenMap.of_isEmpty f) +theorem image_mem_nhds {f : X → Y} (hf : OpenEmbedding f) {s : Set X} {x : X} : + f '' s ∈ 𝓝 (f x) ↔ s ∈ 𝓝 x := by + rw [← hf.map_nhds_eq, mem_map, preimage_image_eq _ hf.inj] + end OpenEmbedding end OpenEmbedding From 00d93f99b5326c8e15772d502ccde43f8f3d2d0d Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 9 Aug 2024 02:35:33 +0000 Subject: [PATCH 140/359] =?UTF-8?q?feat:=20add=20`OrderedSMul=20=E2=84=9D?= =?UTF-8?q?=20A`=20instance=20for=20`StarOrderedRing=20A`=20(#15391)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This instance is already available in the scope `ComplexOrder`, but for `ℝ` that is unnecesary, so we make it a global instance. This instance could be moved earlier (so that we don't make use of `RCLike`), but it's probably not that important to do so. --- Mathlib/Analysis/RCLike/Basic.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 55fb23a8c9e54..5a44f14f45a95 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -813,6 +813,11 @@ lemma _root_.StarModule.instOrderedSMul {A : Type*} [NonUnitalRing A] [StarRing StarModule.smul_lt_smul_of_pos hxy (RCLike.inv_pos_of_pos hc) simpa [smul_smul, inv_mul_cancel hc.ne'] using this +instance {A : Type*} [NonUnitalRing A] [StarRing A] [PartialOrder A] [StarOrderedRing A] + [Module ℝ A] [StarModule ℝ A] [IsScalarTower ℝ A A] [SMulCommClass ℝ A A] : + OrderedSMul ℝ A := + StarModule.instOrderedSMul + scoped[ComplexOrder] attribute [instance] StarModule.instOrderedSMul end Order From bc37ac3c35807d44e9caba747eeebb9171bbf802 Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 9 Aug 2024 03:04:29 +0000 Subject: [PATCH 141/359] chore(Algebra/Ring/SumsOfSquares, Algebra/Ring/Semireal): name convension (#15621) --- Mathlib/Algebra/Ring/Semireal/Defs.lean | 19 ++++--- Mathlib/Algebra/Ring/SumsOfSquares.lean | 66 ++++++++++++++----------- 2 files changed, 49 insertions(+), 36 deletions(-) diff --git a/Mathlib/Algebra/Ring/Semireal/Defs.lean b/Mathlib/Algebra/Ring/Semireal/Defs.lean index 1943e7609eb01..7e26be34cfffb 100644 --- a/Mathlib/Algebra/Ring/Semireal/Defs.lean +++ b/Mathlib/Algebra/Ring/Semireal/Defs.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Florent Schaffhauser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Florent Schaffhauser -/ - import Mathlib.Algebra.Ring.SumsOfSquares import Mathlib.Algebra.Order.Field.Defs @@ -19,7 +18,7 @@ semireal. ## Main declaration -- `isSemireal`: the predicate asserting that a commutative ring `R` is semireal. +- `IsSemireal`: the predicate asserting that a commutative ring `R` is semireal. ## References @@ -31,14 +30,18 @@ variable (R : Type*) /-- A semireal ring is a non-trivial commutative ring (with unit) in which `-1` is *not* a sum of -squares. Note that `-1` does not make sense in a semiring. Below we define the class `isSemiReal R` +squares. Note that `-1` does not make sense in a semiring. Below we define the class `IsSemireal R` for all additive monoid `R` equipped with a multiplication, a multiplicative unit and a negation. -/ @[mk_iff] -class isSemireal [AddMonoid R] [Mul R] [One R] [Neg R] : Prop where - non_trivial : (0 : R) ≠ 1 - neg_one_not_SumSq : ¬isSumSq (-1 : R) +class IsSemireal [AddMonoid R] [Mul R] [One R] [Neg R] : Prop where + non_trivial : (0 : R) ≠ 1 + not_isSumSq_neg_one : ¬IsSumSq (-1 : R) + +@[deprecated (since := "2024-08-09")] alias isSemireal := IsSemireal +@[deprecated (since := "2024-08-09")] alias isSemireal.neg_one_not_SumSq := + IsSemireal.not_isSumSq_neg_one -instance [LinearOrderedField R] : isSemireal R where +instance [LinearOrderedField R] : IsSemireal R where non_trivial := zero_ne_one - neg_one_not_SumSq := fun h ↦ (not_le (α := R)).2 neg_one_lt_zero h.nonneg + not_isSumSq_neg_one := fun h ↦ (not_le (α := R)).2 neg_one_lt_zero h.nonneg diff --git a/Mathlib/Algebra/Ring/SumsOfSquares.lean b/Mathlib/Algebra/Ring/SumsOfSquares.lean index 0d555aa10890b..1eb9bb77744cf 100644 --- a/Mathlib/Algebra/Ring/SumsOfSquares.lean +++ b/Mathlib/Algebra/Ring/SumsOfSquares.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Florent Schaffhauser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Florent Schaffhauser -/ - import Mathlib.Algebra.Ring.Defs import Mathlib.Algebra.Group.Submonoid.Basic import Mathlib.Algebra.Group.Even @@ -13,27 +12,27 @@ import Mathlib.Algebra.Order.Ring.Defs # Sums of squares We introduce sums of squares in a type `R` endowed with an `[Add R]`, `[Zero R]` and `[Mul R]` -instances. Sums of squares in `R` are defined by an inductive predicate `isSumSq : R → Prop`: +instances. Sums of squares in `R` are defined by an inductive predicate `IsSumSq : R → Prop`: `0 : R` is a sum of squares and if `S` is a sum of squares, then for all `a : R`, `a * a + S` is a sum of squares in `R`. ## Main declaration -- The predicate `isSumSq : R → Prop`, defining the property of being a sum of squares in `R`. +- The predicate `IsSumSq : R → Prop`, defining the property of being a sum of squares in `R`. ## Auxiliary declarations -- The type `SumSqIn R` : in an additive monoid with multiplication `R`, we introduce the type -`SumSqIn R` as the submonoid of `R` whose carrier is the subset `{S : R | isSumSq S}`. +- The type `sumSqIn R` : in an additive monoid with multiplication `R`, we introduce the type +`sumSqIn R` as the submonoid of `R` whose carrier is the subset `{S : R | IsSumSq S}`. ## Theorems -- `isSumSq.add`: if `S1` and `S2` are sums of squares in `R`, then so is `S1 + S2`. -- `SumSq.nonneg`: in a linearly ordered semiring `R` with an `[ExistsAddOfLE R]` instance, sums of +- `IsSumSq.add`: if `S1` and `S2` are sums of squares in `R`, then so is `S1 + S2`. +- `IsSumSq.nonneg`: in a linearly ordered semiring `R` with an `[ExistsAddOfLE R]` instance, sums of squares are non-negative. -- `SquaresInSumSq`: a square is a sum of squares. -- `SquaresAddClosure`: the submonoid of `R` generated by the squares in `R` is the set of sums of -squares in `R`. +- `mem_sumSqIn_of_isSquare`: a square is a sum of squares. +- `AddSubmonoid.closure_isSquare`: the submonoid of `R` generated by the squares in `R` is the set +of sums of squares in `R`. -/ @@ -45,57 +44,68 @@ squares is defined by an inductive predicate: `0 : R` is a sum of squares and if squares, then for all `a : R`, `a * a + S` is a sum of squares in `R`. -/ @[mk_iff] -inductive isSumSq [Add R] [Zero R] : R → Prop - | zero : isSumSq 0 - | sq_add (a S : R) (pS : isSumSq S) : isSumSq (a * a + S) +inductive IsSumSq [Add R] [Zero R] : R → Prop + | zero : IsSumSq 0 + | sq_add (a S : R) (pS : IsSumSq S) : IsSumSq (a * a + S) + +@[deprecated (since := "2024-08-09")] alias isSumSq := IsSumSq /-- If `S1` and `S2` are sums of squares in a semiring `R`, then `S1 + S2` is a sum of squares in `R`. -/ -theorem isSumSq.add [AddMonoid R] {S1 S2 : R} (p1 : isSumSq S1) - (p2 : isSumSq S2) : isSumSq (S1 + S2) := by +theorem IsSumSq.add [AddMonoid R] {S1 S2 : R} (p1 : IsSumSq S1) + (p2 : IsSumSq S2) : IsSumSq (S1 + S2) := by induction p1 with | zero => rw [zero_add]; exact p2 - | sq_add a S pS ih => rw [add_assoc]; exact isSumSq.sq_add a (S + S2) ih + | sq_add a S pS ih => rw [add_assoc]; exact IsSumSq.sq_add a (S + S2) ih + +@[deprecated (since := "2024-08-09")] alias isSumSq.add := IsSumSq.add variable (R) in /-- -In an additive monoid with multiplication `R`, the type `SumSqIn R` is the submonoid of sums of +In an additive monoid with multiplication `R`, the type `sumSqIn R` is the submonoid of sums of squares in `R`. -/ -def SumSqIn [AddMonoid R] : AddSubmonoid R where - carrier := {S : R | isSumSq S} - zero_mem' := isSumSq.zero - add_mem' := isSumSq.add +def sumSqIn [AddMonoid R] : AddSubmonoid R where + carrier := {S : R | IsSumSq S} + zero_mem' := IsSumSq.zero + add_mem' := IsSumSq.add + +@[deprecated (since := "2024-08-09")] alias SumSqIn := sumSqIn /-- In an additive monoid with multiplication, every square is a sum of squares. By definition, a square in `R` is a term `x : R` such that `x = y * y` for some `y : R` and in Mathlib this is known as `IsSquare R` (see Mathlib.Algebra.Group.Even). -/ -theorem SquaresInSumSq [AddMonoid R] {x : R} (px : IsSquare x) : x ∈ SumSqIn R := by +theorem mem_sumSqIn_of_isSquare [AddMonoid R] {x : R} (px : IsSquare x) : x ∈ sumSqIn R := by rcases px with ⟨y, py⟩ rw [py, ← AddMonoid.add_zero (y * y)] - exact isSumSq.sq_add _ _ isSumSq.zero + exact IsSumSq.sq_add _ _ IsSumSq.zero + +@[deprecated (since := "2024-08-09")] alias SquaresInSumSq := mem_sumSqIn_of_isSquare -open AddSubmonoid in /-- In an additive monoid with multiplication `R`, the submonoid generated by the squares is the set of sums of squares. -/ -theorem SquaresAddClosure [AddMonoid R] : closure IsSquare = SumSqIn R := by - refine le_antisymm (closure_le.2 (fun x hx ↦ SquaresInSumSq hx)) (fun x hx ↦ ?_) +theorem AddSubmonoid.closure_isSquare [AddMonoid R] : closure {x : R | IsSquare x} = sumSqIn R := by + refine le_antisymm (closure_le.2 (fun x hx ↦ mem_sumSqIn_of_isSquare hx)) (fun x hx ↦ ?_) induction hx with | zero => exact zero_mem _ | sq_add a S _ ih => exact AddSubmonoid.add_mem _ (subset_closure ⟨a, rfl⟩) ih +@[deprecated (since := "2024-08-09")] alias SquaresAddClosure := AddSubmonoid.closure_isSquare + /-- Let `R` be a linearly ordered semiring in which the property `a ≤ b → ∃ c, a + c = b` holds (e.g. `R = ℕ`). If `S : R` is a sum of squares in `R`, then `0 ≤ S`. This is used in `Mathlib.Algebra.Ring.Semireal.Defs` to show that linearly ordered fields are semireal. -/ -theorem isSumSq.nonneg {R : Type*} [LinearOrderedSemiring R] [ExistsAddOfLE R] {S : R} - (pS : isSumSq S) : 0 ≤ S := by +theorem IsSumSq.nonneg {R : Type*} [LinearOrderedSemiring R] [ExistsAddOfLE R] {S : R} + (pS : IsSumSq S) : 0 ≤ S := by induction pS with | zero => simp only [le_refl] | sq_add x S _ ih => apply add_nonneg ?_ ih; simp only [← pow_two x, sq_nonneg] + +@[deprecated (since := "2024-08-09")] alias isSumSq.nonneg := IsSumSq.nonneg From c37e54b849a33204619ff98c136c5a9fd0fe0cf2 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Fri, 9 Aug 2024 04:16:20 +0000 Subject: [PATCH 142/359] chore: update Mathlib dependencies 2024-08-09 (#15633) 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 607daf7296fe1..23a359909f284 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -15,7 +15,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "01ad33937acd996ee99eb74eefb39845e4e4b9f5", + "rev": "71f54425e6fe0fa75f3aef33a2813a7898392222", "name": "Qq", "manifestFile": "lake-manifest.json", "inputRev": "master", From cd6bbdf5b06719a684e6a2968957eeabb8599a68 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 9 Aug 2024 06:50:18 +0000 Subject: [PATCH 143/359] chore(AtTopBot): drop a `Nonempty` assumption (#15634) ... in `atTop_basis'` and `atBot_basis'`. --- Mathlib/Order/Filter/AtTopBot.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index 0f53605218201..692f4d4c39e72 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -240,13 +240,14 @@ lemma atTop_basis_Ioi' [NoMaxOrder α] (a : α) : atTop.HasBasis (a < ·) Ioi := obtain ⟨d, hcd⟩ := exists_gt c exact ⟨d, hac.trans_lt hcd, Ioi_subset_Ioi (hbc.trans hcd.le)⟩ -variable [Nonempty α] - theorem atTop_basis' (a : α) : (@atTop α _).HasBasis (fun x => a ≤ x) Ici := by + have : Nonempty α := ⟨a⟩ refine atTop_basis.to_hasBasis (fun b _ ↦ ?_) fun b _ ↦ ⟨b, trivial, Subset.rfl⟩ obtain ⟨c, hac, hbc⟩ := exists_ge_ge a b exact ⟨c, hac, Ici_subset_Ici.2 hbc⟩ +variable [Nonempty α] + @[instance] lemma atTop_neBot : NeBot (atTop : Filter α) := atTop_basis.neBot_iff.2 fun _ => nonempty_Ici @@ -288,12 +289,12 @@ lemma atBot_basis_Iio [Nonempty α] [NoMinOrder α] : (@atBot α _).HasBasis (fu lemma atBot_basis_Iio' [NoMinOrder α] (a : α) : atBot.HasBasis (· < a) Iio := atTop_basis_Ioi' (α := αᵒᵈ) a +lemma atBot_basis' (a : α) : (@atBot α _).HasBasis (fun x => x ≤ a) Iic := atTop_basis' (α := αᵒᵈ) _ + variable [Nonempty α] lemma atBot_basis : (@atBot α _).HasBasis (fun _ => True) Iic := atTop_basis (α := αᵒᵈ) -lemma atBot_basis' (a : α) : (@atBot α _).HasBasis (fun x => x ≤ a) Iic := atTop_basis' (α := αᵒᵈ) _ - @[instance] lemma atBot_neBot : NeBot (atBot : Filter α) := atTop_neBot (α := αᵒᵈ) @[simp] lemma mem_atBot_sets {s : Set α} : s ∈ (atBot : Filter α) ↔ ∃ a : α, ∀ b ≤ a, b ∈ s := From 6148791722c6e8390e2699ec915b8e3add6b6f03 Mon Sep 17 00:00:00 2001 From: grunweg Date: Fri, 9 Aug 2024 07:41:27 +0000 Subject: [PATCH 144/359] fix(lint-style): correct error message (#15438) Co-authored-by: damiano --- Mathlib/Tactic/Linter/TextBased.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Tactic/Linter/TextBased.lean b/Mathlib/Tactic/Linter/TextBased.lean index caacb32947720..38204f9292cc3 100644 --- a/Mathlib/Tactic/Linter/TextBased.lean +++ b/Mathlib/Tactic/Linter/TextBased.lean @@ -443,7 +443,7 @@ def lintModules (moduleNames : Array String) (mode : OutputSetting) : IO UInt32 | OutputSetting.print style => formatErrors allUnexpectedErrors style if numberErrorFiles > 0 && mode matches OutputSetting.print _ then - IO.println s!"error: found {numberErrorFiles} new style errors\n\ + IO.println s!"error: found {allUnexpectedErrors.size} new style errors\n\ run `lake exe lint-style --update` to ignore all of them" | OutputSetting.update => formatErrors allUnexpectedErrors ErrorFormat.humanReadable From a528f200ec83eae34a93fabb6270e0db1a8be8e4 Mon Sep 17 00:00:00 2001 From: Yudai Yamazaki Date: Fri, 9 Aug 2024 08:00:05 +0000 Subject: [PATCH 145/359] feat(GroupTheory/GroupExtension/Defs): define group extensions (#15049) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Mainly defines: - `structure GroupExtension N E G`: group extensions as short exact sequences `1 → N → E → G → 1` - `structure GroupExtension.Equiv S S'`: equivalences of two group extensions `S : GroupExtension N E G` and `S' : GroupExtension N E' G` - `structure GroupExtension.Splitting S`: splittings of a group extension `S` - `def SemidirectProduct.toGroupExtension (φ : G →* MulAut N)`: the group extension associated with the semidirect product, `1 → N → N ⋊[φ] G → G → 1` --- Mathlib.lean | 1 + Mathlib/GroupTheory/GroupExtension/Defs.lean | 202 +++++++++++++++++++ 2 files changed, 203 insertions(+) create mode 100644 Mathlib/GroupTheory/GroupExtension/Defs.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4d95f16d8fc02..143cc5aed691c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2727,6 +2727,7 @@ import Mathlib.GroupTheory.GroupAction.Ring import Mathlib.GroupTheory.GroupAction.SubMulAction import Mathlib.GroupTheory.GroupAction.SubMulAction.Pointwise import Mathlib.GroupTheory.GroupAction.Support +import Mathlib.GroupTheory.GroupExtension.Defs import Mathlib.GroupTheory.HNNExtension import Mathlib.GroupTheory.Index import Mathlib.GroupTheory.MonoidLocalization.Basic diff --git a/Mathlib/GroupTheory/GroupExtension/Defs.lean b/Mathlib/GroupTheory/GroupExtension/Defs.lean new file mode 100644 index 0000000000000..68610f45d8160 --- /dev/null +++ b/Mathlib/GroupTheory/GroupExtension/Defs.lean @@ -0,0 +1,202 @@ +/- +Copyright (c) 2024 Yudai Yamazaki. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yudai Yamazaki +-/ + +import Mathlib.GroupTheory.GroupAction.ConjAct +import Mathlib.GroupTheory.SemidirectProduct + +/-! +# Group Extensions + +This file defines extensions of multiplicative and additive groups and their associated structures +such as splittings and equivalences. + +## Main definitions + +- `(Add?)GroupExtension N E G`: structure for extensions of `G` by `N` as short exact sequences + `1 → N → E → G → 1` (`0 → N → E → G → 0` for additive groups) +- `(Add?)GroupExtension.Equiv S S'`: structure for equivalences of two group extensions `S` and `S'` + as specific homomorphisms `E → E'` such that each diagram below is commutative + +```text +For multiplicative groups: + ↗︎ E ↘ +1 → N ↓ G → 1 + ↘︎ E' ↗︎️ + +For additive groups: + ↗︎ E ↘ +0 → N ↓ G → 0 + ↘︎ E' ↗︎️ +``` + +- `(Add?)GroupExtension.Splitting S`: structure for splittings of a group extension `S` of `G` by + `N` as section homomorphisms `G → E` +- `SemidirectProduct.toGroupExtension φ`: the multiplicative group extension associated to the + semidirect product coming from `φ : G →* MulAut N`, `1 → N → N ⋊[φ] G → G → 1` + +## TODO + +If `N` is abelian, + +- there is a bijection between `N`-conjugacy classes of + `(SemidirectProduct.toGroupExtension φ).Splitting` and `groupCohomology.H1` + (which will be available in `GroupTheory/GroupExtension/Abelian.lean` to be added in a later PR). +- there is a bijection between equivalence classes of group extensions and `groupCohomology.H2` + (which is also stated as a TODO in `RepresentationTheory/GroupCohomology/LowDegree.lean`). +-/ + +variable (N E G : Type*) + +/-- `AddGroupExtension N E G` is a short exact sequence of additive groups `0 → N → E → G → 0`. -/ +structure AddGroupExtension [AddGroup N] [AddGroup E] [AddGroup G] where + /-- The inclusion homomorphism `N →+ E` -/ + inl : N →+ E + /-- The projection homomorphism `E →+ G` -/ + rightHom : E →+ G + /-- The inclusion map is injective. -/ + inl_injective : Function.Injective inl + /-- The range of the inclusion map is equal to the kernel of the projection map. -/ + range_inl_eq_ker_rightHom : inl.range = rightHom.ker + /-- The projection map is surjective. -/ + rightHom_surjective : Function.Surjective rightHom + +/-- `GroupExtension N E G` is a short exact sequence of groups `1 → N → E → G → 1`. -/ +@[to_additive] +structure GroupExtension [Group N] [Group E] [Group G] where + /-- The inclusion homomorphism `N →* E` -/ + inl : N →* E + /-- The projection homomorphism `E →* G` -/ + rightHom : E →* G + /-- The inclusion map is injective. -/ + inl_injective : Function.Injective inl + /-- The range of the inclusion map is equal to the kernel of the projection map. -/ + range_inl_eq_ker_rightHom : inl.range = rightHom.ker + /-- The projection map is surjective. -/ + rightHom_surjective : Function.Surjective rightHom + +variable {N E G} + +namespace AddGroupExtension + +variable [AddGroup N] [AddGroup E] [AddGroup G] (S : AddGroupExtension N E G) + +/-- `AddGroupExtension`s are equivalent iff there is a homomorphism making a commuting diagram. -/ +structure Equiv {E' : Type*} [AddGroup E'] (S' : AddGroupExtension N E' G) where + /-- The homomorphism -/ + toAddMonoidHom : E →+ E' + /-- The left-hand side of the diagram commutes. -/ + inl_comm : toAddMonoidHom.comp S.inl = S'.inl + /-- The right-hand side of the diagram commutes. -/ + rightHom_comm : S'.rightHom.comp toAddMonoidHom = S.rightHom + +/-- `Splitting` of an additive group extension is a section homomorphism. -/ +structure Splitting where + /-- A section homomorphism -/ + sectionHom : G →+ E + /-- The section is a left inverse of the projection map. -/ + rightHom_comp_sectionHom : S.rightHom.comp sectionHom = AddMonoidHom.id G + +end AddGroupExtension + +namespace GroupExtension + +variable [Group N] [Group E] [Group G] (S : GroupExtension N E G) + +/-- The range of the inclusion map is a normal subgroup. -/ +@[to_additive "The range of the inclusion map is a normal additive subgroup." ] +instance normal_inl_range : S.inl.range.Normal := + S.range_inl_eq_ker_rightHom ▸ S.rightHom.normal_ker + +@[to_additive (attr := simp)] +theorem rightHom_inl (n : N) : S.rightHom (S.inl n) = 1 := by + rw [← MonoidHom.mem_ker, ← S.range_inl_eq_ker_rightHom, MonoidHom.mem_range] + exact exists_apply_eq_apply S.inl n + +@[to_additive (attr := simp)] +theorem rightHom_comp_inl : S.rightHom.comp S.inl = 1 := by + ext n + rw [MonoidHom.one_apply, MonoidHom.comp_apply] + exact S.rightHom_inl n + +/-- `E` acts on `N` by conjugation. -/ +noncomputable def conjAct : E →* MulAut N where + toFun e := (MonoidHom.ofInjective S.inl_injective).trans <| + (MulAut.conjNormal e).trans (MonoidHom.ofInjective S.inl_injective).symm + map_one' := by + ext _ + simp only [map_one, MulEquiv.trans_apply, MulAut.one_apply, MulEquiv.symm_apply_apply] + map_mul' _ _ := by + ext _ + simp only [map_mul, MulEquiv.trans_apply, MulAut.mul_apply, MulEquiv.apply_symm_apply] + +/-- The inclusion and a conjugation commute. -/ +theorem inl_conjAct_comm {e : E} {n : N} : S.inl (S.conjAct e n) = e * S.inl n * e⁻¹ := by + simp only [conjAct, MonoidHom.coe_mk, OneHom.coe_mk, MulEquiv.trans_apply, + MonoidHom.apply_ofInjective_symm] + rfl + +/-- `GroupExtension`s are equivalent iff there is a homomorphism making a commuting diagram. -/ +@[to_additive] +structure Equiv {E' : Type*} [Group E'] (S' : GroupExtension N E' G) where + /-- The homomorphism -/ + toMonoidHom : E →* E' + /-- The left-hand side of the diagram commutes. -/ + inl_comm : toMonoidHom.comp S.inl = S'.inl + /-- The right-hand side of the diagram commutes. -/ + rightHom_comm : S'.rightHom.comp toMonoidHom = S.rightHom + +/-- `Splitting` of a group extension is a section homomorphism. -/ +@[to_additive] +structure Splitting where + /-- A section homomorphism -/ + sectionHom : G →* E + /-- The section is a left inverse of the projection map. -/ + rightHom_comp_sectionHom : S.rightHom.comp sectionHom = MonoidHom.id G + +@[to_additive] +instance : FunLike S.Splitting G E where + coe s := s.sectionHom + coe_injective' := by + intro ⟨_, _⟩ ⟨_, _⟩ h + congr + exact DFunLike.coe_injective h + +@[to_additive] +instance : MonoidHomClass S.Splitting G E where + map_mul s := s.sectionHom.map_mul' + map_one s := s.sectionHom.map_one' + +/-- A splitting of an extension `S` is `N`-conjugate to another iff there exists `n : N` such that +the section homomorphism is a conjugate of the other section homomorphism by `S.inl n`. -/ +@[to_additive + "A splitting of an extension `S` is `N`-conjugate to another iff there exists `n : N` such + that the section homomorphism is a conjugate of the other section homomorphism by `S.inl n`."] +def IsConj (S : GroupExtension N E G) (s s' : S.Splitting) : Prop := + ∃ n : N, s.sectionHom = fun g ↦ S.inl n * s'.sectionHom g * (S.inl n)⁻¹ + +end GroupExtension + +namespace SemidirectProduct + +variable [Group G] [Group N] (φ : G →* MulAut N) + +/-- The group extension associated to the semidirect product -/ +def toGroupExtension : GroupExtension N (N ⋊[φ] G) G where + inl_injective := inl_injective + range_inl_eq_ker_rightHom := range_inl_eq_ker_rightHom + rightHom_surjective := rightHom_surjective + +theorem toGroupExtension_inl : (toGroupExtension φ).inl = SemidirectProduct.inl := rfl + +theorem toGroupExtension_rightHom : (toGroupExtension φ).rightHom = SemidirectProduct.rightHom := + rfl + +/-- A canonical splitting of the group extension associated to the semidirect product -/ +def inr_splitting : (toGroupExtension φ).Splitting where + sectionHom := inr + rightHom_comp_sectionHom := rightHom_comp_inr + +end SemidirectProduct From dc29147da04f4e01dc3fe9b39a80fffe68685e5e Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Fri, 9 Aug 2024 11:40:09 +0000 Subject: [PATCH 146/359] chore (Algebra.Order.Group.Int): split file into unbundled and bundled ordered algebra (#15069) All but ~5 lines of this file did not require bundled ordered algebra classes. We split them off into `Algebra.Order.Group.Unbundled.Int` leaving the `LinearOrderedAddCommGroup` instance. --- Mathlib.lean | 1 + Mathlib/Algebra/EuclideanDomain/Int.lean | 2 +- Mathlib/Algebra/GCDMonoid/Nat.lean | 1 + Mathlib/Algebra/Order/Group/Int.lean | 103 +------------- .../Algebra/Order/Group/Unbundled/Int.lean | 133 ++++++++++++++++++ Mathlib/Algebra/Order/Ring/Abs.lean | 1 + Mathlib/Algebra/Order/Ring/Cast.lean | 1 + Mathlib/Algebra/Order/Ring/Int.lean | 4 +- Mathlib/Data/Int/Interval.lean | 3 +- Mathlib/Data/Int/Star.lean | 1 + 10 files changed, 143 insertions(+), 107 deletions(-) create mode 100644 Mathlib/Algebra/Order/Group/Unbundled/Int.lean diff --git a/Mathlib.lean b/Mathlib.lean index 143cc5aed691c..512d587840374 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -572,6 +572,7 @@ import Mathlib.Algebra.Order.Group.Synonym import Mathlib.Algebra.Order.Group.TypeTags import Mathlib.Algebra.Order.Group.Unbundled.Abs import Mathlib.Algebra.Order.Group.Unbundled.Basic +import Mathlib.Algebra.Order.Group.Unbundled.Int import Mathlib.Algebra.Order.Group.Units import Mathlib.Algebra.Order.GroupWithZero.Canonical import Mathlib.Algebra.Order.GroupWithZero.Synonym diff --git a/Mathlib/Algebra/EuclideanDomain/Int.lean b/Mathlib/Algebra/EuclideanDomain/Int.lean index 43483bf529b23..f67b36b407e2d 100644 --- a/Mathlib/Algebra/EuclideanDomain/Int.lean +++ b/Mathlib/Algebra/EuclideanDomain/Int.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Louis Carlin, Mario Carneiro -/ import Mathlib.Algebra.EuclideanDomain.Defs -import Mathlib.Algebra.Order.Group.Int +import Mathlib.Algebra.Order.Group.Unbundled.Int import Mathlib.Algebra.Ring.Int /-! diff --git a/Mathlib/Algebra/GCDMonoid/Nat.lean b/Mathlib/Algebra/GCDMonoid/Nat.lean index 586b58e945175..c5dadae8f727c 100644 --- a/Mathlib/Algebra/GCDMonoid/Nat.lean +++ b/Mathlib/Algebra/GCDMonoid/Nat.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson -/ import Mathlib.Algebra.GCDMonoid.Basic +import Mathlib.Algebra.Order.Group.Abs import Mathlib.Algebra.Order.Ring.Int import Mathlib.Data.Int.GCD diff --git a/Mathlib/Algebra/Order/Group/Int.lean b/Mathlib/Algebra/Order/Group/Int.lean index e8f3e16b97700..17cd94534280f 100644 --- a/Mathlib/Algebra/Order/Group/Int.lean +++ b/Mathlib/Algebra/Order/Group/Int.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ import Mathlib.Algebra.Group.Int -import Mathlib.Algebra.Order.Group.Abs +import Mathlib.Algebra.Order.Group.Defs /-! # The integers form a linear ordered group @@ -13,14 +13,6 @@ This file contains the linear ordered group instance on the integers. See note [foundational algebra order theory]. -## Recursors - -* `Int.rec`: Sign disjunction. Something is true/defined on `ℤ` if it's true/defined for nonnegative - and for negative values. (Defined in core Lean 3) -* `Int.inductionOn`: Simple growing induction on positive numbers, plus simple decreasing induction - on negative numbers. Note that this recursor is currently only `Prop`-valued. -* `Int.inductionOn'`: Simple growing induction for numbers greater than `b`, plus simple decreasing - induction on numbers less than `b`. -/ -- We should need only a minimal development of sets in order to get here. @@ -32,102 +24,9 @@ open Function Nat namespace Int -theorem natCast_strictMono : StrictMono (· : ℕ → ℤ) := fun _ _ ↦ Int.ofNat_lt.2 - -@[deprecated (since := "2024-05-25")] alias coe_nat_strictMono := natCast_strictMono - instance linearOrderedAddCommGroup : LinearOrderedAddCommGroup ℤ where __ := instLinearOrder __ := instAddCommGroup add_le_add_left _ _ := Int.add_le_add_left -/-! ### Miscellaneous lemmas -/ - -theorem abs_eq_natAbs : ∀ a : ℤ, |a| = natAbs a - | (n : ℕ) => abs_of_nonneg <| ofNat_zero_le _ - | -[_+1] => abs_of_nonpos <| le_of_lt <| negSucc_lt_zero _ - -@[simp, norm_cast] lemma natCast_natAbs (n : ℤ) : (n.natAbs : ℤ) = |n| := n.abs_eq_natAbs.symm - -theorem natAbs_abs (a : ℤ) : natAbs |a| = natAbs a := by rw [abs_eq_natAbs]; rfl - -theorem sign_mul_abs (a : ℤ) : sign a * |a| = a := by - rw [abs_eq_natAbs, sign_mul_natAbs a] - -lemma natAbs_le_self_sq (a : ℤ) : (Int.natAbs a : ℤ) ≤ a ^ 2 := by - rw [← Int.natAbs_sq a, sq] - norm_cast - apply Nat.le_mul_self - -alias natAbs_le_self_pow_two := natAbs_le_self_sq - -lemma le_self_sq (b : ℤ) : b ≤ b ^ 2 := le_trans le_natAbs (natAbs_le_self_sq _) - -alias le_self_pow_two := le_self_sq - -@[norm_cast] lemma abs_natCast (n : ℕ) : |(n : ℤ)| = n := abs_of_nonneg (natCast_nonneg n) - -theorem natAbs_sub_pos_iff {i j : ℤ} : 0 < natAbs (i - j) ↔ i ≠ j := by - rw [natAbs_pos, ne_eq, sub_eq_zero] - -theorem natAbs_sub_ne_zero_iff {i j : ℤ} : natAbs (i - j) ≠ 0 ↔ i ≠ j := - Nat.ne_zero_iff_zero_lt.trans natAbs_sub_pos_iff - -@[simp] -theorem abs_lt_one_iff {a : ℤ} : |a| < 1 ↔ a = 0 := by - rw [← zero_add 1, lt_add_one_iff, abs_nonpos_iff] - -theorem abs_le_one_iff {a : ℤ} : |a| ≤ 1 ↔ a = 0 ∨ a = 1 ∨ a = -1 := by - rw [le_iff_lt_or_eq, abs_lt_one_iff, abs_eq Int.one_nonneg] - -theorem one_le_abs {z : ℤ} (h₀ : z ≠ 0) : 1 ≤ |z| := - add_one_le_iff.mpr (abs_pos.mpr h₀) - -/-! #### `/` -/ - -theorem ediv_eq_zero_of_lt_abs {a b : ℤ} (H1 : 0 ≤ a) (H2 : a < |b|) : a / b = 0 := - match b, |b|, abs_eq_natAbs b, H2 with - | (n : ℕ), _, rfl, H2 => ediv_eq_zero_of_lt H1 H2 - | -[n+1], _, rfl, H2 => neg_injective <| by rw [← Int.ediv_neg]; exact ediv_eq_zero_of_lt H1 H2 - -/-! #### mod -/ - -@[simp] -theorem emod_abs (a b : ℤ) : a % |b| = a % b := - abs_by_cases (fun i => a % i = a % b) rfl (emod_neg _ _) - -theorem emod_lt (a : ℤ) {b : ℤ} (H : b ≠ 0) : a % b < |b| := by - rw [← emod_abs]; exact emod_lt_of_pos _ (abs_pos.2 H) - -/-! ### properties of `/` and `%` -/ - -theorem abs_ediv_le_abs : ∀ a b : ℤ, |a / b| ≤ |a| := - suffices ∀ (a : ℤ) (n : ℕ), |a / n| ≤ |a| from fun a b => - match b, eq_nat_or_neg b with - | _, ⟨n, Or.inl rfl⟩ => this _ _ - | _, ⟨n, Or.inr rfl⟩ => by rw [Int.ediv_neg, abs_neg]; apply this - fun a n => by - rw [abs_eq_natAbs, abs_eq_natAbs] - exact ofNat_le_ofNat_of_le - (match a, n with - | (m : ℕ), n => Nat.div_le_self _ _ - | -[m+1], 0 => Nat.zero_le _ - | -[m+1], n + 1 => Nat.succ_le_succ (Nat.div_le_self _ _)) - -theorem abs_sign_of_nonzero {z : ℤ} (hz : z ≠ 0) : |z.sign| = 1 := by - rw [abs_eq_natAbs, natAbs_sign_of_nonzero hz, Int.ofNat_one] - -protected theorem sign_eq_ediv_abs (a : ℤ) : sign a = a / |a| := - if az : a = 0 then by simp [az] - else (Int.ediv_eq_of_eq_mul_left (mt abs_eq_zero.1 az) (sign_mul_abs _).symm).symm - end Int - -section Group -variable {G : Type*} [Group G] - -@[to_additive (attr := simp) abs_zsmul_eq_zero] -lemma zpow_abs_eq_one (a : G) (n : ℤ) : a ^ |n| = 1 ↔ a ^ n = 1 := by - rw [← Int.natCast_natAbs, zpow_natCast, pow_natAbs_eq_one] - -end Group diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Int.lean b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean new file mode 100644 index 0000000000000..6082988933594 --- /dev/null +++ b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean @@ -0,0 +1,133 @@ +/- +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.Int +import Mathlib.Algebra.Order.Group.Unbundled.Abs + +/-! +# Facts about `ℤ` as an (unbundled) ordered group + +See note [foundational algebra order theory]. + +## Recursors + +* `Int.rec`: Sign disjunction. Something is true/defined on `ℤ` if it's true/defined for nonnegative + and for negative values. (Defined in core Lean 3) +* `Int.inductionOn`: Simple growing induction on positive numbers, plus simple decreasing induction + on negative numbers. Note that this recursor is currently only `Prop`-valued. +* `Int.inductionOn'`: Simple growing induction for numbers greater than `b`, plus simple decreasing + induction on numbers less than `b`. +-/ + +-- We should need only a minimal development of sets in order to get here. +assert_not_exists Set.Subsingleton + +assert_not_exists Ring + +open Function Nat + +namespace Int + +theorem natCast_strictMono : StrictMono (· : ℕ → ℤ) := fun _ _ ↦ Int.ofNat_lt.2 + +@[deprecated (since := "2024-05-25")] alias coe_nat_strictMono := natCast_strictMono + +/-! ### Miscellaneous lemmas -/ + +theorem abs_eq_natAbs : ∀ a : ℤ, |a| = natAbs a + | (n : ℕ) => abs_of_nonneg <| ofNat_zero_le _ + | -[_+1] => abs_of_nonpos <| le_of_lt <| negSucc_lt_zero _ + +@[simp, norm_cast] lemma natCast_natAbs (n : ℤ) : (n.natAbs : ℤ) = |n| := n.abs_eq_natAbs.symm + +theorem natAbs_abs (a : ℤ) : natAbs |a| = natAbs a := by rw [abs_eq_natAbs]; rfl + +theorem sign_mul_abs (a : ℤ) : sign a * |a| = a := by + rw [abs_eq_natAbs, sign_mul_natAbs a] + +lemma natAbs_le_self_sq (a : ℤ) : (Int.natAbs a : ℤ) ≤ a ^ 2 := by + rw [← Int.natAbs_sq a, sq] + norm_cast + apply Nat.le_mul_self + +alias natAbs_le_self_pow_two := natAbs_le_self_sq + +lemma le_self_sq (b : ℤ) : b ≤ b ^ 2 := le_trans le_natAbs (natAbs_le_self_sq _) + +alias le_self_pow_two := le_self_sq + +@[norm_cast] lemma abs_natCast (n : ℕ) : |(n : ℤ)| = n := abs_of_nonneg (natCast_nonneg n) + +theorem natAbs_sub_pos_iff {i j : ℤ} : 0 < natAbs (i - j) ↔ i ≠ j := by + rw [natAbs_pos, ne_eq, sub_eq_zero] + +theorem natAbs_sub_ne_zero_iff {i j : ℤ} : natAbs (i - j) ≠ 0 ↔ i ≠ j := + Nat.ne_zero_iff_zero_lt.trans natAbs_sub_pos_iff + +@[simp] +theorem abs_lt_one_iff {a : ℤ} : |a| < 1 ↔ a = 0 := by + rw [← zero_add 1, lt_add_one_iff, abs_nonpos_iff] + +theorem abs_le_one_iff {a : ℤ} : |a| ≤ 1 ↔ a = 0 ∨ a = 1 ∨ a = -1 := by + rw [le_iff_lt_or_eq, abs_lt_one_iff] + match a with + | (n : ℕ) => simp [abs_eq_natAbs] + | -[n+1] => + simp only [negSucc_ne_zero, abs_eq_natAbs, natAbs_negSucc, succ_eq_add_one, + natCast_add, Nat.cast_ofNat_Int, add_left_eq_self, natCast_eq_zero, false_or, reduceNeg] + rw [negSucc_eq'] + omega + +theorem one_le_abs {z : ℤ} (h₀ : z ≠ 0) : 1 ≤ |z| := + add_one_le_iff.mpr (abs_pos.mpr h₀) + +/-! #### `/` -/ + +theorem ediv_eq_zero_of_lt_abs {a b : ℤ} (H1 : 0 ≤ a) (H2 : a < |b|) : a / b = 0 := + match b, |b|, abs_eq_natAbs b, H2 with + | (n : ℕ), _, rfl, H2 => ediv_eq_zero_of_lt H1 H2 + | -[n+1], _, rfl, H2 => neg_injective <| by rw [← Int.ediv_neg]; exact ediv_eq_zero_of_lt H1 H2 + +/-! #### mod -/ + +@[simp] +theorem emod_abs (a b : ℤ) : a % |b| = a % b := + abs_by_cases (fun i => a % i = a % b) rfl (emod_neg _ _) + +theorem emod_lt (a : ℤ) {b : ℤ} (H : b ≠ 0) : a % b < |b| := by + rw [← emod_abs]; exact emod_lt_of_pos _ (abs_pos.2 H) + +/-! ### properties of `/` and `%` -/ + +theorem abs_ediv_le_abs : ∀ a b : ℤ, |a / b| ≤ |a| := + suffices ∀ (a : ℤ) (n : ℕ), |a / n| ≤ |a| from fun a b => + match b, eq_nat_or_neg b with + | _, ⟨n, Or.inl rfl⟩ => this _ _ + | _, ⟨n, Or.inr rfl⟩ => by rw [Int.ediv_neg, abs_neg]; apply this + fun a n => by + rw [abs_eq_natAbs, abs_eq_natAbs]; + exact ofNat_le_ofNat_of_le + (match a, n with + | (m : ℕ), n => Nat.div_le_self _ _ + | -[m+1], 0 => Nat.zero_le _ + | -[m+1], n + 1 => Nat.succ_le_succ (Nat.div_le_self _ _)) + +theorem abs_sign_of_nonzero {z : ℤ} (hz : z ≠ 0) : |z.sign| = 1 := by + rw [abs_eq_natAbs, natAbs_sign_of_nonzero hz, Int.ofNat_one] + +protected theorem sign_eq_ediv_abs (a : ℤ) : sign a = a / |a| := + if az : a = 0 then by simp [az] + else (Int.ediv_eq_of_eq_mul_left (mt abs_eq_zero.1 az) (sign_mul_abs _).symm).symm + +end Int + +section Group +variable {G : Type*} [Group G] + +@[to_additive (attr := simp) abs_zsmul_eq_zero] +lemma zpow_abs_eq_one (a : G) (n : ℤ) : a ^ |n| = 1 ↔ a ^ n = 1 := by + rw [← Int.natCast_natAbs, zpow_natCast, pow_natAbs_eq_one] + +end Group diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index 224d3010c3071..f834a11db01cb 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -3,6 +3,7 @@ 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 -/ +import Mathlib.Algebra.Order.Group.Abs import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Algebra.Order.Ring.Int import Mathlib.Algebra.Ring.Divisibility.Basic diff --git a/Mathlib/Algebra/Order/Ring/Cast.lean b/Mathlib/Algebra/Order/Ring/Cast.lean index fcc051946ab8c..0b74081308b9e 100644 --- a/Mathlib/Algebra/Order/Ring/Cast.lean +++ b/Mathlib/Algebra/Order/Ring/Cast.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Algebra.Order.Group.Abs import Mathlib.Algebra.Order.Ring.Int import Mathlib.Data.Nat.Cast.Order.Ring diff --git a/Mathlib/Algebra/Order/Ring/Int.lean b/Mathlib/Algebra/Order/Ring/Int.lean index 9888bbcd450bf..19de25e7f6877 100644 --- a/Mathlib/Algebra/Order/Ring/Int.lean +++ b/Mathlib/Algebra/Order/Ring/Int.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 -/ -import Mathlib.Algebra.Ring.Int -import Mathlib.Algebra.Order.Group.Int +import Mathlib.Algebra.Order.Group.Unbundled.Int import Mathlib.Algebra.Order.Ring.Defs +import Mathlib.Algebra.Ring.Int import Mathlib.Data.Set.Basic /-! diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index c7ea171ab05d7..9788ca3705b51 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.CharZero.Lemmas -import Mathlib.Algebra.Order.Group.Int -import Mathlib.Algebra.Ring.Int +import Mathlib.Algebra.Order.Ring.Int import Mathlib.Order.Interval.Finset.Basic /-! diff --git a/Mathlib/Data/Int/Star.lean b/Mathlib/Data/Int/Star.lean index 6f097d94ecbc0..c654adbb78ad6 100644 --- a/Mathlib/Data/Int/Star.lean +++ b/Mathlib/Data/Int/Star.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.Star.Order +import Mathlib.Algebra.Order.Group.Abs import Mathlib.Algebra.Order.Monoid.Submonoid import Mathlib.Algebra.Order.Ring.Basic From 74b685eb4a4f351a394e04504877bc7d133e8205 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 9 Aug 2024 12:27:50 +0000 Subject: [PATCH 147/359] fix: variables in TuringMachine.lean (#15422) This is work by @digama0 that I'm backport from the lean-pr-testing-4814 branch. Co-authored-by: Mario Carneiro --- Mathlib/Computability/TuringMachine.lean | 140 ++++++++++++----------- 1 file changed, 74 insertions(+), 66 deletions(-) diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index dd43164fc11a9..8523938c673cd 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -897,14 +897,13 @@ to remove the infinite tail of blanks.) namespace TM0 - section -- type of tape symbols -variable (Γ : Type*) [Inhabited Γ] +variable (Γ : Type*) -- type of "labels" or TM states -variable (Λ : Type*) [Inhabited Λ] +variable (Λ : Type*) /-- A Turing machine "statement" is just a command to either move left or right, or write a symbol on the tape. -/ @@ -914,7 +913,7 @@ inductive Stmt local notation "Stmt₀" => Stmt Γ -- Porting note (#10750): added this to clean up types. -instance Stmt.inhabited : Inhabited Stmt₀ := +instance Stmt.inhabited [Inhabited Γ] : Inhabited Stmt₀ := ⟨Stmt.write default⟩ /-- A Post-Turing machine with symbol type `Γ` and label type `Λ` @@ -932,7 +931,7 @@ def Machine [Inhabited Λ] := local notation "Machine₀" => Machine Γ Λ -- Porting note (#10750): added this to clean up types. -instance Machine.inhabited : Inhabited Machine₀ := by +instance Machine.inhabited [Inhabited Λ] : Inhabited Machine₀ := by unfold Machine; infer_instance /-- The configuration state of a Turing machine during operation @@ -940,7 +939,7 @@ instance Machine.inhabited : Inhabited Machine₀ := by The tape is represented in the form `(a, L, R)`, meaning the tape looks like `L.rev ++ [a] ++ R` with the machine currently reading the `a`. The lists are automatically extended with blanks as the machine moves around. -/ -structure Cfg where +structure Cfg [Inhabited Γ] where /-- The current machine state. -/ q : Λ /-- The current state of the tape: current symbol, left and right parts. -/ @@ -948,10 +947,13 @@ structure Cfg where local notation "Cfg₀" => Cfg Γ Λ -- Porting note (#10750): added this to clean up types. -instance Cfg.inhabited : Inhabited Cfg₀ := - ⟨⟨default, default⟩⟩ - variable {Γ Λ} +variable [Inhabited Λ] + +section +variable [Inhabited Γ] + +instance Cfg.inhabited : Inhabited Cfg₀ := ⟨⟨default, default⟩⟩ /-- Execution semantics of the Turing machine. -/ def step (M : Machine₀) : Cfg₀ → Option Cfg₀ := @@ -961,12 +963,10 @@ def step (M : Machine₀) : Cfg₀ → Option Cfg₀ := /-- The statement `Reaches M s₁ s₂` means that `s₂` is obtained starting from `s₁` after a finite number of steps from `s₂`. -/ -def Reaches (M : Machine₀) : Cfg₀ → Cfg₀ → Prop := - ReflTransGen fun a b ↦ b ∈ step M a +def Reaches (M : Machine₀) : Cfg₀ → Cfg₀ → Prop := ReflTransGen fun a b ↦ b ∈ step M a /-- The initial configuration. -/ -def init (l : List Γ) : Cfg₀ := - ⟨default, Tape.mk₁ l⟩ +def init (l : List Γ) : Cfg₀ := ⟨default, Tape.mk₁ l⟩ /-- Evaluate a Turing machine on initial input to a final state, if it terminates. -/ @@ -989,6 +989,8 @@ theorem step_supports (M : Machine₀) {S : Set Λ} (ss : Supports M S) : rcases Option.map_eq_some'.1 h₁ with ⟨⟨q', a⟩, h, rfl⟩ exact ss.2 h h₂ +end + theorem univ_supports (M : Machine₀) : Supports M Set.univ := by constructor <;> intros <;> apply Set.mem_univ @@ -1086,7 +1088,7 @@ namespace TM1 section -variable (Γ : Type*) [Inhabited Γ] +variable (Γ : Type*) -- Type of tape symbols variable (Λ : Type*) @@ -1119,12 +1121,11 @@ local notation "Stmt₁" => Stmt Γ Λ σ -- Porting note (#10750): added this open Stmt -instance Stmt.inhabited : Inhabited Stmt₁ := - ⟨halt⟩ +instance Stmt.inhabited : Inhabited Stmt₁ := ⟨halt⟩ /-- The configuration of a TM1 machine is given by the currently evaluating statement, the variable store value, and the tape. -/ -structure Cfg where +structure Cfg [Inhabited Γ] where /-- The statement (if any) which is currently evaluated -/ l : Option Λ /-- The current value of the variable store -/ @@ -1134,13 +1135,13 @@ structure Cfg where local notation "Cfg₁" => Cfg Γ Λ σ -- Porting note (#10750): added this to clean up types. -instance Cfg.inhabited [Inhabited σ] : Inhabited Cfg₁ := +instance Cfg.inhabited [Inhabited Γ] [Inhabited σ] : Inhabited Cfg₁ := ⟨⟨default, default, default⟩⟩ variable {Γ Λ σ} /-- The semantics of TM1 evaluation. -/ -def stepAux : Stmt₁ → σ → Tape Γ → Cfg₁ +def stepAux [Inhabited Γ] : Stmt₁ → σ → Tape Γ → Cfg₁ | move d q, v, T => stepAux q v (T.move d) | write a q, v, T => stepAux q v (T.write (a T.1 v)) | load s q, v, T => stepAux q (s T.1 v) T @@ -1149,7 +1150,7 @@ def stepAux : Stmt₁ → σ → Tape Γ → Cfg₁ | halt, v, T => ⟨none, v, T⟩ /-- The state transition function. -/ -def step (M : Λ → Stmt₁) : Cfg₁ → Option Cfg₁ +def step [Inhabited Γ] (M : Λ → Stmt₁) : Cfg₁ → Option Cfg₁ | ⟨none, _, _⟩ => none | ⟨some l, v, T⟩ => some (stepAux (M l) v T) @@ -1229,6 +1230,8 @@ theorem stmts_supportsStmt {M : Λ → Stmt₁} {S : Finset Λ} {q : Stmt₁} (s forall_eq', exists_imp, and_imp] exact fun l ls h ↦ stmts₁_supportsStmt_mono h (ss.2 _ ls) +variable [Inhabited Γ] + theorem step_supports (M : Λ → Stmt₁) {S : Finset Λ} (ss : Supports M S) : ∀ {c c' : Cfg₁}, c' ∈ step M c → c.l ∈ Finset.insertNone S → c'.l ∈ Finset.insertNone S | ⟨some l₁, v, T⟩, c', h₁, h₂ => by @@ -1280,10 +1283,7 @@ TM1 semantics. namespace TM1to0 - -section - -variable {Γ : Type*} [Inhabited Γ] +variable {Γ : Type*} variable {Λ : Type*} [Inhabited Λ] variable {σ : Type*} [Inhabited σ] @@ -1333,10 +1333,10 @@ def tr : TM0.Machine Γ Λ'₁₀ | (some q, v), s => some (trAux M s q v) /-- Translate configurations from TM1 to TM0. -/ -def trCfg : Cfg₁ → Cfg₁₀ +def trCfg [Inhabited Γ] : Cfg₁ → Cfg₁₀ | ⟨l, v, T⟩ => ⟨(l.map M, v), T⟩ -theorem tr_respects : +theorem tr_respects [Inhabited Γ] : Respects (TM1.step M) (TM0.step (tr M)) fun (c₁ : Cfg₁) (c₂ : Cfg₁₀) ↦ trCfg M c₁ = c₂ := fun_respects.2 fun ⟨l₁, v, T⟩ ↦ by cases' l₁ with l₁; · exact rfl @@ -1352,7 +1352,7 @@ theorem tr_respects : | _ => exact TransGen.single (congr_arg some (congr (congr_arg TM0.Cfg.mk rfl) (Tape.write_self T))) -theorem tr_eval (l : List Γ) : TM0.eval (tr M) l = TM1.eval M l := +theorem tr_eval [Inhabited Γ] (l : List Γ) : TM0.eval (tr M) l = TM1.eval M l := (congr_arg _ (tr_eval' _ _ _ (tr_respects M) ⟨some _, _, _⟩)).trans (by rw [Part.map_eq_map, Part.map_map, TM1.eval] @@ -1413,8 +1413,6 @@ theorem tr_supports {S : Finset Λ} (ss : TM1.Supports M S) : exact Finset.some_mem_insertNone.2 (Finset.mem_biUnion.2 ⟨_, hs _ _, TM1.stmts₁_self⟩) | halt => cases h₁ -end - end TM1to0 /-! @@ -1443,15 +1441,12 @@ finitely long. namespace TM1to1 - open TM1 +variable {Γ : Type*} -section - -variable {Γ : Type*} [Inhabited Γ] - -theorem exists_enc_dec [Finite Γ] : ∃ (n : ℕ) (enc : Γ → Vector Bool n) (dec : Vector Bool n → Γ), - enc default = Vector.replicate n false ∧ ∀ a, dec (enc a) = a := by +theorem exists_enc_dec [Inhabited Γ] [Finite Γ] : + ∃ (n : ℕ) (enc : Γ → Vector Bool n) (dec : 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 let G : Fin n ↪ Fin n → Bool := @@ -1461,8 +1456,7 @@ theorem exists_enc_dec [Finite Γ] : ∃ (n : ℕ) (enc : Γ → Vector Bool n) let enc := H.setValue default (Vector.replicate n false) exact ⟨_, enc, Function.invFun enc, H.setValue_eq _ _, Function.leftInverse_invFun enc.2⟩ -variable {Λ : Type*} [Inhabited Λ] -variable {σ : Type*} [Inhabited σ] +variable {Λ σ : Type*} local notation "Stmt₁" => Stmt Γ Λ σ @@ -1475,7 +1469,7 @@ inductive Λ' local notation "Λ'₁" => @Λ' Γ Λ σ -- Porting note (#10750): added this to clean up types. -instance : Inhabited Λ'₁ := +instance [Inhabited Λ] : Inhabited Λ'₁ := ⟨Λ'.normal default⟩ local notation "Stmt'₁" => Stmt Bool Λ'₁ σ @@ -1544,10 +1538,12 @@ theorem supportsStmt_read {S : Finset Λ'₁} : induction' i with i IH; · exact hf _ constructor <;> apply IH <;> intro <;> apply hf -variable (enc0 : enc default = Vector.replicate n false) +variable (M : Λ → TM1.Stmt Γ Λ σ) -- Porting note: Unfolded `Stmt₁`. section +variable [Inhabited Γ] (enc0 : enc default = Vector.replicate n false) +section variable {enc} /-- The low level tape corresponding to the given tape over alphabet `Γ`. -/ @@ -1566,8 +1562,6 @@ theorem trTape_mk' (L R : ListBlank Γ) : trTape enc0 (Tape.mk' L R) = trTape' e end -variable (M : Λ → TM1.Stmt Γ Λ σ) -- Porting note: Unfolded `Stmt₁`. - /-- The top level program. -/ def tr : Λ'₁ → Stmt'₁ | Λ'.normal l => trNormal dec (M l) @@ -1698,6 +1692,8 @@ theorem tr_respects {enc₀} : trTape'_move_right enc0, trTape_mk'] apply ReflTransGen.refl +end + open scoped Classical variable [Fintype Γ] @@ -1716,7 +1712,8 @@ are the normal states embedded from `S`, plus all write states accessible from t noncomputable def trSupp (S : Finset Λ) : Finset Λ'₁ := S.biUnion fun l ↦ insert (Λ'.normal l) (writes (M l)) -theorem tr_supports {S : Finset Λ} (ss : Supports M S) : Supports (tr enc dec M) (trSupp M S) := +theorem tr_supports [Inhabited Λ] {S : Finset Λ} (ss : Supports M S) : + Supports (tr enc dec M) (trSupp M S) := ⟨Finset.mem_biUnion.2 ⟨_, ss.1, Finset.mem_insert_self _ _⟩, fun q h ↦ by suffices ∀ q, SupportsStmt S q → (∀ q' ∈ writes q, q' ∈ trSupp M S) → SupportsStmt (trSupp M S) (trNormal dec q) ∧ @@ -1759,8 +1756,6 @@ theorem tr_supports {S : Finset Λ} (ss : Supports M S) : Supports (tr enc dec M simp only [writes, Finset.not_mem_empty]; refine ⟨?_, fun _ ↦ False.elim⟩ simp only [SupportsStmt, supportsStmt_move, trNormal]⟩ -end - end TM1to1 /-! @@ -1780,9 +1775,6 @@ unreachable branch). namespace TM0to1 - -section - variable {Γ : Type*} [Inhabited Γ] variable {Λ : Type*} [Inhabited Λ] @@ -1845,8 +1837,6 @@ theorem tr_respects : Respects (TM0.step M) (TM1.step (tr M)) fun a b ↦ trCfg rfl · rfl -end - end TM0to1 /-! @@ -1879,10 +1869,7 @@ as the output stack. namespace TM2 - -section - -variable {K : Type*} [DecidableEq K] +variable {K : Type*} -- Index type of stacks variable (Γ : K → Type*) @@ -1934,6 +1921,9 @@ instance Cfg.inhabited [Inhabited σ] : Inhabited Cfg₂ := variable {Γ Λ σ} +section +variable [DecidableEq K] + /-- The step function for the TM2 model. -/ @[simp] def stepAux : Stmt₂ → σ → (∀ k, List (Γ k)) → Cfg₂ @@ -1955,6 +1945,8 @@ def step (M : Λ → Stmt₂) : Cfg₂ → Option Cfg₂ def Reaches (M : Λ → Stmt₂) : Cfg₂ → Cfg₂ → Prop := ReflTransGen fun a b ↦ b ∈ step M a +end + /-- Given a set `S` of states, `SupportsStmt S q` means that `q` only jumps to states in `S`. -/ def SupportsStmt (S : Finset Λ) : Stmt₂ → Prop | push _ _ q => SupportsStmt S q @@ -1965,6 +1957,7 @@ def SupportsStmt (S : Finset Λ) : Stmt₂ → Prop | goto l => ∀ v, l v ∈ S | halt => True +section open scoped Classical /-- The set of subtree statements in a statement. -/ @@ -2019,6 +2012,8 @@ theorem stmts_trans {M : Λ → Stmt₂} {S : Finset Λ} {q₁ q₂ : Stmt₂} ( forall_eq', exists_imp, and_imp] exact fun l ls h₂ ↦ ⟨_, ls, stmts₁_trans h₂ h₁⟩ +end + variable [Inhabited Λ] /-- Given a TM2 machine `M` and a set `S` of states, `Supports M S` means that all states in @@ -2032,6 +2027,8 @@ theorem stmts_supportsStmt {M : Λ → Stmt₂} {S : Finset Λ} {q : Stmt₂} (s forall_eq', exists_imp, and_imp] exact fun l ls h ↦ stmts₁_supportsStmt_mono h (ss.2 _ ls) +variable [DecidableEq K] + theorem step_supports (M : Λ → Stmt₂) {S : Finset Λ} (ss : Supports M S) : ∀ {c c' : Cfg₂}, c' ∈ step M c → c.l ∈ Finset.insertNone S → c'.l ∈ Finset.insertNone S | ⟨some l₁, v, T⟩, c', h₁, h₂ => by @@ -2056,8 +2053,6 @@ def init (k : K) (L : List (Γ k)) : Cfg₂ := def eval (M : Λ → Stmt₂) (k : K) (L : List (Γ k)) : Part (List (Γ k)) := (Turing.eval (step M) (init k L)).map fun c ↦ c.stk k -end - end TM2 /-! @@ -2101,7 +2096,6 @@ steps to run when emulated in TM1, where `m` is the length of the input. namespace TM2to1 - -- A displaced lemma proved in unnecessary generality theorem stk_nth_val {K : Type*} {Γ : K → Type*} {L : ListBlank (∀ k, Option (Γ k))} {k S} (n) (hL : ListBlank.map (proj k) L = ListBlank.mk (List.map some S).reverse) : @@ -2110,12 +2104,9 @@ theorem stk_nth_val {K : Type*} {Γ : K → Type*} {L : ListBlank (∀ k, Option List.get?_map] cases S.reverse.get? n <;> rfl -section - -variable {K : Type*} [DecidableEq K] +variable {K : Type*} variable {Γ : K → Type*} -variable {Λ : Type*} [Inhabited Λ] -variable {σ : Type*} [Inhabited σ] +variable {Λ σ : Type*} local notation "Stmt₂" => TM2.Stmt Γ Λ σ @@ -2132,7 +2123,7 @@ local notation "Γ'₂₁" => @Γ' K Γ -- Porting note (#10750): added this to instance Γ'.inhabited : Inhabited Γ'₂₁ := ⟨⟨false, fun _ ↦ none⟩⟩ -instance Γ'.fintype [Fintype K] [∀ k, Fintype (Γ k)] : Fintype Γ'₂₁ := +instance Γ'.fintype [DecidableEq K] [Fintype K] [∀ k, Fintype (Γ k)] : Fintype Γ'₂₁ := instFintypeProd _ _ /-- The bottom marker is fixed throughout the calculation, so we use the `addBottom` function @@ -2235,15 +2226,17 @@ local notation "Λ'₂₁" => @Λ' K Γ Λ σ -- Porting note (#10750): added t open Λ' -instance Λ'.inhabited : Inhabited Λ'₂₁ := +instance Λ'.inhabited [Inhabited Λ] : Inhabited Λ'₂₁ := ⟨normal default⟩ local notation "Stmt₂₁" => TM1.Stmt Γ'₂₁ Λ'₂₁ σ - local notation "Cfg₂₁" => TM1.Cfg Γ'₂₁ Λ'₂₁ σ open TM1.Stmt +section +variable [DecidableEq K] + /-- The program corresponding to state transitions at the end of a stack. Here we start out just after the top of the stack, and should end just after the new top of the stack. -/ def trStAct {k : K} (q : Stmt₂₁) : StAct₂ k → Stmt₂₁ @@ -2266,6 +2259,8 @@ theorem step_run {k : K} (q : Stmt₂) (v : σ) (S : ∀ k, List (Γ k)) : ∀ s | StAct.peek f => by unfold stWrite; rw [Function.update_eq_self]; rfl | StAct.pop f => rfl +end + /-- The translation of TM2 statements to TM1 statements. regular actions have direct equivalents, but stack actions are deferred by going to the corresponding `go` state, so that we can find the appropriate stack top. -/ @@ -2282,6 +2277,7 @@ theorem trNormal_run {k : K} (s : StAct₂ k) (q : Stmt₂) : trNormal (stRun s q) = goto fun _ _ ↦ go k s q := by cases s <;> rfl +section open scoped Classical /-- The set of machine states accessible from an initial TM2 statement. -/ @@ -2297,7 +2293,7 @@ theorem trStmts₁_run {k : K} {s : StAct₂ k} {q : Stmt₂} : trStmts₁ (stRun s q) = {go k s q, ret q} ∪ trStmts₁ q := by cases s <;> simp only [trStmts₁] -theorem tr_respects_aux₂ {k : K} {q : Stmt₂₁} {v : σ} {S : ∀ k, List (Γ k)} +theorem tr_respects_aux₂ [DecidableEq K] {k : K} {q : Stmt₂₁} {v : σ} {S : ∀ k, List (Γ k)} {L : ListBlank (∀ k, Option (Γ k))} (hL : ∀ k, L.map (proj k) = ListBlank.mk ((S k).map some).reverse) (o : StAct₂ k) : let v' := stVar v (S k) o @@ -2381,6 +2377,9 @@ theorem tr_respects_aux₂ {k : K} {q : Stmt₂₁} {v : σ} {S : ∀ k, List ( · split_ifs <;> rw [Function.update_noteq h', ← proj_map_nth, hL] rw [Function.update_noteq h'] +end + +variable [DecidableEq K] variable (M : Λ → TM2.Stmt Γ Λ σ) -- Porting note: Unfolded `Stmt₂`. /-- The TM2 emulator machine states written as a TM1 program. @@ -2466,6 +2465,9 @@ theorem tr_respects : Respects (TM2.step M) (TM1.step (tr M)) TrCfg := by | H₄ => exact ⟨_, ⟨_, hT⟩, ReflTransGen.refl⟩ | H₅ => exact ⟨_, ⟨_, hT⟩, ReflTransGen.refl⟩ +section +variable [Inhabited Λ] [Inhabited σ] + theorem trCfg_init (k) (L : List (Γ k)) : TrCfg (TM2.init k L) (TM1.init (trInit k L) : Cfg₂₁) := by rw [(_ : TM1.init _ = _)] · refine ⟨ListBlank.mk (L.reverse.map fun a ↦ update default k (some a)), fun k' ↦ ?_⟩ @@ -2503,6 +2505,12 @@ theorem tr_eval (k) (L : List (Γ k)) {L₁ L₂} (H₁ : L₁ ∈ TM1.eval (tr cases Part.mem_unique h₁ h₃ exact ⟨_, L', by simp only [Tape.mk'_right₀], hT, rfl⟩ +end + +section +open scoped Classical +variable [Inhabited Λ] + /-- The support of a set of TM2 states in the TM2 emulator. -/ noncomputable def trSupp (S : Finset Λ) : Finset Λ'₂₁ := S.biUnion fun l ↦ insert (normal l) (trStmts₁ (M l)) From 4ce28eb31d741f6854600819c941ab06c9d910a6 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Fri, 9 Aug 2024 13:42:47 +0000 Subject: [PATCH 148/359] chore(Algebra): turn `induction'` into `induction` (#15642) This PR turns the (soon to be deprecated?) `induction'` into `induction`. --- Mathlib/Algebra/Algebra/Operations.lean | 7 +++-- Mathlib/Algebra/BigOperators/Group/List.lean | 14 +++++----- .../Computation/Translations.lean | 28 +++++++++++-------- 3 files changed, 27 insertions(+), 22 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 6f96922e603a9..668fb28578148 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -398,9 +398,10 @@ theorem pow_mem_pow {x : A} (hx : x ∈ M) (n : ℕ) : x ^ n ∈ M ^ n := pow_subset_pow _ <| Set.pow_mem_pow hx _ theorem pow_toAddSubmonoid {n : ℕ} (h : n ≠ 0) : (M ^ n).toAddSubmonoid = M.toAddSubmonoid ^ n := by - induction' n with n ih - · exact (h rfl).elim - · rw [pow_succ, pow_succ, mul_toAddSubmonoid] + induction n with + | zero => exact (h rfl).elim + | succ n ih => + rw [pow_succ, pow_succ, mul_toAddSubmonoid] cases n with | zero => rw [pow_zero, pow_zero, one_mul, ← mul_toAddSubmonoid, one_mul] | succ n => rw [ih n.succ_ne_zero] diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 750c88af849a9..9c6cfd3293ff3 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -119,10 +119,9 @@ theorem prod_eq_foldr : ∀ {l : List M}, l.prod = foldr (· * ·) 1 l @[to_additive (attr := simp)] theorem prod_replicate (n : ℕ) (a : M) : (replicate n a).prod = a ^ n := by - induction' n with n ih - · rw [pow_zero] - rfl - · rw [replicate_succ, prod_cons, ih, pow_succ'] + induction n with + | zero => rw [pow_zero]; rfl + | succ n ih => rw [replicate_succ, prod_cons, ih, pow_succ'] @[to_additive sum_eq_card_nsmul] theorem prod_eq_pow_card (l : List M) (m : M) (h : ∀ x ∈ l, x = m) : l.prod = m ^ l.length := by @@ -424,9 +423,10 @@ theorem prod_drop_succ : @[to_additive "Cancellation of a telescoping sum."] theorem prod_range_div' (n : ℕ) (f : ℕ → G) : ((range n).map fun k ↦ f k / f (k + 1)).prod = f 0 / f n := by - induction' n with n h - · exact (div_self' (f 0)).symm - · rw [range_succ, map_append, map_singleton, prod_append, prod_singleton, h, div_mul_div_cancel'] + induction n with + | zero => exact (div_self' (f 0)).symm + | succ n h => + rw [range_succ, map_append, map_singleton, prod_append, prod_singleton, h, div_mul_div_cancel'] lemma prod_rotate_eq_one_of_prod_eq_one : ∀ {l : List G} (_ : l.prod = 1) (n : ℕ), (l.rotate n).prod = 1 diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean b/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean index c14beeefe520e..34c8347407af7 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean @@ -95,10 +95,11 @@ theorem stream_succ_of_some {p : IntFractPair K} (h : IntFractPair.stream v n = /-- The stream of `IntFractPair`s of an integer stops after the first term. -/ theorem stream_succ_of_int (a : ℤ) (n : ℕ) : IntFractPair.stream (a : K) (n + 1) = none := by - induction' n with n ih - · refine IntFractPair.stream_eq_none_of_fr_eq_zero (IntFractPair.stream_zero (a : K)) ?_ + induction n with + | zero => + refine IntFractPair.stream_eq_none_of_fr_eq_zero (IntFractPair.stream_zero (a : K)) ?_ simp only [IntFractPair.of, Int.fract_intCast] - · exact IntFractPair.succ_nth_stream_eq_none_iff.mpr (Or.inl ih) + | succ n ih => exact IntFractPair.succ_nth_stream_eq_none_iff.mpr (Or.inl ih) theorem exists_succ_nth_stream_of_fr_zero {ifp_succ_n : IntFractPair K} (stream_succ_nth_eq : IntFractPair.stream v (n + 1) = some ifp_succ_n) @@ -117,10 +118,12 @@ the inverse of the fractional part of `v`. -/ theorem stream_succ (h : Int.fract v ≠ 0) (n : ℕ) : IntFractPair.stream v (n + 1) = IntFractPair.stream (Int.fract v)⁻¹ n := by - induction' n with n ih - · have H : (IntFractPair.of v).fr = Int.fract v := rfl + induction n with + | zero => + have H : (IntFractPair.of v).fr = Int.fract v := rfl rw [stream_zero, stream_succ_of_some (stream_zero v) (ne_of_eq_of_ne H h), H] - · rcases eq_or_ne (IntFractPair.stream (Int.fract v)⁻¹ n) none with hnone | hsome + | succ n ih => + rcases eq_or_ne (IntFractPair.stream (Int.fract v)⁻¹ n) none with hnone | hsome · rw [hnone] at ih rw [succ_nth_stream_eq_none_iff.mpr (Or.inl hnone), succ_nth_stream_eq_none_iff.mpr (Or.inl ih)] @@ -263,9 +266,9 @@ variable (K) theorem of_s_of_int (a : ℤ) : (of (a : K)).s = Stream'.Seq.nil := haveI h : ∀ n, (of (a : K)).s.get? n = none := by intro n - induction' n with n ih - · rw [of_s_head_aux, stream_succ_of_int, Option.bind] - · exact (of (a : K)).s.prop ih + induction n with + | zero => rw [of_s_head_aux, stream_succ_of_int, Option.bind] + | succ n ih => exact (of (a : K)).s.prop ih Stream'.Seq.ext fun n => (h n).trans (Stream'.Seq.get?_nil n).symm variable {K} (v) @@ -300,9 +303,10 @@ variable (K) (n) are all equal to `a`. -/ theorem convs'_of_int (a : ℤ) : (of (a : K)).convs' n = a := by - induction' n with n - · simp only [zeroth_conv'_eq_h, of_h_eq_floor, floor_intCast, Nat.zero_eq] - · rw [convs', of_h_eq_floor, floor_intCast, add_right_eq_self] + induction n with + | zero => simp only [zeroth_conv'_eq_h, of_h_eq_floor, floor_intCast, Nat.zero_eq] + | succ => + rw [convs', of_h_eq_floor, floor_intCast, add_right_eq_self] exact convs'Aux_succ_none ((of_s_of_int K a).symm ▸ Stream'.Seq.get?_nil 0) _ variable {K} From 7dc099d1860dd3c99ccdd01803d142e1b21aa3aa Mon Sep 17 00:00:00 2001 From: Salvatore Mercuri Date: Fri, 9 Aug 2024 14:31:42 +0000 Subject: [PATCH 149/359] feat(Topology/Instances): the only closed subfield of the field of real numbers is the real numbers itself (#15094) This PR contains the result that the only closed subfield of the field of real numbers is the field of real numbers itself. Split off from the larger PR #13577. --- Mathlib/Topology/Instances/Real.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index 7d74ffc827e2d..764564cb7bbfe 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro import Mathlib.Data.Real.Star import Mathlib.Algebra.Algebra.Basic import Mathlib.Algebra.Periodic +import Mathlib.Topology.Algebra.Order.Archimedean import Mathlib.Topology.Algebra.Order.Field import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Topology.Algebra.Star @@ -128,6 +129,12 @@ instance Real.instCompleteSpace : CompleteSpace ℝ := by theorem Real.totallyBounded_ball (x ε : ℝ) : TotallyBounded (ball x ε) := by rw [Real.ball_eq_Ioo]; apply totallyBounded_Ioo +theorem Real.subfield_eq_of_closed {K : Subfield ℝ} (hc : IsClosed (K : Set ℝ)) : K = ⊤ := by + rw [SetLike.ext'_iff, Subfield.coe_top, ← hc.closure_eq] + refine Rat.denseRange_cast.mono ?_ |>.closure_eq + rintro - ⟨_, rfl⟩ + exact SubfieldClass.ratCast_mem K _ + section theorem closure_of_rat_image_lt {q : ℚ} : From c0ec0639aae39a4c61d38bf19d2f78f050e4d295 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Fri, 9 Aug 2024 14:31:43 +0000 Subject: [PATCH 150/359] feat: add lemma relating Set.encard to tsum over ennreal (#15214) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit two lemmas: `s.tsum (1:ENNReal) = s.encard` and `s.tsum c = s.encard * c`, where `c:ENNReal` and `s : Set α`. Co-authored-by: blizzard_inc --- Mathlib.lean | 1 + Mathlib/Data/Set/Card.lean | 2 + .../Topology/Algebra/InfiniteSum/ENNReal.lean | 41 +++++++++++++++++++ 3 files changed, 44 insertions(+) create mode 100644 Mathlib/Topology/Algebra/InfiniteSum/ENNReal.lean diff --git a/Mathlib.lean b/Mathlib.lean index 512d587840374..c6d9c3843e342 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4262,6 +4262,7 @@ import Mathlib.Topology.Algebra.GroupWithZero import Mathlib.Topology.Algebra.InfiniteSum.Basic import Mathlib.Topology.Algebra.InfiniteSum.Constructions import Mathlib.Topology.Algebra.InfiniteSum.Defs +import Mathlib.Topology.Algebra.InfiniteSum.ENNReal import Mathlib.Topology.Algebra.InfiniteSum.Group import Mathlib.Topology.Algebra.InfiniteSum.GroupCompletion import Mathlib.Topology.Algebra.InfiniteSum.Module diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index 50d28dae9459d..9984499e1f215 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -133,6 +133,8 @@ theorem Finite.exists_encard_eq_coe (h : s.Finite) : ∃ (n : ℕ), s.encard = n @[simp] theorem encard_eq_top_iff : s.encard = ⊤ ↔ s.Infinite := by rw [← not_iff_not, ← Ne, ← lt_top_iff_ne_top, encard_lt_top_iff, not_infinite] +alias ⟨_, encard_eq_top⟩ := encard_eq_top_iff + theorem encard_ne_top_iff : s.encard ≠ ⊤ ↔ s.Finite := by simp diff --git a/Mathlib/Topology/Algebra/InfiniteSum/ENNReal.lean b/Mathlib/Topology/Algebra/InfiniteSum/ENNReal.lean new file mode 100644 index 0000000000000..2e8e9cea6e74e --- /dev/null +++ b/Mathlib/Topology/Algebra/InfiniteSum/ENNReal.lean @@ -0,0 +1,41 @@ +/- +Copyright (c) 2024 Edward van de Meent. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Edward van de Meent +-/ +import Mathlib.Data.Real.ENatENNReal +import Mathlib.Data.Set.Card +import Mathlib.Topology.Instances.ENNReal + +/-! +# Infinite sums of ENNReal and Set.encard + +This file provides lemmas relating sums of constants to the cardinality of the domain of these sums. + +## TODO + ++ Once we have a topology on `ENat`, provide an `ENat` valued version ++ Once we replace `PartENat` entirely with `ENat` (and replace `PartENat.card` with a `ENat.card`), + provide versions which sum over the whole type. +-/ + +open Set Function + +variable {α : Type*} (s : Set α) + +namespace ENNReal + +@[simp] +lemma tsum_set_one_eq : ∑' (_ : s), (1 : ℝ≥0∞) = s.encard := by + obtain (hfin | hinf) := Set.finite_or_infinite s + · lift s to Finset α using hfin + simp [tsum_fintype] + · have : Infinite s := infinite_coe_iff.mpr hinf + rw [tsum_const_eq_top_of_ne_zero one_ne_zero, encard_eq_top hinf, ENat.toENNReal_top] + +@[simp] +lemma tsum_set_const_eq (c : ℝ≥0∞) : ∑' (_:s), (c : ℝ≥0∞) = s.encard * c := by + nth_rw 1 [← one_mul c] + rw [ENNReal.tsum_mul_right,tsum_set_one_eq] + +end ENNReal From 4d65f9ebe7e71258a18c2a22e26cf3cd6db6141b Mon Sep 17 00:00:00 2001 From: FR Date: Fri, 9 Aug 2024 14:59:46 +0000 Subject: [PATCH 151/359] perf: `to_additive (attr := reducible)` -> `abbrev` (#15476) --- Mathlib/Algebra/Group/Action/Defs.lean | 20 +-- Mathlib/Algebra/Group/InjSurj.lean | 122 +++++++++--------- Mathlib/Algebra/Group/MinimalAxioms.lean | 8 +- Mathlib/Algebra/Order/Group/InjSurj.lean | 8 +- Mathlib/Algebra/Order/Monoid/Basic.lean | 18 +-- Mathlib/Analysis/Normed/Group/Basic.lean | 72 +++++------ Mathlib/GroupTheory/EckmannHilton.lean | 8 +- Mathlib/GroupTheory/Exponent.lean | 4 +- Mathlib/GroupTheory/GroupAction/Basic.lean | 4 +- .../GroupTheory/MonoidLocalization/Basic.lean | 12 +- .../GroupTheory/SpecificGroups/KleinFour.lean | 4 +- Mathlib/Logic/Equiv/TransferInstance.lean | 44 +++---- Mathlib/MeasureTheory/Measure/Haar/Basic.lean | 4 +- 13 files changed, 164 insertions(+), 164 deletions(-) diff --git a/Mathlib/Algebra/Group/Action/Defs.lean b/Mathlib/Algebra/Group/Action/Defs.lean index 113c045227580..b96c5a818e280 100644 --- a/Mathlib/Algebra/Group/Action/Defs.lean +++ b/Mathlib/Algebra/Group/Action/Defs.lean @@ -284,9 +284,9 @@ variable (α) -- See note [reducible non-instances] -- Since this is reducible, we make sure to go via -- `SMul.comp.smul` to prevent typeclass inference unfolding too far -@[to_additive (attr := reducible) +@[to_additive "An additive action of `M` on `α` and a function `N → M` induces an additive action of `N` on `α`."] -def comp (g : N → M) : SMul N α where smul := SMul.comp.smul g +abbrev comp (g : N → M) : SMul N α where smul := SMul.comp.smul g variable {α} @@ -391,9 +391,9 @@ variable {M} /-- Pullback a multiplicative action along an injective map respecting `•`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "Pullback an additive action along an injective map respecting `+ᵥ`."] -protected def Function.Injective.mulAction [SMul M β] (f : β → α) (hf : Injective f) +protected abbrev Function.Injective.mulAction [SMul M β] (f : β → α) (hf : Injective f) (smul : ∀ (c : M) (x), f (c • x) = c • f x) : MulAction M β where smul := (· • ·) one_smul x := hf <| (smul _ _).trans <| one_smul _ (f x) @@ -401,9 +401,9 @@ protected def Function.Injective.mulAction [SMul M β] (f : β → α) (hf : Inj /-- Pushforward a multiplicative action along a surjective map respecting `•`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "Pushforward an additive action along a surjective map respecting `+ᵥ`."] -protected def Function.Surjective.mulAction [SMul M β] (f : α → β) (hf : Surjective f) +protected abbrev Function.Surjective.mulAction [SMul M β] (f : α → β) (hf : Surjective f) (smul : ∀ (c : M) (x), f (c • x) = c • f x) : MulAction M β where smul := (· • ·) one_smul := by simp [hf.forall, ← smul] @@ -413,9 +413,9 @@ protected def Function.Surjective.mulAction [SMul M β] (f : α → β) (hf : Su See also `Function.Surjective.distribMulActionLeft` and `Function.Surjective.moduleLeft`. -/ -@[to_additive (attr := reducible) +@[to_additive "Push forward the action of `R` on `M` along a compatible surjective map `f : R →+ S`."] -def Function.Surjective.mulActionLeft {R S M : Type*} [Monoid R] [MulAction R M] [Monoid S] +abbrev Function.Surjective.mulActionLeft {R S M : Type*} [Monoid R] [MulAction R M] [Monoid S] [SMul S M] (f : R →* S) (hf : Surjective f) (hsmul : ∀ (c) (x : M), f c • x = c • x) : MulAction S M where smul := (· • ·) @@ -523,8 +523,8 @@ variable (α) a multiplicative action of `N` on `α`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible)] -def compHom [Monoid N] (g : N →* M) : MulAction N α where +@[to_additive] +abbrev compHom [Monoid N] (g : N →* M) : MulAction N α where smul := SMul.comp.smul g -- Porting note: was `by simp [g.map_one, MulAction.one_smul]` one_smul _ := by simpa [(· • ·)] using MulAction.one_smul .. diff --git a/Mathlib/Algebra/Group/InjSurj.lean b/Mathlib/Algebra/Group/InjSurj.lean index 29049efbaf214..1c5b3b20b480d 100644 --- a/Mathlib/Algebra/Group/InjSurj.lean +++ b/Mathlib/Algebra/Group/InjSurj.lean @@ -51,46 +51,46 @@ variable {M₁ : Type*} {M₂ : Type*} [Mul M₁] /-- A type endowed with `*` is a semigroup, if it admits an injective map that preserves `*` to a semigroup. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) "A type endowed with `+` is an additive semigroup, if it admits an +@[to_additive "A type endowed with `+` is an additive semigroup, if it admits an injective map that preserves `+` to an additive semigroup."] -protected def semigroup [Semigroup M₂] (f : M₁ → M₂) (hf : Injective f) +protected abbrev semigroup [Semigroup M₂] (f : M₁ → M₂) (hf : Injective f) (mul : ∀ x y, f (x * y) = f x * f y) : Semigroup M₁ := { ‹Mul M₁› with mul_assoc := fun x y z => hf <| by erw [mul, mul, mul, mul, mul_assoc] } /-- A type endowed with `*` is a commutative magma, if it admits a surjective map that preserves `*` from a commutative magma. -/ -@[to_additive (attr := reducible) -- See note [reducible non-instances] +@[to_additive -- See note [reducible non-instances] "A type endowed with `+` is an additive commutative semigroup, if it admits a surjective map that preserves `+` from an additive commutative semigroup."] -protected def commMagma [CommMagma M₂] (f : M₁ → M₂) (hf : Injective f) +protected abbrev commMagma [CommMagma M₂] (f : M₁ → M₂) (hf : Injective f) (mul : ∀ x y, f (x * y) = f x * f y) : CommMagma M₁ where mul_comm x y := hf <| by rw [mul, mul, mul_comm] /-- A type endowed with `*` is a commutative semigroup, if it admits an injective map that preserves `*` to a commutative semigroup. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `+` is an additive commutative semigroup,if it admits an injective map that preserves `+` to an additive commutative semigroup."] -protected def commSemigroup [CommSemigroup M₂] (f : M₁ → M₂) (hf : Injective f) +protected abbrev commSemigroup [CommSemigroup M₂] (f : M₁ → M₂) (hf : Injective f) (mul : ∀ x y, f (x * y) = f x * f y) : CommSemigroup M₁ where toSemigroup := hf.semigroup f mul __ := hf.commMagma f mul /-- A type endowed with `*` is a left cancel semigroup, if it admits an injective map that preserves `*` to a left cancel semigroup. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) "A type endowed with `+` is an additive left cancel +@[to_additive "A type endowed with `+` is an additive left cancel semigroup, if it admits an injective map that preserves `+` to an additive left cancel semigroup."] -protected def leftCancelSemigroup [LeftCancelSemigroup M₂] (f : M₁ → M₂) (hf : Injective f) +protected abbrev leftCancelSemigroup [LeftCancelSemigroup M₂] (f : M₁ → M₂) (hf : Injective f) (mul : ∀ x y, f (x * y) = f x * f y) : LeftCancelSemigroup M₁ := { hf.semigroup f mul with mul_left_cancel := fun x y z H => hf <| (mul_right_inj (f x)).1 <| by erw [← mul, ← mul, H] } /-- A type endowed with `*` is a right cancel semigroup, if it admits an injective map that preserves `*` to a right cancel semigroup. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) "A type endowed with `+` is an additive right +@[to_additive "A type endowed with `+` is an additive right cancel semigroup, if it admits an injective map that preserves `+` to an additive right cancel semigroup."] -protected def rightCancelSemigroup [RightCancelSemigroup M₂] (f : M₁ → M₂) (hf : Injective f) +protected abbrev rightCancelSemigroup [RightCancelSemigroup M₂] (f : M₁ → M₂) (hf : Injective f) (mul : ∀ x y, f (x * y) = f x * f y) : RightCancelSemigroup M₁ := { hf.semigroup f mul with mul_right_cancel := fun x y z H => hf <| (mul_left_inj (f y)).1 <| by erw [← mul, ← mul, H] } @@ -99,10 +99,10 @@ variable [One M₁] /-- A type endowed with `1` and `*` is a `MulOneClass`, if it admits an injective map that preserves `1` and `*` to a `MulOneClass`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an `AddZeroClass`, if it admits an injective map that preserves `0` and `+` to an `AddZeroClass`."] -protected def mulOneClass [MulOneClass M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) +protected abbrev mulOneClass [MulOneClass M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) : MulOneClass M₁ := { ‹One M₁›, ‹Mul M₁› with one_mul := fun x => hf <| by erw [mul, one, one_mul], @@ -112,11 +112,11 @@ variable [Pow M₁ ℕ] /-- A type endowed with `1` and `*` is a monoid, if it admits an injective map that preserves `1` and `*` to a monoid. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive monoid, if it admits an injective map that preserves `0` and `+` to an additive monoid. See note [reducible non-instances]."] -protected def monoid [Monoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) +protected abbrev monoid [Monoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : Monoid M₁ := { hf.mulOneClass f one mul, hf.semigroup f mul with npow := fun n x => x ^ n, @@ -137,40 +137,40 @@ protected abbrev addMonoidWithOne {M₁} [Zero M₁] [One M₁] [Add M₁] [SMul /-- A type endowed with `1` and `*` is a left cancel monoid, if it admits an injective map that preserves `1` and `*` to a left cancel monoid. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive left cancel monoid, if it admits an injective map that preserves `0` and `+` to an additive left cancel monoid."] -protected def leftCancelMonoid [LeftCancelMonoid M₂] (f : M₁ → M₂) (hf : Injective f) +protected abbrev leftCancelMonoid [LeftCancelMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : LeftCancelMonoid M₁ := { hf.leftCancelSemigroup f mul, hf.monoid f one mul npow with } /-- A type endowed with `1` and `*` is a right cancel monoid, if it admits an injective map that preserves `1` and `*` to a right cancel monoid. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive left cancel monoid,if it admits an injective map that preserves `0` and `+` to an additive left cancel monoid."] -protected def rightCancelMonoid [RightCancelMonoid M₂] (f : M₁ → M₂) (hf : Injective f) +protected abbrev rightCancelMonoid [RightCancelMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : RightCancelMonoid M₁ := { hf.rightCancelSemigroup f mul, hf.monoid f one mul npow with } /-- A type endowed with `1` and `*` is a cancel monoid, if it admits an injective map that preserves `1` and `*` to a cancel monoid. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive left cancel monoid,if it admits an injective map that preserves `0` and `+` to an additive left cancel monoid."] -protected def cancelMonoid [CancelMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) +protected abbrev cancelMonoid [CancelMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : CancelMonoid M₁ := { hf.leftCancelMonoid f one mul npow, hf.rightCancelMonoid f one mul npow with } /-- A type endowed with `1` and `*` is a commutative monoid, if it admits an injective map that preserves `1` and `*` to a commutative monoid. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive commutative monoid, if it admits an injective map that preserves `0` and `+` to an additive commutative monoid."] -protected def commMonoid [CommMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) +protected abbrev commMonoid [CommMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : CommMonoid M₁ := { hf.monoid f one mul npow, hf.commSemigroup f mul with } @@ -187,20 +187,20 @@ protected abbrev addCommMonoidWithOne {M₁} [Zero M₁] [One M₁] [Add M₁] [ /-- A type endowed with `1` and `*` is a cancel commutative monoid, if it admits an injective map that preserves `1` and `*` to a cancel commutative monoid. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive cancel commutative monoid, if it admits an injective map that preserves `0` and `+` to an additive cancel commutative monoid."] -protected def cancelCommMonoid [CancelCommMonoid M₂] (f : M₁ → M₂) (hf : Injective f) +protected abbrev cancelCommMonoid [CancelCommMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : CancelCommMonoid M₁ := { hf.leftCancelSemigroup f mul, hf.commMonoid f one mul npow with } /-- A type has an involutive inversion if it admits a surjective map that preserves `⁻¹` to a type which has an involutive inversion. See note [reducible non-instances] -/ -@[to_additive (attr := reducible) +@[to_additive "A type has an involutive negation if it admits a surjective map that preserves `-` to a type which has an involutive negation."] -protected def involutiveInv {M₁ : Type*} [Inv M₁] [InvolutiveInv M₂] (f : M₁ → M₂) +protected abbrev involutiveInv {M₁ : Type*} [Inv M₁] [InvolutiveInv M₂] (f : M₁ → M₂) (hf : Injective f) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) : InvolutiveInv M₁ where inv := Inv.inv inv_inv x := hf <| by rw [inv, inv, inv_inv] @@ -209,10 +209,10 @@ variable [Inv M₁] /-- A type endowed with `1` and `⁻¹` is a `InvOneClass`, if it admits an injective map that preserves `1` and `⁻¹` to a `InvOneClass`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and unary `-` is an `NegZeroClass`, if it admits an injective map that preserves `0` and unary `-` to an `NegZeroClass`."] -protected def invOneClass [InvOneClass M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) +protected abbrev invOneClass [InvOneClass M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (inv : ∀ x, f (x⁻¹) = (f x)⁻¹) : InvOneClass M₁ := { ‹One M₁›, ‹Inv M₁› with inv_one := hf <| by erw [inv, one, inv_one] } @@ -221,12 +221,12 @@ variable [Div M₁] [Pow M₁ ℤ] /-- A type endowed with `1`, `*`, `⁻¹`, and `/` is a `DivInvMonoid` if it admits an injective map that preserves `1`, `*`, `⁻¹`, and `/` to a `DivInvMonoid`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) subNegMonoid +@[to_additive subNegMonoid "A type endowed with `0`, `+`, unary `-`, and binary `-` is a `SubNegMonoid` if it admits an injective map that preserves `0`, `+`, unary `-`, and binary `-` to a `SubNegMonoid`. This version takes custom `nsmul` and `zsmul` as `[SMul ℕ M₁]` and `[SMul ℤ M₁]` arguments."] -protected def divInvMonoid [DivInvMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) +protected abbrev divInvMonoid [DivInvMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : DivInvMonoid M₁ := @@ -240,25 +240,25 @@ protected def divInvMonoid [DivInvMonoid M₂] (f : M₁ → M₂) (hf : Injecti /-- A type endowed with `1`, `*`, `⁻¹`, and `/` is a `DivInvOneMonoid` if it admits an injective map that preserves `1`, `*`, `⁻¹`, and `/` to a `DivInvOneMonoid`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) subNegZeroMonoid +@[to_additive subNegZeroMonoid "A type endowed with `0`, `+`, unary `-`, and binary `-` is a `SubNegZeroMonoid` if it admits an injective map that preserves `0`, `+`, unary `-`, and binary `-` to a `SubNegZeroMonoid`. This version takes custom `nsmul` and `zsmul` as `[SMul ℕ M₁]` and `[SMul ℤ M₁]` arguments."] -protected def divInvOneMonoid [DivInvOneMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) - (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) +protected abbrev divInvOneMonoid [DivInvOneMonoid M₂] (f : M₁ → M₂) (hf : Injective f) + (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : DivInvOneMonoid M₁ := { hf.divInvMonoid f one mul inv div npow zpow, hf.invOneClass f one inv with } /-- A type endowed with `1`, `*`, `⁻¹`, and `/` is a `DivisionMonoid` if it admits an injective map that preserves `1`, `*`, `⁻¹`, and `/` to a `DivisionMonoid`. See note [reducible non-instances] -/ -@[to_additive (attr := reducible) subtractionMonoid +@[to_additive subtractionMonoid "A type endowed with `0`, `+`, unary `-`, and binary `-` is a `SubtractionMonoid` if it admits an injective map that preserves `0`, `+`, unary `-`, and binary `-` to a `SubtractionMonoid`. This version takes custom `nsmul` and `zsmul` as `[SMul ℕ M₁]` and `[SMul ℤ M₁]` arguments."] -protected def divisionMonoid [DivisionMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) +protected abbrev divisionMonoid [DivisionMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : DivisionMonoid M₁ := @@ -270,12 +270,12 @@ protected def divisionMonoid [DivisionMonoid M₂] (f : M₁ → M₂) (hf : Inj /-- A type endowed with `1`, `*`, `⁻¹`, and `/` is a `DivisionCommMonoid` if it admits an injective map that preserves `1`, `*`, `⁻¹`, and `/` to a `DivisionCommMonoid`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) subtractionCommMonoid +@[to_additive subtractionCommMonoid "A type endowed with `0`, `+`, unary `-`, and binary `-` is a `SubtractionCommMonoid` if it admits an injective map that preserves `0`, `+`, unary `-`, and binary `-` to a `SubtractionCommMonoid`. This version takes custom `nsmul` and `zsmul` as `[SMul ℕ M₁]` and `[SMul ℤ M₁]` arguments."] -protected def divisionCommMonoid [DivisionCommMonoid M₂] (f : M₁ → M₂) (hf : Injective f) +protected abbrev divisionCommMonoid [DivisionCommMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : DivisionCommMonoid M₁ := @@ -283,10 +283,10 @@ protected def divisionCommMonoid [DivisionCommMonoid M₂] (f : M₁ → M₂) ( /-- A type endowed with `1`, `*` and `⁻¹` is a group, if it admits an injective map that preserves `1`, `*` and `⁻¹` to a group. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive group, if it admits an injective map that preserves `0` and `+` to an additive group."] -protected def group [Group M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) +protected abbrev group [Group M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : Group M₁ := @@ -310,10 +310,10 @@ protected abbrev addGroupWithOne {M₁} [Zero M₁] [One M₁] [Add M₁] [SMul /-- A type endowed with `1`, `*` and `⁻¹` is a commutative group, if it admits an injective map that preserves `1`, `*` and `⁻¹` to a commutative group. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive commutative group, if it admits an injective map that preserves `0` and `+` to an additive commutative group."] -protected def commGroup [CommGroup M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) +protected abbrev commGroup [CommGroup M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : CommGroup M₁ := @@ -344,28 +344,28 @@ variable {M₁ : Type*} {M₂ : Type*} [Mul M₂] /-- A type endowed with `*` is a semigroup, if it admits a surjective map that preserves `*` from a semigroup. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `+` is an additive semigroup, if it admits a surjective map that preserves `+` from an additive semigroup."] -protected def semigroup [Semigroup M₁] (f : M₁ → M₂) (hf : Surjective f) +protected abbrev semigroup [Semigroup M₁] (f : M₁ → M₂) (hf : Surjective f) (mul : ∀ x y, f (x * y) = f x * f y) : Semigroup M₂ := { ‹Mul M₂› with mul_assoc := hf.forall₃.2 fun x y z => by simp only [← mul, mul_assoc] } /-- A type endowed with `*` is a commutative semigroup, if it admits a surjective map that preserves `*` from a commutative semigroup. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `+` is an additive commutative semigroup, if it admits a surjective map that preserves `+` from an additive commutative semigroup."] -protected def commMagma [CommMagma M₁] (f : M₁ → M₂) (hf : Surjective f) +protected abbrev commMagma [CommMagma M₁] (f : M₁ → M₂) (hf : Surjective f) (mul : ∀ x y, f (x * y) = f x * f y) : CommMagma M₂ where mul_comm := hf.forall₂.2 fun x y => by erw [← mul, ← mul, mul_comm] /-- A type endowed with `*` is a commutative semigroup, if it admits a surjective map that preserves `*` from a commutative semigroup. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `+` is an additive commutative semigroup, if it admits a surjective map that preserves `+` from an additive commutative semigroup."] -protected def commSemigroup [CommSemigroup M₁] (f : M₁ → M₂) (hf : Surjective f) +protected abbrev commSemigroup [CommSemigroup M₁] (f : M₁ → M₂) (hf : Surjective f) (mul : ∀ x y, f (x * y) = f x * f y) : CommSemigroup M₂ where toSemigroup := hf.semigroup f mul __ := hf.commMagma f mul @@ -374,10 +374,10 @@ variable [One M₂] /-- A type endowed with `1` and `*` is a `MulOneClass`, if it admits a surjective map that preserves `1` and `*` from a `MulOneClass`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an `AddZeroClass`, if it admits a surjective map that preserves `0` and `+` to an `AddZeroClass`."] -protected def mulOneClass [MulOneClass M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) +protected abbrev mulOneClass [MulOneClass M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) : MulOneClass M₂ := { ‹One M₂›, ‹Mul M₂› with one_mul := hf.forall.2 fun x => by erw [← one, ← mul, one_mul], @@ -387,11 +387,11 @@ variable [Pow M₂ ℕ] /-- A type endowed with `1` and `*` is a monoid, if it admits a surjective map that preserves `1` and `*` to a monoid. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive monoid, if it admits a surjective map that preserves `0` and `+` to an additive monoid. This version takes a custom `nsmul` as a `[SMul ℕ M₂]` argument."] -protected def monoid [Monoid M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) +protected abbrev monoid [Monoid M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : Monoid M₂ := { hf.semigroup f mul, hf.mulOneClass f one mul with npow := fun n x => x ^ n, @@ -415,10 +415,10 @@ protected abbrev addMonoidWithOne {M₂} [Zero M₂] [One M₂] [Add M₂] [SMul /-- A type endowed with `1` and `*` is a commutative monoid, if it admits a surjective map that preserves `1` and `*` from a commutative monoid. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive commutative monoid, if it admits a surjective map that preserves `0` and `+` to an additive commutative monoid."] -protected def commMonoid [CommMonoid M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) +protected abbrev commMonoid [CommMonoid M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : CommMonoid M₂ := { hf.commSemigroup f mul, hf.monoid f one mul npow with } @@ -435,10 +435,10 @@ protected abbrev addCommMonoidWithOne {M₂} [Zero M₂] [One M₂] [Add M₂] [ /-- A type has an involutive inversion if it admits a surjective map that preserves `⁻¹` to a type which has an involutive inversion. See note [reducible non-instances] -/ -@[to_additive (attr := reducible) +@[to_additive "A type has an involutive negation if it admits a surjective map that preserves `-` to a type which has an involutive negation."] -protected def involutiveInv {M₂ : Type*} [Inv M₂] [InvolutiveInv M₁] (f : M₁ → M₂) +protected abbrev involutiveInv {M₂ : Type*} [Inv M₂] [InvolutiveInv M₁] (f : M₁ → M₂) (hf : Surjective f) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) : InvolutiveInv M₂ where inv := Inv.inv inv_inv := hf.forall.2 fun x => by erw [← inv, ← inv, inv_inv] @@ -447,11 +447,11 @@ variable [Inv M₂] [Div M₂] [Pow M₂ ℤ] /-- A type endowed with `1`, `*`, `⁻¹`, and `/` is a `DivInvMonoid` if it admits a surjective map that preserves `1`, `*`, `⁻¹`, and `/` to a `DivInvMonoid`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) subNegMonoid +@[to_additive subNegMonoid "A type endowed with `0`, `+`, unary `-`, and binary `-` is a `SubNegMonoid` if it admits a surjective map that preserves `0`, `+`, unary `-`, and binary `-` to a `SubNegMonoid`."] -protected def divInvMonoid [DivInvMonoid M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) +protected abbrev divInvMonoid [DivInvMonoid M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : DivInvMonoid M₂ := @@ -468,10 +468,10 @@ protected def divInvMonoid [DivInvMonoid M₁] (f : M₁ → M₂) (hf : Surject /-- A type endowed with `1`, `*` and `⁻¹` is a group, if it admits a surjective map that preserves `1`, `*` and `⁻¹` to a group. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive group, if it admits a surjective map that preserves `0` and `+` to an additive group."] -protected def group [Group M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) +protected abbrev group [Group M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : Group M₂ := @@ -497,10 +497,10 @@ protected abbrev addGroupWithOne {M₂} [Zero M₂] [One M₂] [Add M₂] [Neg M /-- A type endowed with `1`, `*`, `⁻¹`, and `/` is a commutative group, if it admits a surjective map that preserves `1`, `*`, `⁻¹`, and `/` from a commutative group. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "A type endowed with `0` and `+` is an additive commutative group, if it admits a surjective map that preserves `0` and `+` to an additive commutative group."] -protected def commGroup [CommGroup M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) +protected abbrev commGroup [CommGroup M₁] (f : M₁ → M₂) (hf : Surjective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : CommGroup M₂ := diff --git a/Mathlib/Algebra/Group/MinimalAxioms.lean b/Mathlib/Algebra/Group/MinimalAxioms.lean index b366185b4c143..3cf4026f974c3 100644 --- a/Mathlib/Algebra/Group/MinimalAxioms.lean +++ b/Mathlib/Algebra/Group/MinimalAxioms.lean @@ -30,12 +30,12 @@ universe u `∀ a, a⁻¹ * a = 1`. Note that this uses the default definitions for `npow`, `zpow` and `div`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "Define an `AddGroup` structure on a Type by proving `∀ a, 0 + a = a` and `∀ a, -a + a = 0`. Note that this uses the default definitions for `nsmul`, `zsmul` and `sub`. See note [reducible non-instances]."] -def Group.ofLeftAxioms {G : Type u} [Mul G] [Inv G] [One G] +abbrev Group.ofLeftAxioms {G : Type u} [Mul G] [Inv G] [One G] (assoc : ∀ a b c : G, (a * b) * c = a * (b * c)) (one_mul : ∀ a : G, 1 * a = a) (mul_left_inv : ∀ a : G, a⁻¹ * a = 1) : Group G := @@ -57,12 +57,12 @@ def Group.ofLeftAxioms {G : Type u} [Mul G] [Inv G] [One G] `∀ a, a * a⁻¹ = 1`. Note that this uses the default definitions for `npow`, `zpow` and `div`. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) +@[to_additive "Define an `AddGroup` structure on a Type by proving `∀ a, a + 0 = a` and `∀ a, a + -a = 0`. Note that this uses the default definitions for `nsmul`, `zsmul` and `sub`. See note [reducible non-instances]."] -def Group.ofRightAxioms {G : Type u} [Mul G] [Inv G] [One G] +abbrev Group.ofRightAxioms {G : Type u} [Mul G] [Inv G] [One G] (assoc : ∀ a b c : G, (a * b) * c = a * (b * c)) (mul_one : ∀ a : G, a * 1 = a) (mul_right_inv : ∀ a : G, a * a⁻¹ = 1) : Group G := diff --git a/Mathlib/Algebra/Order/Group/InjSurj.lean b/Mathlib/Algebra/Order/Group/InjSurj.lean index 7213721e5bd3a..256864357a1f8 100644 --- a/Mathlib/Algebra/Order/Group/InjSurj.lean +++ b/Mathlib/Algebra/Order/Group/InjSurj.lean @@ -15,8 +15,8 @@ variable {α β : Type*} /-- Pullback an `OrderedCommGroup` under an injective map. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) "Pullback an `OrderedAddCommGroup` under an injective map."] -def Function.Injective.orderedCommGroup [OrderedCommGroup α] {β : Type*} [One β] [Mul β] [Inv β] +@[to_additive "Pullback an `OrderedAddCommGroup` under an injective map."] +abbrev Function.Injective.orderedCommGroup [OrderedCommGroup α] {β : Type*} [One β] [Mul β] [Inv β] [Div β] [Pow β ℕ] [Pow β ℤ] (f : β → α) (hf : Function.Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) @@ -27,8 +27,8 @@ def Function.Injective.orderedCommGroup [OrderedCommGroup α] {β : Type*} [One /-- Pullback a `LinearOrderedCommGroup` under an injective map. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) "Pullback a `LinearOrderedAddCommGroup` under an injective map."] -def Function.Injective.linearOrderedCommGroup [LinearOrderedCommGroup α] {β : Type*} [One β] +@[to_additive "Pullback a `LinearOrderedAddCommGroup` under an injective map."] +abbrev Function.Injective.linearOrderedCommGroup [LinearOrderedCommGroup α] {β : Type*} [One β] [Mul β] [Inv β] [Div β] [Pow β ℕ] [Pow β ℤ] [Sup β] [Inf β] (f : β → α) (hf : Function.Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (inv : ∀ x, f x⁻¹ = (f x)⁻¹) (div : ∀ x y, f (x / y) = f x / f y) diff --git a/Mathlib/Algebra/Order/Monoid/Basic.lean b/Mathlib/Algebra/Order/Monoid/Basic.lean index b5234933d23b6..824c12d19dddf 100644 --- a/Mathlib/Algebra/Order/Monoid/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Basic.lean @@ -22,8 +22,8 @@ variable {α : Type u} {β : Type*} /-- Pullback an `OrderedCommMonoid` under an injective map. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) "Pullback an `OrderedAddCommMonoid` under an injective map."] -def Function.Injective.orderedCommMonoid [OrderedCommMonoid α] {β : Type*} [One β] [Mul β] +@[to_additive "Pullback an `OrderedAddCommMonoid` under an injective map."] +abbrev Function.Injective.orderedCommMonoid [OrderedCommMonoid α] {β : Type*} [One β] [Mul β] [Pow β ℕ] (f : β → α) (hf : Function.Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : OrderedCommMonoid β where @@ -34,10 +34,10 @@ def Function.Injective.orderedCommMonoid [OrderedCommMonoid α] {β : Type*} [On /-- Pullback an `OrderedCancelCommMonoid` under an injective map. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) Function.Injective.orderedCancelAddCommMonoid +@[to_additive Function.Injective.orderedCancelAddCommMonoid "Pullback an `OrderedCancelAddCommMonoid` under an injective map."] -def Function.Injective.orderedCancelCommMonoid [OrderedCancelCommMonoid α] [One β] [Mul β] [Pow β ℕ] - (f : β → α) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) +abbrev Function.Injective.orderedCancelCommMonoid [OrderedCancelCommMonoid α] [One β] [Mul β] + [Pow β ℕ] (f : β → α) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) : OrderedCancelCommMonoid β where toOrderedCommMonoid := hf.orderedCommMonoid f one mul npow le_of_mul_le_mul_left a b c (bc : f (a * b) ≤ f (a * c)) := @@ -45,8 +45,8 @@ def Function.Injective.orderedCancelCommMonoid [OrderedCancelCommMonoid α] [One /-- Pullback a `LinearOrderedCommMonoid` under an injective map. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) "Pullback an `OrderedAddCommMonoid` under an injective map."] -def Function.Injective.linearOrderedCommMonoid [LinearOrderedCommMonoid α] {β : Type*} [One β] +@[to_additive "Pullback an `OrderedAddCommMonoid` under an injective map."] +abbrev Function.Injective.linearOrderedCommMonoid [LinearOrderedCommMonoid α] {β : Type*} [One β] [Mul β] [Pow β ℕ] [Sup β] [Inf β] (f : β → α) (hf : Function.Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (sup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (inf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : @@ -56,9 +56,9 @@ def Function.Injective.linearOrderedCommMonoid [LinearOrderedCommMonoid α] {β /-- Pullback a `LinearOrderedCancelCommMonoid` under an injective map. See note [reducible non-instances]. -/ -@[to_additive (attr := reducible) Function.Injective.linearOrderedCancelAddCommMonoid +@[to_additive Function.Injective.linearOrderedCancelAddCommMonoid "Pullback a `LinearOrderedCancelAddCommMonoid` under an injective map."] -def Function.Injective.linearOrderedCancelCommMonoid [LinearOrderedCancelCommMonoid α] [One β] +abbrev Function.Injective.linearOrderedCancelCommMonoid [LinearOrderedCancelCommMonoid α] [One β] [Mul β] [Pow β ℕ] [Sup β] [Inf β] (f : β → α) (hf : Injective f) (one : f 1 = 1) (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index c5a2fba4dda85..5f3c32b121bf3 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -162,11 +162,11 @@ instance (priority := 100) NormedCommGroup.toNormedGroup [NormedCommGroup E] : N /-- Construct a `NormedGroup` from a `SeminormedGroup` satisfying `∀ x, ‖x‖ = 0 → x = 1`. This avoids having to go back to the `(Pseudo)MetricSpace` level when declaring a `NormedGroup` instance as a special case of a more general `SeminormedGroup` instance. -/ -@[to_additive (attr := reducible) "Construct a `NormedAddGroup` from a `SeminormedAddGroup` +@[to_additive "Construct a `NormedAddGroup` from a `SeminormedAddGroup` satisfying `∀ x, ‖x‖ = 0 → x = 0`. This avoids having to go back to the `(Pseudo)MetricSpace` level when declaring a `NormedAddGroup` instance as a special case of a more general `SeminormedAddGroup` instance."] -def NormedGroup.ofSeparation [SeminormedGroup E] (h : ∀ x : E, ‖x‖ = 0 → x = 1) : +abbrev NormedGroup.ofSeparation [SeminormedGroup E] (h : ∀ x : E, ‖x‖ = 0 → x = 1) : NormedGroup E where dist_eq := ‹SeminormedGroup E›.dist_eq toMetricSpace := @@ -182,19 +182,19 @@ def NormedGroup.ofSeparation [SeminormedGroup E] (h : ∀ x : E, ‖x‖ = 0 → `∀ x, ‖x‖ = 0 → x = 1`. This avoids having to go back to the `(Pseudo)MetricSpace` level when declaring a `NormedCommGroup` instance as a special case of a more general `SeminormedCommGroup` instance. -/ -@[to_additive (attr := reducible) "Construct a `NormedAddCommGroup` from a +@[to_additive "Construct a `NormedAddCommGroup` from a `SeminormedAddCommGroup` satisfying `∀ x, ‖x‖ = 0 → x = 0`. This avoids having to go back to the `(Pseudo)MetricSpace` level when declaring a `NormedAddCommGroup` instance as a special case of a more general `SeminormedAddCommGroup` instance."] -def NormedCommGroup.ofSeparation [SeminormedCommGroup E] (h : ∀ x : E, ‖x‖ = 0 → x = 1) : +abbrev NormedCommGroup.ofSeparation [SeminormedCommGroup E] (h : ∀ x : E, ‖x‖ = 0 → x = 1) : NormedCommGroup E := { ‹SeminormedCommGroup E›, NormedGroup.ofSeparation h with } -- See note [reducible non-instances] /-- Construct a seminormed group from a multiplication-invariant distance. -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a seminormed group from a translation-invariant distance."] -def SeminormedGroup.ofMulDist [Norm E] [Group E] [PseudoMetricSpace E] +abbrev SeminormedGroup.ofMulDist [Norm E] [Group E] [PseudoMetricSpace E] (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : SeminormedGroup E where dist_eq x y := by @@ -204,9 +204,9 @@ def SeminormedGroup.ofMulDist [Norm E] [Group E] [PseudoMetricSpace E] -- See note [reducible non-instances] /-- Construct a seminormed group from a multiplication-invariant pseudodistance. -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a seminormed group from a translation-invariant pseudodistance."] -def SeminormedGroup.ofMulDist' [Norm E] [Group E] [PseudoMetricSpace E] +abbrev SeminormedGroup.ofMulDist' [Norm E] [Group E] [PseudoMetricSpace E] (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : SeminormedGroup E where dist_eq x y := by @@ -216,9 +216,9 @@ def SeminormedGroup.ofMulDist' [Norm E] [Group E] [PseudoMetricSpace E] -- See note [reducible non-instances] /-- Construct a seminormed group from a multiplication-invariant pseudodistance. -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a seminormed group from a translation-invariant pseudodistance."] -def SeminormedCommGroup.ofMulDist [Norm E] [CommGroup E] [PseudoMetricSpace E] +abbrev SeminormedCommGroup.ofMulDist [Norm E] [CommGroup E] [PseudoMetricSpace E] (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : SeminormedCommGroup E := { SeminormedGroup.ofMulDist h₁ h₂ with @@ -226,9 +226,9 @@ def SeminormedCommGroup.ofMulDist [Norm E] [CommGroup E] [PseudoMetricSpace E] -- See note [reducible non-instances] /-- Construct a seminormed group from a multiplication-invariant pseudodistance. -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a seminormed group from a translation-invariant pseudodistance."] -def SeminormedCommGroup.ofMulDist' [Norm E] [CommGroup E] [PseudoMetricSpace E] +abbrev SeminormedCommGroup.ofMulDist' [Norm E] [CommGroup E] [PseudoMetricSpace E] (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : SeminormedCommGroup E := { SeminormedGroup.ofMulDist' h₁ h₂ with @@ -236,27 +236,27 @@ def SeminormedCommGroup.ofMulDist' [Norm E] [CommGroup E] [PseudoMetricSpace E] -- See note [reducible non-instances] /-- Construct a normed group from a multiplication-invariant distance. -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a normed group from a translation-invariant distance."] -def NormedGroup.ofMulDist [Norm E] [Group E] [MetricSpace E] (h₁ : ∀ x : E, ‖x‖ = dist x 1) +abbrev NormedGroup.ofMulDist [Norm E] [Group E] [MetricSpace E] (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : NormedGroup E := { SeminormedGroup.ofMulDist h₁ h₂ with eq_of_dist_eq_zero := eq_of_dist_eq_zero } -- See note [reducible non-instances] /-- Construct a normed group from a multiplication-invariant pseudodistance. -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a normed group from a translation-invariant pseudodistance."] -def NormedGroup.ofMulDist' [Norm E] [Group E] [MetricSpace E] (h₁ : ∀ x : E, ‖x‖ = dist x 1) +abbrev NormedGroup.ofMulDist' [Norm E] [Group E] [MetricSpace E] (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : NormedGroup E := { SeminormedGroup.ofMulDist' h₁ h₂ with eq_of_dist_eq_zero := eq_of_dist_eq_zero } -- See note [reducible non-instances] /-- Construct a normed group from a multiplication-invariant pseudodistance. -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a normed group from a translation-invariant pseudodistance."] -def NormedCommGroup.ofMulDist [Norm E] [CommGroup E] [MetricSpace E] +abbrev NormedCommGroup.ofMulDist [Norm E] [CommGroup E] [MetricSpace E] (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : NormedCommGroup E := { NormedGroup.ofMulDist h₁ h₂ with @@ -264,9 +264,9 @@ def NormedCommGroup.ofMulDist [Norm E] [CommGroup E] [MetricSpace E] -- See note [reducible non-instances] /-- Construct a normed group from a multiplication-invariant pseudodistance. -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a normed group from a translation-invariant pseudodistance."] -def NormedCommGroup.ofMulDist' [Norm E] [CommGroup E] [MetricSpace E] +abbrev NormedCommGroup.ofMulDist' [Norm E] [CommGroup E] [MetricSpace E] (h₁ : ∀ x : E, ‖x‖ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : NormedCommGroup E := { NormedGroup.ofMulDist' h₁ h₂ with @@ -277,12 +277,12 @@ def NormedCommGroup.ofMulDist' [Norm E] [CommGroup E] [MetricSpace E] pseudometric space structure from the seminorm properties. Note that in most cases this instance creates bad definitional equalities (e.g., it does not take into account a possibly existing `UniformSpace` instance on `E`). -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a seminormed group from a seminorm, i.e., registering the pseudodistance and the pseudometric space structure from the seminorm properties. Note that in most cases this instance creates bad definitional equalities (e.g., it does not take into account a possibly existing `UniformSpace` instance on `E`)."] -def GroupSeminorm.toSeminormedGroup [Group E] (f : GroupSeminorm E) : SeminormedGroup E where +abbrev GroupSeminorm.toSeminormedGroup [Group E] (f : GroupSeminorm E) : SeminormedGroup E where dist x y := f (x / y) norm := f dist_eq x y := rfl @@ -297,12 +297,12 @@ def GroupSeminorm.toSeminormedGroup [Group E] (f : GroupSeminorm E) : Seminormed pseudometric space structure from the seminorm properties. Note that in most cases this instance creates bad definitional equalities (e.g., it does not take into account a possibly existing `UniformSpace` instance on `E`). -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a seminormed group from a seminorm, i.e., registering the pseudodistance and the pseudometric space structure from the seminorm properties. Note that in most cases this instance creates bad definitional equalities (e.g., it does not take into account a possibly existing `UniformSpace` instance on `E`)."] -def GroupSeminorm.toSeminormedCommGroup [CommGroup E] (f : GroupSeminorm E) : +abbrev GroupSeminorm.toSeminormedCommGroup [CommGroup E] (f : GroupSeminorm E) : SeminormedCommGroup E := { f.toSeminormedGroup with mul_comm := mul_comm } @@ -312,12 +312,12 @@ def GroupSeminorm.toSeminormedCommGroup [CommGroup E] (f : GroupSeminorm E) : structure from the norm properties. Note that in most cases this instance creates bad definitional equalities (e.g., it does not take into account a possibly existing `UniformSpace` instance on `E`). -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a normed group from a norm, i.e., registering the distance and the metric space structure from the norm properties. Note that in most cases this instance creates bad definitional equalities (e.g., it does not take into account a possibly existing `UniformSpace` instance on `E`)."] -def GroupNorm.toNormedGroup [Group E] (f : GroupNorm E) : NormedGroup E := +abbrev GroupNorm.toNormedGroup [Group E] (f : GroupNorm E) : NormedGroup E := { f.toGroupSeminorm.toSeminormedGroup with eq_of_dist_eq_zero := fun h => div_eq_one.1 <| eq_one_of_map_eq_zero f h } @@ -326,12 +326,12 @@ def GroupNorm.toNormedGroup [Group E] (f : GroupNorm E) : NormedGroup E := structure from the norm properties. Note that in most cases this instance creates bad definitional equalities (e.g., it does not take into account a possibly existing `UniformSpace` instance on `E`). -/ -@[to_additive (attr := reducible) +@[to_additive "Construct a normed group from a norm, i.e., registering the distance and the metric space structure from the norm properties. Note that in most cases this instance creates bad definitional equalities (e.g., it does not take into account a possibly existing `UniformSpace` instance on `E`)."] -def GroupNorm.toNormedCommGroup [CommGroup E] (f : GroupNorm E) : NormedCommGroup E := +abbrev GroupNorm.toNormedCommGroup [CommGroup E] (f : GroupNorm E) : NormedCommGroup E := { f.toNormedGroup with mul_comm := mul_comm } @@ -902,9 +902,9 @@ variable [FunLike 𝓕 E F] -- See note [reducible non-instances] /-- A group homomorphism from a `Group` to a `SeminormedGroup` induces a `SeminormedGroup` structure on the domain. -/ -@[to_additive (attr := reducible) "A group homomorphism from an `AddGroup` to a +@[to_additive "A group homomorphism from an `AddGroup` to a `SeminormedAddGroup` induces a `SeminormedAddGroup` structure on the domain."] -def SeminormedGroup.induced [Group E] [SeminormedGroup F] [MonoidHomClass 𝓕 E F] (f : 𝓕) : +abbrev SeminormedGroup.induced [Group E] [SeminormedGroup F] [MonoidHomClass 𝓕 E F] (f : 𝓕) : SeminormedGroup E := { PseudoMetricSpace.induced f toPseudoMetricSpace with -- Porting note: needed to add the instance explicitly, and `‹PseudoMetricSpace F›` failed @@ -914,9 +914,9 @@ def SeminormedGroup.induced [Group E] [SeminormedGroup F] [MonoidHomClass 𝓕 E -- See note [reducible non-instances] /-- A group homomorphism from a `CommGroup` to a `SeminormedGroup` induces a `SeminormedCommGroup` structure on the domain. -/ -@[to_additive (attr := reducible) "A group homomorphism from an `AddCommGroup` to a +@[to_additive "A group homomorphism from an `AddCommGroup` to a `SeminormedAddGroup` induces a `SeminormedAddCommGroup` structure on the domain."] -def SeminormedCommGroup.induced +abbrev SeminormedCommGroup.induced [CommGroup E] [SeminormedGroup F] [MonoidHomClass 𝓕 E F] (f : 𝓕) : SeminormedCommGroup E := { SeminormedGroup.induced E F f with @@ -925,9 +925,9 @@ def SeminormedCommGroup.induced -- See note [reducible non-instances]. /-- An injective group homomorphism from a `Group` to a `NormedGroup` induces a `NormedGroup` structure on the domain. -/ -@[to_additive (attr := reducible) "An injective group homomorphism from an `AddGroup` to a +@[to_additive "An injective group homomorphism from an `AddGroup` to a `NormedAddGroup` induces a `NormedAddGroup` structure on the domain."] -def NormedGroup.induced +abbrev NormedGroup.induced [Group E] [NormedGroup F] [MonoidHomClass 𝓕 E F] (f : 𝓕) (h : Injective f) : NormedGroup E := { SeminormedGroup.induced E F f, MetricSpace.induced f h _ with } @@ -935,9 +935,9 @@ def NormedGroup.induced -- See note [reducible non-instances]. /-- An injective group homomorphism from a `CommGroup` to a `NormedGroup` induces a `NormedCommGroup` structure on the domain. -/ -@[to_additive (attr := reducible) "An injective group homomorphism from a `CommGroup` to a +@[to_additive "An injective group homomorphism from a `CommGroup` to a `NormedCommGroup` induces a `NormedCommGroup` structure on the domain."] -def NormedCommGroup.induced [CommGroup E] [NormedGroup F] [MonoidHomClass 𝓕 E F] (f : 𝓕) +abbrev NormedCommGroup.induced [CommGroup E] [NormedGroup F] [MonoidHomClass 𝓕 E F] (f : 𝓕) (h : Injective f) : NormedCommGroup E := { SeminormedGroup.induced E F f, MetricSpace.induced f h _ with mul_comm := mul_comm } diff --git a/Mathlib/GroupTheory/EckmannHilton.lean b/Mathlib/GroupTheory/EckmannHilton.lean index 85be56414aa01..9c6b5b653e7f9 100644 --- a/Mathlib/GroupTheory/EckmannHilton.lean +++ b/Mathlib/GroupTheory/EckmannHilton.lean @@ -78,10 +78,10 @@ theorem mul_assoc : Std.Associative m₂ := /-- If a type carries a unital magma structure that distributes over a unital binary operation, then the magma structure is a commutative monoid. -/ -@[to_additive (attr := reducible) +@[to_additive "If a type carries a unital additive magma structure that distributes over a unital binary operation, then the additive magma structure is a commutative additive monoid."] -def commMonoid [h : MulOneClass X] +abbrev commMonoid [h : MulOneClass X] (distrib : ∀ a b c d, ((a * b) c * d) = (a c) * b d) : CommMonoid X := { h with mul_comm := (mul_comm h₁ MulOneClass.isUnital distrib).comm, @@ -89,10 +89,10 @@ def commMonoid [h : MulOneClass X] /-- If a type carries a group structure that distributes over a unital binary operation, then the group is commutative. -/ -@[to_additive (attr := reducible) +@[to_additive "If a type carries an additive group structure that distributes over a unital binary operation, then the additive group is commutative."] -def commGroup [G : Group X] +abbrev commGroup [G : Group X] (distrib : ∀ a b c d, ((a * b) c * d) = (a c) * b d) : CommGroup X := { EckmannHilton.commMonoid h₁ distrib, G with .. } diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index dcdd5fdff4dee..801c4d8d5a2c0 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -615,8 +615,8 @@ lemma mul_comm_of_exponent_two [IsCancelMul G] (hG : Monoid.exponent G = 2) (a b Commute.of_orderOf_dvd_two (fun g => hG ▸ Monoid.order_dvd_exponent g) a b /-- Any cancellative monoid of exponent two is abelian. -/ -@[to_additive (attr := reducible) "Any additive group of exponent two is abelian."] -def commMonoidOfExponentTwo [IsCancelMul G] (hG : Monoid.exponent G = 2) : CommMonoid G where +@[to_additive "Any additive group of exponent two is abelian."] +abbrev commMonoidOfExponentTwo [IsCancelMul G] (hG : Monoid.exponent G = 2) : CommMonoid G where mul_comm := mul_comm_of_exponent_two hG end Monoid diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index a9941862041dd..02cadaba567a8 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -428,9 +428,9 @@ theorem image_inter_image_iff (U V : Set α) : variable (G α) /-- The quotient by `MulAction.orbitRel`, given a name to enable dot notation. -/ -@[to_additive (attr := reducible) +@[to_additive "The quotient by `AddAction.orbitRel`, given a name to enable dot notation."] -def orbitRel.Quotient : Type _ := +abbrev orbitRel.Quotient : Type _ := _root_.Quotient <| orbitRel G α /-- An action is pretransitive if and only if the quotient by `MulAction.orbitRel` is a diff --git a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean index 8089e5d60ce04..fee25caf6f462 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean @@ -1003,10 +1003,10 @@ variable (x : M) /-- Given `x : M`, the type of `CommMonoid` homomorphisms `f : M →* N` such that `N` is isomorphic to the Localization of `M` at the Submonoid generated by `x`. -/ -@[to_additive (attr := reducible) +@[to_additive "Given `x : M`, the type of `AddCommMonoid` homomorphisms `f : M →+ N` such that `N` is isomorphic to the localization of `M` at the AddSubmonoid generated by `x`."] -def AwayMap (N' : Type*) [CommMonoid N'] := LocalizationMap (powers x) N' +abbrev AwayMap (N' : Type*) [CommMonoid N'] := LocalizationMap (powers x) N' variable (F : AwayMap x N) @@ -1389,9 +1389,9 @@ section Away variable (x : M) /-- Given `x : M`, the Localization of `M` at the Submonoid generated by `x`, as a quotient. -/ -@[to_additive (attr := reducible) +@[to_additive "Given `x : M`, the Localization of `M` at the Submonoid generated by `x`, as a quotient."] -def Away := +abbrev Away := Localization (Submonoid.powers x) /-- Given `x : M`, `invSelf` is `x⁻¹` in the Localization (as a quotient type) of `M` at the @@ -1404,10 +1404,10 @@ def Away.invSelf : Away x := /-- Given `x : M`, the natural hom sending `y : M`, `M` a `CommMonoid`, to the equivalence class of `(y, 1)` in the Localization of `M` at the Submonoid generated by `x`. -/ -@[to_additive (attr := reducible) +@[to_additive "Given `x : M`, the natural hom sending `y : M`, `M` an `AddCommMonoid`, to the equivalence class of `(y, 0)` in the Localization of `M` at the Submonoid generated by `x`."] -def Away.monoidOf : Submonoid.LocalizationMap.AwayMap x (Away x) := +abbrev Away.monoidOf : Submonoid.LocalizationMap.AwayMap x (Away x) := Localization.monoidOf (Submonoid.powers x) -- @[simp] -- Porting note (#10618): simp can prove thisrove this diff --git a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean index 618817e7c30e8..b902f89c46909 100644 --- a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean +++ b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean @@ -151,9 +151,9 @@ def mulEquiv' (e : G₁ ≃ G₂) (he : e 1 = 1) (h : Monoid.exponent G₂ = 2) /-- Any two `IsKleinFour` groups are isomorphic via any equivalence which sends the identity of one group to the identity of the other. -/ -@[to_additive (attr := reducible) "Any two `IsAddKleinFour` groups are isomorphic via any +@[to_additive "Any two `IsAddKleinFour` groups are isomorphic via any equivalence which sends the identity of one group to the identity of the other."] -def mulEquiv [IsKleinFour G₂] (e : G₁ ≃ G₂) (he : e 1 = 1) : G₁ ≃* G₂ := +abbrev mulEquiv [IsKleinFour G₂] (e : G₁ ≃ G₂) (he : e 1 = 1) : G₁ ≃* G₂ := mulEquiv' e he exponent_two /-- Any two `IsKleinFour` groups are isomorphic. -/ diff --git a/Mathlib/Logic/Equiv/TransferInstance.lean b/Mathlib/Logic/Equiv/TransferInstance.lean index 81d0ee3a67923..8f6b53e102e82 100644 --- a/Mathlib/Logic/Equiv/TransferInstance.lean +++ b/Mathlib/Logic/Equiv/TransferInstance.lean @@ -40,8 +40,8 @@ section Instances variable (e : α ≃ β) /-- Transfer `One` across an `Equiv` -/ -@[to_additive (attr := reducible) "Transfer `Zero` across an `Equiv`"] -protected def one [One β] : One α := +@[to_additive "Transfer `Zero` across an `Equiv`"] +protected abbrev one [One β] : One α := ⟨e.symm 1⟩ @[to_additive] @@ -55,8 +55,8 @@ noncomputable instance [Small.{v} α] [One α] : One (Shrink.{v} α) := (equivShrink α).symm.one /-- Transfer `Mul` across an `Equiv` -/ -@[to_additive (attr := reducible) "Transfer `Add` across an `Equiv`"] -protected def mul [Mul β] : Mul α := +@[to_additive "Transfer `Add` across an `Equiv`"] +protected abbrev mul [Mul β] : Mul α := ⟨fun x y => e.symm (e x * e y)⟩ @[to_additive] @@ -70,8 +70,8 @@ noncomputable instance [Small.{v} α] [Mul α] : Mul (Shrink.{v} α) := (equivShrink α).symm.mul /-- Transfer `Div` across an `Equiv` -/ -@[to_additive (attr := reducible) "Transfer `Sub` across an `Equiv`"] -protected def div [Div β] : Div α := +@[to_additive "Transfer `Sub` across an `Equiv`"] +protected abbrev div [Div β] : Div α := ⟨fun x y => e.symm (e x / e y)⟩ @[to_additive] @@ -87,8 +87,8 @@ noncomputable instance [Small.{v} α] [Div α] : Div (Shrink.{v} α) := -- Porting note: this should be called `inv`, -- but we already have an `Equiv.inv` (which perhaps should move to `Perm.inv`?) /-- Transfer `Inv` across an `Equiv` -/ -@[to_additive (attr := reducible) "Transfer `Neg` across an `Equiv`"] -protected def Inv [Inv β] : Inv α := +@[to_additive "Transfer `Neg` across an `Equiv`"] +protected abbrev Inv [Inv β] : Inv α := ⟨fun x => e.symm (e x)⁻¹⟩ @[to_additive] @@ -190,8 +190,8 @@ noncomputable def _root_.Shrink.ringEquiv [Small.{v} α] [Ring α] : Shrink.{v} (equivShrink α).symm.ringEquiv /-- Transfer `Semigroup` across an `Equiv` -/ -@[to_additive (attr := reducible) "Transfer `add_semigroup` across an `Equiv`"] -protected def semigroup [Semigroup β] : Semigroup α := by +@[to_additive "Transfer `add_semigroup` across an `Equiv`"] +protected abbrev semigroup [Semigroup β] : Semigroup α := by let mul := e.mul apply e.injective.semigroup _; intros; exact e.apply_symm_apply _ @@ -210,8 +210,8 @@ noncomputable instance [Small.{v} α] [SemigroupWithZero α] : SemigroupWithZero (equivShrink α).symm.semigroupWithZero /-- Transfer `CommSemigroup` across an `Equiv` -/ -@[to_additive (attr := reducible) "Transfer `AddCommSemigroup` across an `Equiv`"] -protected def commSemigroup [CommSemigroup β] : CommSemigroup α := by +@[to_additive "Transfer `AddCommSemigroup` across an `Equiv`"] +protected abbrev commSemigroup [CommSemigroup β] : CommSemigroup α := by let mul := e.mul apply e.injective.commSemigroup _; intros; exact e.apply_symm_apply _ @@ -229,8 +229,8 @@ noncomputable instance [Small.{v} α] [MulZeroClass α] : MulZeroClass (Shrink.{ (equivShrink α).symm.mulZeroClass /-- Transfer `MulOneClass` across an `Equiv` -/ -@[to_additive (attr := reducible) "Transfer `AddZeroClass` across an `Equiv`"] -protected def mulOneClass [MulOneClass β] : MulOneClass α := by +@[to_additive "Transfer `AddZeroClass` across an `Equiv`"] +protected abbrev mulOneClass [MulOneClass β] : MulOneClass α := by let one := e.one let mul := e.mul apply e.injective.mulOneClass _ <;> intros <;> exact e.apply_symm_apply _ @@ -250,8 +250,8 @@ noncomputable instance [Small.{v} α] [MulZeroOneClass α] : MulZeroOneClass (Sh (equivShrink α).symm.mulZeroOneClass /-- Transfer `Monoid` across an `Equiv` -/ -@[to_additive (attr := reducible) "Transfer `AddMonoid` across an `Equiv`"] -protected def monoid [Monoid β] : Monoid α := by +@[to_additive "Transfer `AddMonoid` across an `Equiv`"] +protected abbrev monoid [Monoid β] : Monoid α := by let one := e.one let mul := e.mul let pow := e.pow ℕ @@ -262,8 +262,8 @@ noncomputable instance [Small.{v} α] [Monoid α] : Monoid (Shrink.{v} α) := (equivShrink α).symm.monoid /-- Transfer `CommMonoid` across an `Equiv` -/ -@[to_additive (attr := reducible) "Transfer `AddCommMonoid` across an `Equiv`"] -protected def commMonoid [CommMonoid β] : CommMonoid α := by +@[to_additive "Transfer `AddCommMonoid` across an `Equiv`"] +protected abbrev commMonoid [CommMonoid β] : CommMonoid α := by let one := e.one let mul := e.mul let pow := e.pow ℕ @@ -274,8 +274,8 @@ noncomputable instance [Small.{v} α] [CommMonoid α] : CommMonoid (Shrink.{v} (equivShrink α).symm.commMonoid /-- Transfer `Group` across an `Equiv` -/ -@[to_additive (attr := reducible) "Transfer `AddGroup` across an `Equiv`"] -protected def group [Group β] : Group α := by +@[to_additive "Transfer `AddGroup` across an `Equiv`"] +protected abbrev group [Group β] : Group α := by let one := e.one let mul := e.mul let inv := e.Inv @@ -289,8 +289,8 @@ noncomputable instance [Small.{v} α] [Group α] : Group (Shrink.{v} α) := (equivShrink α).symm.group /-- Transfer `CommGroup` across an `Equiv` -/ -@[to_additive (attr := reducible) "Transfer `AddCommGroup` across an `Equiv`"] -protected def commGroup [CommGroup β] : CommGroup α := by +@[to_additive "Transfer `AddCommGroup` across an `Equiv`"] +protected abbrev commGroup [CommGroup β] : CommGroup α := by let one := e.one let mul := e.mul let inv := e.Inv diff --git a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean index f14d89823fcd7..58a67caecf0c9 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean @@ -576,9 +576,9 @@ instance isHaarMeasure_haarMeasure (K₀ : PositiveCompacts G) : IsHaarMeasure ( · simp only [haarMeasure_self, ne_eq, ENNReal.one_ne_top, not_false_eq_true] /-- `haar` is some choice of a Haar measure, on a locally compact group. -/ -@[to_additive (attr := reducible) +@[to_additive "`addHaar` is some choice of a Haar measure, on a locally compact additive group."] -noncomputable def haar [LocallyCompactSpace G] : Measure G := +noncomputable abbrev haar [LocallyCompactSpace G] : Measure G := haarMeasure <| Classical.arbitrary _ /-! Steinhaus theorem: if `E` has positive measure, then `E / E` contains a neighborhood of zero. From 6e0804c6c717befb33fa6fb8b12a12e0c1ae6ea1 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 9 Aug 2024 14:59:48 +0000 Subject: [PATCH 152/359] chore: backports for leanprover/lean4#4814 (part 39) (#15636) --- Mathlib/Data/DFinsupp/Basic.lean | 50 +++++++++++++++++++---------- Mathlib/Data/DFinsupp/Notation.lean | 2 +- 2 files changed, 34 insertions(+), 18 deletions(-) diff --git a/Mathlib/Data/DFinsupp/Basic.lean b/Mathlib/Data/DFinsupp/Basic.lean index 2c12c922b8012..1715e4660a60f 100644 --- a/Mathlib/Data/DFinsupp/Basic.lean +++ b/Mathlib/Data/DFinsupp/Basic.lean @@ -479,8 +479,6 @@ theorem subtypeDomain_sub [∀ i, AddGroup (β i)] {p : ι → Prop} [DecidableP end FilterAndSubtypeDomain -variable [DecidableEq ι] - section Basic variable [∀ i, Zero (β i)] @@ -489,6 +487,9 @@ theorem finite_support (f : Π₀ i, β i) : Set.Finite { i | f i ≠ 0 } := Trunc.induction_on f.support' fun xs ↦ xs.1.finite_toSet.subset fun i H ↦ ((xs.prop i).resolve_right H) +section DecidableEq +variable [DecidableEq ι] + /-- Create an element of `Π₀ i, β i` from a finset `s` and a function `x` defined on this `Finset`. -/ def mk (s : Finset ι) (x : ∀ i : (↑s : Set ι), β (i : ι)) : Π₀ i, β i := @@ -515,6 +516,8 @@ theorem mk_injective (s : Finset ι) : Function.Injective (@mk ι β _ _ s) := b dsimp only [mk_apply, Subtype.coe_mk] at h1 simpa only [dif_pos hi] using h1 +end DecidableEq + instance unique [∀ i, Subsingleton (β i)] : Unique (Π₀ i, β i) := DFunLike.coe_injective.unique @@ -534,6 +537,8 @@ def equivFunOnFintype [Fintype ι] : (Π₀ i, β i) ≃ ∀ i, β i where theorem equivFunOnFintype_symm_coe [Fintype ι] (f : Π₀ i, β i) : equivFunOnFintype.symm f = f := Equiv.symm_apply_apply _ _ +variable [DecidableEq ι] + /-- The function `single i b : Π₀ i, β i` sends `i` to `b` and all other points to `0`. -/ def single (i : ι) (b : β i) : Π₀ i, β i := @@ -757,6 +762,9 @@ end Update end Basic +section DecidableEq +variable [DecidableEq ι] + section AddMonoid variable [∀ i, AddZeroClass (β i)] @@ -1154,6 +1162,8 @@ instance [∀ i, Zero (β i)] [∀ i, DecidableEq (β i)] : DecidableEq (Π₀ i rw [hf, hg], by rintro rfl; simp⟩ +end DecidableEq + section Equiv open Finset @@ -1194,8 +1204,9 @@ theorem comapDomain_smul [Monoid γ] [∀ i, AddMonoid (β i)] [∀ i, DistribMu rw [smul_apply, comapDomain_apply, smul_apply, comapDomain_apply] @[simp] -theorem comapDomain_single [DecidableEq κ] [∀ i, Zero (β i)] (h : κ → ι) (hh : Function.Injective h) - (k : κ) (x : β (h k)) : comapDomain h hh (single (h k) x) = single k x := by +theorem comapDomain_single [DecidableEq ι] [DecidableEq κ] [∀ i, Zero (β i)] (h : κ → ι) + (hh : Function.Injective h) (k : κ) (x : β (h k)) : + comapDomain h hh (single (h k) x) = single k x := by ext i rw [comapDomain_apply] obtain rfl | hik := Decidable.eq_or_ne i k @@ -1286,6 +1297,8 @@ instance distribMulAction₂ [Monoid γ] [∀ i j, AddMonoid (δ i j)] [∀ i j, DistribMulAction γ (δ i j)] : DistribMulAction γ (Π₀ (i : ι) (j : α i), δ i j) := @DFinsupp.distribMulAction ι _ (fun i => Π₀ j, δ i j) _ _ _ +variable [DecidableEq ι] + /-- The natural map between `Π₀ (i : Σ i, α i), δ i.1 i.2` and `Π₀ i (j : α i), δ i j`. -/ def sigmaCurry [∀ i j, Zero (δ i j)] (f : Π₀ (i : Σ _, _), δ i.1 i.2) : Π₀ (i) (j), δ i j where @@ -1363,33 +1376,32 @@ def sigmaUncurry [∀ i j, Zero (δ i j)] [DecidableEq ι] (f : Π₀ (i) (j), /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[simp] -theorem sigmaUncurry_apply [∀ i j, Zero (δ i j)] [DecidableEq ι] +theorem sigmaUncurry_apply [∀ i j, Zero (δ i j)] (f : Π₀ (i) (j), δ i j) (i : ι) (j : α i) : sigmaUncurry f ⟨i, j⟩ = f i j := rfl /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[simp] -theorem sigmaUncurry_zero [∀ i j, Zero (δ i j)] [DecidableEq ι] : +theorem sigmaUncurry_zero [∀ i j, Zero (δ i j)] : sigmaUncurry (0 : Π₀ (i) (j), δ i j) = 0 := rfl /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[simp] -theorem sigmaUncurry_add [∀ i j, AddZeroClass (δ i j)] [DecidableEq ι] - (f g : Π₀ (i) (j), δ i j) : +theorem sigmaUncurry_add [∀ i j, AddZeroClass (δ i j)] (f g : Π₀ (i) (j), δ i j) : sigmaUncurry (f + g) = sigmaUncurry f + sigmaUncurry g := DFunLike.coe_injective rfl /- ./././Mathport/Syntax/Translate/Expr.lean:107:6: warning: expanding binder group (i j) -/ @[simp] -theorem sigmaUncurry_smul [Monoid γ] [∀ i j, AddMonoid (δ i j)] [DecidableEq ι] +theorem sigmaUncurry_smul [Monoid γ] [∀ i j, AddMonoid (δ i j)] [∀ i j, DistribMulAction γ (δ i j)] (r : γ) (f : Π₀ (i) (j), δ i j) : sigmaUncurry (r • f) = r • sigmaUncurry f := DFunLike.coe_injective rfl @[simp] -theorem sigmaUncurry_single [∀ i j, Zero (δ i j)] [DecidableEq ι] [∀ i, DecidableEq (α i)] +theorem sigmaUncurry_single [∀ i j, Zero (δ i j)] [∀ i, DecidableEq (α i)] (i) (j : α i) (x : δ i j) : sigmaUncurry (single i (single j x : Π₀ j : α i, δ i j)) = single ⟨i, j⟩ (by exact x) := by ext ⟨i', j'⟩ @@ -1493,6 +1505,8 @@ end Equiv section ProdAndSum +variable [DecidableEq ι] + /-- `DFinsupp.prod f g` is the product of `g i (f i)` over the support of `f`. -/ @[to_additive "`sum f g` is the sum of `g i (f i)` over the support of `f`."] def prod [∀ i, Zero (β i)] [∀ (i) (x : β i), Decidable (x ≠ 0)] [CommMonoid γ] (f : Π₀ i, β i) @@ -1553,9 +1567,10 @@ theorem prod_comm {ι₁ ι₂ : Sort _} {β₁ : ι₁ → Type*} {β₂ : ι Finset.prod_comm @[simp] -theorem sum_apply {ι₁ : Type u₁} [DecidableEq ι₁] {β₁ : ι₁ → Type v₁} [∀ i₁, Zero (β₁ i₁)] - [∀ (i) (x : β₁ i), Decidable (x ≠ 0)] [∀ i, AddCommMonoid (β i)] {f : Π₀ i₁, β₁ i₁} - {g : ∀ i₁, β₁ i₁ → Π₀ i, β i} {i₂ : ι} : (f.sum g) i₂ = f.sum fun i₁ b => g i₁ b i₂ := +theorem sum_apply {ι} {β : ι → Type v} {ι₁ : Type u₁} [DecidableEq ι₁] {β₁ : ι₁ → Type v₁} + [∀ i₁, Zero (β₁ i₁)] [∀ (i) (x : β₁ i), Decidable (x ≠ 0)] [∀ i, AddCommMonoid (β i)] + {f : Π₀ i₁, β₁ i₁} {g : ∀ i₁, β₁ i₁ → Π₀ i, β i} {i₂ : ι} : + (f.sum g) i₂ = f.sum fun i₁ b => g i₁ b i₂ := map_sum (evalAddMonoidHom i₂) _ f.support theorem support_sum {ι₁ : Type u₁} [DecidableEq ι₁] {β₁ : ι₁ → Type v₁} [∀ i₁, Zero (β₁ i₁)] @@ -1874,13 +1889,14 @@ theorem prod_subtypeDomain_index [∀ i, Zero (β i)] [∀ (i) (x : β i), Decid (hp : ∀ x ∈ v.support, p x) : (v.subtypeDomain p).prod (fun i b => h i b) = v.prod h := by refine Finset.prod_bij (fun p _ ↦ p) ?_ ?_ ?_ ?_ <;> aesop -theorem subtypeDomain_sum [∀ i, AddCommMonoid (β i)] {s : Finset γ} {h : γ → Π₀ i, β i} - {p : ι → Prop} [DecidablePred p] : +theorem subtypeDomain_sum {ι} {β : ι → Type v} [∀ i, AddCommMonoid (β i)] {s : Finset γ} + {h : γ → Π₀ i, β i} {p : ι → Prop} [DecidablePred p] : (∑ c ∈ s, h c).subtypeDomain p = ∑ c ∈ s, (h c).subtypeDomain p := map_sum (subtypeDomainAddMonoidHom β p) _ s -theorem subtypeDomain_finsupp_sum {δ : γ → Type x} [DecidableEq γ] [∀ c, Zero (δ c)] - [∀ (c) (x : δ c), Decidable (x ≠ 0)] [∀ i, AddCommMonoid (β i)] {p : ι → Prop} [DecidablePred p] +theorem subtypeDomain_finsupp_sum {ι} {β : ι → Type v} {δ : γ → Type x} [DecidableEq γ] + [∀ c, Zero (δ c)] [∀ (c) (x : δ c), Decidable (x ≠ 0)] + [∀ i, AddCommMonoid (β i)] {p : ι → Prop} [DecidablePred p] {s : Π₀ c, δ c} {h : ∀ c, δ c → Π₀ i, β i} : (s.sum h).subtypeDomain p = s.sum fun c d => (h c d).subtypeDomain p := subtypeDomain_sum diff --git a/Mathlib/Data/DFinsupp/Notation.lean b/Mathlib/Data/DFinsupp/Notation.lean index c071bc38b5c71..7da5e6203d43a 100644 --- a/Mathlib/Data/DFinsupp/Notation.lean +++ b/Mathlib/Data/DFinsupp/Notation.lean @@ -42,7 +42,7 @@ def elabUpdate₀ : Elab.Term.TermElab Elab.Term.tryPostponeIfNoneOrMVar ty? let .some ty := ty? | Elab.throwUnsupportedSyntax let_expr DFinsupp _ _ _ := ← Meta.withReducible (Meta.whnf ty) | Elab.throwUnsupportedSyntax - Elab.Term.elabTerm (← `(DFinsupp.update $i $f $x)) ty? + Elab.Term.elabTerm (← `(DFinsupp.update $f $i $x)) ty? | _ => fun _ => Elab.throwUnsupportedSyntax /-- Unexpander for the `fun₀ | i => x` notation. -/ From 96a6efb931bacb29b366aa07146e1ac6a90b28d4 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Fri, 9 Aug 2024 15:52:59 +0000 Subject: [PATCH 153/359] feat(Analysis/SpecialFunctions/Complex/LogBounds): Add some log bounds (#15376) Add some basic bounds on complex logs that are useful for bounding infinite sums/products. Co-authored-by: Chris Birkbeck --- .../SpecialFunctions/Complex/LogBounds.lean | 25 +++++++++++++++++++ 1 file changed, 25 insertions(+) diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean b/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean index 266cce7326149..111f476afb799 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean @@ -178,6 +178,31 @@ lemma norm_log_one_add_sub_self_le {z : ℂ} (hz : ‖z‖ < 1) : · simp [logTaylor_succ, logTaylor_zero, sub_eq_add_neg] · norm_num +lemma norm_log_one_add_le {z : ℂ} (hz : ‖z‖ < 1) : + ‖log (1 + z)‖ ≤ ‖z‖ ^ 2 * (1 - ‖z‖)⁻¹ / 2 + ‖z‖ := by + rw [Eq.symm (sub_add_cancel (log (1 + z)) z)] + apply le_trans (norm_add_le _ _) + exact add_le_add_right (Complex.norm_log_one_add_sub_self_le hz) ‖z‖ + +/--For `‖z‖ ≤ 1/2`, the complex logarithm is bounded by `(3/2) * ‖z‖`. -/ +lemma norm_log_one_add_half_le_self {z : ℂ} (hz : ‖z‖ ≤ 1/2) : ‖(log (1 + z))‖ ≤ (3/2) * ‖z‖ := by + apply le_trans (norm_log_one_add_le (lt_of_le_of_lt hz one_half_lt_one)) + have hz3 : (1 - ‖z‖)⁻¹ ≤ 2 := by + rw [inv_eq_one_div, div_le_iff] + · linarith + · linarith + have hz4 : ‖z‖^2 * (1 - ‖z‖)⁻¹ / 2 ≤ ‖z‖/2 * 2 / 2 := by + gcongr + · rw [inv_nonneg] + linarith + · rw [sq, div_eq_mul_one_div] + apply mul_le_mul (by simp only [norm_eq_abs, mul_one, le_refl]) + (by simpa only [norm_eq_abs, one_div] using hz) (norm_nonneg z) + (by simp only [norm_eq_abs, mul_one, apply_nonneg]) + simp only [isUnit_iff_ne_zero, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, + IsUnit.div_mul_cancel] at hz4 + linarith + /-- The difference of `log (1-z)⁻¹` and its `(n+1)`st Taylor polynomial can be bounded in terms of `‖z‖`. -/ lemma norm_log_one_sub_inv_add_logTaylor_neg_le (n : ℕ) {z : ℂ} (hz : ‖z‖ < 1) : From f66b157f546ac7b4703b1957abf198d770fa1ea6 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 9 Aug 2024 17:52:17 +0000 Subject: [PATCH 154/359] chore: cleaning up deprecations (#15637) --- Mathlib.lean | 3 ++- Mathlib/Data/String/Defs.lean | 4 ---- Mathlib/Deprecated/Aliases.lean | 17 +++++++++++++++++ Mathlib/{Data => Deprecated}/HashMap.lean | 0 4 files changed, 19 insertions(+), 5 deletions(-) create mode 100644 Mathlib/Deprecated/Aliases.lean rename Mathlib/{Data => Deprecated}/HashMap.lean (100%) diff --git a/Mathlib.lean b/Mathlib.lean index c6d9c3843e342..0fe959c964f95 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2158,7 +2158,6 @@ import Mathlib.Data.FunLike.Basic import Mathlib.Data.FunLike.Embedding import Mathlib.Data.FunLike.Equiv import Mathlib.Data.FunLike.Fintype -import Mathlib.Data.HashMap import Mathlib.Data.Holor import Mathlib.Data.Int.AbsoluteValue import Mathlib.Data.Int.Align @@ -2533,7 +2532,9 @@ import Mathlib.Data.ZMod.Module import Mathlib.Data.ZMod.Parity import Mathlib.Data.ZMod.Quotient import Mathlib.Data.ZMod.Units +import Mathlib.Deprecated.Aliases import Mathlib.Deprecated.Group +import Mathlib.Deprecated.HashMap import Mathlib.Deprecated.Ring import Mathlib.Deprecated.Subfield import Mathlib.Deprecated.Subgroup diff --git a/Mathlib/Data/String/Defs.lean b/Mathlib/Data/String/Defs.lean index 3e121b543fc6f..0df9b21527948 100644 --- a/Mathlib/Data/String/Defs.lean +++ b/Mathlib/Data/String/Defs.lean @@ -3,8 +3,6 @@ Copyright (c) 2019 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Keeley Hoek, Floris van Doorn, Chris Bailey -/ -import Batteries.Data.String.Basic -import Batteries.Tactic.Alias /-! # Definitions for `String` @@ -44,8 +42,6 @@ then reassembles the string by intercalating the separator token `c` over the ma def mapTokens (c : Char) (f : String → String) : String → String := intercalate (singleton c) ∘ List.map f ∘ (·.split (· = c)) -@[deprecated (since := "2024-06-04")] alias getRest := dropPrefix? - /-- Produce the head character from the string `s`, if `s` is not empty, otherwise `'A'`. -/ def head (s : String) : Char := s.iter.curr diff --git a/Mathlib/Deprecated/Aliases.lean b/Mathlib/Deprecated/Aliases.lean new file mode 100644 index 0000000000000..3ccfb1ced2c07 --- /dev/null +++ b/Mathlib/Deprecated/Aliases.lean @@ -0,0 +1,17 @@ +/- +Copyright (c) 2024 Kim Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison +-/ +import Batteries.Tactic.Alias +import Batteries.Data.String.Basic + +/-! +Deprecated aliases can be dumped here if they are no longer used in Mathlib, +to avoid needing their imports if they are otherwise unnecessary. +-/ +namespace String + +@[deprecated (since := "2024-06-04")] alias getRest := dropPrefix? + +end String diff --git a/Mathlib/Data/HashMap.lean b/Mathlib/Deprecated/HashMap.lean similarity index 100% rename from Mathlib/Data/HashMap.lean rename to Mathlib/Deprecated/HashMap.lean From 23c8f16ffb2ebdf9a6307a772bce6c682ade9c69 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Dupuis?= Date: Fri, 9 Aug 2024 19:53:23 +0000 Subject: [PATCH 155/359] feat(MeasureTheory): add `ae_restrict_of_forall_mem` (#15625) I found myself repeatedly needing this lemma. --- Mathlib/MeasureTheory/Measure/Restrict.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index a149c75366f44..09321be0ee129 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -561,6 +561,10 @@ theorem ae_restrict_mem₀ (hs : NullMeasurableSet s μ) : ∀ᵐ x ∂μ.restri theorem ae_restrict_mem (hs : MeasurableSet s) : ∀ᵐ x ∂μ.restrict s, x ∈ s := ae_restrict_mem₀ hs.nullMeasurableSet +theorem ae_restrict_of_forall_mem {μ : Measure α} {s : Set α} + (hs : MeasurableSet s) {p : α → Prop} (h : ∀ x ∈ s, p x) : ∀ᵐ (x : α) ∂μ.restrict s, p x := + (ae_restrict_mem hs).mono h + theorem ae_restrict_of_ae {s : Set α} {p : α → Prop} (h : ∀ᵐ x ∂μ, p x) : ∀ᵐ x ∂μ.restrict s, p x := h.filter_mono (ae_mono Measure.restrict_le_self) From 82bb2e0547f310182b9dd8d76cd3b80f07562d63 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 9 Aug 2024 20:24:59 +0000 Subject: [PATCH 156/359] chore: add a space after `notation3` for prettier printing (#15515) The change means that `notation3 ...` gets pretty-printed with a space after `notation3`, instead of collapsing the following whitespaces into nothing. Discovered in the process of writing a "pretty-formatted" linter. --- Mathlib/Mathport/Notation.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Mathport/Notation.lean b/Mathlib/Mathport/Notation.lean index 16ae3cad942e3..fd02609f63fab 100644 --- a/Mathlib/Mathport/Notation.lean +++ b/Mathlib/Mathport/Notation.lean @@ -458,8 +458,8 @@ This command can be used in mathlib4 but it has an uncertain future and was crea for backward compatibility. -/ elab (name := notation3) doc:(docComment)? attrs?:(Parser.Term.attributes)? attrKind:Term.attrKind - "notation3" prec?:(precedence)? name?:(namedName)? prio?:(namedPrio)? pp?:(prettyPrintOpt)? - ppSpace items:(notation3Item)+ " => " val:term : command => do + "notation3" prec?:(precedence)? name?:(namedName)? prio?:(namedPrio)? + pp?:(ppSpace prettyPrintOpt)? items:(ppSpace notation3Item)+ " => " val:term : command => do -- We use raw `Name`s for variables. This maps variable names back to the -- identifiers that appear in `items` let mut boundIdents : HashMap Name Ident := {} From d8159b91c85b2c530b1d8b8d0182f9d99207c9e2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 9 Aug 2024 22:59:12 +0000 Subject: [PATCH 157/359] feat: define ergodic actions (#14952) Also prove that a function is ergodic iff the corresponding `IterateMulAct` action is ergodic. --- Mathlib.lean | 1 + Mathlib/Dynamics/Ergodic/Action/Basic.lean | 96 +++++++++++++++++++ .../GroupTheory/GroupAction/IterateAct.lean | 14 ++- Mathlib/MeasureTheory/Group/Action.lean | 19 +++- Mathlib/MeasureTheory/Group/Arithmetic.lean | 23 +++++ .../MeasurableSpace/Instances.lean | 13 ++- 6 files changed, 158 insertions(+), 8 deletions(-) create mode 100644 Mathlib/Dynamics/Ergodic/Action/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0fe959c964f95..9d01a48ae0c6a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2544,6 +2544,7 @@ import Mathlib.Dynamics.BirkhoffSum.Average import Mathlib.Dynamics.BirkhoffSum.Basic import Mathlib.Dynamics.BirkhoffSum.NormedSpace import Mathlib.Dynamics.Circle.RotationNumber.TranslationNumber +import Mathlib.Dynamics.Ergodic.Action.Basic import Mathlib.Dynamics.Ergodic.AddCircle import Mathlib.Dynamics.Ergodic.Conservative import Mathlib.Dynamics.Ergodic.Ergodic diff --git a/Mathlib/Dynamics/Ergodic/Action/Basic.lean b/Mathlib/Dynamics/Ergodic/Action/Basic.lean new file mode 100644 index 0000000000000..06740aa5a6bad --- /dev/null +++ b/Mathlib/Dynamics/Ergodic/Action/Basic.lean @@ -0,0 +1,96 @@ +/- +Copyright (c) 2024 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.MeasureTheory.Group.AEStabilizer +import Mathlib.Dynamics.Ergodic.Ergodic + +/-! +# Ergodic group actions + +A group action of `G` on a space `α` with measure `μ` is called *ergodic*, +if for any (null) measurable set `s`, +if it is a.e.-invariant under each scalar multiplication `(g • ·)`, `g : G`, +then it is either null or conull. +-/ + +open Set Filter MeasureTheory MulAction +open scoped Pointwise + +/-- +An additive group action of `G` on a space `α` with measure `μ` is called *ergodic*, +if for any (null) measurable set `s`, +if it is a.e.-invariant under each scalar addition `(g +ᵥ ·)`, `g : G`, +then it is either null or conull. +-/ +class ErgodicVAdd (G α : Type*) [VAdd G α] {_ : MeasurableSpace α} (μ : Measure α) + extends VAddInvariantMeasure G α μ : Prop where + aeconst_of_forall_preimage_vadd_ae_eq {s : Set α} : MeasurableSet s → + (∀ g : G, (g +ᵥ ·) ⁻¹' s =ᵐ[μ] s) → EventuallyConst s (ae μ) + +/-- +A group action of `G` on a space `α` with measure `μ` is called *ergodic*, +if for any (null) measurable set `s`, +if it is a.e.-invariant under each scalar multiplication `(g • ·)`, `g : G`, +then it is either null or conull. +-/ +@[to_additive, mk_iff] +class ErgodicSMul (G α : Type*) [SMul G α] {_ : MeasurableSpace α} (μ : Measure α) + extends SMulInvariantMeasure G α μ : Prop where + aeconst_of_forall_preimage_smul_ae_eq {s : Set α} : MeasurableSet s → + (∀ g : G, (g • ·) ⁻¹' s =ᵐ[μ] s) → EventuallyConst s (ae μ) + +attribute [to_additive] ergodicSMul_iff + +namespace MeasureTheory + +variable (G : Type*) {α : Type*} {m : MeasurableSpace α} {μ : Measure α} + +@[to_additive] +theorem aeconst_of_forall_preimage_smul_ae_eq [SMul G α] [ErgodicSMul G α μ] {s : Set α} + (hm : NullMeasurableSet s μ) (h : ∀ g : G, (g • ·) ⁻¹' s =ᵐ[μ] s) : + EventuallyConst s (ae μ) := by + rcases hm with ⟨t, htm, hst⟩ + refine .congr ?_ hst.symm + refine ErgodicSMul.aeconst_of_forall_preimage_smul_ae_eq htm fun g : G ↦ ?_ + refine .trans (.trans ?_ (h g)) hst + exact tendsto_smul_ae _ _ hst.symm + +section Group + +variable [Group G] [MulAction G α] [ErgodicSMul G α μ] {s : Set α} + +@[to_additive] +theorem aeconst_of_forall_smul_ae_eq (hm : NullMeasurableSet s μ) (h : ∀ g : G, g • s =ᵐ[μ] s) : + EventuallyConst s (ae μ) := + aeconst_of_forall_preimage_smul_ae_eq G hm fun g ↦ by + simpa only [preimage_smul] using h g⁻¹ + +@[to_additive] +theorem _root_.MulAction.aeconst_of_aestabilizer_eq_top + (hm : NullMeasurableSet s μ) (h : aestabilizer G μ s = ⊤) : EventuallyConst s (ae μ) := + aeconst_of_forall_smul_ae_eq G hm <| (Subgroup.eq_top_iff' _).1 h + +end Group + +theorem _root_.ErgodicSMul.of_aestabilizer [Group G] [MulAction G α] [SMulInvariantMeasure G α μ] + (h : ∀ s, MeasurableSet s → aestabilizer G μ s = ⊤ → EventuallyConst s (ae μ)) : + ErgodicSMul G α μ := + ⟨fun hm hs ↦ h _ hm <| (Subgroup.eq_top_iff' _).2 fun g ↦ by + simpa only [preimage_smul_inv] using hs g⁻¹⟩ + +theorem ergodicSMul_iterateMulAct {f : α → α} (hf : Measurable f) : + ErgodicSMul (IterateMulAct f) α μ ↔ Ergodic f μ := by + simp only [ergodicSMul_iff, smulInvariantMeasure_iterateMulAct, hf] + refine ⟨fun ⟨h₁, h₂⟩ ↦ ⟨h₁, ⟨?_⟩⟩, fun h ↦ ⟨h.1, ?_⟩⟩ + · intro s hm hs + rw [← eventuallyConst_set'] + refine h₂ hm fun n ↦ ?_ + nth_rewrite 2 [← Function.IsFixedPt.preimage_iterate hs n.val] + rfl + · intro s hm hs + rw [eventuallyConst_set'] + exact h.quasiErgodic.ae_empty_or_univ' hm <| hs (.mk 1) + +end MeasureTheory diff --git a/Mathlib/GroupTheory/GroupAction/IterateAct.lean b/Mathlib/GroupTheory/GroupAction/IterateAct.lean index 9259861a738b6..912fcd1b2bd0f 100644 --- a/Mathlib/GroupTheory/GroupAction/IterateAct.lean +++ b/Mathlib/GroupTheory/GroupAction/IterateAct.lean @@ -3,8 +3,8 @@ Copyright (c) 2024 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Algebra.Ring.Nat import Mathlib.Algebra.Group.Action.Defs +import Mathlib.Data.Countable.Defs /-! # Monoid action by iterates of a map @@ -31,16 +31,20 @@ namespace IterateMulAct variable {α : Type*} {f : α → α} +@[to_additive] +instance instCountable : Countable (IterateMulAct f) := + Function.Injective.countable fun _ _ ↦ IterateMulAct.ext + @[to_additive] instance instCommMonoid : CommMonoid (IterateMulAct f) where one := ⟨0⟩ mul m n := ⟨m.1 + n.1⟩ - mul_assoc a b c := by ext; apply add_assoc - one_mul _ := by ext; apply zero_add + mul_assoc a b c := by ext; apply Nat.add_assoc + one_mul _ := by ext; apply Nat.zero_add mul_one _ := rfl - mul_comm _ _ := by ext; apply add_comm + mul_comm _ _ := by ext; apply Nat.add_comm npow n a := ⟨n * a.val⟩ - npow_zero _ := by ext; apply zero_mul + npow_zero _ := by ext; apply Nat.zero_mul npow_succ n a := by ext; apply Nat.succ_mul @[to_additive] diff --git a/Mathlib/MeasureTheory/Group/Action.lean b/Mathlib/MeasureTheory/Group/Action.lean index 5623eac81131a..51654f035648f 100644 --- a/Mathlib/MeasureTheory/Group/Action.lean +++ b/Mathlib/MeasureTheory/Group/Action.lean @@ -39,11 +39,13 @@ class VAddInvariantMeasure (M α : Type*) [VAdd M α] {_ : MeasurableSpace α} ( /-- A measure `μ : Measure α` is invariant under a multiplicative action of `M` on `α` if for any measurable set `s : Set α` and `c : M`, the measure of its preimage under `fun x => c • x` is equal to the measure of `s`. -/ -@[to_additive] +@[to_additive, mk_iff smulInvariantMeasure_iff] class SMulInvariantMeasure (M α : Type*) [SMul M α] {_ : MeasurableSpace α} (μ : Measure α) : Prop where measure_preimage_smul : ∀ (c : M) ⦃s : Set α⦄, MeasurableSet s → μ ((fun x => c • x) ⁻¹' s) = μ s +attribute [to_additive] smulInvariantMeasure_iff + namespace SMulInvariantMeasure @[to_additive] @@ -167,6 +169,21 @@ theorem map_smul : map (c • ·) μ = μ := end MeasurableSMul +@[to_additive] +theorem MeasurePreserving.smulInvariantMeasure_iterateMulAct + {f : α → α} {_ : MeasurableSpace α} {μ : Measure α} (hf : MeasurePreserving f μ μ) : + SMulInvariantMeasure (IterateMulAct f) α μ := + ⟨fun n _s hs ↦ (hf.iterate n.val).measure_preimage hs.nullMeasurableSet⟩ + +@[to_additive] +theorem smulInvariantMeasure_iterateMulAct + {f : α → α} {_ : MeasurableSpace α} {μ : Measure α} (hf : Measurable f) : + SMulInvariantMeasure (IterateMulAct f) α μ ↔ MeasurePreserving f μ μ := + ⟨fun _ ↦ + have := hf.measurableSMul₂_iterateMulAct + measurePreserving_smul (IterateMulAct.mk (f := f) 1) μ, + MeasurePreserving.smulInvariantMeasure_iterateMulAct⟩ + section SMulHomClass universe uM uN uα uβ diff --git a/Mathlib/MeasureTheory/Group/Arithmetic.lean b/Mathlib/MeasureTheory/Group/Arithmetic.lean index 16333e859b450..6fd334466abb6 100644 --- a/Mathlib/MeasureTheory/Group/Arithmetic.lean +++ b/Mathlib/MeasureTheory/Group/Arithmetic.lean @@ -606,6 +606,29 @@ instance SubNegMonoid.measurableSMul_int₂ (M : Type*) [SubNegMonoid M] [Measur end SMul +section IterateMulAct + +variable {α : Type*} {_ : MeasurableSpace α} {f : α → α} + +@[to_additive] +theorem Measurable.measurableSMul₂_iterateMulAct (h : Measurable f) : + MeasurableSMul₂ (IterateMulAct f) α where + measurable_smul := + suffices Measurable fun p : α × IterateMulAct f ↦ f^[p.2.val] p.1 from this.comp measurable_swap + measurable_from_prod_countable fun n ↦ h.iterate n.val + +@[to_additive (attr := simp)] +theorem measurableSMul_iterateMulAct : MeasurableSMul (IterateMulAct f) α ↔ Measurable f := + ⟨fun _ ↦ measurable_const_smul (IterateMulAct.mk (f := f) 1), fun h ↦ + have := h.measurableSMul₂_iterateMulAct; inferInstance⟩ + +@[to_additive (attr := simp)] +theorem measurableSMul₂_iterateMulAct : MeasurableSMul₂ (IterateMulAct f) α ↔ Measurable f := + ⟨fun _ ↦ measurableSMul_iterateMulAct.mp inferInstance, + Measurable.measurableSMul₂_iterateMulAct⟩ + +end IterateMulAct + section MulAction variable {M β α : Type*} [MeasurableSpace M] [MeasurableSpace β] [Monoid M] [MulAction M β] diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Instances.lean b/Mathlib/MeasureTheory/MeasurableSpace/Instances.lean index 0536130df662f..2de7f4465974c 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Instances.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Instances.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.MeasureTheory.MeasurableSpace.Defs +import Mathlib.GroupTheory.GroupAction.IterateAct import Mathlib.Data.Rat.Init import Mathlib.Data.ZMod.Defs @@ -32,8 +33,16 @@ instance Int.instMeasurableSpace : MeasurableSpace ℤ := ⊤ instance Rat.instMeasurableSpace : MeasurableSpace ℚ := ⊤ -instance Subsingleton.measurableSingletonClass {α} [MeasurableSpace α] [Subsingleton α] : - MeasurableSingletonClass α := by +@[to_additive] +instance IterateMulAct.instMeasurableSpace {α : Type*} {f : α → α} : + MeasurableSpace (IterateMulAct f) := ⊤ + +@[to_additive] +instance IterateMulAct.instDiscreteMeasurableSpace {α : Type*} {f : α → α} : + DiscreteMeasurableSpace (IterateMulAct f) := inferInstance + +instance (priority := 100) Subsingleton.measurableSingletonClass + {α} [MeasurableSpace α] [Subsingleton α] : MeasurableSingletonClass α := by refine ⟨fun i => ?_⟩ convert MeasurableSet.univ simp [Set.eq_univ_iff_forall, eq_iff_true_of_subsingleton] From e15449d2b90e79bb184efc57eece4196818683f8 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 10 Aug 2024 03:36:09 +0000 Subject: [PATCH 158/359] chore: split LinearAlgebra.Basis (#15639) --- Mathlib.lean | 4 +- Mathlib/Algebra/DirectSum/Module.lean | 2 +- Mathlib/Algebra/Polynomial/Roots.lean | 1 + .../Normed/Operator/LinearIsometry.lean | 2 +- Mathlib/LinearAlgebra/AffineSpace/Basis.lean | 1 - Mathlib/LinearAlgebra/Alternating/Basic.lean | 1 + Mathlib/LinearAlgebra/Basis/Basic.lean | 563 +++++++++++++++ Mathlib/LinearAlgebra/Basis/Bilinear.lean | 2 +- Mathlib/LinearAlgebra/Basis/Cardinality.lean | 155 ++++ .../{Basis.lean => Basis/Defs.lean} | 662 +----------------- Mathlib/LinearAlgebra/Basis/Flag.lean | 2 +- Mathlib/LinearAlgebra/Basis/VectorSpace.lean | 1 - Mathlib/LinearAlgebra/BilinearForm/Hom.lean | 2 +- Mathlib/LinearAlgebra/FreeAlgebra.lean | 2 +- Mathlib/LinearAlgebra/FreeModule/Basic.lean | 1 + Mathlib/LinearAlgebra/Matrix/DotProduct.lean | 1 + Mathlib/LinearAlgebra/Multilinear/Basis.lean | 2 +- Mathlib/LinearAlgebra/SesquilinearForm.lean | 2 +- Mathlib/LinearAlgebra/StdBasis.lean | 3 +- Mathlib/RingTheory/Adjoin/Basic.lean | 3 +- Mathlib/RingTheory/AlgebraTower.lean | 2 +- Mathlib/RingTheory/Finiteness.lean | 1 + Mathlib/RingTheory/Ideal/Basis.lean | 2 +- Mathlib/RingTheory/Localization/Module.lean | 2 +- .../MvPolynomial/WeightedHomogeneous.lean | 1 + 25 files changed, 743 insertions(+), 677 deletions(-) create mode 100644 Mathlib/LinearAlgebra/Basis/Basic.lean create mode 100644 Mathlib/LinearAlgebra/Basis/Cardinality.lean rename Mathlib/LinearAlgebra/{Basis.lean => Basis/Defs.lean} (51%) diff --git a/Mathlib.lean b/Mathlib.lean index 9d01a48ae0c6a..6b53163fcfd2d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2841,8 +2841,10 @@ import Mathlib.LinearAlgebra.AffineSpace.Slope import Mathlib.LinearAlgebra.Alternating.Basic import Mathlib.LinearAlgebra.Alternating.DomCoprod import Mathlib.LinearAlgebra.AnnihilatingPolynomial -import Mathlib.LinearAlgebra.Basis +import Mathlib.LinearAlgebra.Basis.Basic import Mathlib.LinearAlgebra.Basis.Bilinear +import Mathlib.LinearAlgebra.Basis.Cardinality +import Mathlib.LinearAlgebra.Basis.Defs import Mathlib.LinearAlgebra.Basis.Flag import Mathlib.LinearAlgebra.Basis.VectorSpace import Mathlib.LinearAlgebra.BilinearForm.Basic diff --git a/Mathlib/Algebra/DirectSum/Module.lean b/Mathlib/Algebra/DirectSum/Module.lean index 69dfbc7dc4fe0..5d0336e5ed1e4 100644 --- a/Mathlib/Algebra/DirectSum/Module.lean +++ b/Mathlib/Algebra/DirectSum/Module.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau -/ import Mathlib.Algebra.DirectSum.Basic import Mathlib.LinearAlgebra.DFinsupp -import Mathlib.LinearAlgebra.Basis +import Mathlib.LinearAlgebra.Basis.Defs /-! # Direct sum of modules diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index cba7f3b9fdacf..06b3d6fd29716 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -6,6 +6,7 @@ Authors: Chris Hughes, Johannes Hölzl, Scott Morrison, Jens Wagemaker, Johan Co import Mathlib.Algebra.Polynomial.RingDivision import Mathlib.RingTheory.Localization.FractionRing +import Mathlib.SetTheory.Cardinal.Basic /-! # Theory of univariate polynomials diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index 6c8da81a65094..e6b3a063b3a61 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -8,7 +8,7 @@ import Mathlib.Analysis.Normed.Group.Constructions import Mathlib.Analysis.Normed.Group.Submodule import Mathlib.Analysis.Normed.Group.Uniform import Mathlib.Topology.Algebra.Module.Basic -import Mathlib.LinearAlgebra.Basis +import Mathlib.LinearAlgebra.Basis.Defs /-! # (Semi-)linear isometries diff --git a/Mathlib/LinearAlgebra/AffineSpace/Basis.lean b/Mathlib/LinearAlgebra/AffineSpace/Basis.lean index 7ce59eedb5493..b3c2e7715e81e 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Basis.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Basis.lean @@ -5,7 +5,6 @@ Authors: Oliver Nash -/ import Mathlib.LinearAlgebra.AffineSpace.Independent import Mathlib.LinearAlgebra.AffineSpace.Pointwise -import Mathlib.LinearAlgebra.Basis /-! # Affine bases and barycentric coordinates diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index d838fd0ac45a3..90b2803f45d8e 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -6,6 +6,7 @@ Authors: Eric Wieser, Zhangir Azerbayev import Mathlib.GroupTheory.Perm.Sign import Mathlib.Data.Fintype.Perm import Mathlib.LinearAlgebra.Multilinear.Basis +import Mathlib.LinearAlgebra.LinearIndependent /-! # Alternating Maps diff --git a/Mathlib/LinearAlgebra/Basis/Basic.lean b/Mathlib/LinearAlgebra/Basis/Basic.lean new file mode 100644 index 0000000000000..3ec32691a84b3 --- /dev/null +++ b/Mathlib/LinearAlgebra/Basis/Basic.lean @@ -0,0 +1,563 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro, Alexander Bentkamp +-/ +import Mathlib.LinearAlgebra.Basis.Defs +import Mathlib.Algebra.BigOperators.Finsupp +import Mathlib.Algebra.BigOperators.Finprod +import Mathlib.Data.Fintype.BigOperators +import Mathlib.LinearAlgebra.LinearIndependent + +/-! +# Bases + +Further results on bases in modules and vector spaces. + +-/ + + +noncomputable section + +universe u + +open Function Set Submodule + +variable {ι : Type*} {ι' : Type*} {R : Type*} {R₂ : Type*} {K : Type*} +variable {M : Type*} {M' M'' : Type*} {V : Type u} {V' : Type*} + +section Module + +variable [Semiring R] +variable [AddCommMonoid M] [Module R M] [AddCommMonoid M'] [Module R M'] + + +namespace Basis + +variable (b b₁ : Basis ι R M) (i : ι) (c : R) (x : M) + +section Coord + +theorem coe_sumCoords_eq_finsum : (b.sumCoords : M → R) = fun m => ∑ᶠ i, b.coord i m := by + ext m + simp only [Basis.sumCoords, Basis.coord, Finsupp.lapply_apply, LinearMap.id_coe, + LinearEquiv.coe_coe, Function.comp_apply, Finsupp.coe_lsum, LinearMap.coe_comp, + finsum_eq_sum _ (b.repr m).finite_support, Finsupp.sum, Finset.finite_toSet_toFinset, id, + Finsupp.fun_support_eq] + +end Coord + +protected theorem linearIndependent : LinearIndependent R b := + linearIndependent_iff.mpr fun l hl => + calc + l = b.repr (Finsupp.total _ _ _ b l) := (b.repr_total l).symm + _ = 0 := by rw [hl, LinearEquiv.map_zero] + +protected theorem ne_zero [Nontrivial R] (i) : b i ≠ 0 := + b.linearIndependent.ne_zero i + +section Prod + +variable (b' : Basis ι' R M') + +/-- `Basis.prod` maps an `ι`-indexed basis for `M` and an `ι'`-indexed basis for `M'` +to an `ι ⊕ ι'`-index basis for `M × M'`. +For the specific case of `R × R`, see also `Basis.finTwoProd`. -/ +protected def prod : Basis (ι ⊕ ι') R (M × M') := + ofRepr ((b.repr.prod b'.repr).trans (Finsupp.sumFinsuppLEquivProdFinsupp R).symm) + +@[simp] +theorem prod_repr_inl (x) (i) : (b.prod b').repr x (Sum.inl i) = b.repr x.1 i := + rfl + +@[simp] +theorem prod_repr_inr (x) (i) : (b.prod b').repr x (Sum.inr i) = b'.repr x.2 i := + rfl + +theorem prod_apply_inl_fst (i) : (b.prod b' (Sum.inl i)).1 = b i := + b.repr.injective <| by + ext j + simp only [Basis.prod, Basis.coe_ofRepr, LinearEquiv.symm_trans_apply, LinearEquiv.prod_symm, + LinearEquiv.prod_apply, b.repr.apply_symm_apply, LinearEquiv.symm_symm, repr_self, + Equiv.toFun_as_coe, Finsupp.fst_sumFinsuppLEquivProdFinsupp] + apply Finsupp.single_apply_left Sum.inl_injective + +theorem prod_apply_inr_fst (i) : (b.prod b' (Sum.inr i)).1 = 0 := + b.repr.injective <| by + ext i + simp only [Basis.prod, Basis.coe_ofRepr, LinearEquiv.symm_trans_apply, LinearEquiv.prod_symm, + LinearEquiv.prod_apply, b.repr.apply_symm_apply, LinearEquiv.symm_symm, repr_self, + Equiv.toFun_as_coe, Finsupp.fst_sumFinsuppLEquivProdFinsupp, LinearEquiv.map_zero, + Finsupp.zero_apply] + apply Finsupp.single_eq_of_ne Sum.inr_ne_inl + +theorem prod_apply_inl_snd (i) : (b.prod b' (Sum.inl i)).2 = 0 := + b'.repr.injective <| by + ext j + simp only [Basis.prod, Basis.coe_ofRepr, LinearEquiv.symm_trans_apply, LinearEquiv.prod_symm, + LinearEquiv.prod_apply, b'.repr.apply_symm_apply, LinearEquiv.symm_symm, repr_self, + Equiv.toFun_as_coe, Finsupp.snd_sumFinsuppLEquivProdFinsupp, LinearEquiv.map_zero, + Finsupp.zero_apply] + apply Finsupp.single_eq_of_ne Sum.inl_ne_inr + +theorem prod_apply_inr_snd (i) : (b.prod b' (Sum.inr i)).2 = b' i := + b'.repr.injective <| by + ext i + simp only [Basis.prod, Basis.coe_ofRepr, LinearEquiv.symm_trans_apply, LinearEquiv.prod_symm, + LinearEquiv.prod_apply, b'.repr.apply_symm_apply, LinearEquiv.symm_symm, repr_self, + Equiv.toFun_as_coe, Finsupp.snd_sumFinsuppLEquivProdFinsupp] + apply Finsupp.single_apply_left Sum.inr_injective + +@[simp] +theorem prod_apply (i) : + b.prod b' i = Sum.elim (LinearMap.inl R M M' ∘ b) (LinearMap.inr R M M' ∘ b') i := by + ext <;> cases i <;> + simp only [prod_apply_inl_fst, Sum.elim_inl, LinearMap.inl_apply, prod_apply_inr_fst, + Sum.elim_inr, LinearMap.inr_apply, prod_apply_inl_snd, prod_apply_inr_snd, Function.comp] + +end Prod + +section NoZeroSMulDivisors + +-- Can't be an instance because the basis can't be inferred. +protected theorem noZeroSMulDivisors [NoZeroDivisors R] (b : Basis ι R M) : + NoZeroSMulDivisors R M := + ⟨fun {c x} hcx => by + exact or_iff_not_imp_right.mpr fun hx => by + rw [← b.total_repr x, ← LinearMap.map_smul] at hcx + have := linearIndependent_iff.mp b.linearIndependent (c • b.repr x) hcx + rw [smul_eq_zero] at this + exact this.resolve_right fun hr => hx (b.repr.map_eq_zero_iff.mp hr)⟩ + +protected theorem smul_eq_zero [NoZeroDivisors R] (b : Basis ι R M) {c : R} {x : M} : + c • x = 0 ↔ c = 0 ∨ x = 0 := + @smul_eq_zero _ _ _ _ _ b.noZeroSMulDivisors _ _ + +theorem eq_bot_of_rank_eq_zero [NoZeroDivisors R] (b : Basis ι R M) (N : Submodule R M) + (rank_eq : ∀ {m : ℕ} (v : Fin m → N), LinearIndependent R ((↑) ∘ v : Fin m → M) → m = 0) : + N = ⊥ := by + rw [Submodule.eq_bot_iff] + intro x hx + contrapose! rank_eq with x_ne + refine ⟨1, fun _ => ⟨x, hx⟩, ?_, one_ne_zero⟩ + rw [Fintype.linearIndependent_iff] + rintro g sum_eq i + cases' i with _ hi + simp only [Function.const_apply, Fin.default_eq_zero, Submodule.coe_mk, Finset.univ_unique, + Function.comp_const, Finset.sum_singleton] at sum_eq + convert (b.smul_eq_zero.mp sum_eq).resolve_right x_ne + +end NoZeroSMulDivisors + +section Singleton + +theorem basis_singleton_iff {R M : Type*} [Ring R] [Nontrivial R] [AddCommGroup M] [Module R M] + [NoZeroSMulDivisors R M] (ι : Type*) [Unique ι] : + Nonempty (Basis ι R M) ↔ ∃ x ≠ 0, ∀ y : M, ∃ r : R, r • x = y := by + constructor + · rintro ⟨b⟩ + refine ⟨b default, b.linearIndependent.ne_zero _, ?_⟩ + simpa [span_singleton_eq_top_iff, Set.range_unique] using b.span_eq + · rintro ⟨x, nz, w⟩ + refine ⟨ofRepr <| LinearEquiv.symm + { toFun := fun f => f default • x + invFun := fun y => Finsupp.single default (w y).choose + left_inv := fun f => Finsupp.unique_ext ?_ + right_inv := fun y => ?_ + map_add' := fun y z => ?_ + map_smul' := fun c y => ?_ }⟩ + · simp [Finsupp.add_apply, add_smul] + · simp only [Finsupp.coe_smul, Pi.smul_apply, RingHom.id_apply] + rw [← smul_assoc] + · refine smul_left_injective _ nz ?_ + simp only [Finsupp.single_eq_same] + exact (w (f default • x)).choose_spec + · simp only [Finsupp.single_eq_same] + exact (w y).choose_spec + +end Singleton + +end Basis + +end Module + +section Module + +open LinearMap + +variable {v : ι → M} +variable [Ring R] [CommRing R₂] [AddCommGroup M] [AddCommGroup M'] [AddCommGroup M''] +variable [Module R M] [Module R₂ M] [Module R M'] [Module R M''] +variable {c d : R} {x y : M} +variable (b : Basis ι R M) + +namespace Basis + +/-- Any basis is a maximal linear independent set. +-/ +theorem maximal [Nontrivial R] (b : Basis ι R M) : b.linearIndependent.Maximal := fun w hi h => by + -- If `w` is strictly bigger than `range b`, + apply le_antisymm h + -- then choose some `x ∈ w \ range b`, + intro x p + by_contra q + -- and write it in terms of the basis. + have e := b.total_repr x + -- This then expresses `x` as a linear combination + -- of elements of `w` which are in the range of `b`, + let u : ι ↪ w := + ⟨fun i => ⟨b i, h ⟨i, rfl⟩⟩, fun i i' r => + b.injective (by simpa only [Subtype.mk_eq_mk] using r)⟩ + simp_rw [Finsupp.total_apply] at e + change ((b.repr x).sum fun (i : ι) (a : R) ↦ a • (u i : M)) = ((⟨x, p⟩ : w) : M) at e + rw [← Finsupp.sum_embDomain (f := u) (g := fun x r ↦ r • (x : M)), ← Finsupp.total_apply] at e + -- Now we can contradict the linear independence of `hi` + refine hi.total_ne_of_not_mem_support _ ?_ e + simp only [Finset.mem_map, Finsupp.support_embDomain] + rintro ⟨j, -, W⟩ + simp only [u, Embedding.coeFn_mk, Subtype.mk_eq_mk] at W + apply q ⟨j, W⟩ + +section Mk + +variable (hli : LinearIndependent R v) (hsp : ⊤ ≤ span R (range v)) + +/-- A linear independent family of vectors spanning the whole module is a basis. -/ +protected noncomputable def mk : Basis ι R M := + .ofRepr + { hli.repr.comp (LinearMap.id.codRestrict _ fun _ => hsp Submodule.mem_top) with + invFun := Finsupp.total _ _ _ v + left_inv := fun x => hli.total_repr ⟨x, _⟩ + right_inv := fun _ => hli.repr_eq rfl } + +@[simp] +theorem mk_repr : (Basis.mk hli hsp).repr x = hli.repr ⟨x, hsp Submodule.mem_top⟩ := + rfl + +theorem mk_apply (i : ι) : Basis.mk hli hsp i = v i := + show Finsupp.total _ _ _ v _ = v i by simp + +@[simp] +theorem coe_mk : ⇑(Basis.mk hli hsp) = v := + funext (mk_apply _ _) + +variable {hli hsp} + +/-- Given a basis, the `i`th element of the dual basis evaluates to 1 on the `i`th element of the +basis. -/ +theorem mk_coord_apply_eq (i : ι) : (Basis.mk hli hsp).coord i (v i) = 1 := + show hli.repr ⟨v i, Submodule.subset_span (mem_range_self i)⟩ i = 1 by simp [hli.repr_eq_single i] + +/-- Given a basis, the `i`th element of the dual basis evaluates to 0 on the `j`th element of the +basis if `j ≠ i`. -/ +theorem mk_coord_apply_ne {i j : ι} (h : j ≠ i) : (Basis.mk hli hsp).coord i (v j) = 0 := + show hli.repr ⟨v j, Submodule.subset_span (mem_range_self j)⟩ i = 0 by + simp [hli.repr_eq_single j, h] + +/-- Given a basis, the `i`th element of the dual basis evaluates to the Kronecker delta on the +`j`th element of the basis. -/ +theorem mk_coord_apply [DecidableEq ι] {i j : ι} : + (Basis.mk hli hsp).coord i (v j) = if j = i then 1 else 0 := by + rcases eq_or_ne j i with h | h + · simp only [h, if_true, eq_self_iff_true, mk_coord_apply_eq i] + · simp only [h, if_false, mk_coord_apply_ne h] + +end Mk + +section Span + +variable (hli : LinearIndependent R v) + +/-- A linear independent family of vectors is a basis for their span. -/ +protected noncomputable def span : Basis ι R (span R (range v)) := + Basis.mk (linearIndependent_span hli) <| by + intro x _ + have : ∀ i, v i ∈ span R (range v) := fun i ↦ subset_span (Set.mem_range_self _) + have h₁ : (((↑) : span R (range v) → M) '' range fun i => ⟨v i, this i⟩) = range v := by + simp only [SetLike.coe_sort_coe, ← Set.range_comp] + rfl + have h₂ : map (Submodule.subtype (span R (range v))) (span R (range fun i => ⟨v i, this i⟩)) = + span R (range v) := by + rw [← span_image, Submodule.coeSubtype] + -- Porting note: why doesn't `rw [h₁]` work here? + exact congr_arg _ h₁ + have h₃ : (x : M) ∈ map (Submodule.subtype (span R (range v))) + (span R (Set.range fun i => Subtype.mk (v i) (this i))) := by + rw [h₂] + apply Subtype.mem x + rcases mem_map.1 h₃ with ⟨y, hy₁, hy₂⟩ + have h_x_eq_y : x = y := by + rw [Subtype.ext_iff, ← hy₂] + simp + rwa [h_x_eq_y] + +protected theorem span_apply (i : ι) : (Basis.span hli i : M) = v i := + congr_arg ((↑) : span R (range v) → M) <| Basis.mk_apply _ _ _ + +end Span + +theorem groupSMul_span_eq_top {G : Type*} [Group G] [DistribMulAction G R] [DistribMulAction G M] + [IsScalarTower G R M] {v : ι → M} (hv : Submodule.span R (Set.range v) = ⊤) {w : ι → G} : + Submodule.span R (Set.range (w • v)) = ⊤ := by + rw [eq_top_iff] + intro j hj + rw [← hv] at hj + rw [Submodule.mem_span] at hj ⊢ + refine fun p hp => hj p fun u hu => ?_ + obtain ⟨i, rfl⟩ := hu + have : ((w i)⁻¹ • (1 : R)) • w i • v i ∈ p := p.smul_mem ((w i)⁻¹ • (1 : R)) (hp ⟨i, rfl⟩) + rwa [smul_one_smul, inv_smul_smul] at this + +/-- Given a basis `v` and a map `w` such that for all `i`, `w i` are elements of a group, +`groupSMul` provides the basis corresponding to `w • v`. -/ +def groupSMul {G : Type*} [Group G] [DistribMulAction G R] [DistribMulAction G M] + [IsScalarTower G R M] [SMulCommClass G R M] (v : Basis ι R M) (w : ι → G) : Basis ι R M := + Basis.mk (LinearIndependent.group_smul v.linearIndependent w) (groupSMul_span_eq_top v.span_eq).ge + +theorem groupSMul_apply {G : Type*} [Group G] [DistribMulAction G R] [DistribMulAction G M] + [IsScalarTower G R M] [SMulCommClass G R M] {v : Basis ι R M} {w : ι → G} (i : ι) : + v.groupSMul w i = (w • (v : ι → M)) i := + mk_apply (LinearIndependent.group_smul v.linearIndependent w) + (groupSMul_span_eq_top v.span_eq).ge i + +theorem units_smul_span_eq_top {v : ι → M} (hv : Submodule.span R (Set.range v) = ⊤) {w : ι → Rˣ} : + Submodule.span R (Set.range (w • v)) = ⊤ := + groupSMul_span_eq_top hv + +/-- Given a basis `v` and a map `w` such that for all `i`, `w i` is a unit, `unitsSMul` +provides the basis corresponding to `w • v`. -/ +def unitsSMul (v : Basis ι R M) (w : ι → Rˣ) : Basis ι R M := + Basis.mk (LinearIndependent.units_smul v.linearIndependent w) + (units_smul_span_eq_top v.span_eq).ge + +theorem unitsSMul_apply {v : Basis ι R M} {w : ι → Rˣ} (i : ι) : unitsSMul v w i = w i • v i := + mk_apply (LinearIndependent.units_smul v.linearIndependent w) + (units_smul_span_eq_top v.span_eq).ge i + +@[simp] +theorem coord_unitsSMul (e : Basis ι R₂ M) (w : ι → R₂ˣ) (i : ι) : + (unitsSMul e w).coord i = (w i)⁻¹ • e.coord i := by + classical + apply e.ext + intro j + trans ((unitsSMul e w).coord i) ((w j)⁻¹ • (unitsSMul e w) j) + · congr + simp [Basis.unitsSMul, ← mul_smul] + simp only [Basis.coord_apply, LinearMap.smul_apply, Basis.repr_self, Units.smul_def, + map_smul, Finsupp.single_apply] + split_ifs with h <;> simp [h] + +@[simp] +theorem repr_unitsSMul (e : Basis ι R₂ M) (w : ι → R₂ˣ) (v : M) (i : ι) : + (e.unitsSMul w).repr v i = (w i)⁻¹ • e.repr v i := + congr_arg (fun f : M →ₗ[R₂] R₂ => f v) (e.coord_unitsSMul w i) + +/-- A version of `unitsSMul` that uses `IsUnit`. -/ +def isUnitSMul (v : Basis ι R M) {w : ι → R} (hw : ∀ i, IsUnit (w i)) : Basis ι R M := + unitsSMul v fun i => (hw i).unit + +theorem isUnitSMul_apply {v : Basis ι R M} {w : ι → R} (hw : ∀ i, IsUnit (w i)) (i : ι) : + v.isUnitSMul hw i = w i • v i := + unitsSMul_apply i + +section Fin + +/-- Let `b` be a basis for a submodule `N` of `M`. If `y : M` is linear independent of `N` +and `y` and `N` together span the whole of `M`, then there is a basis for `M` +whose basis vectors are given by `Fin.cons y b`. -/ +noncomputable def mkFinCons {n : ℕ} {N : Submodule R M} (y : M) (b : Basis (Fin n) R N) + (hli : ∀ (c : R), ∀ x ∈ N, c • y + x = 0 → c = 0) (hsp : ∀ z : M, ∃ c : R, z + c • y ∈ N) : + Basis (Fin (n + 1)) R M := + have span_b : Submodule.span R (Set.range (N.subtype ∘ b)) = N := by + rw [Set.range_comp, Submodule.span_image, b.span_eq, Submodule.map_subtype_top] + Basis.mk (v := Fin.cons y (N.subtype ∘ b)) + ((b.linearIndependent.map' N.subtype (Submodule.ker_subtype _)).fin_cons' _ _ + (by + rintro c ⟨x, hx⟩ hc + rw [span_b] at hx + exact hli c x hx hc)) + fun x _ => by + rw [Fin.range_cons, Submodule.mem_span_insert', span_b] + exact hsp x + +@[simp] +theorem coe_mkFinCons {n : ℕ} {N : Submodule R M} (y : M) (b : Basis (Fin n) R N) + (hli : ∀ (c : R), ∀ x ∈ N, c • y + x = 0 → c = 0) (hsp : ∀ z : M, ∃ c : R, z + c • y ∈ N) : + (mkFinCons y b hli hsp : Fin (n + 1) → M) = Fin.cons y ((↑) ∘ b) := by + -- Porting note: without `unfold`, Lean can't reuse the proofs included in the definition + -- `mkFinCons` + unfold mkFinCons + exact coe_mk (v := Fin.cons y (N.subtype ∘ b)) _ _ + +/-- Let `b` be a basis for a submodule `N ≤ O`. If `y ∈ O` is linear independent of `N` +and `y` and `N` together span the whole of `O`, then there is a basis for `O` +whose basis vectors are given by `Fin.cons y b`. -/ +noncomputable def mkFinConsOfLE {n : ℕ} {N O : Submodule R M} (y : M) (yO : y ∈ O) + (b : Basis (Fin n) R N) (hNO : N ≤ O) (hli : ∀ (c : R), ∀ x ∈ N, c • y + x = 0 → c = 0) + (hsp : ∀ z ∈ O, ∃ c : R, z + c • y ∈ N) : Basis (Fin (n + 1)) R O := + mkFinCons ⟨y, yO⟩ (b.map (Submodule.comapSubtypeEquivOfLe hNO).symm) + (fun c x hc hx => hli c x (Submodule.mem_comap.mp hc) (congr_arg ((↑) : O → M) hx)) + fun z => hsp z z.2 + +@[simp] +theorem coe_mkFinConsOfLE {n : ℕ} {N O : Submodule R M} (y : M) (yO : y ∈ O) (b : Basis (Fin n) R N) + (hNO : N ≤ O) (hli : ∀ (c : R), ∀ x ∈ N, c • y + x = 0 → c = 0) + (hsp : ∀ z ∈ O, ∃ c : R, z + c • y ∈ N) : + (mkFinConsOfLE y yO b hNO hli hsp : Fin (n + 1) → O) = + Fin.cons ⟨y, yO⟩ (Submodule.inclusion hNO ∘ b) := + coe_mkFinCons _ _ _ _ + +/-- The basis of `R × R` given by the two vectors `(1, 0)` and `(0, 1)`. -/ +protected def finTwoProd (R : Type*) [Semiring R] : Basis (Fin 2) R (R × R) := + Basis.ofEquivFun (LinearEquiv.finTwoArrow R R).symm + +@[simp] +theorem finTwoProd_zero (R : Type*) [Semiring R] : Basis.finTwoProd R 0 = (1, 0) := by + simp [Basis.finTwoProd, LinearEquiv.finTwoArrow] + +@[simp] +theorem finTwoProd_one (R : Type*) [Semiring R] : Basis.finTwoProd R 1 = (0, 1) := by + simp [Basis.finTwoProd, LinearEquiv.finTwoArrow] + +@[simp] +theorem coe_finTwoProd_repr {R : Type*} [Semiring R] (x : R × R) : + ⇑((Basis.finTwoProd R).repr x) = ![x.fst, x.snd] := + rfl + +end Fin + +end Basis + +end Module + +section Induction + +variable [Ring R] [IsDomain R] +variable [AddCommGroup M] [Module R M] {b : ι → M} + +/-- If `N` is a submodule with finite rank, do induction on adjoining a linear independent +element to a submodule. -/ +def Submodule.inductionOnRankAux (b : Basis ι R M) (P : Submodule R M → Sort*) + (ih : ∀ N : Submodule R M, + (∀ N' ≤ N, ∀ x ∈ N, (∀ (c : R), ∀ y ∈ N', c • x + y = (0 : M) → c = 0) → P N') → P N) + (n : ℕ) (N : Submodule R M) + (rank_le : ∀ {m : ℕ} (v : Fin m → N), LinearIndependent R ((↑) ∘ v : Fin m → M) → m ≤ n) : + P N := by + haveI : DecidableEq M := Classical.decEq M + have Pbot : P ⊥ := by + apply ih + intro N _ x x_mem x_ortho + exfalso + rw [mem_bot] at x_mem + simpa [x_mem] using x_ortho 1 0 N.zero_mem + induction' n with n rank_ih generalizing N + · suffices N = ⊥ by rwa [this] + apply Basis.eq_bot_of_rank_eq_zero b _ fun m hv => Nat.le_zero.mp (rank_le _ hv) + apply ih + intro N' N'_le x x_mem x_ortho + apply rank_ih + intro m v hli + refine Nat.succ_le_succ_iff.mp (rank_le (Fin.cons ⟨x, x_mem⟩ fun i => ⟨v i, N'_le (v i).2⟩) ?_) + convert hli.fin_cons' x _ ?_ + · ext i + refine Fin.cases ?_ ?_ i <;> simp + · intro c y hcy + refine x_ortho c y (Submodule.span_le.mpr ?_ y.2) hcy + rintro _ ⟨z, rfl⟩ + exact (v z).2 + +end Induction + +/-- An element of a non-unital-non-associative algebra is in the center exactly when it commutes +with the basis elements. -/ +lemma Basis.mem_center_iff {A} + [Semiring R] [NonUnitalNonAssocSemiring A] + [Module R A] [SMulCommClass R A A] [SMulCommClass R R A] [IsScalarTower R A A] + (b : Basis ι R A) {z : A} : + z ∈ Set.center A ↔ + (∀ i, Commute (b i) z) ∧ ∀ i j, + z * (b i * b j) = (z * b i) * b j + ∧ (b i * z) * b j = b i * (z * b j) + ∧ (b i * b j) * z = b i * (b j * z) := by + constructor + · intro h + constructor + · intro i + apply (h.1 (b i)).symm + · intros + exact ⟨h.2 _ _, ⟨h.3 _ _, h.4 _ _⟩⟩ + · intro h + rw [center, mem_setOf_eq] + constructor + case comm => + intro y + rw [← b.total_repr y, Finsupp.total_apply, Finsupp.sum, Finset.sum_mul, Finset.mul_sum] + simp_rw [mul_smul_comm, smul_mul_assoc, (h.1 _).eq] + case left_assoc => + intro c d + rw [← b.total_repr c, ← b.total_repr d, Finsupp.total_apply, Finsupp.total_apply, Finsupp.sum, + Finsupp.sum, Finset.sum_mul, Finset.mul_sum, Finset.mul_sum, Finset.mul_sum] + simp_rw [smul_mul_assoc, Finset.mul_sum, Finset.sum_mul, mul_smul_comm, Finset.mul_sum, + Finset.smul_sum, smul_mul_assoc, mul_smul_comm, (h.2 _ _).1, + (@SMulCommClass.smul_comm R R A)] + rw [Finset.sum_comm] + case mid_assoc => + intro c d + rw [← b.total_repr c, ← b.total_repr d, Finsupp.total_apply, Finsupp.total_apply, Finsupp.sum, + Finsupp.sum, Finset.sum_mul, Finset.mul_sum, Finset.mul_sum, Finset.mul_sum] + simp_rw [smul_mul_assoc, Finset.sum_mul, mul_smul_comm, smul_mul_assoc, (h.2 _ _).2.1] + case right_assoc => + intro c d + rw [← b.total_repr c, ← b.total_repr d, Finsupp.total_apply, Finsupp.total_apply, Finsupp.sum, + Finsupp.sum, Finset.sum_mul] + simp_rw [smul_mul_assoc, Finset.mul_sum, Finset.sum_mul, mul_smul_comm, Finset.mul_sum, + Finset.smul_sum, smul_mul_assoc, mul_smul_comm, Finset.sum_mul, smul_mul_assoc, + (h.2 _ _).2.2] + +section RestrictScalars + +variable {S : Type*} [CommRing R] [Ring S] [Nontrivial S] [AddCommGroup M] +variable [Algebra R S] [Module S M] [Module R M] +variable [IsScalarTower R S M] [NoZeroSMulDivisors R S] (b : Basis ι S M) +variable (R) + +open Submodule + +/-- Let `b` be an `S`-basis of `M`. Let `R` be a CommRing such that `Algebra R S` has no zero smul +divisors, then the submodule of `M` spanned by `b` over `R` admits `b` as an `R`-basis. -/ +noncomputable def Basis.restrictScalars : Basis ι R (span R (Set.range b)) := + Basis.span (b.linearIndependent.restrict_scalars (smul_left_injective R one_ne_zero)) + +@[simp] +theorem Basis.restrictScalars_apply (i : ι) : (b.restrictScalars R i : M) = b i := by + simp only [Basis.restrictScalars, Basis.span_apply] + +@[simp] +theorem Basis.restrictScalars_repr_apply (m : span R (Set.range b)) (i : ι) : + algebraMap R S ((b.restrictScalars R).repr m i) = b.repr m i := by + suffices + Finsupp.mapRange.linearMap (Algebra.linearMap R S) ∘ₗ (b.restrictScalars R).repr.toLinearMap = + ((b.repr : M →ₗ[S] ι →₀ S).restrictScalars R).domRestrict _ + by exact DFunLike.congr_fun (LinearMap.congr_fun this m) i + refine Basis.ext (b.restrictScalars R) fun _ => ?_ + simp only [LinearMap.coe_comp, LinearEquiv.coe_toLinearMap, Function.comp_apply, map_one, + Basis.repr_self, Finsupp.mapRange.linearMap_apply, Finsupp.mapRange_single, + Algebra.linearMap_apply, LinearMap.domRestrict_apply, LinearEquiv.coe_coe, + Basis.restrictScalars_apply, LinearMap.coe_restrictScalars] + +/-- Let `b` be an `S`-basis of `M`. Then `m : M` lies in the `R`-module spanned by `b` iff all the +coordinates of `m` on the basis `b` are in `R` (see `Basis.mem_span` for the case `R = S`). -/ +theorem Basis.mem_span_iff_repr_mem (m : M) : + m ∈ span R (Set.range b) ↔ ∀ i, b.repr m i ∈ Set.range (algebraMap R S) := by + refine + ⟨fun hm i => ⟨(b.restrictScalars R).repr ⟨m, hm⟩ i, b.restrictScalars_repr_apply R ⟨m, hm⟩ i⟩, + fun h => ?_⟩ + rw [← b.total_repr m, Finsupp.total_apply S _] + refine sum_mem fun i _ => ?_ + obtain ⟨_, h⟩ := h i + simp_rw [← h, algebraMap_smul] + exact smul_mem _ _ (subset_span (Set.mem_range_self i)) + +end RestrictScalars + +assert_not_exists Ordinal diff --git a/Mathlib/LinearAlgebra/Basis/Bilinear.lean b/Mathlib/LinearAlgebra/Basis/Bilinear.lean index 565bab1514e72..dd756b57e5821 100644 --- a/Mathlib/LinearAlgebra/Basis/Bilinear.lean +++ b/Mathlib/LinearAlgebra/Basis/Bilinear.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Moritz Doll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ -import Mathlib.LinearAlgebra.Basis import Mathlib.LinearAlgebra.BilinearMap +import Mathlib.LinearAlgebra.Basis.Defs /-! # Lemmas about bilinear maps with a basis over each argument diff --git a/Mathlib/LinearAlgebra/Basis/Cardinality.lean b/Mathlib/LinearAlgebra/Basis/Cardinality.lean new file mode 100644 index 0000000000000..74b34a45465d4 --- /dev/null +++ b/Mathlib/LinearAlgebra/Basis/Cardinality.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro, Alexander Bentkamp, Kim Morrison +-/ +import Mathlib.LinearAlgebra.Basis.Basic +import Mathlib.SetTheory.Cardinal.Cofinality + +/-! +# Results relating bases and cardinality. +-/ + +section Finite + +open Basis Cardinal Set Submodule + +universe u v v' v'' u₁' w w' + +variable {R : Type u} {M M₁ : Type v} {M' : Type v'} {ι : Type w} +variable [Ring R] [AddCommGroup M] [AddCommGroup M'] [AddCommGroup M₁] [Nontrivial R] +variable [Module R M] [Module R M'] [Module R M₁] + +-- One might hope that a finite spanning set implies that any linearly independent set is finite. +-- While this is true over a division ring +-- (simply because any linearly independent set can be extended to a basis), +-- or more generally over a ring satisfying the strong rank condition +-- (which covers all commutative rings; see `LinearIndependent.finite_of_le_span_finite`), +-- this is not true in general. +-- For example, the left ideal generated by the variables in a noncommutative polynomial ring +-- (`FreeAlgebra R ι`) in infinitely many variables (indexed by `ι`) is free +-- with an infinite basis (consisting of the variables). +-- As another example, for any commutative ring R, the ring of column-finite matrices +-- `Module.End R (ℕ →₀ R)` is isomorphic to `ℕ → Module.End R (ℕ →₀ R)` as a module over itself, +-- which also clearly contains an infinite linearly independent set. +/-- +Over any nontrivial ring, the existence of a finite spanning set implies that any basis is finite. +-/ +lemma basis_finite_of_finite_spans (w : Set M) (hw : w.Finite) (s : span R w = ⊤) {ι : Type w} + (b : Basis ι R M) : Finite ι := by + classical + haveI := hw.to_subtype + cases nonempty_fintype w + -- We'll work by contradiction, assuming `ι` is infinite. + rw [← not_infinite_iff_finite] + intro i + -- Let `S` be the union of the supports of `x ∈ w` expressed as linear combinations of `b`. + -- This is a finite set since `w` is finite. + let S : Finset ι := Finset.univ.sup fun x : w => (b.repr x).support + let bS : Set M := b '' S + have h : ∀ x ∈ w, x ∈ span R bS := by + intro x m + rw [← b.total_repr x, Finsupp.span_image_eq_map_total, Submodule.mem_map] + use b.repr x + simp only [and_true_iff, eq_self_iff_true, Finsupp.mem_supported] + rw [Finset.coe_subset, ← Finset.le_iff_subset] + exact Finset.le_sup (f := fun x : w ↦ (b.repr ↑x).support) (Finset.mem_univ (⟨x, m⟩ : w)) + -- Thus this finite subset of the basis elements spans the entire module. + have k : span R bS = ⊤ := eq_top_iff.2 (le_trans s.ge (span_le.2 h)) + -- Now there is some `x : ι` not in `S`, since `ι` is infinite. + obtain ⟨x, nm⟩ := Infinite.exists_not_mem_finset S + -- However it must be in the span of the finite subset, + have k' : b x ∈ span R bS := by + rw [k] + exact mem_top + -- giving the desire contradiction. + exact b.linearIndependent.not_mem_span_image nm k' + +-- From [Les familles libres maximales d'un module ont-elles le meme cardinal?][lazarus1973] +/-- Over any ring `R`, if `b` is a basis for a module `M`, +and `s` is a maximal linearly independent set, +then the union of the supports of `x ∈ s` (when written out in the basis `b`) is all of `b`. +-/ +theorem union_support_maximal_linearIndependent_eq_range_basis {ι : Type w} (b : Basis ι R M) + {κ : Type w'} (v : κ → M) (i : LinearIndependent R v) (m : i.Maximal) : + ⋃ k, ((b.repr (v k)).support : Set ι) = Set.univ := by + -- If that's not the case, + by_contra h + simp only [← Ne.eq_def, ne_univ_iff_exists_not_mem, mem_iUnion, not_exists_not, + Finsupp.mem_support_iff, Finset.mem_coe] at h + -- We have some basis element `b b'` which is not in the support of any of the `v i`. + obtain ⟨b', w⟩ := h + -- Using this, we'll construct a linearly independent family strictly larger than `v`, + -- by also using this `b b'`. + let v' : Option κ → M := fun o => o.elim (b b') v + have r : range v ⊆ range v' := by + rintro - ⟨k, rfl⟩ + use some k + simp only [v', Option.elim_some] + have r' : b b' ∉ range v := by + rintro ⟨k, p⟩ + simpa [w] using congr_arg (fun m => (b.repr m) b') p + have r'' : range v ≠ range v' := by + intro e + have p : b b' ∈ range v' := by + use none + simp only [v', Option.elim_none] + rw [← e] at p + exact r' p + -- The key step in the proof is checking that this strictly larger family is linearly independent. + have i' : LinearIndependent R ((↑) : range v' → M) := by + apply LinearIndependent.to_subtype_range + rw [linearIndependent_iff] + intro l z + rw [Finsupp.total_option] at z + simp only [v', Option.elim'] at z + change _ + Finsupp.total κ M R v l.some = 0 at z + -- We have some linear combination of `b b'` and the `v i`, which we want to show is trivial. + -- We'll first show the coefficient of `b b'` is zero, + -- by expressing the `v i` in the basis `b`, and using that the `v i` have no `b b'` term. + have l₀ : l none = 0 := by + rw [← eq_neg_iff_add_eq_zero] at z + replace z := neg_eq_iff_eq_neg.mpr z + apply_fun fun x => b.repr x b' at z + simp only [repr_self, map_smul, mul_one, Finsupp.single_eq_same, Pi.neg_apply, + Finsupp.smul_single', map_neg, Finsupp.coe_neg] at z + erw [DFunLike.congr_fun (Finsupp.apply_total R (b.repr : M →ₗ[R] ι →₀ R) v l.some) b'] at z + simpa [Finsupp.total_apply, w] using z + -- Then all the other coefficients are zero, because `v` is linear independent. + have l₁ : l.some = 0 := by + rw [l₀, zero_smul, zero_add] at z + exact linearIndependent_iff.mp i _ z + -- Finally we put those facts together to show the linear combination is trivial. + ext (_ | a) + · simp only [l₀, Finsupp.coe_zero, Pi.zero_apply] + · erw [DFunLike.congr_fun l₁ a] + simp only [Finsupp.coe_zero, Pi.zero_apply] + rw [LinearIndependent.Maximal] at m + specialize m (range v') i' r + exact r'' m + +/-- Over any ring `R`, if `b` is an infinite basis for a module `M`, +and `s` is a maximal linearly independent set, +then the cardinality of `b` is bounded by the cardinality of `s`. +-/ +theorem infinite_basis_le_maximal_linearIndependent' {ι : Type w} (b : Basis ι R M) [Infinite ι] + {κ : Type w'} (v : κ → M) (i : LinearIndependent R v) (m : i.Maximal) : + Cardinal.lift.{w'} #ι ≤ Cardinal.lift.{w} #κ := by + let Φ := fun k : κ => (b.repr (v k)).support + have w₁ : #ι ≤ #(Set.range Φ) := by + apply Cardinal.le_range_of_union_finset_eq_top + exact union_support_maximal_linearIndependent_eq_range_basis b v i m + have w₂ : Cardinal.lift.{w'} #(Set.range Φ) ≤ Cardinal.lift.{w} #κ := Cardinal.mk_range_le_lift + exact (Cardinal.lift_le.mpr w₁).trans w₂ + +-- (See `infinite_basis_le_maximal_linearIndependent'` for the more general version +-- where the index types can live in different universes.) +/-- Over any ring `R`, if `b` is an infinite basis for a module `M`, +and `s` is a maximal linearly independent set, +then the cardinality of `b` is bounded by the cardinality of `s`. +-/ +theorem infinite_basis_le_maximal_linearIndependent {ι : Type w} (b : Basis ι R M) [Infinite ι] + {κ : Type w} (v : κ → M) (i : LinearIndependent R v) (m : i.Maximal) : #ι ≤ #κ := + Cardinal.lift_le.mp (infinite_basis_le_maximal_linearIndependent' b v i m) + +end Finite diff --git a/Mathlib/LinearAlgebra/Basis.lean b/Mathlib/LinearAlgebra/Basis/Defs.lean similarity index 51% rename from Mathlib/LinearAlgebra/Basis.lean rename to Mathlib/LinearAlgebra/Basis/Defs.lean index e3321fdbb56bd..639fb724332cf 100644 --- a/Mathlib/LinearAlgebra/Basis.lean +++ b/Mathlib/LinearAlgebra/Basis/Defs.lean @@ -3,12 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Alexander Bentkamp -/ -import Mathlib.Algebra.BigOperators.Finsupp -import Mathlib.Algebra.BigOperators.Finprod -import Mathlib.Data.Fintype.BigOperators import Mathlib.LinearAlgebra.Finsupp -import Mathlib.LinearAlgebra.LinearIndependent -import Mathlib.SetTheory.Cardinal.Cofinality /-! # Bases @@ -203,13 +198,6 @@ noncomputable def sumCoords : M →ₗ[R] R := theorem coe_sumCoords : (b.sumCoords : M → R) = fun m => (b.repr m).sum fun _ => id := rfl -theorem coe_sumCoords_eq_finsum : (b.sumCoords : M → R) = fun m => ∑ᶠ i, b.coord i m := by - ext m - simp only [Basis.sumCoords, Basis.coord, Finsupp.lapply_apply, LinearMap.id_coe, - LinearEquiv.coe_coe, Function.comp_apply, Finsupp.coe_lsum, LinearMap.coe_comp, - finsum_eq_sum _ (b.repr m).finite_support, Finsupp.sum, Finset.finite_toSet_toFinset, id, - Finsupp.fun_support_eq] - @[simp high] theorem coe_sumCoords_of_fintype [Fintype ι] : (b.sumCoords : M → R) = ∑ i, b.coord i := by ext m @@ -525,15 +513,6 @@ end Fintype end Reindex -protected theorem linearIndependent : LinearIndependent R b := - linearIndependent_iff.mpr fun l hl => - calc - l = b.repr (Finsupp.total _ _ _ b l) := (b.repr_total l).symm - _ = 0 := by rw [hl, LinearEquiv.map_zero] - -protected theorem ne_zero [Nontrivial R] (i) : b i ≠ 0 := - b.linearIndependent.ne_zero i - protected theorem mem_span (x : M) : x ∈ span R (range b) := span_mono (image_subset_range _ _) (mem_span_repr_support b x) @@ -652,99 +631,6 @@ theorem map_equiv (b : Basis ι R M) (b' : Basis ι' R M') (e : ι ≃ ι') : end Equiv -section Prod - -variable (b' : Basis ι' R M') - -/-- `Basis.prod` maps an `ι`-indexed basis for `M` and an `ι'`-indexed basis for `M'` -to an `ι ⊕ ι'`-index basis for `M × M'`. -For the specific case of `R × R`, see also `Basis.finTwoProd`. -/ -protected def prod : Basis (ι ⊕ ι') R (M × M') := - ofRepr ((b.repr.prod b'.repr).trans (Finsupp.sumFinsuppLEquivProdFinsupp R).symm) - -@[simp] -theorem prod_repr_inl (x) (i) : (b.prod b').repr x (Sum.inl i) = b.repr x.1 i := - rfl - -@[simp] -theorem prod_repr_inr (x) (i) : (b.prod b').repr x (Sum.inr i) = b'.repr x.2 i := - rfl - -theorem prod_apply_inl_fst (i) : (b.prod b' (Sum.inl i)).1 = b i := - b.repr.injective <| by - ext j - simp only [Basis.prod, Basis.coe_ofRepr, LinearEquiv.symm_trans_apply, LinearEquiv.prod_symm, - LinearEquiv.prod_apply, b.repr.apply_symm_apply, LinearEquiv.symm_symm, repr_self, - Equiv.toFun_as_coe, Finsupp.fst_sumFinsuppLEquivProdFinsupp] - apply Finsupp.single_apply_left Sum.inl_injective - -theorem prod_apply_inr_fst (i) : (b.prod b' (Sum.inr i)).1 = 0 := - b.repr.injective <| by - ext i - simp only [Basis.prod, Basis.coe_ofRepr, LinearEquiv.symm_trans_apply, LinearEquiv.prod_symm, - LinearEquiv.prod_apply, b.repr.apply_symm_apply, LinearEquiv.symm_symm, repr_self, - Equiv.toFun_as_coe, Finsupp.fst_sumFinsuppLEquivProdFinsupp, LinearEquiv.map_zero, - Finsupp.zero_apply] - apply Finsupp.single_eq_of_ne Sum.inr_ne_inl - -theorem prod_apply_inl_snd (i) : (b.prod b' (Sum.inl i)).2 = 0 := - b'.repr.injective <| by - ext j - simp only [Basis.prod, Basis.coe_ofRepr, LinearEquiv.symm_trans_apply, LinearEquiv.prod_symm, - LinearEquiv.prod_apply, b'.repr.apply_symm_apply, LinearEquiv.symm_symm, repr_self, - Equiv.toFun_as_coe, Finsupp.snd_sumFinsuppLEquivProdFinsupp, LinearEquiv.map_zero, - Finsupp.zero_apply] - apply Finsupp.single_eq_of_ne Sum.inl_ne_inr - -theorem prod_apply_inr_snd (i) : (b.prod b' (Sum.inr i)).2 = b' i := - b'.repr.injective <| by - ext i - simp only [Basis.prod, Basis.coe_ofRepr, LinearEquiv.symm_trans_apply, LinearEquiv.prod_symm, - LinearEquiv.prod_apply, b'.repr.apply_symm_apply, LinearEquiv.symm_symm, repr_self, - Equiv.toFun_as_coe, Finsupp.snd_sumFinsuppLEquivProdFinsupp] - apply Finsupp.single_apply_left Sum.inr_injective - -@[simp] -theorem prod_apply (i) : - b.prod b' i = Sum.elim (LinearMap.inl R M M' ∘ b) (LinearMap.inr R M M' ∘ b') i := by - ext <;> cases i <;> - simp only [prod_apply_inl_fst, Sum.elim_inl, LinearMap.inl_apply, prod_apply_inr_fst, - Sum.elim_inr, LinearMap.inr_apply, prod_apply_inl_snd, prod_apply_inr_snd, Function.comp] - -end Prod - -section NoZeroSMulDivisors - --- Can't be an instance because the basis can't be inferred. -protected theorem noZeroSMulDivisors [NoZeroDivisors R] (b : Basis ι R M) : - NoZeroSMulDivisors R M := - ⟨fun {c x} hcx => by - exact or_iff_not_imp_right.mpr fun hx => by - rw [← b.total_repr x, ← LinearMap.map_smul] at hcx - have := linearIndependent_iff.mp b.linearIndependent (c • b.repr x) hcx - rw [smul_eq_zero] at this - exact this.resolve_right fun hr => hx (b.repr.map_eq_zero_iff.mp hr)⟩ - -protected theorem smul_eq_zero [NoZeroDivisors R] (b : Basis ι R M) {c : R} {x : M} : - c • x = 0 ↔ c = 0 ∨ x = 0 := - @smul_eq_zero _ _ _ _ _ b.noZeroSMulDivisors _ _ - -theorem eq_bot_of_rank_eq_zero [NoZeroDivisors R] (b : Basis ι R M) (N : Submodule R M) - (rank_eq : ∀ {m : ℕ} (v : Fin m → N), LinearIndependent R ((↑) ∘ v : Fin m → M) → m = 0) : - N = ⊥ := by - rw [Submodule.eq_bot_iff] - intro x hx - contrapose! rank_eq with x_ne - refine ⟨1, fun _ => ⟨x, hx⟩, ?_, one_ne_zero⟩ - rw [Fintype.linearIndependent_iff] - rintro g sum_eq i - cases' i with _ hi - simp only [Function.const_apply, Fin.default_eq_zero, Submodule.coe_mk, Finset.univ_unique, - Function.comp_const, Finset.sum_singleton] at sum_eq - convert (b.smul_eq_zero.mp sum_eq).resolve_right x_ne - -end NoZeroSMulDivisors - section Singleton /-- `Basis.singleton ι R` is the basis sending the unique element of `ι` to `1 : R`. -/ @@ -765,30 +651,6 @@ theorem singleton_apply (ι R : Type*) [Unique ι] [Semiring R] (i) : Basis.sing theorem singleton_repr (ι R : Type*) [Unique ι] [Semiring R] (x i) : (Basis.singleton ι R).repr x i = x := by simp [Basis.singleton, Unique.eq_default i] -theorem basis_singleton_iff {R M : Type*} [Ring R] [Nontrivial R] [AddCommGroup M] [Module R M] - [NoZeroSMulDivisors R M] (ι : Type*) [Unique ι] : - Nonempty (Basis ι R M) ↔ ∃ x ≠ 0, ∀ y : M, ∃ r : R, r • x = y := by - constructor - · rintro ⟨b⟩ - refine ⟨b default, b.linearIndependent.ne_zero _, ?_⟩ - simpa [span_singleton_eq_top_iff, Set.range_unique] using b.span_eq - · rintro ⟨x, nz, w⟩ - refine ⟨ofRepr <| LinearEquiv.symm - { toFun := fun f => f default • x - invFun := fun y => Finsupp.single default (w y).choose - left_inv := fun f => Finsupp.unique_ext ?_ - right_inv := fun y => ?_ - map_add' := fun y z => ?_ - map_smul' := fun c y => ?_ }⟩ - · simp [Finsupp.add_apply, add_smul] - · simp only [Finsupp.coe_smul, Pi.smul_apply, RingHom.id_apply] - rw [← smul_assoc] - · refine smul_left_injective _ nz ?_ - simp only [Finsupp.single_eq_same] - exact (w (f default • x)).choose_spec - · simp only [Finsupp.single_eq_same] - exact (w y).choose_spec - end Singleton section Empty @@ -971,525 +833,5 @@ end Basis end CommSemiring -section Module - -open LinearMap - -variable {v : ι → M} -variable [Ring R] [CommRing R₂] [AddCommGroup M] [AddCommGroup M'] [AddCommGroup M''] -variable [Module R M] [Module R₂ M] [Module R M'] [Module R M''] -variable {c d : R} {x y : M} -variable (b : Basis ι R M) - -namespace Basis - -/-- Any basis is a maximal linear independent set. --/ -theorem maximal [Nontrivial R] (b : Basis ι R M) : b.linearIndependent.Maximal := fun w hi h => by - -- If `w` is strictly bigger than `range b`, - apply le_antisymm h - -- then choose some `x ∈ w \ range b`, - intro x p - by_contra q - -- and write it in terms of the basis. - have e := b.total_repr x - -- This then expresses `x` as a linear combination - -- of elements of `w` which are in the range of `b`, - let u : ι ↪ w := - ⟨fun i => ⟨b i, h ⟨i, rfl⟩⟩, fun i i' r => - b.injective (by simpa only [Subtype.mk_eq_mk] using r)⟩ - simp_rw [Finsupp.total_apply] at e - change ((b.repr x).sum fun (i : ι) (a : R) ↦ a • (u i : M)) = ((⟨x, p⟩ : w) : M) at e - rw [← Finsupp.sum_embDomain (f := u) (g := fun x r ↦ r • (x : M)), ← Finsupp.total_apply] at e - -- Now we can contradict the linear independence of `hi` - refine hi.total_ne_of_not_mem_support _ ?_ e - simp only [Finset.mem_map, Finsupp.support_embDomain] - rintro ⟨j, -, W⟩ - simp only [u, Embedding.coeFn_mk, Subtype.mk_eq_mk] at W - apply q ⟨j, W⟩ - -section Mk - -variable (hli : LinearIndependent R v) (hsp : ⊤ ≤ span R (range v)) - -/-- A linear independent family of vectors spanning the whole module is a basis. -/ -protected noncomputable def mk : Basis ι R M := - .ofRepr - { hli.repr.comp (LinearMap.id.codRestrict _ fun _ => hsp Submodule.mem_top) with - invFun := Finsupp.total _ _ _ v - left_inv := fun x => hli.total_repr ⟨x, _⟩ - right_inv := fun _ => hli.repr_eq rfl } - -@[simp] -theorem mk_repr : (Basis.mk hli hsp).repr x = hli.repr ⟨x, hsp Submodule.mem_top⟩ := - rfl - -theorem mk_apply (i : ι) : Basis.mk hli hsp i = v i := - show Finsupp.total _ _ _ v _ = v i by simp - -@[simp] -theorem coe_mk : ⇑(Basis.mk hli hsp) = v := - funext (mk_apply _ _) - -variable {hli hsp} - -/-- Given a basis, the `i`th element of the dual basis evaluates to 1 on the `i`th element of the -basis. -/ -theorem mk_coord_apply_eq (i : ι) : (Basis.mk hli hsp).coord i (v i) = 1 := - show hli.repr ⟨v i, Submodule.subset_span (mem_range_self i)⟩ i = 1 by simp [hli.repr_eq_single i] - -/-- Given a basis, the `i`th element of the dual basis evaluates to 0 on the `j`th element of the -basis if `j ≠ i`. -/ -theorem mk_coord_apply_ne {i j : ι} (h : j ≠ i) : (Basis.mk hli hsp).coord i (v j) = 0 := - show hli.repr ⟨v j, Submodule.subset_span (mem_range_self j)⟩ i = 0 by - simp [hli.repr_eq_single j, h] - -/-- Given a basis, the `i`th element of the dual basis evaluates to the Kronecker delta on the -`j`th element of the basis. -/ -theorem mk_coord_apply [DecidableEq ι] {i j : ι} : - (Basis.mk hli hsp).coord i (v j) = if j = i then 1 else 0 := by - rcases eq_or_ne j i with h | h - · simp only [h, if_true, eq_self_iff_true, mk_coord_apply_eq i] - · simp only [h, if_false, mk_coord_apply_ne h] - -end Mk - -section Span - -variable (hli : LinearIndependent R v) - -/-- A linear independent family of vectors is a basis for their span. -/ -protected noncomputable def span : Basis ι R (span R (range v)) := - Basis.mk (linearIndependent_span hli) <| by - intro x _ - have : ∀ i, v i ∈ span R (range v) := fun i ↦ subset_span (Set.mem_range_self _) - have h₁ : (((↑) : span R (range v) → M) '' range fun i => ⟨v i, this i⟩) = range v := by - simp only [SetLike.coe_sort_coe, ← Set.range_comp] - rfl - have h₂ : map (Submodule.subtype (span R (range v))) (span R (range fun i => ⟨v i, this i⟩)) = - span R (range v) := by - rw [← span_image, Submodule.coeSubtype] - -- Porting note: why doesn't `rw [h₁]` work here? - exact congr_arg _ h₁ - have h₃ : (x : M) ∈ map (Submodule.subtype (span R (range v))) - (span R (Set.range fun i => Subtype.mk (v i) (this i))) := by - rw [h₂] - apply Subtype.mem x - rcases mem_map.1 h₃ with ⟨y, hy₁, hy₂⟩ - have h_x_eq_y : x = y := by - rw [Subtype.ext_iff, ← hy₂] - simp - rwa [h_x_eq_y] - -protected theorem span_apply (i : ι) : (Basis.span hli i : M) = v i := - congr_arg ((↑) : span R (range v) → M) <| Basis.mk_apply _ _ _ - -end Span - -theorem groupSMul_span_eq_top {G : Type*} [Group G] [DistribMulAction G R] [DistribMulAction G M] - [IsScalarTower G R M] {v : ι → M} (hv : Submodule.span R (Set.range v) = ⊤) {w : ι → G} : - Submodule.span R (Set.range (w • v)) = ⊤ := by - rw [eq_top_iff] - intro j hj - rw [← hv] at hj - rw [Submodule.mem_span] at hj ⊢ - refine fun p hp => hj p fun u hu => ?_ - obtain ⟨i, rfl⟩ := hu - have : ((w i)⁻¹ • (1 : R)) • w i • v i ∈ p := p.smul_mem ((w i)⁻¹ • (1 : R)) (hp ⟨i, rfl⟩) - rwa [smul_one_smul, inv_smul_smul] at this - -/-- Given a basis `v` and a map `w` such that for all `i`, `w i` are elements of a group, -`groupSMul` provides the basis corresponding to `w • v`. -/ -def groupSMul {G : Type*} [Group G] [DistribMulAction G R] [DistribMulAction G M] - [IsScalarTower G R M] [SMulCommClass G R M] (v : Basis ι R M) (w : ι → G) : Basis ι R M := - Basis.mk (LinearIndependent.group_smul v.linearIndependent w) (groupSMul_span_eq_top v.span_eq).ge - -theorem groupSMul_apply {G : Type*} [Group G] [DistribMulAction G R] [DistribMulAction G M] - [IsScalarTower G R M] [SMulCommClass G R M] {v : Basis ι R M} {w : ι → G} (i : ι) : - v.groupSMul w i = (w • (v : ι → M)) i := - mk_apply (LinearIndependent.group_smul v.linearIndependent w) - (groupSMul_span_eq_top v.span_eq).ge i - -theorem units_smul_span_eq_top {v : ι → M} (hv : Submodule.span R (Set.range v) = ⊤) {w : ι → Rˣ} : - Submodule.span R (Set.range (w • v)) = ⊤ := - groupSMul_span_eq_top hv - -/-- Given a basis `v` and a map `w` such that for all `i`, `w i` is a unit, `unitsSMul` -provides the basis corresponding to `w • v`. -/ -def unitsSMul (v : Basis ι R M) (w : ι → Rˣ) : Basis ι R M := - Basis.mk (LinearIndependent.units_smul v.linearIndependent w) - (units_smul_span_eq_top v.span_eq).ge - -theorem unitsSMul_apply {v : Basis ι R M} {w : ι → Rˣ} (i : ι) : unitsSMul v w i = w i • v i := - mk_apply (LinearIndependent.units_smul v.linearIndependent w) - (units_smul_span_eq_top v.span_eq).ge i - -@[simp] -theorem coord_unitsSMul (e : Basis ι R₂ M) (w : ι → R₂ˣ) (i : ι) : - (unitsSMul e w).coord i = (w i)⁻¹ • e.coord i := by - classical - apply e.ext - intro j - trans ((unitsSMul e w).coord i) ((w j)⁻¹ • (unitsSMul e w) j) - · congr - simp [Basis.unitsSMul, ← mul_smul] - simp only [Basis.coord_apply, LinearMap.smul_apply, Basis.repr_self, Units.smul_def, - map_smul, Finsupp.single_apply] - split_ifs with h <;> simp [h] - -@[simp] -theorem repr_unitsSMul (e : Basis ι R₂ M) (w : ι → R₂ˣ) (v : M) (i : ι) : - (e.unitsSMul w).repr v i = (w i)⁻¹ • e.repr v i := - congr_arg (fun f : M →ₗ[R₂] R₂ => f v) (e.coord_unitsSMul w i) - -/-- A version of `unitsSMul` that uses `IsUnit`. -/ -def isUnitSMul (v : Basis ι R M) {w : ι → R} (hw : ∀ i, IsUnit (w i)) : Basis ι R M := - unitsSMul v fun i => (hw i).unit - -theorem isUnitSMul_apply {v : Basis ι R M} {w : ι → R} (hw : ∀ i, IsUnit (w i)) (i : ι) : - v.isUnitSMul hw i = w i • v i := - unitsSMul_apply i - -section Fin - -/-- Let `b` be a basis for a submodule `N` of `M`. If `y : M` is linear independent of `N` -and `y` and `N` together span the whole of `M`, then there is a basis for `M` -whose basis vectors are given by `Fin.cons y b`. -/ -noncomputable def mkFinCons {n : ℕ} {N : Submodule R M} (y : M) (b : Basis (Fin n) R N) - (hli : ∀ (c : R), ∀ x ∈ N, c • y + x = 0 → c = 0) (hsp : ∀ z : M, ∃ c : R, z + c • y ∈ N) : - Basis (Fin (n + 1)) R M := - have span_b : Submodule.span R (Set.range (N.subtype ∘ b)) = N := by - rw [Set.range_comp, Submodule.span_image, b.span_eq, Submodule.map_subtype_top] - Basis.mk (v := Fin.cons y (N.subtype ∘ b)) - ((b.linearIndependent.map' N.subtype (Submodule.ker_subtype _)).fin_cons' _ _ - (by - rintro c ⟨x, hx⟩ hc - rw [span_b] at hx - exact hli c x hx hc)) - fun x _ => by - rw [Fin.range_cons, Submodule.mem_span_insert', span_b] - exact hsp x - -@[simp] -theorem coe_mkFinCons {n : ℕ} {N : Submodule R M} (y : M) (b : Basis (Fin n) R N) - (hli : ∀ (c : R), ∀ x ∈ N, c • y + x = 0 → c = 0) (hsp : ∀ z : M, ∃ c : R, z + c • y ∈ N) : - (mkFinCons y b hli hsp : Fin (n + 1) → M) = Fin.cons y ((↑) ∘ b) := by - -- Porting note: without `unfold`, Lean can't reuse the proofs included in the definition - -- `mkFinCons` - unfold mkFinCons - exact coe_mk (v := Fin.cons y (N.subtype ∘ b)) _ _ - -/-- Let `b` be a basis for a submodule `N ≤ O`. If `y ∈ O` is linear independent of `N` -and `y` and `N` together span the whole of `O`, then there is a basis for `O` -whose basis vectors are given by `Fin.cons y b`. -/ -noncomputable def mkFinConsOfLE {n : ℕ} {N O : Submodule R M} (y : M) (yO : y ∈ O) - (b : Basis (Fin n) R N) (hNO : N ≤ O) (hli : ∀ (c : R), ∀ x ∈ N, c • y + x = 0 → c = 0) - (hsp : ∀ z ∈ O, ∃ c : R, z + c • y ∈ N) : Basis (Fin (n + 1)) R O := - mkFinCons ⟨y, yO⟩ (b.map (Submodule.comapSubtypeEquivOfLe hNO).symm) - (fun c x hc hx => hli c x (Submodule.mem_comap.mp hc) (congr_arg ((↑) : O → M) hx)) - fun z => hsp z z.2 - -@[simp] -theorem coe_mkFinConsOfLE {n : ℕ} {N O : Submodule R M} (y : M) (yO : y ∈ O) (b : Basis (Fin n) R N) - (hNO : N ≤ O) (hli : ∀ (c : R), ∀ x ∈ N, c • y + x = 0 → c = 0) - (hsp : ∀ z ∈ O, ∃ c : R, z + c • y ∈ N) : - (mkFinConsOfLE y yO b hNO hli hsp : Fin (n + 1) → O) = - Fin.cons ⟨y, yO⟩ (Submodule.inclusion hNO ∘ b) := - coe_mkFinCons _ _ _ _ - -/-- The basis of `R × R` given by the two vectors `(1, 0)` and `(0, 1)`. -/ -protected def finTwoProd (R : Type*) [Semiring R] : Basis (Fin 2) R (R × R) := - Basis.ofEquivFun (LinearEquiv.finTwoArrow R R).symm - -@[simp] -theorem finTwoProd_zero (R : Type*) [Semiring R] : Basis.finTwoProd R 0 = (1, 0) := by - simp [Basis.finTwoProd, LinearEquiv.finTwoArrow] - -@[simp] -theorem finTwoProd_one (R : Type*) [Semiring R] : Basis.finTwoProd R 1 = (0, 1) := by - simp [Basis.finTwoProd, LinearEquiv.finTwoArrow] - -@[simp] -theorem coe_finTwoProd_repr {R : Type*} [Semiring R] (x : R × R) : - ⇑((Basis.finTwoProd R).repr x) = ![x.fst, x.snd] := - rfl - -end Fin - -end Basis - -end Module - -section Induction - -variable [Ring R] [IsDomain R] -variable [AddCommGroup M] [Module R M] {b : ι → M} - -/-- If `N` is a submodule with finite rank, do induction on adjoining a linear independent -element to a submodule. -/ -def Submodule.inductionOnRankAux (b : Basis ι R M) (P : Submodule R M → Sort*) - (ih : ∀ N : Submodule R M, - (∀ N' ≤ N, ∀ x ∈ N, (∀ (c : R), ∀ y ∈ N', c • x + y = (0 : M) → c = 0) → P N') → P N) - (n : ℕ) (N : Submodule R M) - (rank_le : ∀ {m : ℕ} (v : Fin m → N), LinearIndependent R ((↑) ∘ v : Fin m → M) → m ≤ n) : - P N := by - haveI : DecidableEq M := Classical.decEq M - have Pbot : P ⊥ := by - apply ih - intro N _ x x_mem x_ortho - exfalso - rw [mem_bot] at x_mem - simpa [x_mem] using x_ortho 1 0 N.zero_mem - induction' n with n rank_ih generalizing N - · suffices N = ⊥ by rwa [this] - apply Basis.eq_bot_of_rank_eq_zero b _ fun m hv => Nat.le_zero.mp (rank_le _ hv) - apply ih - intro N' N'_le x x_mem x_ortho - apply rank_ih - intro m v hli - refine Nat.succ_le_succ_iff.mp (rank_le (Fin.cons ⟨x, x_mem⟩ fun i => ⟨v i, N'_le (v i).2⟩) ?_) - convert hli.fin_cons' x _ ?_ - · ext i - refine Fin.cases ?_ ?_ i <;> simp - · intro c y hcy - refine x_ortho c y (Submodule.span_le.mpr ?_ y.2) hcy - rintro _ ⟨z, rfl⟩ - exact (v z).2 - -end Induction - -/-- An element of a non-unital-non-associative algebra is in the center exactly when it commutes -with the basis elements. -/ -lemma Basis.mem_center_iff {A} - [Semiring R] [NonUnitalNonAssocSemiring A] - [Module R A] [SMulCommClass R A A] [SMulCommClass R R A] [IsScalarTower R A A] - (b : Basis ι R A) {z : A} : - z ∈ Set.center A ↔ - (∀ i, Commute (b i) z) ∧ ∀ i j, - z * (b i * b j) = (z * b i) * b j - ∧ (b i * z) * b j = b i * (z * b j) - ∧ (b i * b j) * z = b i * (b j * z) := by - constructor - · intro h - constructor - · intro i - apply (h.1 (b i)).symm - · intros - exact ⟨h.2 _ _, ⟨h.3 _ _, h.4 _ _⟩⟩ - · intro h - rw [center, mem_setOf_eq] - constructor - case comm => - intro y - rw [← b.total_repr y, Finsupp.total_apply, Finsupp.sum, Finset.sum_mul, Finset.mul_sum] - simp_rw [mul_smul_comm, smul_mul_assoc, (h.1 _).eq] - case left_assoc => - intro c d - rw [← b.total_repr c, ← b.total_repr d, Finsupp.total_apply, Finsupp.total_apply, Finsupp.sum, - Finsupp.sum, Finset.sum_mul, Finset.mul_sum, Finset.mul_sum, Finset.mul_sum] - simp_rw [smul_mul_assoc, Finset.mul_sum, Finset.sum_mul, mul_smul_comm, Finset.mul_sum, - Finset.smul_sum, smul_mul_assoc, mul_smul_comm, (h.2 _ _).1, - (@SMulCommClass.smul_comm R R A)] - rw [Finset.sum_comm] - case mid_assoc => - intro c d - rw [← b.total_repr c, ← b.total_repr d, Finsupp.total_apply, Finsupp.total_apply, Finsupp.sum, - Finsupp.sum, Finset.sum_mul, Finset.mul_sum, Finset.mul_sum, Finset.mul_sum] - simp_rw [smul_mul_assoc, Finset.sum_mul, mul_smul_comm, smul_mul_assoc, (h.2 _ _).2.1] - case right_assoc => - intro c d - rw [← b.total_repr c, ← b.total_repr d, Finsupp.total_apply, Finsupp.total_apply, Finsupp.sum, - Finsupp.sum, Finset.sum_mul] - simp_rw [smul_mul_assoc, Finset.mul_sum, Finset.sum_mul, mul_smul_comm, Finset.mul_sum, - Finset.smul_sum, smul_mul_assoc, mul_smul_comm, Finset.sum_mul, smul_mul_assoc, - (h.2 _ _).2.2] - -section RestrictScalars - -variable {S : Type*} [CommRing R] [Ring S] [Nontrivial S] [AddCommGroup M] -variable [Algebra R S] [Module S M] [Module R M] -variable [IsScalarTower R S M] [NoZeroSMulDivisors R S] (b : Basis ι S M) -variable (R) - -open Submodule - -/-- Let `b` be an `S`-basis of `M`. Let `R` be a CommRing such that `Algebra R S` has no zero smul -divisors, then the submodule of `M` spanned by `b` over `R` admits `b` as an `R`-basis. -/ -noncomputable def Basis.restrictScalars : Basis ι R (span R (Set.range b)) := - Basis.span (b.linearIndependent.restrict_scalars (smul_left_injective R one_ne_zero)) - -@[simp] -theorem Basis.restrictScalars_apply (i : ι) : (b.restrictScalars R i : M) = b i := by - simp only [Basis.restrictScalars, Basis.span_apply] - -@[simp] -theorem Basis.restrictScalars_repr_apply (m : span R (Set.range b)) (i : ι) : - algebraMap R S ((b.restrictScalars R).repr m i) = b.repr m i := by - suffices - Finsupp.mapRange.linearMap (Algebra.linearMap R S) ∘ₗ (b.restrictScalars R).repr.toLinearMap = - ((b.repr : M →ₗ[S] ι →₀ S).restrictScalars R).domRestrict _ - by exact DFunLike.congr_fun (LinearMap.congr_fun this m) i - refine Basis.ext (b.restrictScalars R) fun _ => ?_ - simp only [LinearMap.coe_comp, LinearEquiv.coe_toLinearMap, Function.comp_apply, map_one, - Basis.repr_self, Finsupp.mapRange.linearMap_apply, Finsupp.mapRange_single, - Algebra.linearMap_apply, LinearMap.domRestrict_apply, LinearEquiv.coe_coe, - Basis.restrictScalars_apply, LinearMap.coe_restrictScalars] - -/-- Let `b` be an `S`-basis of `M`. Then `m : M` lies in the `R`-module spanned by `b` iff all the -coordinates of `m` on the basis `b` are in `R` (see `Basis.mem_span` for the case `R = S`). -/ -theorem Basis.mem_span_iff_repr_mem (m : M) : - m ∈ span R (Set.range b) ↔ ∀ i, b.repr m i ∈ Set.range (algebraMap R S) := by - refine - ⟨fun hm i => ⟨(b.restrictScalars R).repr ⟨m, hm⟩ i, b.restrictScalars_repr_apply R ⟨m, hm⟩ i⟩, - fun h => ?_⟩ - rw [← b.total_repr m, Finsupp.total_apply S _] - refine sum_mem fun i _ => ?_ - obtain ⟨_, h⟩ := h i - simp_rw [← h, algebraMap_smul] - exact smul_mem _ _ (subset_span (Set.mem_range_self i)) - -end RestrictScalars - -section Finite - -open Basis Cardinal - -universe v v' v'' u₁' w w' - -variable {R : Type u} {M M₁ : Type v} {M' : Type v'} {ι : Type w} -variable [Ring R] [AddCommGroup M] [AddCommGroup M'] [AddCommGroup M₁] [Nontrivial R] -variable [Module R M] [Module R M'] [Module R M₁] - --- One might hope that a finite spanning set implies that any linearly independent set is finite. --- While this is true over a division ring --- (simply because any linearly independent set can be extended to a basis), --- or more generally over a ring satisfying the strong rank condition --- (which covers all commutative rings; see `LinearIndependent.finite_of_le_span_finite`), --- this is not true in general. --- For example, the left ideal generated by the variables in a noncommutative polynomial ring --- (`FreeAlgebra R ι`) in infinitely many variables (indexed by `ι`) is free --- with an infinite basis (consisting of the variables). --- As another example, for any commutative ring R, the ring of column-finite matrices --- `Module.End R (ℕ →₀ R)` is isomorphic to `ℕ → Module.End R (ℕ →₀ R)` as a module over itself, --- which also clearly contains an infinite linearly independent set. -/-- -Over any nontrivial ring, the existence of a finite spanning set implies that any basis is finite. --/ -lemma basis_finite_of_finite_spans (w : Set M) (hw : w.Finite) (s : span R w = ⊤) {ι : Type w} - (b : Basis ι R M) : Finite ι := by - classical - haveI := hw.to_subtype - cases nonempty_fintype w - -- We'll work by contradiction, assuming `ι` is infinite. - rw [← not_infinite_iff_finite] - intro i - -- Let `S` be the union of the supports of `x ∈ w` expressed as linear combinations of `b`. - -- This is a finite set since `w` is finite. - let S : Finset ι := Finset.univ.sup fun x : w => (b.repr x).support - let bS : Set M := b '' S - have h : ∀ x ∈ w, x ∈ span R bS := by - intro x m - rw [← b.total_repr x, Finsupp.span_image_eq_map_total, Submodule.mem_map] - use b.repr x - simp only [and_true_iff, eq_self_iff_true, Finsupp.mem_supported] - rw [Finset.coe_subset, ← Finset.le_iff_subset] - exact Finset.le_sup (f := fun x : w ↦ (b.repr ↑x).support) (Finset.mem_univ (⟨x, m⟩ : w)) - -- Thus this finite subset of the basis elements spans the entire module. - have k : span R bS = ⊤ := eq_top_iff.2 (le_trans s.ge (span_le.2 h)) - -- Now there is some `x : ι` not in `S`, since `ι` is infinite. - obtain ⟨x, nm⟩ := Infinite.exists_not_mem_finset S - -- However it must be in the span of the finite subset, - have k' : b x ∈ span R bS := by - rw [k] - exact mem_top - -- giving the desire contradiction. - exact b.linearIndependent.not_mem_span_image nm k' - --- From [Les familles libres maximales d'un module ont-elles le meme cardinal?][lazarus1973] -/-- Over any ring `R`, if `b` is a basis for a module `M`, -and `s` is a maximal linearly independent set, -then the union of the supports of `x ∈ s` (when written out in the basis `b`) is all of `b`. --/ -theorem union_support_maximal_linearIndependent_eq_range_basis {ι : Type w} (b : Basis ι R M) - {κ : Type w'} (v : κ → M) (i : LinearIndependent R v) (m : i.Maximal) : - ⋃ k, ((b.repr (v k)).support : Set ι) = Set.univ := by - -- If that's not the case, - by_contra h - simp only [← Ne.eq_def, ne_univ_iff_exists_not_mem, mem_iUnion, not_exists_not, - Finsupp.mem_support_iff, Finset.mem_coe] at h - -- We have some basis element `b b'` which is not in the support of any of the `v i`. - obtain ⟨b', w⟩ := h - -- Using this, we'll construct a linearly independent family strictly larger than `v`, - -- by also using this `b b'`. - let v' : Option κ → M := fun o => o.elim (b b') v - have r : range v ⊆ range v' := by - rintro - ⟨k, rfl⟩ - use some k - simp only [v', Option.elim_some] - have r' : b b' ∉ range v := by - rintro ⟨k, p⟩ - simpa [w] using congr_arg (fun m => (b.repr m) b') p - have r'' : range v ≠ range v' := by - intro e - have p : b b' ∈ range v' := by - use none - simp only [v', Option.elim_none] - rw [← e] at p - exact r' p - -- The key step in the proof is checking that this strictly larger family is linearly independent. - have i' : LinearIndependent R ((↑) : range v' → M) := by - apply LinearIndependent.to_subtype_range - rw [linearIndependent_iff] - intro l z - rw [Finsupp.total_option] at z - simp only [v', Option.elim'] at z - change _ + Finsupp.total κ M R v l.some = 0 at z - -- We have some linear combination of `b b'` and the `v i`, which we want to show is trivial. - -- We'll first show the coefficient of `b b'` is zero, - -- by expressing the `v i` in the basis `b`, and using that the `v i` have no `b b'` term. - have l₀ : l none = 0 := by - rw [← eq_neg_iff_add_eq_zero] at z - replace z := neg_eq_iff_eq_neg.mpr z - apply_fun fun x => b.repr x b' at z - simp only [repr_self, map_smul, mul_one, Finsupp.single_eq_same, Pi.neg_apply, - Finsupp.smul_single', map_neg, Finsupp.coe_neg] at z - erw [DFunLike.congr_fun (Finsupp.apply_total R (b.repr : M →ₗ[R] ι →₀ R) v l.some) b'] at z - simpa [Finsupp.total_apply, w] using z - -- Then all the other coefficients are zero, because `v` is linear independent. - have l₁ : l.some = 0 := by - rw [l₀, zero_smul, zero_add] at z - exact linearIndependent_iff.mp i _ z - -- Finally we put those facts together to show the linear combination is trivial. - ext (_ | a) - · simp only [l₀, Finsupp.coe_zero, Pi.zero_apply] - · erw [DFunLike.congr_fun l₁ a] - simp only [Finsupp.coe_zero, Pi.zero_apply] - rw [LinearIndependent.Maximal] at m - specialize m (range v') i' r - exact r'' m - -/-- Over any ring `R`, if `b` is an infinite basis for a module `M`, -and `s` is a maximal linearly independent set, -then the cardinality of `b` is bounded by the cardinality of `s`. --/ -theorem infinite_basis_le_maximal_linearIndependent' {ι : Type w} (b : Basis ι R M) [Infinite ι] - {κ : Type w'} (v : κ → M) (i : LinearIndependent R v) (m : i.Maximal) : - Cardinal.lift.{w'} #ι ≤ Cardinal.lift.{w} #κ := by - let Φ := fun k : κ => (b.repr (v k)).support - have w₁ : #ι ≤ #(Set.range Φ) := by - apply Cardinal.le_range_of_union_finset_eq_top - exact union_support_maximal_linearIndependent_eq_range_basis b v i m - have w₂ : Cardinal.lift.{w'} #(Set.range Φ) ≤ Cardinal.lift.{w} #κ := Cardinal.mk_range_le_lift - exact (Cardinal.lift_le.mpr w₁).trans w₂ - --- (See `infinite_basis_le_maximal_linearIndependent'` for the more general version --- where the index types can live in different universes.) -/-- Over any ring `R`, if `b` is an infinite basis for a module `M`, -and `s` is a maximal linearly independent set, -then the cardinality of `b` is bounded by the cardinality of `s`. --/ -theorem infinite_basis_le_maximal_linearIndependent {ι : Type w} (b : Basis ι R M) [Infinite ι] - {κ : Type w} (v : κ → M) (i : LinearIndependent R v) (m : i.Maximal) : #ι ≤ #κ := - Cardinal.lift_le.mp (infinite_basis_le_maximal_linearIndependent' b v i m) - -end Finite +assert_not_exists LinearIndependent +assert_not_exists Cardinal diff --git a/Mathlib/LinearAlgebra/Basis/Flag.lean b/Mathlib/LinearAlgebra/Basis/Flag.lean index ccc0731224c54..4ffe11d49a060 100644 --- a/Mathlib/LinearAlgebra/Basis/Flag.lean +++ b/Mathlib/LinearAlgebra/Basis/Flag.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Patrick Massot -/ -import Mathlib.LinearAlgebra.Basis +import Mathlib.LinearAlgebra.Basis.Cardinality import Mathlib.LinearAlgebra.Dual import Mathlib.Data.Fin.FlagRange diff --git a/Mathlib/LinearAlgebra/Basis/VectorSpace.lean b/Mathlib/LinearAlgebra/Basis/VectorSpace.lean index bcf4a1e2dedee..5ffcc1cc36b43 100644 --- a/Mathlib/LinearAlgebra/Basis/VectorSpace.lean +++ b/Mathlib/LinearAlgebra/Basis/VectorSpace.lean @@ -3,7 +3,6 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Alexander Bentkamp -/ -import Mathlib.LinearAlgebra.Basis import Mathlib.LinearAlgebra.FreeModule.Basic import Mathlib.LinearAlgebra.LinearPMap import Mathlib.LinearAlgebra.Projection diff --git a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean index 1b4cb8e10bb56..9e99d5b63ae90 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean @@ -5,8 +5,8 @@ Authors: Andreas Swerdlow, Kexing Ying -/ import Mathlib.LinearAlgebra.BilinearMap import Mathlib.LinearAlgebra.BilinearForm.Basic -import Mathlib.LinearAlgebra.Basis import Mathlib.Algebra.Algebra.Bilinear +import Mathlib.LinearAlgebra.Basis.Defs /-! # Bilinear form and linear maps diff --git a/Mathlib/LinearAlgebra/FreeAlgebra.lean b/Mathlib/LinearAlgebra/FreeAlgebra.lean index a8c41f4c2823e..5b221a8d639dd 100644 --- a/Mathlib/LinearAlgebra/FreeAlgebra.lean +++ b/Mathlib/LinearAlgebra/FreeAlgebra.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import Mathlib.LinearAlgebra.Basis +import Mathlib.LinearAlgebra.Basis.Cardinality import Mathlib.Algebra.FreeAlgebra import Mathlib.LinearAlgebra.FinsuppVectorSpace import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition diff --git a/Mathlib/LinearAlgebra/FreeModule/Basic.lean b/Mathlib/LinearAlgebra/FreeModule/Basic.lean index 50482a630872d..d93fc0610a6b4 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ import Mathlib.Data.Finsupp.Fintype +import Mathlib.LinearAlgebra.Basis.Basic import Mathlib.LinearAlgebra.TensorProduct.Basis /-! diff --git a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean index 2d3e7dd8b3acf..2c95ece083bf7 100644 --- a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean +++ b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen import Mathlib.Algebra.Star.Order import Mathlib.Data.Matrix.Basic import Mathlib.LinearAlgebra.StdBasis +import Mathlib.Algebra.Order.BigOperators.Group.Finset /-! # Dot product of two vectors diff --git a/Mathlib/LinearAlgebra/Multilinear/Basis.lean b/Mathlib/LinearAlgebra/Multilinear/Basis.lean index a88d7fa503423..e35c07a3fd4da 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basis.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basis.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ -import Mathlib.LinearAlgebra.Basis import Mathlib.LinearAlgebra.Multilinear.Basic +import Mathlib.LinearAlgebra.Basis.Defs /-! # Multilinear maps in relation to bases. diff --git a/Mathlib/LinearAlgebra/SesquilinearForm.lean b/Mathlib/LinearAlgebra/SesquilinearForm.lean index faece3e4cba33..0226a9d6fc42b 100644 --- a/Mathlib/LinearAlgebra/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/SesquilinearForm.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Andreas Swerdlow. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andreas Swerdlow -/ -import Mathlib.LinearAlgebra.Basis import Mathlib.LinearAlgebra.BilinearMap +import Mathlib.LinearAlgebra.Basis.Basic /-! # Sesquilinear maps diff --git a/Mathlib/LinearAlgebra/StdBasis.lean b/Mathlib/LinearAlgebra/StdBasis.lean index f3a3045b9ec79..97b36e6f8dbd0 100644 --- a/Mathlib/LinearAlgebra/StdBasis.lean +++ b/Mathlib/LinearAlgebra/StdBasis.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Data.Matrix.Basis -import Mathlib.LinearAlgebra.Basis import Mathlib.LinearAlgebra.Pi +import Mathlib.LinearAlgebra.LinearIndependent +import Mathlib.LinearAlgebra.Basis.Defs /-! # The standard basis diff --git a/Mathlib/RingTheory/Adjoin/Basic.lean b/Mathlib/RingTheory/Adjoin/Basic.lean index 74d1dcbf2f39d..2e497804ccf8a 100644 --- a/Mathlib/RingTheory/Adjoin/Basic.lean +++ b/Mathlib/RingTheory/Adjoin/Basic.lean @@ -6,11 +6,10 @@ Authors: Kenny Lau import Mathlib.Algebra.Algebra.Operations import Mathlib.Algebra.Algebra.Subalgebra.Prod import Mathlib.Algebra.Algebra.Subalgebra.Tower -import Mathlib.LinearAlgebra.Basis import Mathlib.LinearAlgebra.Prod import Mathlib.LinearAlgebra.Finsupp -import Mathlib.LinearAlgebra.Prod import Mathlib.Algebra.Module.Submodule.EqLocus +import Mathlib.LinearAlgebra.Basis.Defs /-! # Adjoining elements to form subalgebras diff --git a/Mathlib/RingTheory/AlgebraTower.lean b/Mathlib/RingTheory/AlgebraTower.lean index 48c238fa1bb95..ac5c55707ae21 100644 --- a/Mathlib/RingTheory/AlgebraTower.lean +++ b/Mathlib/RingTheory/AlgebraTower.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau -/ import Mathlib.Algebra.Algebra.Tower import Mathlib.Algebra.Module.BigOperators -import Mathlib.LinearAlgebra.Basis +import Mathlib.LinearAlgebra.Basis.Basic /-! # Towers of algebras diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index c0a3dcf78bc80..fae5966bb346c 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -10,6 +10,7 @@ import Mathlib.LinearAlgebra.StdBasis import Mathlib.GroupTheory.Finiteness import Mathlib.RingTheory.Ideal.Maps import Mathlib.RingTheory.Nilpotent.Defs +import Mathlib.LinearAlgebra.Basis.Cardinality /-! # Finiteness conditions in commutative algebra diff --git a/Mathlib/RingTheory/Ideal/Basis.lean b/Mathlib/RingTheory/Ideal/Basis.lean index 6c6efe1b3f352..70d55368a9142 100644 --- a/Mathlib/RingTheory/Ideal/Basis.lean +++ b/Mathlib/RingTheory/Ideal/Basis.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import Mathlib.Algebra.Algebra.Bilinear -import Mathlib.LinearAlgebra.Basis import Mathlib.RingTheory.Ideal.Basic +import Mathlib.LinearAlgebra.Basis.Defs /-! # The basis of ideals diff --git a/Mathlib/RingTheory/Localization/Module.lean b/Mathlib/RingTheory/Localization/Module.lean index 2af79f7be7e56..1c96e3fd8daf3 100644 --- a/Mathlib/RingTheory/Localization/Module.lean +++ b/Mathlib/RingTheory/Localization/Module.lean @@ -3,10 +3,10 @@ Copyright (c) 2022 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Junyan Xu, Anne Baanen -/ -import Mathlib.LinearAlgebra.Basis import Mathlib.Algebra.Module.LocalizedModule import Mathlib.RingTheory.Localization.FractionRing import Mathlib.RingTheory.Localization.Integer +import Mathlib.LinearAlgebra.Basis.Basic /-! # Modules / vector spaces over localizations / fraction fields diff --git a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean index c906097c8e36e..2f0b07d5dd242 100644 --- a/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean +++ b/Mathlib/RingTheory/MvPolynomial/WeightedHomogeneous.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 María Inés de Frutos-Fernández. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández -/ +import Mathlib.Algebra.BigOperators.Finprod import Mathlib.Algebra.DirectSum.Decomposition import Mathlib.Algebra.GradedMonoid import Mathlib.Algebra.MvPolynomial.Basic From 6a9aca0a5a11df04f3e97930d9ee7fc9d7a8827c Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Sat, 10 Aug 2024 06:09:17 +0000 Subject: [PATCH 159/359] chore(CategoryTheory): tidy file Sites/Discrete (#15658) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Fix typos in docstrings and use `isIso_of_reflects_iso` instead of `ReflectsIsomorphisms.reflects` --- Mathlib/CategoryTheory/Sites/Discrete.lean | 18 +++++++++--------- 1 file changed, 9 insertions(+), 9 deletions(-) diff --git a/Mathlib/CategoryTheory/Sites/Discrete.lean b/Mathlib/CategoryTheory/Sites/Discrete.lean index e2a50fdc19dd8..6998b9657e496 100644 --- a/Mathlib/CategoryTheory/Sites/Discrete.lean +++ b/Mathlib/CategoryTheory/Sites/Discrete.lean @@ -94,8 +94,9 @@ variable (A) in /-- The constant sheaf functor commutes up to isomorphism with any equivalence of sheaf categories. -This is an auxiliary definition used to prove `Sheaf.isDiscrete_iff` below, which says that the -property of a sheaf of being a discrete object is invariant under equivalence of sheaf categories. +This is an auxiliary definition used to prove `Sheaf.isDiscrete_iff_of_equivalence` below, which +says that the property of a sheaf of being a discrete object is invariant under equivalence of +sheaf categories. -/ noncomputable def equivCommuteConstant : let e : Sheaf J A ≌ Sheaf K A := @@ -110,8 +111,9 @@ variable (A) in /-- The constant sheaf functor commutes up to isomorphism with any equivalence of sheaf categories. -This is an auxiliary definition used to prove `Sheaf.isDiscrete_iff` below, which says that the -property of a sheaf of being a discrete object is invariant under equivalence of sheaf categories. +This is an auxiliary definition used to prove `Sheaf.isDiscrete_iff_of_equivalence` below, which +says that the property of a sheaf of being a discrete object is invariant under equivalence of +sheaf categories. -/ noncomputable def equivCommuteConstant' : let e : Sheaf J A ≌ Sheaf K A := @@ -258,15 +260,13 @@ lemma sheafCompose_reflects_discrete [(sheafCompose J U).ReflectsIsomorphisms] sheafToPresheaf_map, f, ← constantSheafAdj_counit_w] exact inferInstanceAs (IsIso (_ ≫ ((sheafToPresheaf J B).map ((constantSheafAdj J B ht).counit.app ((sheafCompose J U).obj F))))) - have : IsIso f := by - apply ReflectsIsomorphisms.reflects (sheafToPresheaf J B) _ - apply ReflectsIsomorphisms.reflects (sheafCompose J U) _ + have := isIso_of_reflects_iso f (sheafToPresheaf J B) + exact isIso_of_reflects_iso _ (sheafCompose J U) variable [(constantSheaf J A).Full] [(constantSheaf J A).Faithful] [(constantSheaf J B).Full] [(constantSheaf J B).Faithful] -instance [h : F.IsDiscrete J ht] : - ((sheafCompose J U).obj F).IsDiscrete J ht := by +instance [h : F.IsDiscrete J ht] : ((sheafCompose J U).obj F).IsDiscrete J ht := by rw [isDiscrete_iff_mem_essImage] at h ⊢ obtain ⟨Y, ⟨i⟩⟩ := h exact ⟨U.obj Y, ⟨(fullyFaithfulSheafToPresheaf _ _).preimageIso From a43394c090100a2ef35302ed549e9e60cac51516 Mon Sep 17 00:00:00 2001 From: adomasbaliuka <52975890+adomasbaliuka@users.noreply.github.com> Date: Sat, 10 Aug 2024 08:13:06 +0000 Subject: [PATCH 160/359] feat: Generalize statements about deriv of MulLog (#14594) Uses non-differentiability to generalize statements about deriv of MulLog --- .../SpecialFunctions/Log/NegMulLog.lean | 61 ++++++++++++++++--- Mathlib/Topology/Basic.lean | 8 +++ 2 files changed, 61 insertions(+), 8 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/Log/NegMulLog.lean b/Mathlib/Analysis/SpecialFunctions/Log/NegMulLog.lean index 4508ffb3b2d92..87070fe6bbc9d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/NegMulLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/NegMulLog.lean @@ -52,18 +52,49 @@ lemma hasDerivAt_mul_log {x : ℝ} (hx : x ≠ 0) : HasDerivAt (fun x ↦ x * lo refine DifferentiableOn.differentiableAt differentiableOn_mul_log ?_ simp [hx] -lemma deriv2_mul_log {x : ℝ} (hx : x ≠ 0) : deriv^[2] (fun x ↦ x * log x) x = x⁻¹ := by +open Filter in +private lemma tendsto_deriv_mul_log_nhdsWithin_zero : + Tendsto (deriv (fun x ↦ x * log x)) (𝓝[>] 0) atBot := by + have : (deriv (fun x ↦ x * log x)) =ᶠ[𝓝[>] 0] (fun x ↦ log x + 1) := by + apply eventuallyEq_nhdsWithin_of_eqOn + intro x hx + rw [deriv_mul_log] + simp only [Set.mem_Ioi, ne_eq] + exact ne_of_gt hx + simp only [tendsto_congr' this, tendsto_atBot_add_const_right, tendsto_log_nhdsWithin_zero_right] + +/-- At `x=0`, `(fun x ↦ x * log x)` is not differentiable +(but note that it is continuous, see `continuous_mul_log`). -/ +lemma not_DifferentiableAt_log_mul_zero : + ¬ DifferentiableAt ℝ (fun x ↦ x * log x) 0 := fun h ↦ + (not_differentiableWithinAt_of_deriv_tendsto_atBot_Ioi (fun (x:ℝ) ↦ x * log x) (a:=0)) + tendsto_deriv_mul_log_nhdsWithin_zero + (h.differentiableWithinAt (s:=(Set.Ioi 0))) + +/-- Not differentiable, hence `deriv` has junk value zero. -/ +lemma deriv_mul_log_zero : deriv (fun x ↦ x * log x) 0 = 0 := + deriv_zero_of_not_differentiableAt not_DifferentiableAt_log_mul_zero + +lemma not_continuousAt_deriv_mul_log_zero : + ¬ ContinuousAt (deriv (fun (x : ℝ) ↦ x * log x)) 0 := + not_continuousAt_of_tendsto tendsto_deriv_mul_log_nhdsWithin_zero nhdsWithin_le_nhds (by simp) + +lemma deriv2_mul_log (x : ℝ) : deriv^[2] (fun x ↦ x * log x) x = x⁻¹ := by simp only [Function.iterate_succ, Function.iterate_zero, Function.id_comp, Function.comp_apply] - suffices ∀ᶠ y in (𝓝 x), deriv (fun x ↦ x * log x) y = log y + 1 by - refine (Filter.EventuallyEq.deriv_eq this).trans ?_ - rw [deriv_add_const, deriv_log x] - filter_upwards [eventually_ne_nhds hx] with y hy using deriv_mul_log hy + by_cases hx : x ≠ 0 + · suffices ∀ᶠ y in (𝓝 x), deriv (fun x ↦ x * log x) y = log y + 1 by + refine (Filter.EventuallyEq.deriv_eq this).trans ?_ + rw [deriv_add_const, deriv_log x] + filter_upwards [eventually_ne_nhds hx] with y hy using deriv_mul_log hy + · rw [show x = 0 by simp_all only [ne_eq, Decidable.not_not], inv_zero] + exact deriv_zero_of_not_differentiableAt + (fun h ↦ not_continuousAt_deriv_mul_log_zero h.continuousAt) lemma strictConvexOn_mul_log : StrictConvexOn ℝ (Set.Ici (0 : ℝ)) (fun x ↦ x * log x) := by refine strictConvexOn_of_deriv2_pos (convex_Ici 0) (continuous_mul_log.continuousOn) ?_ intro x hx simp only [Set.nonempty_Iio, interior_Ici', Set.mem_Ioi] at hx - rw [deriv2_mul_log hx.ne'] + rw [deriv2_mul_log] positivity lemma convexOn_mul_log : ConvexOn ℝ (Set.Ici (0 : ℝ)) (fun x ↦ x * log x) := @@ -106,6 +137,20 @@ lemma continuous_negMulLog : Continuous negMulLog := by lemma differentiableOn_negMulLog : DifferentiableOn ℝ negMulLog {0}ᶜ := by simpa only [negMulLog_eq_neg] using differentiableOn_mul_log.neg +lemma differentiableAt_negMulLog_iff {x : ℝ} : DifferentiableAt ℝ negMulLog x ↔ x ≠ 0 := by + constructor + · unfold negMulLog + intro h eq0 + simp only [neg_mul, differentiableAt_neg_iff, eq0] at h + exact not_DifferentiableAt_log_mul_zero h + · intro hx + have : x ∈ ({0} : Set ℝ)ᶜ := by + simp_all only [ne_eq, Set.mem_compl_iff, Set.mem_singleton_iff, not_false_eq_true] + have := differentiableOn_negMulLog x this + apply DifferentiableWithinAt.differentiableAt (s := {0}ᶜ) + <;> simp_all only [ne_eq, Set.mem_compl_iff, Set.mem_singleton_iff, not_false_eq_true, + compl_singleton_mem_nhds_iff] + lemma deriv_negMulLog {x : ℝ} (hx : x ≠ 0) : deriv negMulLog x = - log x - 1 := by rw [negMulLog_eq_neg, deriv.neg, deriv_mul_log hx] ring @@ -115,9 +160,9 @@ lemma hasDerivAt_negMulLog {x : ℝ} (hx : x ≠ 0) : HasDerivAt negMulLog (- lo refine DifferentiableOn.differentiableAt differentiableOn_negMulLog ?_ simp [hx] -lemma deriv2_negMulLog {x : ℝ} (hx : x ≠ 0) : deriv^[2] negMulLog x = - x⁻¹ := by +lemma deriv2_negMulLog (x : ℝ) : deriv^[2] negMulLog x = - x⁻¹ := by rw [negMulLog_eq_neg] - have h := deriv2_mul_log hx + have h := deriv2_mul_log simp only [Function.iterate_succ, Function.iterate_zero, Function.id_comp, Function.comp_apply, deriv.neg', differentiableAt_id', differentiableAt_log_iff, ne_eq] at h ⊢ rw [h] diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index 3558c44c57f83..4b628fa42662d 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -1362,6 +1362,14 @@ theorem ContinuousAt.eventually_mem {f : X → Y} {x : X} (hf : ContinuousAt f x (hs : s ∈ 𝓝 (f x)) : ∀ᶠ y in 𝓝 x, f y ∈ s := hf hs +/-- If a function ``f` tends to somewhere other than `𝓝 (f x)` at `x`, +then `f` is not continuous at `x` +-/ +lemma not_continuousAt_of_tendsto {f : X → Y} {l₁ : Filter X} {l₂ : Filter Y} {x : X} + (hf : Tendsto f l₁ l₂) [l₁.NeBot] (hl₁ : l₁ ≤ 𝓝 x) (hl₂ : Disjoint (𝓝 (f x)) l₂) : + ¬ ContinuousAt f x := fun cont ↦ + (cont.mono_left hl₁).not_tendsto hl₂ hf + /-- Deprecated, please use `not_mem_tsupport_iff_eventuallyEq` instead. -/ @[deprecated (since := "2024-01-15")] theorem eventuallyEq_zero_nhds {M₀} [Zero M₀] {f : X → M₀} : From c1f09cd7292a891b7b64759c1c0a9616c906bb64 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sat, 10 Aug 2024 09:16:38 +0000 Subject: [PATCH 161/359] chore: deprecate MulEquiv.map_mul, AddEquiv.map_add (#15615) --- Mathlib/Algebra/Group/Aut.lean | 8 ++++---- Mathlib/Algebra/Group/Conj.lean | 4 ++-- Mathlib/Algebra/Group/Equiv/Basic.lean | 9 ++++++--- Mathlib/Algebra/Group/Equiv/TypeTags.lean | 16 ++++++++-------- Mathlib/Algebra/Group/Opposite.lean | 4 ++-- Mathlib/Algebra/Group/Prod.lean | 2 +- Mathlib/Algebra/Group/Submonoid/Membership.lean | 2 +- Mathlib/CategoryTheory/Conj.lean | 10 +++++----- Mathlib/Data/DFinsupp/Basic.lean | 4 ++-- Mathlib/Data/Finsupp/Multiset.lean | 2 +- Mathlib/Data/Matrix/Basic.lean | 2 +- Mathlib/Deprecated/Group.lean | 6 +++--- Mathlib/GroupTheory/Commensurable.lean | 2 +- .../GroupTheory/MonoidLocalization/Basic.lean | 4 ++-- Mathlib/GroupTheory/Subgroup/Center.lean | 2 +- .../RingTheory/UniqueFactorizationDomain.lean | 2 +- 16 files changed, 41 insertions(+), 38 deletions(-) diff --git a/Mathlib/Algebra/Group/Aut.lean b/Mathlib/Algebra/Group/Aut.lean index 3329102ff1e06..ff3f00ba4f303 100644 --- a/Mathlib/Algebra/Group/Aut.lean +++ b/Mathlib/Algebra/Group/Aut.lean @@ -106,8 +106,8 @@ instance applyMulDistribMulAction {M} [Monoid M] : MulDistribMulAction (MulAut M smul := (· <| ·) one_smul _ := rfl mul_smul _ _ _ := rfl - smul_one := MulEquiv.map_one - smul_mul := MulEquiv.map_mul + smul_one := map_one + smul_mul := map_mul @[simp] protected theorem smul_def {M} [Monoid M] (f : MulAut M) (a : M) : f • a = f a := @@ -211,8 +211,8 @@ def toPerm : AddAut A →* Equiv.Perm A where This generalizes `Function.End.applyMulAction`. -/ instance applyDistribMulAction {A} [AddMonoid A] : DistribMulAction (AddAut A) A where smul := (· <| ·) - smul_zero := AddEquiv.map_zero - smul_add := AddEquiv.map_add + smul_zero := map_zero + smul_add := map_add one_smul _ := rfl mul_smul _ _ _ := rfl diff --git a/Mathlib/Algebra/Group/Conj.lean b/Mathlib/Algebra/Group/Conj.lean index 398cbc87f6573..f73a41e75af25 100644 --- a/Mathlib/Algebra/Group/Conj.lean +++ b/Mathlib/Algebra/Group/Conj.lean @@ -86,11 +86,11 @@ theorem isConj_iff {a b : α} : IsConj a b ↔ ∃ c : α, c * a * c⁻¹ = b := -- Porting note: not in simp NF. -- @[simp] theorem conj_inv {a b : α} : (b * a * b⁻¹)⁻¹ = b * a⁻¹ * b⁻¹ := - ((MulAut.conj b).map_inv a).symm + (map_inv (MulAut.conj b) a).symm @[simp] theorem conj_mul {a b c : α} : b * a * b⁻¹ * (b * c * b⁻¹) = b * (a * c) * b⁻¹ := - ((MulAut.conj b).map_mul a c).symm + (map_mul (MulAut.conj b) a c).symm @[simp] theorem conj_pow {i : ℕ} {a b : α} : (a * b * a⁻¹) ^ i = a * b ^ i * a⁻¹ := by diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 0a4884f19e160..f39dd56b339f8 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -215,6 +215,9 @@ theorem coe_toMulHom {f : M ≃* N} : (f.toMulHom : M → N) = f := rfl protected theorem map_mul (f : M ≃* N) : ∀ x y, f (x * y) = f x * f y := _root_.map_mul f +attribute [deprecated map_mul (since := "2024-08-08")] MulEquiv.map_mul +attribute [deprecated map_add (since := "2024-08-08")] AddEquiv.map_add + /-- Makes a multiplicative isomorphism from a bijection which preserves multiplication. -/ @[to_additive "Makes an additive isomorphism from a bijection which preserves addition."] def mk' (f : M ≃ N) (h : ∀ x y, f (x * y) = f x * f y) : M ≃* N := ⟨f, h⟩ @@ -245,7 +248,7 @@ https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/!4.234183. @[to_additive] lemma symm_map_mul {M N : Type*} [Mul M] [Mul N] (h : M ≃* N) (x y : N) : h.symm (x * y) = h.symm x * h.symm y := - (h.toMulHom.inverse h.toEquiv.symm h.left_inv h.right_inv).map_mul x y + map_mul (h.toMulHom.inverse h.toEquiv.symm h.left_inv h.right_inv) x y /-- The inverse of an isomorphism is an isomorphism. -/ @[to_additive (attr := symm) "The inverse of an isomorphism is an isomorphism."] @@ -301,7 +304,7 @@ theorem refl_symm : (refl M).symm = refl M := rfl def trans (h1 : M ≃* N) (h2 : N ≃* P) : M ≃* P := { h1.toEquiv.trans h2.toEquiv with map_mul' := fun x y => show h2 (h1 (x * y)) = h2 (h1 x) * h2 (h1 y) by - rw [h1.map_mul, h2.map_mul] } + rw [map_mul, map_mul] } /-- `e.symm` is a right inverse of `e`, written as `e (e.symm y) = y`. -/ @[to_additive (attr := simp) "`e.symm` is a right inverse of `e`, written as `e (e.symm y) = y`."] @@ -543,7 +546,7 @@ def piCongrRight {η : Type*} {Ms Ns : η → Type*} [∀ j, Mul (Ms j)] [∀ j, { Equiv.piCongrRight fun j => (es j).toEquiv with toFun := fun x j => es j (x j), invFun := fun x j => (es j).symm (x j), - map_mul' := fun x y => funext fun j => (es j).map_mul (x j) (y j) } + map_mul' := fun x y => funext fun j => map_mul (es j) (x j) (y j) } @[to_additive (attr := simp)] theorem piCongrRight_refl {η : Type*} {Ms : η → Type*} [∀ j, Mul (Ms j)] : diff --git a/Mathlib/Algebra/Group/Equiv/TypeTags.lean b/Mathlib/Algebra/Group/Equiv/TypeTags.lean index a8b6f3258db96..812c4308708bd 100644 --- a/Mathlib/Algebra/Group/Equiv/TypeTags.lean +++ b/Mathlib/Algebra/Group/Equiv/TypeTags.lean @@ -22,13 +22,13 @@ def AddEquiv.toMultiplicative [AddZeroClass G] [AddZeroClass H] : invFun := AddMonoidHom.toMultiplicative f.symm.toAddMonoidHom left_inv := f.left_inv right_inv := f.right_inv - map_mul' := f.map_add } + map_mul' := map_add f } invFun f := { toFun := AddMonoidHom.toMultiplicative.symm f.toMonoidHom invFun := AddMonoidHom.toMultiplicative.symm f.symm.toMonoidHom left_inv := f.left_inv right_inv := f.right_inv - map_add' := f.map_mul } + map_add' := map_mul f } left_inv x := by ext; rfl right_inv x := by ext; rfl @@ -41,13 +41,13 @@ def MulEquiv.toAdditive [MulOneClass G] [MulOneClass H] : invFun := MonoidHom.toAdditive f.symm.toMonoidHom left_inv := f.left_inv right_inv := f.right_inv - map_add' := f.map_mul } + map_add' := map_mul f } invFun f := { toFun := MonoidHom.toAdditive.symm f.toAddMonoidHom invFun := MonoidHom.toAdditive.symm f.symm.toAddMonoidHom left_inv := f.left_inv right_inv := f.right_inv - map_mul' := f.map_add } + map_mul' := map_add f } left_inv x := by ext; rfl right_inv x := by ext; rfl @@ -60,13 +60,13 @@ def AddEquiv.toMultiplicative' [MulOneClass G] [AddZeroClass H] : invFun := AddMonoidHom.toMultiplicative'' f.symm.toAddMonoidHom left_inv := f.left_inv right_inv := f.right_inv - map_mul' := f.map_add } + map_mul' := map_add f } invFun f := { toFun := AddMonoidHom.toMultiplicative'.symm f.toMonoidHom invFun := AddMonoidHom.toMultiplicative''.symm f.symm.toMonoidHom left_inv := f.left_inv right_inv := f.right_inv - map_add' := f.map_mul } + map_add' := map_mul f } left_inv x := by ext; rfl right_inv x := by ext; rfl @@ -84,13 +84,13 @@ def AddEquiv.toMultiplicative'' [AddZeroClass G] [MulOneClass H] : invFun := AddMonoidHom.toMultiplicative' f.symm.toAddMonoidHom left_inv := f.left_inv right_inv := f.right_inv - map_mul' := f.map_add } + map_mul' := map_add f } invFun f := { toFun := AddMonoidHom.toMultiplicative''.symm f.toMonoidHom invFun := AddMonoidHom.toMultiplicative'.symm f.symm.toMonoidHom left_inv := f.left_inv right_inv := f.right_inv - map_add' := f.map_mul } + map_add' := map_mul f } left_inv x := by ext; rfl right_inv x := by ext; rfl diff --git a/Mathlib/Algebra/Group/Opposite.lean b/Mathlib/Algebra/Group/Opposite.lean index d9778571e9a0c..55cb4466fa95e 100644 --- a/Mathlib/Algebra/Group/Opposite.lean +++ b/Mathlib/Algebra/Group/Opposite.lean @@ -572,12 +572,12 @@ def MulEquiv.op {α β} [Mul α] [Mul β] : α ≃* β ≃ (αᵐᵒᵖ ≃* β { toFun := MulOpposite.op ∘ f ∘ unop, invFun := MulOpposite.op ∘ f.symm ∘ unop, left_inv := fun x => unop_injective (f.symm_apply_apply x.unop), right_inv := fun x => unop_injective (f.apply_symm_apply x.unop), - map_mul' := fun x y => unop_injective (f.map_mul y.unop x.unop) } + map_mul' := fun x y => unop_injective (map_mul f y.unop x.unop) } invFun f := { toFun := unop ∘ f ∘ MulOpposite.op, invFun := unop ∘ f.symm ∘ MulOpposite.op, left_inv := fun x => by simp, right_inv := fun x => by simp, - map_mul' := fun x y => congr_arg unop (f.map_mul (MulOpposite.op y) (MulOpposite.op x)) } + map_mul' := fun x y => congr_arg unop (map_mul f (MulOpposite.op y) (MulOpposite.op x)) } left_inv _ := rfl right_inv _ := rfl diff --git a/Mathlib/Algebra/Group/Prod.lean b/Mathlib/Algebra/Group/Prod.lean index 57a99fa3a0389..8467b413ead24 100644 --- a/Mathlib/Algebra/Group/Prod.lean +++ b/Mathlib/Algebra/Group/Prod.lean @@ -612,7 +612,7 @@ end @[to_additive prodCongr "Product of additive isomorphisms; the maps come from `Equiv.prodCongr`."] def prodCongr (f : M ≃* M') (g : N ≃* N') : M × N ≃* M' × N' := { f.toEquiv.prodCongr g.toEquiv with - map_mul' := fun _ _ => Prod.ext (f.map_mul _ _) (g.map_mul _ _) } + map_mul' := fun _ _ => Prod.ext (map_mul f _ _) (map_mul g _ _) } /-- Multiplying by the trivial monoid doesn't change the structure. -/ @[to_additive uniqueProd "Multiplying by the trivial monoid doesn't change the structure."] diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index 077f178f38d11..e3f2d53159647 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -482,7 +482,7 @@ def powLogEquiv [DecidableEq M] {n : M} (h : Function.Injective fun m : ℕ => n theorem log_mul [DecidableEq M] {n : M} (h : Function.Injective fun m : ℕ => n ^ m) (x y : powers (n : M)) : log (x * y) = log x + log y := - (powLogEquiv h).symm.map_mul x y + map_mul (powLogEquiv h).symm x y theorem log_pow_int_eq_self {x : ℤ} (h : 1 < x.natAbs) (m : ℕ) : log (pow x m) = m := (powLogEquiv (Int.pow_right_injective h)).symm_apply_apply _ diff --git a/Mathlib/CategoryTheory/Conj.lean b/Mathlib/CategoryTheory/Conj.lean index aae647372fcd2..5246270f857c8 100644 --- a/Mathlib/CategoryTheory/Conj.lean +++ b/Mathlib/CategoryTheory/Conj.lean @@ -41,11 +41,11 @@ theorem conj_apply (f : End X) : α.conj f = α.inv ≫ f ≫ α.hom := @[simp] theorem conj_comp (f g : End X) : α.conj (f ≫ g) = α.conj f ≫ α.conj g := - α.conj.map_mul g f + map_mul α.conj g f @[simp] theorem conj_id : α.conj (𝟙 X) = 𝟙 Y := - α.conj.map_one + map_one α.conj @[simp] theorem refl_conj (f : End X) : (Iso.refl X).conj f = f := by @@ -85,7 +85,7 @@ theorem trans_conjAut {Z : C} (β : Y ≅ Z) (f : Aut X) : /- Porting note (#10618): removed `@[simp]`; simp can prove this -/ theorem conjAut_mul (f g : Aut X) : α.conjAut (f * g) = α.conjAut f * α.conjAut g := - α.conjAut.map_mul f g + map_mul α.conjAut f g @[simp] theorem conjAut_trans (f g : Aut X) : α.conjAut (f ≪≫ g) = α.conjAut f ≪≫ α.conjAut g := @@ -93,11 +93,11 @@ theorem conjAut_trans (f g : Aut X) : α.conjAut (f ≪≫ g) = α.conjAut f ≪ /- Porting note (#10618): removed `@[simp]`; simp can prove this -/ theorem conjAut_pow (f : Aut X) (n : ℕ) : α.conjAut (f ^ n) = α.conjAut f ^ n := - α.conjAut.toMonoidHom.map_pow f n + map_pow α.conjAut f n /- Porting note (#10618): removed `@[simp]`; simp can prove this -/ theorem conjAut_zpow (f : Aut X) (n : ℤ) : α.conjAut (f ^ n) = α.conjAut f ^ n := - α.conjAut.toMonoidHom.map_zpow f n + map_zpow α.conjAut f n end Iso diff --git a/Mathlib/Data/DFinsupp/Basic.lean b/Mathlib/Data/DFinsupp/Basic.lean index 1715e4660a60f..4b0b99b3d2e74 100644 --- a/Mathlib/Data/DFinsupp/Basic.lean +++ b/Mathlib/Data/DFinsupp/Basic.lean @@ -1834,12 +1834,12 @@ theorem comp_liftAddHom {δ : Type*} [∀ i, AddZeroClass (β i)] [AddCommMonoid @[simp] theorem sumAddHom_zero [∀ i, AddZeroClass (β i)] [AddCommMonoid γ] : (sumAddHom fun i => (0 : β i →+ γ)) = 0 := - (liftAddHom (β := β) : (∀ i, β i →+ γ) ≃+ _).map_zero + map_zero (liftAddHom (β := β)) @[simp] theorem sumAddHom_add [∀ i, AddZeroClass (β i)] [AddCommMonoid γ] (g : ∀ i, β i →+ γ) (h : ∀ i, β i →+ γ) : (sumAddHom fun i => g i + h i) = sumAddHom g + sumAddHom h := - (liftAddHom (β := β)).map_add _ _ + map_add (liftAddHom (β := β)) _ _ @[simp] theorem sumAddHom_singleAddHom [∀ i, AddCommMonoid (β i)] : diff --git a/Mathlib/Data/Finsupp/Multiset.lean b/Mathlib/Data/Finsupp/Multiset.lean index 2be4d8e3cb641..4d4b1f940f9a8 100644 --- a/Mathlib/Data/Finsupp/Multiset.lean +++ b/Mathlib/Data/Finsupp/Multiset.lean @@ -147,7 +147,7 @@ theorem toFinsupp_apply (s : Multiset α) (a : α) : toFinsupp s a = s.count a : theorem toFinsupp_zero : toFinsupp (0 : Multiset α) = 0 := _root_.map_zero _ theorem toFinsupp_add (s t : Multiset α) : toFinsupp (s + t) = toFinsupp s + toFinsupp t := - toFinsupp.map_add s t + _root_.map_add toFinsupp s t @[simp] theorem toFinsupp_singleton (a : α) : toFinsupp ({a} : Multiset α) = Finsupp.single a 1 := by diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index 3ad3b1225f011..29c4b9256a7d2 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -1263,7 +1263,7 @@ def mapMatrix (f : α ≃+ β) : Matrix m n α ≃+ Matrix m n β := { f.toEquiv.mapMatrix with toFun := fun M => M.map f invFun := fun M => M.map f.symm - map_add' := Matrix.map_add f f.map_add } + map_add' := Matrix.map_add f (map_add f) } @[simp] theorem mapMatrix_refl : (AddEquiv.refl α).mapMatrix = AddEquiv.refl (Matrix m n α) := diff --git a/Mathlib/Deprecated/Group.lean b/Mathlib/Deprecated/Group.lean index 656630e6d541c..20326690cf9f3 100644 --- a/Mathlib/Deprecated/Group.lean +++ b/Mathlib/Deprecated/Group.lean @@ -122,7 +122,7 @@ variable {M : Type*} {N : Type*} [MulOneClass M] [MulOneClass N] /-- A multiplicative isomorphism preserves multiplication (deprecated). -/ @[to_additive "An additive isomorphism preserves addition (deprecated)."] theorem isMulHom (h : M ≃* N) : IsMulHom h := - ⟨h.map_mul⟩ + ⟨map_mul h⟩ /-- A multiplicative bijection between two monoids is a monoid hom (deprecated -- use `MulEquiv.toMonoidHom`). -/ @@ -130,7 +130,7 @@ theorem isMulHom (h : M ≃* N) : IsMulHom h := "An additive bijection between two additive monoids is an additive monoid hom (deprecated). "] theorem isMonoidHom (h : M ≃* N) : IsMonoidHom h := - { map_mul := h.map_mul + { map_mul := map_mul h map_one := h.map_one } end MulEquiv @@ -216,7 +216,7 @@ theorem MonoidHom.isGroupHom {G H : Type*} {_ : Group G} {_ : Group H} (f : G @[to_additive] theorem MulEquiv.isGroupHom {G H : Type*} {_ : Group G} {_ : Group H} (h : G ≃* H) : IsGroupHom h := - { map_mul := h.map_mul } + { map_mul := map_mul h } /-- Construct `IsGroupHom` from its only hypothesis. -/ @[to_additive "Construct `IsAddGroupHom` from its only hypothesis."] diff --git a/Mathlib/GroupTheory/Commensurable.lean b/Mathlib/GroupTheory/Commensurable.lean index bc2c93f90a19f..b084767be576a 100644 --- a/Mathlib/GroupTheory/Commensurable.lean +++ b/Mathlib/GroupTheory/Commensurable.lean @@ -57,7 +57,7 @@ def quotConjEquiv (H K : Subgroup G) (g : ConjAct G) : Quotient.congr (K.equivSMul g).toEquiv fun a b => by dsimp rw [← Quotient.eq'', ← Quotient.eq'', QuotientGroup.eq, QuotientGroup.eq, - Subgroup.mem_subgroupOf, Subgroup.mem_subgroupOf, ← MulEquiv.map_inv, ← MulEquiv.map_mul, + Subgroup.mem_subgroupOf, Subgroup.mem_subgroupOf, ← map_inv, ← map_mul, Subgroup.equivSMul_apply_coe] exact Subgroup.smul_mem_pointwise_smul_iff.symm diff --git a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean index fee25caf6f462..61f01685b604e 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Basic.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Basic.lean @@ -1133,7 +1133,7 @@ def ofMulEquivOfLocalizations (k : N ≃* P) : LocalizationMap S P := (fun v ↦ let ⟨z, hz⟩ := k.toEquiv.surjective v let ⟨x, hx⟩ := f.surj z - ⟨x, show v * k _ = k _ by rw [← hx, k.map_mul, ← hz]; rfl⟩) + ⟨x, show v * k _ = k _ by rw [← hx, map_mul, ← hz]; rfl⟩) fun x y ↦ (k.apply_eq_iff_eq.trans f.eq_iff_exists).1 @[to_additive (attr := simp)] @@ -1214,7 +1214,7 @@ def ofMulEquivOfDom {k : P ≃* M} (H : T.map k.toMonoidHom = S) : LocalizationM fun ⟨c, hc⟩ ↦ let ⟨d, hd⟩ := k.toEquiv.surjective c ⟨⟨d, H' ▸ show k d ∈ S from hd.symm ▸ c.2⟩, by - erw [← hd, ← k.map_mul, ← k.map_mul] at hc; exact k.toEquiv.injective hc⟩ + erw [← hd, ← map_mul k, ← map_mul k] at hc; exact k.toEquiv.injective hc⟩ @[to_additive (attr := simp)] theorem ofMulEquivOfDom_apply {k : P ≃* M} (H : T.map k.toMonoidHom = S) (x) : diff --git a/Mathlib/GroupTheory/Subgroup/Center.lean b/Mathlib/GroupTheory/Subgroup/Center.lean index 0f10d4480d50a..e0be786c9f89a 100644 --- a/Mathlib/GroupTheory/Subgroup/Center.lean +++ b/Mathlib/GroupTheory/Subgroup/Center.lean @@ -74,7 +74,7 @@ instance centerCharacteristic : (center G).Characteristic := by refine characteristic_iff_comap_le.mpr fun ϕ g hg => ?_ rw [mem_center_iff] intro h - rw [← ϕ.injective.eq_iff, ϕ.map_mul, ϕ.map_mul] + rw [← ϕ.injective.eq_iff, map_mul, map_mul] exact (hg.comm (ϕ h)).symm theorem _root_.CommGroup.center_eq_top {G : Type*} [CommGroup G] : center G = ⊤ := by diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index ca7bf9b421ff0..99569ecf6d0d0 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -374,7 +374,7 @@ theorem MulEquiv.uniqueFactorizationMonoid (e : α ≃* β) (hα : UniqueFactori he ▸ e.prime_iff.1 (hp c hc), Units.map e.toMonoidHom u, by - erw [Multiset.prod_hom, ← e.map_mul, h] + erw [Multiset.prod_hom, ← map_mul e, h] simp⟩ theorem MulEquiv.uniqueFactorizationMonoid_iff (e : α ≃* β) : From ee220ded614597f6ff66a559a47fb3dbe137c11f Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Sat, 10 Aug 2024 12:31:39 +0000 Subject: [PATCH 162/359] perf(Algebra.Ring.BooleanRing): scope simp theorems with weak keys but complex parameters (#15611) The keys of these are the form `HAdd.hAdd _ _ _ _ _ _` etc. but they require a typeclass search for `BooleanRing` which extends `Ring`. As they are only meant to apply to `BooleanRing`s, we add a namespace and scope them to `BooleanRing`. --- Mathlib/Algebra/Ring/BooleanRing.lean | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Ring/BooleanRing.lean b/Mathlib/Algebra/Ring/BooleanRing.lean index 5b74dec89b33c..761b053c9354f 100644 --- a/Mathlib/Algebra/Ring/BooleanRing.lean +++ b/Mathlib/Algebra/Ring/BooleanRing.lean @@ -48,18 +48,16 @@ class BooleanRing (α) extends Ring α where /-- Multiplication in a boolean ring is idempotent. -/ mul_self : ∀ a : α, a * a = a -section BooleanRing +namespace BooleanRing variable [BooleanRing α] (a b : α) instance : Std.IdempotentOp (α := α) (· * ·) := ⟨BooleanRing.mul_self⟩ -@[simp] -theorem mul_self : a * a = a := - BooleanRing.mul_self _ +attribute [scoped simp] mul_self -@[simp] +@[scoped simp] theorem add_self : a + a = 0 := by have : a + a = a + a + (a + a) := calc @@ -68,7 +66,7 @@ theorem add_self : a + a = 0 := by _ = a + a + (a + a) := by rw [mul_self] rwa [self_eq_add_left] at this -@[simp] +@[scoped simp] theorem neg_eq : -a = a := calc -a = -a + 0 := by rw [add_zero] @@ -90,14 +88,14 @@ theorem mul_add_mul : a * b + b * a = 0 := by _ = a + b + (a * b + b * a) := by abel rwa [self_eq_add_right] at this -@[simp] +@[scoped simp] theorem sub_eq_add : a - b = a + b := by rw [sub_eq_add_neg, add_right_inj, neg_eq] @[simp] theorem mul_one_add_self : a * (1 + a) = 0 := by rw [mul_add, mul_one, mul_self, add_self] -- Note [lower instance priority] -instance (priority := 100) BooleanRing.toCommRing : CommRing α := +instance (priority := 100) toCommRing : CommRing α := { (inferInstance : BooleanRing α) with mul_comm := fun a b => by rw [← add_eq_zero', mul_add_mul] } @@ -236,6 +234,8 @@ scoped[BooleanAlgebraOfBooleanRing] attribute [instance] BooleanRing.toBooleanAl end BooleanRing +open BooleanRing + instance : BooleanAlgebra (AsBoolAlg α) := @BooleanRing.toBooleanAlgebra α _ From ac1484f2267e5d4757bcc02e57c0d84d0cfee51e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 10 Aug 2024 15:38:38 +0000 Subject: [PATCH 163/359] chore: cleanup Init imports (#15638) --- Mathlib/Algebra/Group/Defs.lean | 1 + Mathlib/Init/Data/List/Basic.lean | 2 +- Mathlib/Init/Data/List/Lemmas.lean | 16 +--------------- Mathlib/Init/Data/Nat/Lemmas.lean | 1 - Mathlib/Lean/Expr/Basic.lean | 1 - Mathlib/Tactic/Core.lean | 2 ++ Mathlib/Tactic/ExtractLets.lean | 1 + Mathlib/Tactic/WLOG.lean | 1 + test/cases.lean | 1 + test/meta.lean | 2 ++ 10 files changed, 10 insertions(+), 18 deletions(-) diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index fbd37bae38782..74d035dc670c9 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -10,6 +10,7 @@ import Mathlib.Tactic.TypeStar import Mathlib.Tactic.Simps.Basic import Mathlib.Tactic.ToAdditive import Mathlib.Util.AssertExists +import Batteries.Logic /-! # Typeclasses for (semi)groups and monoids diff --git a/Mathlib/Init/Data/List/Basic.lean b/Mathlib/Init/Data/List/Basic.lean index 9d280ebec530f..f0b5deabfcfcc 100644 --- a/Mathlib/Init/Data/List/Basic.lean +++ b/Mathlib/Init/Data/List/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ import Mathlib.Data.Nat.Notation -import Batteries.Data.List.Basic +import Batteries.Tactic.Alias /-! # Note about `Mathlib/Init/` diff --git a/Mathlib/Init/Data/List/Lemmas.lean b/Mathlib/Init/Data/List/Lemmas.lean index ab9fa9af0a2c0..d3326d2be2dba 100644 --- a/Mathlib/Init/Data/List/Lemmas.lean +++ b/Mathlib/Init/Data/List/Lemmas.lean @@ -3,8 +3,8 @@ Copyright (c) 2014 Parikshit Khanna. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn -/ -import Batteries.Data.List.Lemmas import Mathlib.Tactic.Cases +import Batteries.Logic /-! # Note about `Mathlib/Init/` @@ -25,14 +25,6 @@ namespace List open Nat -/-! append -/ - -/-! length -/ - -/-! map -/ - -/-! bind -/ - /-! mem -/ theorem mem_cons_eq (a y : α) (l : List α) : (a ∈ y :: l) = (a = y ∨ a ∈ l) := @@ -46,18 +38,12 @@ theorem not_exists_mem_nil (p : α → Prop) : ¬∃ x ∈ @nil α, p x := @[deprecated (since := "2024-03-23")] alias not_bex_nil := not_exists_mem_nil @[deprecated (since := "2024-03-23")] alias bex_cons := exists_mem_cons -/-! list subset -/ --- This is relying on an automatically generated instance name from Batteries. - /-! sublists -/ alias length_le_of_sublist := Sublist.length_le -/-! filter -/ - /-! map_accumr -/ - section MapAccumr variable {φ : Type w₁} {σ : Type w₂} diff --git a/Mathlib/Init/Data/Nat/Lemmas.lean b/Mathlib/Init/Data/Nat/Lemmas.lean index aa856ee56343c..d580fa8f4aed0 100644 --- a/Mathlib/Init/Data/Nat/Lemmas.lean +++ b/Mathlib/Init/Data/Nat/Lemmas.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad -/ import Batteries.Data.Nat.Lemmas -import Batteries.WF import Mathlib.Util.AssertExists import Mathlib.Data.Nat.Notation diff --git a/Mathlib/Lean/Expr/Basic.lean b/Mathlib/Lean/Expr/Basic.lean index b7a64f8e19344..a21111425b25f 100644 --- a/Mathlib/Lean/Expr/Basic.lean +++ b/Mathlib/Lean/Expr/Basic.lean @@ -8,7 +8,6 @@ import Lean.Meta.Tactic.Rewrite import Batteries.Lean.Expr import Batteries.Data.Rat.Basic import Batteries.Data.List.Basic -import Batteries.Logic /-! # Additional operations on Expr and related types diff --git a/Mathlib/Tactic/Core.lean b/Mathlib/Tactic/Core.lean index 0b0fb7514a433..bc9892963dcee 100644 --- a/Mathlib/Tactic/Core.lean +++ b/Mathlib/Tactic/Core.lean @@ -4,7 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Aurélien Saue, Mario Carneiro -/ import Lean.Elab.PreDefinition.Basic +import Lean.Elab.Tactic.ElabTerm import Lean.Util.Paths +import Lean.Meta.Tactic.Intro import Mathlib.Lean.Expr.Basic import Batteries.Tactic.OpenPrivate diff --git a/Mathlib/Tactic/ExtractLets.lean b/Mathlib/Tactic/ExtractLets.lean index ddb3c4fd15aaf..7026754495f4a 100644 --- a/Mathlib/Tactic/ExtractLets.lean +++ b/Mathlib/Tactic/ExtractLets.lean @@ -5,6 +5,7 @@ Authors: Kyle Miller -/ import Mathlib.Lean.Expr.Basic import Mathlib.Tactic.Basic +import Batteries.Tactic.Lint.Misc /-! # The `extract_lets at` tactic diff --git a/Mathlib/Tactic/WLOG.lean b/Mathlib/Tactic/WLOG.lean index 5f2974426cc5e..54a4007761d67 100644 --- a/Mathlib/Tactic/WLOG.lean +++ b/Mathlib/Tactic/WLOG.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Johan Commelin, Reid Barton, Thomas Murrills -/ import Mathlib.Tactic.Core +import Lean.Meta.Tactic.Cases /-! diff --git a/test/cases.lean b/test/cases.lean index 7984014c12102..b199a9b8748e9 100644 --- a/test/cases.lean +++ b/test/cases.lean @@ -1,3 +1,4 @@ +import Batteries.Logic import Mathlib.Tactic.Cases import Mathlib.Init.Logic import Mathlib.Data.Nat.Notation diff --git a/test/meta.lean b/test/meta.lean index d28e69f1fedd1..5cd408d0b195a 100644 --- a/test/meta.lean +++ b/test/meta.lean @@ -1,3 +1,5 @@ +import Lean.Elab.Tactic.ElabTerm +import Lean.Elab.Tactic.Rfl import Mathlib.Lean.Expr.Basic import Mathlib.Lean.Meta.Basic import Mathlib.Tactic.Relation.Rfl From 3dd071bc2260b3cf9a71863d0dee1242fec41522 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Sat, 10 Aug 2024 17:43:53 +0000 Subject: [PATCH 164/359] chore: turn some `induction'` to `induction` (#15663) This PR turns some (soon to be deprecated?) `induction'` into `induction`. --- Mathlib/Algebra/Algebra/Basic.lean | 7 +-- Mathlib/Algebra/Associated/Basic.lean | 52 +++++++++++-------- .../Algebra/BigOperators/Group/Finset.lean | 7 +-- Mathlib/Algebra/Group/Int.lean | 2 +- Mathlib/Algebra/Group/Nat.lean | 2 +- Mathlib/Algebra/Group/Semiconj/Defs.lean | 8 +-- Mathlib/Algebra/Lie/Sl2.lean | 19 ++++--- Mathlib/Algebra/Module/Defs.lean | 6 +-- .../Algebra/Module/Submodule/LinearMap.lean | 7 +-- Mathlib/Algebra/MvPolynomial/Variables.lean | 7 +-- Mathlib/Algebra/Order/CauSeq/Basic.lean | 6 +-- Mathlib/Algebra/Order/Group/Nat.lean | 6 +-- Mathlib/Algebra/Polynomial/Basic.lean | 13 ++--- Mathlib/Algebra/Polynomial/Eval.lean | 7 ++- Mathlib/Algebra/Polynomial/FieldDivision.lean | 8 +-- Mathlib/Algebra/Polynomial/Smeval.lean | 6 +-- Mathlib/Algebra/TrivSqZeroExt.lean | 7 +-- Mathlib/RingTheory/Ideal/Basic.lean | 8 +-- Mathlib/RingTheory/MvPowerSeries/Basic.lean | 6 +-- .../RingTheory/Polynomial/Hermite/Basic.lean | 38 ++++++++------ Mathlib/RingTheory/Polynomial/Pochhammer.lean | 34 ++++++------ 21 files changed, 142 insertions(+), 114 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 8e377b67365e8..96aac9668dc57 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -123,9 +123,10 @@ theorem mul_sub_algebraMap_commutes [Ring A] [Algebra R A] (x : A) (r : R) : theorem mul_sub_algebraMap_pow_commutes [Ring A] [Algebra R A] (x : A) (r : R) (n : ℕ) : x * (x - algebraMap R A r) ^ n = (x - algebraMap R A r) ^ n * x := by - induction' n with n ih - · simp - · rw [pow_succ', ← mul_assoc, mul_sub_algebraMap_commutes, mul_assoc, ih, ← mul_assoc] + induction n with + | zero => simp + | succ n ih => + rw [pow_succ', ← mul_assoc, mul_sub_algebraMap_commutes, mul_assoc, ih, ← mul_assoc] end CommSemiring diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index 37fd87f4d697b..368bb8846445f 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -68,15 +68,17 @@ theorem not_dvd_mul {a b : α} (ha : ¬ p ∣ a) (hb : ¬ p ∣ b) : ¬ p ∣ a hp.dvd_mul.not.mpr <| not_or.mpr ⟨ha, hb⟩ theorem dvd_of_dvd_pow (hp : Prime p) {a : α} {n : ℕ} (h : p ∣ a ^ n) : p ∣ a := by - induction' n with n ih - · rw [pow_zero] at h + induction n with + | zero => + rw [pow_zero] at h have := isUnit_of_dvd_one h have := not_unit hp contradiction - rw [pow_succ'] at h - cases' dvd_or_dvd hp h with dvd_a dvd_pow - · assumption - exact ih dvd_pow + | succ n ih => + rw [pow_succ'] at h + cases' dvd_or_dvd hp h with dvd_a dvd_pow + · assumption + · exact ih dvd_pow theorem dvd_pow_iff_dvd {a : α} {n : ℕ} (hn : n ≠ 0) : p ∣ a ^ n ↔ p ∣ a := ⟨hp.dvd_of_dvd_pow, (dvd_pow · hn)⟩ @@ -123,10 +125,12 @@ theorem Prime.left_dvd_or_dvd_right_of_dvd_mul [CancelCommMonoidWithZero α] {p theorem Prime.pow_dvd_of_dvd_mul_left [CancelCommMonoidWithZero α] {p a b : α} (hp : Prime p) (n : ℕ) (h : ¬p ∣ a) (h' : p ^ n ∣ a * b) : p ^ n ∣ b := by - induction' n with n ih - · rw [pow_zero] + induction n with + | zero => + rw [pow_zero] exact one_dvd b - · obtain ⟨c, rfl⟩ := ih (dvd_trans (pow_dvd_pow p n.le_succ) h') + | succ n ih => + obtain ⟨c, rfl⟩ := ih (dvd_trans (pow_dvd_pow p n.le_succ) h') rw [pow_succ] apply mul_dvd_mul_left _ ((hp.dvd_or_dvd _).resolve_left h) rwa [← mul_dvd_mul_iff_left (pow_ne_zero n hp.ne_zero), ← pow_succ, mul_left_comm] @@ -493,9 +497,9 @@ theorem Associated.mul_mul [CommMonoid α] {a₁ a₂ b₁ b₂ : α} (h₁ : a₁ ~ᵤ b₁) (h₂ : a₂ ~ᵤ b₂) : a₁ * a₂ ~ᵤ b₁ * b₂ := (h₁.mul_right _).trans (h₂.mul_left _) theorem Associated.pow_pow [CommMonoid α] {a b : α} {n : ℕ} (h : a ~ᵤ b) : a ^ n ~ᵤ b ^ n := by - induction' n with n ih - · simp [Associated.refl] - convert h.mul_mul ih <;> rw [pow_succ'] + induction n with + | zero => simp [Associated.refl] + | succ n ih => convert h.mul_mul ih <;> rw [pow_succ'] protected theorem Associated.dvd [Monoid α] {a b : α} : a ~ᵤ b → a ∣ b := fun ⟨u, hu⟩ => ⟨u, hu.symm⟩ @@ -1130,17 +1134,19 @@ theorem pow_injective_of_not_unit [CancelCommMonoidWithZero α] {q : α} (hq : theorem dvd_prime_pow [CancelCommMonoidWithZero α] {p q : α} (hp : Prime p) (n : ℕ) : q ∣ p ^ n ↔ ∃ i ≤ n, Associated q (p ^ i) := by - induction' n with n ih generalizing q - · simp [← isUnit_iff_dvd_one, associated_one_iff_isUnit] - refine ⟨fun h => ?_, fun ⟨i, hi, hq⟩ => hq.dvd.trans (pow_dvd_pow p hi)⟩ - rw [pow_succ'] at h - rcases hp.left_dvd_or_dvd_right_of_dvd_mul h with (⟨q, rfl⟩ | hno) - · rw [mul_dvd_mul_iff_left hp.ne_zero, ih] at h - rcases h with ⟨i, hi, hq⟩ - refine ⟨i + 1, Nat.succ_le_succ hi, (hq.mul_left p).trans ?_⟩ - rw [pow_succ'] - · obtain ⟨i, hi, hq⟩ := ih.mp hno - exact ⟨i, hi.trans n.le_succ, hq⟩ + induction n generalizing q with + | zero => + simp [← isUnit_iff_dvd_one, associated_one_iff_isUnit] + | succ n ih => + refine ⟨fun h => ?_, fun ⟨i, hi, hq⟩ => hq.dvd.trans (pow_dvd_pow p hi)⟩ + rw [pow_succ'] at h + rcases hp.left_dvd_or_dvd_right_of_dvd_mul h with (⟨q, rfl⟩ | hno) + · rw [mul_dvd_mul_iff_left hp.ne_zero, ih] at h + rcases h with ⟨i, hi, hq⟩ + refine ⟨i + 1, Nat.succ_le_succ hi, (hq.mul_left p).trans ?_⟩ + rw [pow_succ'] + · obtain ⟨i, hi, hq⟩ := ih.mp hno + exact ⟨i, hi.trans n.le_succ, hq⟩ end CancelCommMonoidWithZero diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 16c37fd1f67ba..f6842570c540f 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1504,9 +1504,10 @@ lemma prod_powersetCard (n : ℕ) (s : Finset α) (f : ℕ → β) : @[to_additive] theorem prod_flip {n : ℕ} (f : ℕ → β) : (∏ r ∈ range (n + 1), f (n - r)) = ∏ k ∈ range (n + 1), f k := by - induction' n with n ih - · rw [prod_range_one, prod_range_one] - · rw [prod_range_succ', prod_range_succ _ (Nat.succ n)] + induction n with + | zero => rw [prod_range_one, prod_range_one] + | succ n ih => + rw [prod_range_succ', prod_range_succ _ (Nat.succ n)] simp [← ih] @[to_additive] diff --git a/Mathlib/Algebra/Group/Int.lean b/Mathlib/Algebra/Group/Int.lean index 2a7f13beea4d7..fbdbbfb08b060 100644 --- a/Mathlib/Algebra/Group/Int.lean +++ b/Mathlib/Algebra/Group/Int.lean @@ -207,7 +207,7 @@ lemma even_sub : Even (m - n) ↔ (Even m ↔ Even n) := by simp [sub_eq_add_neg simp [even_iff, h₁, h₂, Int.mul_emod] @[parity_simps] lemma even_pow {n : ℕ} : Even (m ^ n) ↔ Even m ∧ n ≠ 0 := by - induction' n with n ih <;> simp [*, even_mul, pow_succ]; tauto + induction n <;> simp [*, even_mul, pow_succ]; tauto lemma even_pow' {n : ℕ} (h : n ≠ 0) : Even (m ^ n) ↔ Even m := even_pow.trans <| and_iff_left h diff --git a/Mathlib/Algebra/Group/Nat.lean b/Mathlib/Algebra/Group/Nat.lean index 93791bbb0e667..c9678bda255d1 100644 --- a/Mathlib/Algebra/Group/Nat.lean +++ b/Mathlib/Algebra/Group/Nat.lean @@ -123,7 +123,7 @@ lemma two_not_dvd_two_mul_sub_one : ∀ {n}, 0 < n → ¬2 ∣ 2 * n - 1 /-- If `m` and `n` are natural numbers, then the natural number `m^n` is even if and only if `m` is even and `n` is positive. -/ @[parity_simps] lemma even_pow : Even (m ^ n) ↔ Even m ∧ n ≠ 0 := by - induction' n with n ih <;> simp (config := { contextual := true }) [*, pow_succ', even_mul] + induction n <;> simp (config := { contextual := true }) [*, pow_succ', even_mul] lemma even_pow' (h : n ≠ 0) : Even (m ^ n) ↔ Even m := even_pow.trans <| and_iff_left h diff --git a/Mathlib/Algebra/Group/Semiconj/Defs.lean b/Mathlib/Algebra/Group/Semiconj/Defs.lean index f20bc77af2491..bcc5e310fabbb 100644 --- a/Mathlib/Algebra/Group/Semiconj/Defs.lean +++ b/Mathlib/Algebra/Group/Semiconj/Defs.lean @@ -104,10 +104,12 @@ variable [Monoid M] @[to_additive (attr := simp)] theorem pow_right {a x y : M} (h : SemiconjBy a x y) (n : ℕ) : SemiconjBy a (x ^ n) (y ^ n) := by - induction' n with n ih - · rw [pow_zero, pow_zero] + induction n with + | zero => + rw [pow_zero, pow_zero] exact SemiconjBy.one_right _ - · rw [pow_succ, pow_succ] + | succ n ih => + rw [pow_succ, pow_succ] exact ih.mul_right h end Monoid diff --git a/Mathlib/Algebra/Lie/Sl2.lean b/Mathlib/Algebra/Lie/Sl2.lean index 765625e1d8442..e0da8e26ab1ff 100644 --- a/Mathlib/Algebra/Lie/Sl2.lean +++ b/Mathlib/Algebra/Lie/Sl2.lean @@ -99,9 +99,10 @@ local notation "ψ" n => ((toEnd R L M f) ^ n) m lemma lie_h_pow_toEnd_f (n : ℕ) : ⁅h, ψ n⁆ = (μ - 2 * n) • ψ n := by - induction' n with n ih - · simpa using P.lie_h - · rw [pow_succ', LinearMap.mul_apply, toEnd_apply_apply, Nat.cast_add, Nat.cast_one, + induction n with + | zero => simpa using P.lie_h + | succ n ih => + rw [pow_succ', LinearMap.mul_apply, toEnd_apply_apply, Nat.cast_add, Nat.cast_one, leibniz_lie h, t.lie_lie_smul_f R, ← neg_smul, ih, lie_smul, smul_lie, ← add_smul] congr ring @@ -116,11 +117,13 @@ lemma lie_f_pow_toEnd_f (P : HasPrimitiveVectorWith t m μ) (n : ℕ) : lemma lie_e_pow_succ_toEnd_f (n : ℕ) : ⁅e, ψ (n + 1)⁆ = ((n + 1) * (μ - n)) • ψ n := by - induction' n with n ih - · simp only [zero_add, pow_one, toEnd_apply_apply, Nat.cast_zero, sub_zero, one_mul, - pow_zero, LinearMap.one_apply, leibniz_lie e, t.lie_e_f, P.lie_e, P.lie_h, lie_zero, - add_zero] - · rw [pow_succ', LinearMap.mul_apply, toEnd_apply_apply, leibniz_lie e, t.lie_e_f, + induction n with + | zero => + simp only [zero_add, pow_one, toEnd_apply_apply, Nat.cast_zero, sub_zero, one_mul, + pow_zero, LinearMap.one_apply, leibniz_lie e, t.lie_e_f, P.lie_e, P.lie_h, lie_zero, + add_zero] + | succ n ih => + rw [pow_succ', LinearMap.mul_apply, toEnd_apply_apply, leibniz_lie e, t.lie_e_f, lie_h_pow_toEnd_f P, ih, lie_smul, lie_f_pow_toEnd_f P, ← add_smul, Nat.cast_add, Nat.cast_one] congr diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index 86d77ca8b94dd..058a63cdb7344 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -327,9 +327,9 @@ variable (R) /-- `nsmul` is equal to any other module structure via a cast. -/ lemma Nat.cast_smul_eq_nsmul (n : ℕ) (b : M) : (n : R) • b = n • b := by - induction' n with n ih - · rw [Nat.cast_zero, zero_smul, zero_smul] - · rw [Nat.cast_succ, add_smul, add_smul, one_smul, ih, one_smul] + induction n with + | zero => rw [Nat.cast_zero, zero_smul, zero_smul] + | succ n ih => rw [Nat.cast_succ, add_smul, add_smul, one_smul, ih, one_smul] /-- `nsmul` is equal to any other module structure via a cast. -/ -- See note [no_index around OfNat.ofNat] diff --git a/Mathlib/Algebra/Module/Submodule/LinearMap.lean b/Mathlib/Algebra/Module/Submodule/LinearMap.lean index e66221d177479..4ff64ac81d3b1 100644 --- a/Mathlib/Algebra/Module/Submodule/LinearMap.lean +++ b/Mathlib/Algebra/Module/Submodule/LinearMap.lean @@ -229,9 +229,10 @@ variable {f' : M →ₗ[R] M} theorem pow_apply_mem_of_forall_mem {p : Submodule R M} (n : ℕ) (h : ∀ x ∈ p, f' x ∈ p) (x : M) (hx : x ∈ p) : (f' ^ n) x ∈ p := by - induction' n with n ih generalizing x - · simpa - · simpa only [iterate_succ, coe_comp, Function.comp_apply, restrict_apply] using ih _ (h _ hx) + induction n generalizing x with + | zero => simpa + | succ n ih => + simpa only [iterate_succ, coe_comp, Function.comp_apply, restrict_apply] using ih _ (h _ hx) theorem pow_restrict {p : Submodule R M} (n : ℕ) (h : ∀ x ∈ p, f' x ∈ p) (h' := pow_apply_mem_of_forall_mem n h) : diff --git a/Mathlib/Algebra/MvPolynomial/Variables.lean b/Mathlib/Algebra/MvPolynomial/Variables.lean index b018b177ef9e9..f431fe8b5f796 100644 --- a/Mathlib/Algebra/MvPolynomial/Variables.lean +++ b/Mathlib/Algebra/MvPolynomial/Variables.lean @@ -117,9 +117,10 @@ theorem vars_one : (1 : MvPolynomial σ R).vars = ∅ := theorem vars_pow (φ : MvPolynomial σ R) (n : ℕ) : (φ ^ n).vars ⊆ φ.vars := by classical - induction' n with n ih - · simp - · rw [pow_succ'] + induction n with + | zero => simp + | succ n ih => + rw [pow_succ'] apply Finset.Subset.trans (vars_mul _ _) exact Finset.union_subset (Finset.Subset.refl _) ih diff --git a/Mathlib/Algebra/Order/CauSeq/Basic.lean b/Mathlib/Algebra/Order/CauSeq/Basic.lean index baf08077c3cf0..3688f98cebaca 100644 --- a/Mathlib/Algebra/Order/CauSeq/Basic.lean +++ b/Mathlib/Algebra/Order/CauSeq/Basic.lean @@ -516,9 +516,9 @@ theorem smul_equiv_smul {G : Type*} [SMul G β] [IsScalarTower G β β] {f1 f2 : mul_equiv_mul (const_equiv.mpr <| Eq.refl <| c • (1 : β)) hf theorem pow_equiv_pow {f1 f2 : CauSeq β abv} (hf : f1 ≈ f2) (n : ℕ) : f1 ^ n ≈ f2 ^ n := by - induction' n with n ih - · simp only [Nat.zero_eq, pow_zero, Setoid.refl] - · simpa only [pow_succ'] using mul_equiv_mul hf ih + induction n with + | zero => simp only [Nat.zero_eq, pow_zero, Setoid.refl] + | succ n ih => simpa only [pow_succ'] using mul_equiv_mul hf ih end Ring diff --git a/Mathlib/Algebra/Order/Group/Nat.lean b/Mathlib/Algebra/Order/Group/Nat.lean index ae2f1465b7ed6..ce5bf4a2722b7 100644 --- a/Mathlib/Algebra/Order/Group/Nat.lean +++ b/Mathlib/Algebra/Order/Group/Nat.lean @@ -39,9 +39,9 @@ instance instLinearOrderedCancelAddCommMonoid : LinearOrderedCancelAddCommMonoid instance instOrderedSub : OrderedSub ℕ := by refine ⟨fun m n k ↦ ?_⟩ - induction' n with n ih generalizing k - · simp - · simp only [sub_succ, pred_le_iff, ih, succ_add, add_succ] + induction n generalizing k with + | zero => simp + | succ n ih => simp only [sub_succ, pred_le_iff, ih, succ_add, add_succ] /-! ### Miscellaneous lemmas -/ diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index fbfa554def41c..095de48696935 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -478,9 +478,9 @@ theorem monomial_one_one_eq_X : monomial 1 (1 : R) = X := rfl theorem monomial_one_right_eq_X_pow (n : ℕ) : monomial n (1 : R) = X ^ n := by - induction' n with n ih - · simp [monomial_zero_one] - · rw [pow_succ, ← ih, ← monomial_one_one_eq_X, monomial_mul_monomial, mul_one] + induction n with + | zero => simp [monomial_zero_one] + | succ n ih => rw [pow_succ, ← ih, ← monomial_one_one_eq_X, monomial_mul_monomial, mul_one] @[simp] theorem toFinsupp_X : X.toFinsupp = Finsupp.single 1 (1 : R) := @@ -496,9 +496,10 @@ theorem X_mul : X * p = p * X := by simp [AddMonoidAlgebra.mul_apply, AddMonoidAlgebra.sum_single_index, add_comm] theorem X_pow_mul {n : ℕ} : X ^ n * p = p * X ^ n := by - induction' n with n ih - · simp - · conv_lhs => rw [pow_succ] + induction n with + | zero => simp + | succ n ih => + conv_lhs => rw [pow_succ] rw [mul_assoc, X_mul, ← mul_assoc, ih, mul_assoc, ← pow_succ] /-- Prefer putting constants to the left of `X`. diff --git a/Mathlib/Algebra/Polynomial/Eval.lean b/Mathlib/Algebra/Polynomial/Eval.lean index a2dd15cfd8666..026dd4a013229 100644 --- a/Mathlib/Algebra/Polynomial/Eval.lean +++ b/Mathlib/Algebra/Polynomial/Eval.lean @@ -104,10 +104,9 @@ def eval₂AddMonoidHom : R[X] →+ S where @[simp] theorem eval₂_natCast (n : ℕ) : (n : R[X]).eval₂ f x = n := by - induction' n with n ih - -- Porting note: `Nat.zero_eq` is required. - · simp only [eval₂_zero, Nat.cast_zero, Nat.zero_eq] - · rw [n.cast_succ, eval₂_add, ih, eval₂_one, n.cast_succ] + induction n with + | zero => simp only [eval₂_zero, Nat.cast_zero] + | succ n ih => rw [n.cast_succ, eval₂_add, ih, eval₂_one, n.cast_succ] @[deprecated (since := "2024-04-17")] alias eval₂_nat_cast := eval₂_natCast diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index cfabc410defc9..f77638479e29a 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -93,10 +93,12 @@ theorem lt_rootMultiplicity_of_isRoot_iterate_derivative_of_mem_nonZeroDivisors' n < p.rootMultiplicity t := by apply lt_rootMultiplicity_of_isRoot_iterate_derivative_of_mem_nonZeroDivisors h hroot clear hroot - induction' n with n ih - · simp only [Nat.zero_eq, Nat.factorial_zero, Nat.cast_one] + induction n with + | zero => + simp only [Nat.zero_eq, Nat.factorial_zero, Nat.cast_one] exact Submonoid.one_mem _ - · rw [Nat.factorial_succ, Nat.cast_mul, mul_mem_nonZeroDivisors] + | succ n ih => + rw [Nat.factorial_succ, Nat.cast_mul, mul_mem_nonZeroDivisors] exact ⟨hnzd _ le_rfl n.succ_ne_zero, ih fun m h ↦ hnzd m (h.trans n.le_succ)⟩ theorem lt_rootMultiplicity_iff_isRoot_iterate_derivative_of_mem_nonZeroDivisors diff --git a/Mathlib/Algebra/Polynomial/Smeval.lean b/Mathlib/Algebra/Polynomial/Smeval.lean index 4f3e4012ca414..6db5f23f5b21e 100644 --- a/Mathlib/Algebra/Polynomial/Smeval.lean +++ b/Mathlib/Algebra/Polynomial/Smeval.lean @@ -107,9 +107,9 @@ theorem smeval_add : (p + q).smeval x = p.smeval x + q.smeval x := by · rw [smul_pow, smul_pow, smul_pow, add_smul] theorem smeval_natCast (n : ℕ) : (n : R[X]).smeval x = n • x ^ 0 := by - induction' n with n ih - · simp only [smeval_zero, Nat.cast_zero, Nat.zero_eq, zero_smul] - · rw [n.cast_succ, smeval_add, ih, smeval_one, ← add_nsmul] + induction n with + | zero => simp only [smeval_zero, Nat.cast_zero, Nat.zero_eq, zero_smul] + | succ n ih => rw [n.cast_succ, smeval_add, ih, smeval_one, ← add_nsmul] @[deprecated (since := "2024-04-17")] alias smeval_nat_cast := smeval_natCast diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index a265c07f161ca..56213cf92612d 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -585,9 +585,10 @@ theorem snd_pow_of_smul_comm [Monoid R] [AddMonoid M] [DistribMulAction R M] where aux : ∀ n : ℕ, x.snd <• x.fst ^ n = x.fst ^ n •> x.snd := by intro n - induction' n with n ih - · simp - · rw [pow_succ, op_mul, mul_smul, mul_smul, ← h, smul_comm (_ : R) (op x.fst) x.snd, ih] + induction n with + | zero => simp + | succ n ih => + rw [pow_succ, op_mul, mul_smul, mul_smul, ← h, smul_comm (_ : R) (op x.fst) x.snd, ih] theorem snd_pow_of_smul_comm' [Monoid R] [AddMonoid M] [DistribMulAction R M] [DistribMulAction Rᵐᵒᵖ M] [SMulCommClass R Rᵐᵒᵖ M] (x : tsze R M) (n : ℕ) diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 689fa8e0c1472..bb631626fa74b 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -221,10 +221,12 @@ theorem IsPrime.mem_or_mem_of_mul_eq_zero {I : Ideal α} (hI : I.IsPrime) {x y : theorem IsPrime.mem_of_pow_mem {I : Ideal α} (hI : I.IsPrime) {r : α} (n : ℕ) (H : r ^ n ∈ I) : r ∈ I := by - induction' n with n ih - · rw [pow_zero] at H + induction n with + | zero => + rw [pow_zero] at H exact (mt (eq_top_iff_one _).2 hI.1).elim H - · rw [pow_succ] at H + | succ n ih => + rw [pow_succ] at H exact Or.casesOn (hI.mem_or_mem H) ih id theorem not_isPrime_iff {I : Ideal α} : diff --git a/Mathlib/RingTheory/MvPowerSeries/Basic.lean b/Mathlib/RingTheory/MvPowerSeries/Basic.lean index 626b2270b0471..f05450e9a3df8 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Basic.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Basic.lean @@ -384,9 +384,9 @@ theorem X_def (s : σ) : X s = monomial R (single s 1) 1 := rfl theorem X_pow_eq (s : σ) (n : ℕ) : (X s : MvPowerSeries σ R) ^ n = monomial R (single s n) 1 := by - induction' n with n ih - · simp - · rw [pow_succ, ih, Finsupp.single_add, X, monomial_mul_monomial, one_mul] + induction n with + | zero => simp + | succ n ih => rw [pow_succ, ih, Finsupp.single_add, X, monomial_mul_monomial, one_mul] theorem coeff_X_pow [DecidableEq σ] (m : σ →₀ ℕ) (s : σ) (n : ℕ) : coeff R m ((X s : MvPowerSeries σ R) ^ n) = if m = single s n then 1 else 0 := by diff --git a/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean b/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean index e4599fc2385e4..a4a78d9406f1c 100644 --- a/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean @@ -35,7 +35,6 @@ This file defines `Polynomial.hermite n`, the `n`th probabilists' Hermite polyno -/ - noncomputable section open Polynomial @@ -53,9 +52,9 @@ theorem hermite_succ (n : ℕ) : hermite (n + 1) = X * hermite n - derivative (h rw [hermite] theorem hermite_eq_iterate (n : ℕ) : hermite n = (fun p => X * p - derivative p)^[n] 1 := by - induction' n with n ih - · rfl - · rw [Function.iterate_succ_apply', ← ih, hermite_succ] + induction n with + | zero => rfl + | succ n ih => rw [Function.iterate_succ_apply', ← ih, hermite_succ] @[simp] theorem hermite_zero : hermite 0 = C 1 := @@ -83,17 +82,18 @@ theorem coeff_hermite_succ_succ (n k : ℕ) : coeff (hermite (n + 1)) (k + 1) = theorem coeff_hermite_of_lt {n k : ℕ} (hnk : n < k) : coeff (hermite n) k = 0 := by obtain ⟨k, rfl⟩ := Nat.exists_eq_add_of_lt hnk clear hnk - induction' n with n ih generalizing k - · apply coeff_C - · have : n + k + 1 + 2 = n + (k + 2) + 1 := by ring - rw [coeff_hermite_succ_succ, add_right_comm, this, ih k, ih (k + 2), - mul_zero, sub_zero] + induction n generalizing k with + | zero => exact coeff_C + | succ n ih => + have : n + k + 1 + 2 = n + (k + 2) + 1 := by ring + rw [coeff_hermite_succ_succ, add_right_comm, this, ih k, ih (k + 2), mul_zero, sub_zero] @[simp] theorem coeff_hermite_self (n : ℕ) : coeff (hermite n) n = 1 := by - induction' n with n ih - · apply coeff_C - · rw [coeff_hermite_succ_succ, ih, coeff_hermite_of_lt, mul_zero, sub_zero] + induction n with + | zero => exact coeff_C + | succ n ih => + rw [coeff_hermite_succ_succ, ih, coeff_hermite_of_lt, mul_zero, sub_zero] simp @[simp] @@ -116,13 +116,17 @@ theorem hermite_monic (n : ℕ) : (hermite n).Monic := leadingCoeff_hermite n theorem coeff_hermite_of_odd_add {n k : ℕ} (hnk : Odd (n + k)) : coeff (hermite n) k = 0 := by - induction' n with n ih generalizing k - · rw [zero_add k] at hnk + induction n generalizing k with + | zero => + rw [zero_add k] at hnk exact coeff_hermite_of_lt hnk.pos - · cases' k with k - · rw [Nat.succ_add_eq_add_succ] at hnk + | succ n ih => + cases k with + | zero => + rw [Nat.succ_add_eq_add_succ] at hnk rw [coeff_hermite_succ_zero, ih hnk, neg_zero] - · rw [coeff_hermite_succ_succ, ih, ih, mul_zero, sub_zero] + | succ k => + rw [coeff_hermite_succ_succ, ih, ih, mul_zero, sub_zero] · rwa [Nat.succ_add_eq_add_succ] at hnk · rw [(by rw [Nat.succ_add, Nat.add_succ] : n.succ + k.succ = n + k + 2)] at hnk exact (Nat.odd_add.mp hnk).mpr even_two diff --git a/Mathlib/RingTheory/Polynomial/Pochhammer.lean b/Mathlib/RingTheory/Polynomial/Pochhammer.lean index ee2f504310659..63e7f8d15a3bd 100644 --- a/Mathlib/RingTheory/Polynomial/Pochhammer.lean +++ b/Mathlib/RingTheory/Polynomial/Pochhammer.lean @@ -76,9 +76,9 @@ variable {S} {T : Type v} [Semiring T] @[simp] theorem ascPochhammer_map (f : S →+* T) (n : ℕ) : (ascPochhammer S n).map f = ascPochhammer T n := by - induction' n with n ih - · simp - · simp [ih, ascPochhammer_succ_left, map_comp] + induction n with + | zero => simp + | succ n ih => simp [ih, ascPochhammer_succ_left, map_comp] theorem ascPochhammer_eval₂ (f : S →+* T) (n : ℕ) (t : T) : (ascPochhammer T n).eval t = (ascPochhammer S n).eval₂ f t := by @@ -116,9 +116,10 @@ theorem ascPochhammer_succ_right (n : ℕ) : apply_fun Polynomial.map (algebraMap ℕ S) at h simpa only [ascPochhammer_map, Polynomial.map_mul, Polynomial.map_add, map_X, Polynomial.map_natCast] using h - induction' n with n ih - · simp - · conv_lhs => + induction n with + | zero => simp + | succ n ih => + conv_lhs => rw [ascPochhammer_succ_left, ih, mul_comp, ← mul_assoc, ← ascPochhammer_succ_left, add_comp, X_comp, natCast_comp, add_assoc, add_comm (1 : ℕ[X]), ← Nat.cast_succ] @@ -175,10 +176,12 @@ section StrictOrderedSemiring variable {S : Type*} [StrictOrderedSemiring S] theorem ascPochhammer_pos (n : ℕ) (s : S) (h : 0 < s) : 0 < (ascPochhammer S n).eval s := by - induction' n with n ih - · simp only [Nat.zero_eq, ascPochhammer_zero, eval_one] + induction n with + | zero => + simp only [Nat.zero_eq, ascPochhammer_zero, eval_one] exact zero_lt_one - · rw [ascPochhammer_succ_right, mul_add, eval_add, ← Nat.cast_comm, eval_natCast_mul, eval_mul_X, + | succ n ih => + rw [ascPochhammer_succ_right, mul_add, eval_add, ← Nat.cast_comm, eval_natCast_mul, eval_mul_X, Nat.cast_comm, ← mul_add] exact mul_pos ih (lt_of_lt_of_le h (le_add_of_nonneg_right (Nat.cast_nonneg n))) @@ -252,9 +255,9 @@ variable {R} {T : Type v} [Ring T] @[simp] theorem descPochhammer_map (f : R →+* T) (n : ℕ) : (descPochhammer R n).map f = descPochhammer T n := by - induction' n with n ih - · simp - · simp [ih, descPochhammer_succ_left, map_comp] + induction n with + | zero => simp + | succ n ih => simp [ih, descPochhammer_succ_left, map_comp] end @[simp, norm_cast] @@ -281,9 +284,10 @@ theorem descPochhammer_succ_right (n : ℕ) : apply_fun Polynomial.map (algebraMap ℤ R) at h simpa [descPochhammer_map, Polynomial.map_mul, Polynomial.map_add, map_X, Polynomial.map_intCast] using h - induction' n with n ih - · simp [descPochhammer] - · conv_lhs => + induction n with + | zero => simp [descPochhammer] + | succ n ih => + conv_lhs => rw [descPochhammer_succ_left, ih, mul_comp, ← mul_assoc, ← descPochhammer_succ_left, sub_comp, X_comp, natCast_comp] rw [Nat.cast_add, Nat.cast_one, sub_add_eq_sub_sub_swap] From 0077925cbead60622b55b430eb5109f7365954c8 Mon Sep 17 00:00:00 2001 From: Scott Carnahan <128885296+ScottCarnahan@users.noreply.github.com> Date: Sat, 10 Aug 2024 20:14:37 +0000 Subject: [PATCH 165/359] feat (LinearAlgebra/RootSystem) : infinitely many roots when two are parallel (#15582) This PR has a proof that if two roots are linearly independent and "parallel" (meaning their pairings multiply to 4), then the root system containing them is infinite. This is mostly done by refactoring an existing proof. --- Mathlib/Algebra/Module/Torsion.lean | 19 ++++++++ Mathlib/LinearAlgebra/Reflection.lean | 54 +++++++++++++-------- Mathlib/LinearAlgebra/RootSystem/Basic.lean | 22 +++++++-- 3 files changed, 71 insertions(+), 24 deletions(-) diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index e0b05d0b86998..4fe308cff7cb6 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -860,3 +860,22 @@ theorem isTorsion_iff_isTorsion_int [AddCommGroup M] : exact ⟨_, Int.natAbs_pos.2 (nonZeroDivisors.coe_ne_zero n), natAbs_nsmul_eq_zero.2 hn⟩ end AddMonoid + +section InfiniteRange + +@[simp] +lemma infinite_range_add_smul_iff + [AddCommGroup M] [Ring R] [Module R M] [Infinite R] [NoZeroSMulDivisors R M] (x y : M) : + (Set.range <| fun r : R ↦ x + r • y).Infinite ↔ y ≠ 0 := by + refine ⟨fun h hy ↦ by simp [hy] at h, fun h ↦ Set.infinite_range_of_injective fun r s hrs ↦ ?_⟩ + rw [add_right_inj] at hrs + exact smul_left_injective _ h hrs + +@[simp] +lemma infinite_range_add_nsmul_iff [AddCommGroup M] [NoZeroSMulDivisors ℤ M] (x y : M) : + (Set.range <| fun n : ℕ ↦ x + n • y).Infinite ↔ y ≠ 0 := by + refine ⟨fun h hy ↦ by simp [hy] at h, fun h ↦ Set.infinite_range_of_injective fun r s hrs ↦ ?_⟩ + rw [add_right_inj, ← natCast_zsmul, ← natCast_zsmul] at hrs + simpa using smul_left_injective _ h hrs + +end InfiniteRange diff --git a/Mathlib/LinearAlgebra/Reflection.lean b/Mathlib/LinearAlgebra/Reflection.lean index a6fec257edbdc..e6bde244d8722 100644 --- a/Mathlib/LinearAlgebra/Reflection.lean +++ b/Mathlib/LinearAlgebra/Reflection.lean @@ -174,33 +174,45 @@ lemma Dual.eq_of_preReflection_mapsTo' [CharZero R] [NoZeroSMulDivisors R M] variable {y} variable {g : Dual R M} +/-- Composite of reflections in "parallel" hyperplanes is a shear (special case). -/ +lemma reflection_reflection_iterate + (hfx : f x = 2) (hgy : g y = 2) (hgxfy : f y * g x = 4) (n : ℕ) : + ((reflection hgy).trans (reflection hfx))^[n] y = y + n • (f y • x - (2 : R) • y) := by + induction n with + | zero => simp + | succ n ih => + have hz : ∀ z : M, f y • g x • z = 2 • 2 • z := by + intro z + rw [smul_smul, hgxfy, smul_smul, ← Nat.cast_smul_eq_nsmul R (2 * 2), Nat.cast_eq_ofNat] + simp only [iterate_succ', comp_apply, ih, two_smul, smul_sub, smul_add, map_add, + LinearEquiv.trans_apply, reflection_apply_self, map_neg, reflection_apply, neg_sub, map_sub, + map_nsmul, map_smul, smul_neg, hz, add_smul] + abel + +lemma infinite_range_reflection_reflection_iterate_iff [NoZeroSMulDivisors ℤ M] + (hfx : f x = 2) (hgy : g y = 2) (hgxfy : f y * g x = 4) : + (range <| fun n ↦ ((reflection hgy).trans (reflection hfx))^[n] y).Infinite ↔ + f y • x ≠ (2 : R) • y := by + simp only [reflection_reflection_iterate hfx hgy hgxfy, infinite_range_add_nsmul_iff, sub_ne_zero] + lemma eq_of_mapsTo_reflection_of_mem [NoZeroSMulDivisors ℤ M] {Φ : Set M} (hΦ : Φ.Finite) (hfx : f x = 2) (hgy : g y = 2) (hgx : g x = 2) (hfy : f y = 2) (hxfΦ : MapsTo (preReflection x f) Φ Φ) (hygΦ : MapsTo (preReflection y g) Φ Φ) (hyΦ : y ∈ Φ) : x = y := by - rw [← finite_coe_iff] at hΦ - set sxy : M ≃ₗ[R] M := (Module.reflection hgy).trans (Module.reflection hfx) - have hb : BijOn sxy Φ Φ := - (bijOn_reflection_of_mapsTo hfx hxfΦ).comp (bijOn_reflection_of_mapsTo hgy hygΦ) - have hsxy : ∀ n : ℕ, (sxy^[n]) y = y + (2 * n : ℤ) • (x - y) := by - intro n - induction n with - | zero => simp - | succ n ih => - simp only [iterate_succ', comp_apply, ih, zsmul_sub, map_add, LinearEquiv.trans_apply, - reflection_apply_self, map_neg, reflection_apply, hfy, two_smul, neg_sub, map_sub, - map_zsmul, hgx, smul_neg, smul_add, Nat.cast_succ, mul_add, mul_one, add_smul, sxy] - abel - set f' : ℕ → Φ := fun n ↦ ⟨(sxy^[n]) y, by - rw [← IsFixedPt.image_iterate hb.image_eq n]; exact mem_image_of_mem _ hyΦ⟩ - have : ¬ Injective f' := not_injective_infinite_finite f' - contrapose! this - intros n m hnm - rw [Subtype.mk_eq_mk, hsxy, hsxy, add_right_inj, ← sub_eq_zero, ← sub_smul, smul_eq_zero, - sub_eq_zero, sub_eq_zero] at hnm - simpa using hnm.resolve_right this + suffices h : f y • x = (2 : R) • y by + rw [hfy, two_smul R x, two_smul R y, ← two_zsmul, ← two_zsmul] at h + exact smul_right_injective _ two_ne_zero h + rw [← not_infinite] at hΦ + contrapose! hΦ + apply ((infinite_range_reflection_reflection_iterate_iff hfx hgy + (by rw [hfy, hgx]; norm_cast)).mpr hΦ).mono + rw [range_subset_iff] + intro n + rw [← IsFixedPt.image_iterate ((bijOn_reflection_of_mapsTo hfx hxfΦ).comp + (bijOn_reflection_of_mapsTo hgy hygΦ)).image_eq n] + exact mem_image_of_mem _ hyΦ lemma injOn_dualMap_subtype_span_range_range {ι : Type*} [NoZeroSMulDivisors ℤ M] {r : ι ↪ M} {c : ι → Dual R M} (hfin : (range r).Finite) diff --git a/Mathlib/LinearAlgebra/RootSystem/Basic.lean b/Mathlib/LinearAlgebra/RootSystem/Basic.lean index a6376901bcdf1..9c4a0351f4b44 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Basic.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Basic.lean @@ -25,10 +25,7 @@ This file contains basic results for root systems and root data. ## TODO -* Derived properties of pairs, e.g., (ultra)parallel linearly independent pairs generate infinite - dihedral groups. * Properties of Weyl group (faithful action on roots, finiteness for finite `ι`) -* Conditions for existence of Weyl-invariant form (e.g., finiteness). -/ @@ -74,6 +71,25 @@ protected def equiv_of_mapsTo : end reflection_perm +lemma infinite_of_linearly_independent_coxeterWeight_four [CharZero R] [NoZeroSMulDivisors ℤ M] + (P : RootPairing ι R M N) (i j : ι) (hl : LinearIndependent R ![P.root i, P.root j]) + (hc : P.coxeterWeight i j = 4) : Infinite ι := by + refine (infinite_range_iff (Embedding.injective P.root)).mp (Infinite.mono ?_ + ((infinite_range_reflection_reflection_iterate_iff (P.coroot_root_two i) + (P.coroot_root_two j) ?_).mpr ?_)) + · rw [range_subset_iff] + intro n + rw [← IsFixedPt.image_iterate ((bijOn_reflection_of_mapsTo (P.coroot_root_two i) + (P.mapsTo_reflection_root i)).comp (bijOn_reflection_of_mapsTo (P.coroot_root_two j) + (P.mapsTo_reflection_root j))).image_eq n] + exact mem_image_of_mem _ (mem_range_self j) + · rw [coroot_root_eq_pairing, coroot_root_eq_pairing, ← hc, mul_comm, coxeterWeight] + · rw [LinearIndependent.pair_iff] at hl + specialize hl (P.pairing j i) (-2) + simp only [neg_smul, neg_eq_zero, OfNat.ofNat_ne_zero, and_false, imp_false] at hl + rw [ne_eq, coroot_root_eq_pairing, ← sub_eq_zero, sub_eq_add_neg] + exact hl + variable [Finite ι] (P : RootPairing ι R M N) (i j : ι) /-- Even though the roots may not span, coroots are distinguished by their pairing with the From 19796eea79b36c81ba5b1de34f5500e0348c84a2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Dupuis?= Date: Sat, 10 Aug 2024 22:22:37 +0000 Subject: [PATCH 166/359] =?UTF-8?q?feat(ContinuousMapZero):=20add=20norm?= =?UTF-8?q?=20structure=20on=20`C(X,=20R)=E2=82=80`=20(#15664)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds a norm structure on `C(X, R)₀`. This is needed to work with integral representations of functions defined via the non-unital continuous functional calculus. --- .../ContinuousFunction/ContinuousMapZero.lean | 41 +++++++++++++++++++ 1 file changed, 41 insertions(+) diff --git a/Mathlib/Topology/ContinuousFunction/ContinuousMapZero.lean b/Mathlib/Topology/ContinuousFunction/ContinuousMapZero.lean index f02fe2b8cc5f7..76f9a54f3d595 100644 --- a/Mathlib/Topology/ContinuousFunction/ContinuousMapZero.lean +++ b/Mathlib/Topology/ContinuousFunction/ContinuousMapZero.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ import Mathlib.Topology.ContinuousFunction.Algebra +import Mathlib.Topology.ContinuousFunction.Compact /-! # Continuous maps sending zero to zero @@ -196,6 +197,23 @@ lemma coe_toContinuousMapHom [StarRing R] [ContinuousStar R] : ⇑(toContinuousMapHom (X := X) (R := R)) = (↑) := rfl +/-- The coercion `C(X, R)₀ → C(X, R)` bundled as a continuous linear map. -/ +@[simps] +def toContinuousMapCLM (M : Type*) [Semiring M] [Module M R] [ContinuousConstSMul M R] : + C(X, R)₀ →L[M] C(X, R) where + toFun f := f + map_add' _ _ := rfl + map_smul' _ _ := rfl + +/-- The evaluation at a point, as a continuous linear map from `C(X, R)₀` to `R`. -/ +def evalCLM (𝕜 : Type*) {R : Type*} [CompactSpace X] [NormedField 𝕜] [NormedCommRing R] + [NormedSpace 𝕜 R] (x : X) : C(X, R)₀ →L[𝕜] R := + (ContinuousMap.evalCLM 𝕜 x).comp (toContinuousMapCLM 𝕜 : C(X, R)₀ →L[𝕜] C(X, R)) + +@[simp] +lemma evalCLM_apply {𝕜 : Type*} {R : Type*} [CompactSpace X] [NormedField 𝕜] [NormedCommRing R] + [NormedSpace 𝕜 R] (x : X) (f : C(X, R)₀) : evalCLM 𝕜 x f = f x := rfl + /-- Coercion to a function as an `AddMonoidHom`. Similar to `ContinuousMap.coeFnAddMonoidHom`. -/ def coeFnAddMonoidHom : C(X, R)₀ →+ X → R where toFun f := f @@ -309,4 +327,27 @@ def nonUnitalStarAlgHom_postcomp (φ : R →⋆ₙₐ[M] S) (hφ : Continuous φ end CompHoms +section Norm + +variable {α : Type*} {𝕜 : Type*} {R : Type*} [TopologicalSpace α] [CompactSpace α] [Zero α] + +noncomputable instance [MetricSpace R] [Zero R]: MetricSpace C(α, R)₀ := + ContinuousMapZero.uniformEmbedding_toContinuousMap.comapMetricSpace _ + +noncomputable instance [NormedAddCommGroup R] : Norm C(α, R)₀ where + norm f := ‖(f : C(α, R))‖ + +noncomputable instance [NormedCommRing R] : NonUnitalNormedCommRing C(α, R)₀ where + dist_eq f g := NormedAddGroup.dist_eq (f : C(α, R)) g + norm_mul f g := NormedRing.norm_mul (f : C(α, R)) g + mul_comm f g := mul_comm f g + +instance [NormedField 𝕜] [NormedCommRing R] [NormedAlgebra 𝕜 R] : NormedSpace 𝕜 C(α, R)₀ where + norm_smul_le r f := norm_smul_le r (f : C(α, R)) + +instance [NormedCommRing R] [StarRing R] [CStarRing R] : CStarRing C(α, R)₀ where + norm_mul_self_le f := CStarRing.norm_mul_self_le (f : C(α, R)) + +end Norm + end ContinuousMapZero From 06bc0d1ba3813a9d3209f55ce74af5dc04866514 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Sun, 11 Aug 2024 03:09:08 +0000 Subject: [PATCH 167/359] chore(CategoryTheory/Monoidal): fix duplicate lemmas (#15682) --- Mathlib/CategoryTheory/GradedObject/Monoidal.lean | 2 +- Mathlib/CategoryTheory/Monoidal/Category.lean | 4 ---- 2 files changed, 1 insertion(+), 5 deletions(-) diff --git a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean index 6875157be3ba4..88a32f2bf3bcd 100644 --- a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean +++ b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean @@ -421,7 +421,7 @@ lemma pentagon_inv : MonoidalCategory.whiskerLeft_comp_assoc, ← ιTensorObj₃_eq_assoc X₁ X₂ (tensorObj X₃ X₄) i₁ i₂ (i₃ + i₄) j (by rw [← add_assoc, h]) (i₂ + i₃ + i₄) (by rw [add_assoc]), - ιTensorObj₃_associator_inv_assoc, whiskerLeft_whiskerLeft_associator_inv_assoc, + ιTensorObj₃_associator_inv_assoc, associator_inv_naturality_right_assoc, ιTensorObj₃'_eq_assoc X₁ X₂ (tensorObj X₃ X₄) i₁ i₂ (i₃ + i₄) j (by rw [← add_assoc, h]) _ rfl, whisker_exchange_assoc, ← ιTensorObj₃_eq_assoc (tensorObj X₁ X₂) X₃ X₄ (i₁ + i₂) i₃ i₄ j h _ rfl, diff --git a/Mathlib/CategoryTheory/Monoidal/Category.lean b/Mathlib/CategoryTheory/Monoidal/Category.lean index e560f1dfca4aa..53371440fcf7e 100644 --- a/Mathlib/CategoryTheory/Monoidal/Category.lean +++ b/Mathlib/CategoryTheory/Monoidal/Category.lean @@ -279,10 +279,6 @@ theorem tensorHom_def' {X₁ Y₁ X₂ Y₂ : C} (f : X₁ ⟶ Y₁) (g : X₂ f ⊗ g = X₁ ◁ g ≫ f ▷ Y₂ := whisker_exchange f g ▸ tensorHom_def f g -@[reassoc] -lemma whiskerLeft_whiskerLeft_associator_inv (X Y : C) {Z₁ Z₂ : C} (f : Z₁ ⟶ Z₂) : - X ◁ Y ◁ f ≫ (α_ _ _ _).inv = (α_ _ _ _).inv ≫ _ ◁ f := by simp - end MonoidalCategory open scoped MonoidalCategory From bc89bc5b0f98eb6bc2ad73bdb276e89a9fa5a14c Mon Sep 17 00:00:00 2001 From: grunweg Date: Sun, 11 Aug 2024 04:28:48 +0000 Subject: [PATCH 168/359] chore(SetTheory/Ordinal): remove bare open Classical (#15670) --- Mathlib/SetTheory/Cardinal/Basic.lean | 24 ++++++++++++-------- Mathlib/SetTheory/Cardinal/Cofinality.lean | 5 +---- Mathlib/SetTheory/Cardinal/Ordinal.lean | 20 ++++++++++------- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 26 ++++++++++++---------- Mathlib/SetTheory/Ordinal/Basic.lean | 4 +--- Mathlib/SetTheory/Ordinal/Exponential.lean | 7 ++---- 6 files changed, 45 insertions(+), 41 deletions(-) diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 080fe52fe5888..bc9c1b0e031a5 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -80,7 +80,6 @@ Cantor's theorem, König's theorem, Konig's theorem assert_not_exists Field assert_not_exists Module -open scoped Classical open Mathlib (Vector) open Function Set Order @@ -566,8 +565,9 @@ instance canonicallyOrderedCommSemiring : CanonicallyOrderedCommSemiring Cardina add_le_add_left := fun a b => add_le_add_left exists_add_of_le := fun {a b} => inductionOn₂ a b fun α β ⟨⟨f, hf⟩⟩ => - have : α ⊕ ((range f)ᶜ : Set β) ≃ β := - (Equiv.sumCongr (Equiv.ofInjective f hf) (Equiv.refl _)).trans <| + have : α ⊕ ((range f)ᶜ : Set β) ≃ β := by + classical + exact (Equiv.sumCongr (Equiv.ofInjective f hf) (Equiv.refl _)).trans <| Equiv.Set.sumCompl (range f) ⟨#(↥(range f)ᶜ), mk_congr this.symm⟩ le_self_add := fun a b => (add_zero a).ge.trans <| add_le_add_left (Cardinal.zero_le _) _ @@ -1784,8 +1784,9 @@ theorem mk_eq_nat_iff_fintype {n : ℕ} : #α = n ↔ ∃ h : Fintype α, @Finty exact ⟨t, eq_univ_iff_forall.2 ht, hn⟩ theorem mk_union_add_mk_inter {α : Type u} {S T : Set α} : - #(S ∪ T : Set α) + #(S ∩ T : Set α) = #S + #T := - Quot.sound ⟨Equiv.Set.unionSumInter S T⟩ + #(S ∪ T : Set α) + #(S ∩ T : Set α) = #S + #T := by + classical + exact Quot.sound ⟨Equiv.Set.unionSumInter S T⟩ /-- The cardinality of a union is at most the sum of the cardinalities of the two sets. -/ @@ -1793,8 +1794,9 @@ theorem mk_union_le {α : Type u} (S T : Set α) : #(S ∪ T : Set α) ≤ #S + @mk_union_add_mk_inter α S T ▸ self_le_add_right #(S ∪ T : Set α) #(S ∩ T : Set α) theorem mk_union_of_disjoint {α : Type u} {S T : Set α} (H : Disjoint S T) : - #(S ∪ T : Set α) = #S + #T := - Quot.sound ⟨Equiv.Set.union H.le_bot⟩ + #(S ∪ T : Set α) = #S + #T := by + classical + exact Quot.sound ⟨Equiv.Set.union H.le_bot⟩ theorem mk_insert {α : Type u} {s : Set α} {a : α} (h : a ∉ s) : #(insert a s : Set α) = #s + 1 := by @@ -1806,8 +1808,9 @@ theorem mk_insert_le {α : Type u} {s : Set α} {a : α} : #(insert a s : Set α · simp only [insert_eq_of_mem h, self_le_add_right] · rw [mk_insert h] -theorem mk_sum_compl {α} (s : Set α) : #s + #(sᶜ : Set α) = #α := - mk_congr (Equiv.Set.sumCompl s) +theorem mk_sum_compl {α} (s : Set α) : #s + #(sᶜ : Set α) = #α := by + classical + exact mk_congr (Equiv.Set.sumCompl s) theorem mk_le_mk_of_subset {α} {s t : Set α} (h : s ⊆ t) : #s ≤ #t := ⟨Set.embeddingOfSubset s t h⟩ @@ -1816,6 +1819,7 @@ theorem mk_le_iff_forall_finset_subset_card_le {α : Type u} {n : ℕ} {t : Set #t ≤ n ↔ ∀ s : Finset α, (s : Set α) ⊆ t → s.card ≤ n := by refine ⟨fun H s hs ↦ by simpa using (mk_le_mk_of_subset hs).trans H, fun H ↦ ?_⟩ apply card_le_of (fun s ↦ ?_) + classical let u : Finset α := s.image Subtype.val have : u.card = s.card := Finset.card_image_of_injOn Subtype.coe_injective.injOn rw [← this] @@ -1902,6 +1906,7 @@ theorem two_le_iff' (x : α) : (2 : Cardinal) ≤ #α ↔ ∃ y : α, y ≠ x := rw [two_le_iff, ← nontrivial_iff, nontrivial_iff_exists_ne x] theorem mk_eq_two_iff : #α = 2 ↔ ∃ x y : α, x ≠ y ∧ ({x, y} : Set α) = univ := by + classical simp only [← @Nat.cast_two Cardinal, mk_eq_nat_iff_finset, Finset.card_eq_two] constructor · rintro ⟨t, ht, x, y, hne, rfl⟩ @@ -1920,6 +1925,7 @@ theorem mk_eq_two_iff' (x : α) : #α = 2 ↔ ∃! y, y ≠ x := by theorem exists_not_mem_of_length_lt {α : Type*} (l : List α) (h : ↑l.length < #α) : ∃ z : α, z ∉ l := by + classical contrapose! h calc #α = #(Set.univ : Set α) := mk_univ.symm diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index c8063908425a6..aa8f521ccd377 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -42,13 +42,10 @@ cofinality, regular cardinals, limits cardinals, inaccessible cardinals, infinite pigeonhole principle -/ - noncomputable section open Function Cardinal Set Order - -open scoped Classical -open Cardinal Ordinal +open scoped Ordinal universe u v w diff --git a/Mathlib/SetTheory/Cardinal/Ordinal.lean b/Mathlib/SetTheory/Cardinal/Ordinal.lean index 5cfcdcf416e0e..cccfc9d1719b0 100644 --- a/Mathlib/SetTheory/Cardinal/Ordinal.lean +++ b/Mathlib/SetTheory/Cardinal/Ordinal.lean @@ -48,8 +48,6 @@ noncomputable section open Function Set Cardinal Equiv Order Ordinal -open scoped Classical - universe u v w namespace Cardinal @@ -436,6 +434,7 @@ theorem mul_eq_self {c : Cardinal} (h : ℵ₀ ≤ c) : c * c = c := by refine Acc.recOn (Cardinal.lt_wf.apply c) (fun c _ => Quotient.inductionOn c fun α IH ol => ?_) h -- consider the minimal well-order `r` on `α` (a type with cardinality `c`). rcases ord_eq α with ⟨r, wo, e⟩ + classical letI := linearOrderOfSTO r haveI : IsWellOrder α (· < ·) := wo -- Define an order `s` on `α × α` by writing `(a, b) < (c, d)` if `max a b < max c d`, or @@ -1066,8 +1065,9 @@ theorem mk_list_le_max (α : Type u) : #(List α) ≤ max ℵ₀ #α := by apply le_max_right @[simp] -theorem mk_finset_of_infinite (α : Type u) [Infinite α] : #(Finset α) = #α := - Eq.symm <| +theorem mk_finset_of_infinite (α : Type u) [Infinite α] : #(Finset α) = #α := by + classical + exact Eq.symm <| le_antisymm (mk_le_of_injective fun _ _ => Finset.singleton_inj.1) <| calc #(Finset α) ≤ #(List α) := mk_le_of_surjective List.toFinset_surjective @@ -1108,16 +1108,18 @@ theorem mk_finsupp_of_infinite' (α β : Type u) [Nonempty α] [Zero β] [Infini theorem mk_finsupp_nat (α : Type u) [Nonempty α] : #(α →₀ ℕ) = max #α ℵ₀ := by simp @[simp] -theorem mk_multiset_of_nonempty (α : Type u) [Nonempty α] : #(Multiset α) = max #α ℵ₀ := - Multiset.toFinsupp.toEquiv.cardinal_eq.trans (mk_finsupp_nat α) +theorem mk_multiset_of_nonempty (α : Type u) [Nonempty α] : #(Multiset α) = max #α ℵ₀ := by + classical + exact Multiset.toFinsupp.toEquiv.cardinal_eq.trans (mk_finsupp_nat α) theorem mk_multiset_of_infinite (α : Type u) [Infinite α] : #(Multiset α) = #α := by simp theorem mk_multiset_of_isEmpty (α : Type u) [IsEmpty α] : #(Multiset α) = 1 := Multiset.toFinsupp.toEquiv.cardinal_eq.trans (by simp) -theorem mk_multiset_of_countable (α : Type u) [Countable α] [Nonempty α] : #(Multiset α) = ℵ₀ := - Multiset.toFinsupp.toEquiv.cardinal_eq.trans (by simp) +theorem mk_multiset_of_countable (α : Type u) [Countable α] [Nonempty α] : #(Multiset α) = ℵ₀ := by + classical + exact Multiset.toFinsupp.toEquiv.cardinal_eq.trans (by simp) theorem mk_bounded_set_le_of_infinite (α : Type u) [Infinite α] (c : Cardinal) : #{ t : Set α // #t ≤ c } ≤ #α ^ c := by @@ -1129,6 +1131,7 @@ theorem mk_bounded_set_le_of_infinite (α : Type u) [Infinite α] (c : Cardinal) refine le_trans (mk_preimage_of_injective _ _ fun x y => Sum.inl.inj) ?_ apply mk_range_le rintro ⟨s, ⟨g⟩⟩ + classical use fun y => if h : ∃ x : s, g x = y then Sum.inl (Classical.choose h).val else Sum.inr (ULift.up 0) apply Subtype.eq; ext x @@ -1223,6 +1226,7 @@ end compl theorem extend_function {α β : Type*} {s : Set α} (f : s ↪ β) (h : Nonempty ((sᶜ : Set α) ≃ ((range f)ᶜ : Set β))) : ∃ g : α ≃ β, ∀ x : s, g x = f x := by + classical have := h; cases' this with g let h : α ≃ β := (Set.sumCompl (s : Set α)).symm.trans diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 23a23ba8334f5..42eadc412c344 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -55,9 +55,7 @@ assert_not_exists Module noncomputable section open Function Cardinal Set Equiv Order - -open scoped Classical -open Cardinal Ordinal +open scoped Ordinal universe u v w @@ -150,7 +148,7 @@ theorem right_eq_zero_of_add_eq_zero {a b : Ordinal} (h : a + b = 0) : b = 0 := /-! ### The predecessor of an ordinal -/ - +open Classical in /-- The ordinal predecessor of `o` is `o'` if `o = succ o'`, and `o` otherwise. -/ def pred (o : Ordinal) : Ordinal := @@ -161,8 +159,9 @@ theorem pred_succ (o) : pred (succ o) = o := by have h : ∃ a, succ o = succ a := ⟨_, rfl⟩ simpa only [pred, dif_pos h] using (succ_injective <| Classical.choose_spec h).symm -theorem pred_le_self (o) : pred o ≤ o := - if h : ∃ a, o = succ a then by +theorem pred_le_self (o) : pred o ≤ o := by + classical + exact if h : ∃ a, o = succ a then by let ⟨a, e⟩ := h rw [e, pred_succ]; exact le_succ a else by rw [pred, dif_neg h] @@ -187,8 +186,9 @@ theorem succ_pred_iff_is_succ {o} : succ (pred o) = o ↔ ∃ a, o = succ a := theorem succ_lt_of_not_succ {o b : Ordinal} (h : ¬∃ a, o = succ a) : succ b < o ↔ b < o := ⟨(lt_succ b).trans, fun l => lt_of_le_of_ne (succ_le_of_lt l) fun e => h ⟨_, e.symm⟩⟩ -theorem lt_pred {a b} : a < pred b ↔ succ a < b := - if h : ∃ a, b = succ a then by +theorem lt_pred {a b} : a < pred b ↔ succ a < b := by + classical + exact if h : ∃ a, b = succ a then by let ⟨c, e⟩ := h rw [e, pred_succ, succ_lt_succ_iff] else by simp only [pred, dif_neg h, succ_lt_of_not_succ h] @@ -204,8 +204,9 @@ theorem lift_is_succ {o : Ordinal.{v}} : (∃ a, lift.{u} o = succ a) ↔ ∃ a, fun ⟨a, h⟩ => ⟨lift.{u} a, by simp only [h, lift_succ]⟩⟩ @[simp] -theorem lift_pred (o : Ordinal.{v}) : lift.{u} (pred o) = pred (lift.{u} o) := - if h : ∃ a, o = succ a then by cases' h with a e; simp only [e, pred_succ, lift_succ] +theorem lift_pred (o : Ordinal.{v}) : lift.{u} (pred o) = pred (lift.{u} o) := by + classical + exact if h : ∃ a, o = succ a then by cases' h with a e; simp only [e, pred_succ, lift_succ] else by rw [pred_eq_iff_not_succ.2 h, pred_eq_iff_not_succ.2 (mt lift_is_succ.1 h)] /-! ### Limit ordinals -/ @@ -264,8 +265,9 @@ theorem IsLimit.nat_lt {o : Ordinal} (h : IsLimit o) : ∀ n : ℕ, (n : Ordinal | 0 => h.pos | n + 1 => h.2 _ (IsLimit.nat_lt h n) -theorem zero_or_succ_or_limit (o : Ordinal) : o = 0 ∨ (∃ a, o = succ a) ∨ IsLimit o := - if o0 : o = 0 then Or.inl o0 +theorem zero_or_succ_or_limit (o : Ordinal) : o = 0 ∨ (∃ a, o = succ a) ∨ IsLimit o := by + classical + exact if o0 : o = 0 then Or.inl o0 else if h : ∃ a, o = succ a then Or.inr (Or.inl h) else Or.inr <| Or.inr ⟨o0, fun _a => (succ_lt_of_not_succ h).2⟩ diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 6266f0e4b0ba6..5aa1b13790de0 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -59,9 +59,7 @@ assert_not_exists Field noncomputable section open Function Cardinal Set Equiv Order - -open scoped Classical -open Cardinal InitialSeg +open scoped Cardinal InitialSeg universe u v w diff --git a/Mathlib/SetTheory/Ordinal/Exponential.lean b/Mathlib/SetTheory/Ordinal/Exponential.lean index 3fa7d53f8ab2a..7c5c895e19db7 100644 --- a/Mathlib/SetTheory/Ordinal/Exponential.lean +++ b/Mathlib/SetTheory/Ordinal/Exponential.lean @@ -12,13 +12,10 @@ related by the lemma `Ordinal.opow_le_iff_le_log : b ^ c ≤ x ↔ c ≤ log b x `b`, `c`. -/ - noncomputable section -open Function Cardinal Set Equiv Order - -open scoped Classical -open Cardinal Ordinal +open Function Set Equiv Order +open scoped Cardinal Ordinal universe u v w From ce6b0a984cb0cdf3d5915ec352d7f8f0d7a8af36 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 11 Aug 2024 05:10:28 +0000 Subject: [PATCH 169/359] refactor: Unify the different versions of `mul_eq_one` (#14996) Replace `mul_eq_one_iff` and `Associates.mul_eq_one_iff` by a single lemma `mul_eq_one`. Rename `mul_eq_one_iff'` to `mul_eq_one_iff_of_one_le`. Similarly for the additive versions. --- Archive/Wiedijk100Theorems/BallotProblem.lean | 2 +- Mathlib/Algebra/Associated/Basic.lean | 21 +++++++------------ Mathlib/Algebra/BigOperators/Associated.lean | 4 ++-- .../Order/BigOperators/Group/Finset.lean | 2 +- .../Algebra/Order/Monoid/Canonical/Defs.lean | 13 ++++++------ .../Algebra/Order/Monoid/Unbundled/Basic.lean | 5 ++++- .../Algebra/Order/Ring/Unbundled/Basic.lean | 2 +- Mathlib/Algebra/Polynomial/Monic.lean | 2 +- Mathlib/Algebra/Polynomial/RingDivision.lean | 2 +- Mathlib/Algebra/Quaternion.lean | 3 ++- .../DoldKan/FunctorGamma.lean | 2 +- Mathlib/Analysis/BoundedVariation.lean | 2 +- Mathlib/Analysis/Convex/Basic.lean | 3 ++- Mathlib/Analysis/MeanInequalitiesPow.lean | 4 ++-- .../Enumerative/Composition.lean | 4 ++-- Mathlib/Data/Finsupp/Antidiagonal.lean | 2 +- Mathlib/Data/Nat/Factorization/Basic.lean | 2 +- Mathlib/Data/Nat/WithBot.lean | 2 +- Mathlib/Data/Ordmap/Ordset.lean | 12 +++++------ .../Angle/Unoriented/RightAngle.lean | 2 +- Mathlib/LinearAlgebra/QuadraticForm/Prod.lean | 2 +- .../MeasureTheory/Covering/Besicovitch.lean | 2 +- .../MeasureTheory/Decomposition/Jordan.lean | 4 ++-- .../Decomposition/SignedLebesgue.lean | 2 +- Mathlib/MeasureTheory/Measure/Restrict.lean | 2 +- Mathlib/NumberTheory/Bernoulli.lean | 2 +- Mathlib/NumberTheory/Zsqrtd/Basic.lean | 2 +- Mathlib/RingTheory/Coprime/Basic.lean | 2 +- Mathlib/RingTheory/GradedAlgebra/Basic.lean | 6 +++--- Mathlib/RingTheory/PowerSeries/Order.lean | 2 +- .../RingTheory/UniqueFactorizationDomain.lean | 2 +- Mathlib/SetTheory/Cardinal/Divisibility.lean | 2 +- 32 files changed, 61 insertions(+), 60 deletions(-) diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index 8b45732a49ddb..56bcc026d1228 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -365,7 +365,7 @@ theorem ballot_problem : exacts [Nat.cast_le.2 qp.le, ENNReal.natCast_ne_top _] rwa [ENNReal.toReal_eq_toReal (measure_lt_top _ _).ne] at this simp only [Ne, ENNReal.div_eq_top, tsub_eq_zero_iff_le, Nat.cast_le, not_le, - add_eq_zero_iff, Nat.cast_eq_zero, ENNReal.add_eq_top, ENNReal.natCast_ne_top, or_self_iff, + add_eq_zero, Nat.cast_eq_zero, ENNReal.add_eq_top, ENNReal.natCast_ne_top, or_self_iff, not_false_iff, and_true_iff] push_neg exact ⟨fun _ _ => by linarith, (tsub_le_self.trans_lt (ENNReal.natCast_ne_top p).lt_top).ne⟩ diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index 368bb8846445f..cfa5e0393c92a 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -844,20 +844,15 @@ theorem mk_pow (a : α) (n : ℕ) : Associates.mk (a ^ n) = Associates.mk a ^ n theorem dvd_eq_le : ((· ∣ ·) : Associates α → Associates α → Prop) = (· ≤ ·) := rfl -theorem mul_eq_one_iff {x y : Associates α} : x * y = 1 ↔ x = 1 ∧ y = 1 := - Iff.intro - (Quotient.inductionOn₂ x y fun a b h => - have : a * b ~ᵤ 1 := Quotient.exact h - ⟨Quotient.sound <| associated_one_of_associated_mul_one this, - Quotient.sound <| associated_one_of_associated_mul_one (b := a) (by rwa [mul_comm])⟩) - (by simp (config := { contextual := true })) - -theorem units_eq_one (u : (Associates α)ˣ) : u = 1 := - Units.ext (mul_eq_one_iff.1 u.val_inv).1 - instance uniqueUnits : Unique (Associates α)ˣ where - default := 1 - uniq := Associates.units_eq_one + uniq := by + rintro ⟨a, b, hab, hba⟩ + revert hab hba + exact Quotient.inductionOn₂ a b $ fun a b hab hba ↦ Units.ext $ Quotient.sound $ + associated_one_of_associated_mul_one $ Quotient.exact hab + +@[deprecated (since := "2024-07-22")] alias mul_eq_one_iff := mul_eq_one +@[deprecated (since := "2024-07-22")] protected alias units_eq_one := Subsingleton.elim @[simp] theorem coe_unit_eq_one (u : (Associates α)ˣ) : (u : Associates α) = 1 := by diff --git a/Mathlib/Algebra/BigOperators/Associated.lean b/Mathlib/Algebra/BigOperators/Associated.lean index c0dc9dbb0050f..fd06fc15dc080 100644 --- a/Mathlib/Algebra/BigOperators/Associated.lean +++ b/Mathlib/Algebra/BigOperators/Associated.lean @@ -129,7 +129,7 @@ theorem rel_associated_iff_map_eq_map {p q : Multiset α} : theorem prod_eq_one_iff {p : Multiset (Associates α)} : p.prod = 1 ↔ ∀ a ∈ p, (a : Associates α) = 1 := Multiset.induction_on p (by simp) - (by simp (config := { contextual := true }) [mul_eq_one_iff, or_imp, forall_and]) + (by simp (config := { contextual := true }) [mul_eq_one, or_imp, forall_and]) theorem prod_le_prod {p q : Multiset (Associates α)} (h : p ≤ q) : p.prod ≤ q.prod := by haveI := Classical.decEq (Associates α) @@ -146,7 +146,7 @@ variable [CancelCommMonoidWithZero α] theorem exists_mem_multiset_le_of_prime {s : Multiset (Associates α)} {p : Associates α} (hp : Prime p) : p ≤ s.prod → ∃ a ∈ s, p ≤ a := - Multiset.induction_on s (fun ⟨d, Eq⟩ => (hp.ne_one (mul_eq_one_iff.1 Eq.symm).1).elim) + Multiset.induction_on s (fun ⟨d, Eq⟩ => (hp.ne_one (mul_eq_one.1 Eq.symm).1).elim) fun a s ih h => have : p ≤ a * s.prod := by simpa using h match Prime.le_or_le hp this with diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index a37fbfc93fced..ce3eb2fbc4a57 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -147,7 +147,7 @@ theorem prod_eq_one_iff_of_one_le' : (fun _ ↦ ⟨fun _ _ h ↦ False.elim (Finset.not_mem_empty _ h), fun _ ↦ rfl⟩) ?_ intro a s ha ih H have : ∀ i ∈ s, 1 ≤ f i := fun _ ↦ H _ ∘ mem_insert_of_mem - rw [prod_insert ha, mul_eq_one_iff' (H _ <| mem_insert_self _ _) (one_le_prod' this), + rw [prod_insert ha, mul_eq_one_iff_of_one_le (H _ <| mem_insert_self _ _) (one_le_prod' this), forall_mem_insert, ih this] @[to_additive sum_eq_zero_iff_of_nonpos] diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean index b5e980fc71b9d..0646b0b825b56 100644 --- a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean @@ -3,6 +3,7 @@ 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.Units import Mathlib.Algebra.Order.Monoid.Defs import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE import Mathlib.Algebra.NeZero @@ -110,11 +111,11 @@ theorem one_le (a : α) : 1 ≤ a := theorem bot_eq_one : (⊥ : α) = 1 := le_antisymm bot_le (one_le ⊥) ---TODO: This is a special case of `mul_eq_one`. We need the instance --- `CanonicallyOrderedCommMonoid α → Unique αˣ` -@[to_additive (attr := simp)] -theorem mul_eq_one_iff : a * b = 1 ↔ a = 1 ∧ b = 1 := - mul_eq_one_iff' (one_le _) (one_le _) +@[to_additive] instance CanonicallyOrderedCommMonoid.toUniqueUnits : Unique αˣ where + uniq a := Units.ext ((mul_eq_one_iff_of_one_le (α := α) (one_le _) $ one_le _).1 a.mul_inv).1 + +@[deprecated (since := "2024-07-24")] alias mul_eq_one_iff := mul_eq_one +@[deprecated (since := "2024-07-24")] alias add_eq_zero_iff := add_eq_zero @[to_additive (attr := simp)] theorem le_one_iff_eq_one : a ≤ 1 ↔ a = 1 := @@ -129,7 +130,7 @@ theorem eq_one_or_one_lt (a : α) : a = 1 ∨ 1 < a := (one_le a).eq_or_lt.imp_l @[to_additive (attr := simp) add_pos_iff] theorem one_lt_mul_iff : 1 < a * b ↔ 1 < a ∨ 1 < b := by - simp only [one_lt_iff_ne_one, Ne, mul_eq_one_iff, not_and_or] + simp only [one_lt_iff_ne_one, Ne, mul_eq_one, not_and_or] @[to_additive] theorem exists_one_lt_mul_of_lt (h : a < b) : ∃ (c : _) (_ : 1 < c), a * c = b := by diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean index 53a7cbf0f18a3..05333081038c7 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean @@ -954,7 +954,7 @@ section PartialOrder variable [PartialOrder α] @[to_additive] -theorem mul_eq_one_iff' [CovariantClass α α (· * ·) (· ≤ ·)] +theorem mul_eq_one_iff_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : a * b = 1 ↔ a = 1 ∧ b = 1 := Iff.intro @@ -969,6 +969,9 @@ theorem mul_eq_one_iff' [CovariantClass α α (· * ·) (· ≤ ·)] -- `fun ⟨ha', hb'⟩ => by rw [ha', hb', mul_one]`, -- had its `to_additive`-ization fail due to some bug +@[deprecated (since := "2024-07-24")] alias mul_eq_one_iff' := mul_eq_one_iff_of_one_le +@[deprecated (since := "2024-07-24")] alias add_eq_zero_iff' := add_eq_zero_iff_of_nonneg + section Left variable [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean index 9e5c499c5722b..4ec7c3d1885a2 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean @@ -915,7 +915,7 @@ lemma mul_self_add_mul_self_eq_zero [IsRightCancelAdd α] [NoZeroDivisors α] [ZeroLEOneClass α] [ExistsAddOfLE α] [PosMulMono α] [CovariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· < ·)] : a * a + b * b = 0 ↔ a = 0 ∧ b = 0 := by - rw [add_eq_zero_iff', mul_self_eq_zero (M₀ := α), mul_self_eq_zero (M₀ := α)] <;> + rw [add_eq_zero_iff_of_nonneg, mul_self_eq_zero (M₀ := α), mul_self_eq_zero (M₀ := α)] <;> apply mul_self_nonneg lemma eq_zero_of_mul_self_add_mul_self_eq_zero [IsRightCancelAdd α] [NoZeroDivisors α] diff --git a/Mathlib/Algebra/Polynomial/Monic.lean b/Mathlib/Algebra/Polynomial/Monic.lean index 07cb29792f10d..9783d0e918278 100644 --- a/Mathlib/Algebra/Polynomial/Monic.lean +++ b/Mathlib/Algebra/Polynomial/Monic.lean @@ -220,7 +220,7 @@ theorem Monic.eq_one_of_isUnit (hm : Monic p) (hpu : IsUnit p) : p = 1 := by nontriviality R obtain ⟨q, h⟩ := hpu.exists_right_inv have := hm.natDegree_mul' (right_ne_zero_of_mul_eq_one h) - rw [h, natDegree_one, eq_comm, add_eq_zero_iff] at this + rw [h, natDegree_one, eq_comm, add_eq_zero] at this exact hm.natDegree_eq_zero_iff_eq_one.mp this.1 theorem Monic.isUnit_iff (hm : p.Monic) : IsUnit p ↔ p = 1 := diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index 83504114522a2..32336ed620bbc 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -179,7 +179,7 @@ theorem natDegree_eq_zero_of_isUnit (h : IsUnit p) : natDegree p = 0 := by nontriviality R obtain ⟨q, hq⟩ := h.exists_right_inv have := natDegree_mul (left_ne_zero_of_mul_eq_one hq) (right_ne_zero_of_mul_eq_one hq) - rw [hq, natDegree_one, eq_comm, add_eq_zero_iff] at this + rw [hq, natDegree_one, eq_comm, add_eq_zero] at this exact this.1 theorem degree_eq_zero_of_isUnit [Nontrivial R] (h : IsUnit p) : degree p = 0 := diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 177560a1d9ab8..86c899ed938f5 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -1170,7 +1170,8 @@ variable [LinearOrderedCommRing R] {a : ℍ[R]} @[simp] theorem normSq_eq_zero : normSq a = 0 ↔ a = 0 := by refine ⟨fun h => ?_, fun h => h.symm ▸ normSq.map_zero⟩ - rw [normSq_def', add_eq_zero_iff', add_eq_zero_iff', add_eq_zero_iff'] at h + rw [normSq_def', add_eq_zero_iff_of_nonneg, add_eq_zero_iff_of_nonneg, add_eq_zero_iff_of_nonneg] + at h · exact ext a 0 (pow_eq_zero h.1.1.1) (pow_eq_zero h.1.1.2) (pow_eq_zero h.1.2) (pow_eq_zero h.2) all_goals apply_rules [sq_nonneg, add_nonneg] diff --git a/Mathlib/AlgebraicTopology/DoldKan/FunctorGamma.lean b/Mathlib/AlgebraicTopology/DoldKan/FunctorGamma.lean index 91eb4fd4ec16c..72511e140c0f6 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/FunctorGamma.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/FunctorGamma.lean @@ -148,7 +148,7 @@ theorem mapMono_comp (i' : Δ'' ⟶ Δ') (i : Δ' ⟶ Δ) [Mono i'] [Mono i] : have eq : Δ.len = Δ''.len + (k + k' + 2) := by omega rw [mapMono_eq_zero K (i' ≫ i) _ _]; rotate_left · by_contra h - simp only [self_eq_add_right, h, add_eq_zero_iff, and_false] at eq + simp only [self_eq_add_right, h, add_eq_zero, and_false] at eq · by_contra h simp only [h.1, add_right_inj] at eq omega diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index b8c93daf2c989..d902df8fafe19 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -292,7 +292,7 @@ theorem add_point (f : α → E) {s : Set α} {x : α} (hx : x ∈ s) (u : ℕ apply Finset.sum_congr rfl fun i _hi => ?_ dsimp only [w] simp only [← Npos, Nat.not_lt_zero, Nat.add_succ_sub_one, add_zero, if_false, - add_eq_zero_iff, Nat.one_ne_zero, false_and_iff, Nat.succ_add_sub_one, zero_add] + add_eq_zero, Nat.one_ne_zero, false_and_iff, Nat.succ_add_sub_one, zero_add] rw [add_comm 1 i] _ = ∑ i ∈ Finset.Ico 1 (n + 1), edist (f (w (i + 1))) (f (w i)) := by rw [Finset.range_eq_Ico] diff --git a/Mathlib/Analysis/Convex/Basic.lean b/Mathlib/Analysis/Convex/Basic.lean index bd775685ea4b8..89b70137948b7 100644 --- a/Mathlib/Analysis/Convex/Basic.lean +++ b/Mathlib/Analysis/Convex/Basic.lean @@ -511,7 +511,8 @@ theorem Convex.exists_mem_add_smul_eq (h : Convex 𝕜 s) {x y : E} {p q : 𝕜} rcases _root_.em (p = 0 ∧ q = 0) with (⟨rfl, rfl⟩ | hpq) · use x, hx simp - · replace hpq : 0 < p + q := (add_nonneg hp hq).lt_of_ne' (mt (add_eq_zero_iff' hp hq).1 hpq) + · replace hpq : 0 < p + q := + (add_nonneg hp hq).lt_of_ne' (mt (add_eq_zero_iff_of_nonneg hp hq).1 hpq) refine ⟨_, convex_iff_div.1 h hx hy hp hq hpq, ?_⟩ simp only [smul_add, smul_smul, mul_div_cancel₀ _ hpq.ne'] diff --git a/Mathlib/Analysis/MeanInequalitiesPow.lean b/Mathlib/Analysis/MeanInequalitiesPow.lean index 91c0a4ed1cb78..1eb4445cc0366 100644 --- a/Mathlib/Analysis/MeanInequalitiesPow.lean +++ b/Mathlib/Analysis/MeanInequalitiesPow.lean @@ -162,8 +162,8 @@ private theorem add_rpow_le_one_of_add_le_one {p : ℝ} (a b : ℝ≥0) (hab : a theorem add_rpow_le_rpow_add {p : ℝ} (a b : ℝ≥0) (hp1 : 1 ≤ p) : a ^ p + b ^ p ≤ (a + b) ^ p := by have hp_pos : 0 < p := by positivity by_cases h_zero : a + b = 0 - · simp [add_eq_zero_iff.mp h_zero, hp_pos.ne'] - have h_nonzero : ¬(a = 0 ∧ b = 0) := by rwa [add_eq_zero_iff] at h_zero + · simp [add_eq_zero.mp h_zero, hp_pos.ne'] + have h_nonzero : ¬(a = 0 ∧ b = 0) := by rwa [add_eq_zero] at h_zero have h_add : a / (a + b) + b / (a + b) = 1 := by rw [div_add_div_same, div_self h_zero] have h := add_rpow_le_one_of_add_le_one (a / (a + b)) (b / (a + b)) h_add.le hp1 rw [div_rpow a (a + b), div_rpow b (a + b)] at h diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index ac4b0794cb1d6..87ab6914334a6 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -758,10 +758,10 @@ def compositionAsSetEquiv (n : ℕ) : CompositionAsSet n ≃ Finset (Fin (n - 1) apply (Nat.succ_pred_eq_of_pos _).symm exact (zero_le i.val).trans_lt (i.2.trans_le (Nat.sub_le n 1)) simp only [add_comm, Fin.ext_iff, Fin.val_zero, Fin.val_last, exists_prop, Set.toFinset_setOf, - Finset.mem_univ, forall_true_left, Finset.mem_filter, add_eq_zero_iff, and_false, + Finset.mem_univ, forall_true_left, Finset.mem_filter, add_eq_zero, and_false, add_left_inj, false_or, true_and] erw [Set.mem_setOf_eq] - simp [this, false_or_iff, add_right_inj, add_eq_zero_iff, one_ne_zero, false_and_iff, + simp [this, false_or_iff, add_right_inj, add_eq_zero, one_ne_zero, false_and_iff, Fin.val_mk] constructor · intro h diff --git a/Mathlib/Data/Finsupp/Antidiagonal.lean b/Mathlib/Data/Finsupp/Antidiagonal.lean index 3a5bdbb67f655..d3cc6f813d245 100644 --- a/Mathlib/Data/Finsupp/Antidiagonal.lean +++ b/Mathlib/Data/Finsupp/Antidiagonal.lean @@ -63,7 +63,7 @@ theorem antidiagonal_single (a : α) (n : ℕ) : simp_rw [single_apply, Finsupp.add_apply] at h ⊢ obtain rfl | hai := Decidable.eq_or_ne a i · exact ⟨if_pos rfl, if_pos rfl⟩ - · simp_rw [if_neg hai, _root_.add_eq_zero_iff] at h ⊢ + · simp_rw [if_neg hai, add_eq_zero] at h ⊢ exact h.imp Eq.symm Eq.symm · rintro ⟨a, b, rfl, rfl, rfl⟩ exact (single_add _ _ _).symm diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 3c1c6c7359aaa..3b99f2f2baaac 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -40,7 +40,7 @@ theorem factorization_eq_zero_iff_remainder {p r : ℕ} (i : ℕ) (pp : p.Prime) refine ⟨pp, ?_, ?_⟩ · rwa [← Nat.dvd_add_iff_right (dvd_mul_right p i)] · contrapose! hr0 - exact (add_eq_zero_iff.mp hr0).2 + exact (add_eq_zero.1 hr0).2 /-- The only numbers with empty prime factorization are `0` and `1` -/ theorem factorization_eq_zero_iff' (n : ℕ) : n.factorization = 0 ↔ n = 0 ∨ n = 1 := by diff --git a/Mathlib/Data/Nat/WithBot.lean b/Mathlib/Data/Nat/WithBot.lean index cdcd0dfca6075..125eeae18ef2b 100644 --- a/Mathlib/Data/Nat/WithBot.lean +++ b/Mathlib/Data/Nat/WithBot.lean @@ -28,7 +28,7 @@ theorem add_eq_zero_iff {n m : WithBot ℕ} : n + m = 0 ↔ n = 0 ∧ m = 0 := b · simp [WithBot.bot_add] cases m · simp [WithBot.add_bot] - simp [← WithBot.coe_add, _root_.add_eq_zero_iff] + simp [← WithBot.coe_add, add_eq_zero_iff_of_nonneg] theorem add_eq_one_iff {n m : WithBot ℕ} : n + m = 1 ↔ n = 0 ∧ m = 1 ∨ n = 1 ∧ m = 0 := by cases n diff --git a/Mathlib/Data/Ordmap/Ordset.lean b/Mathlib/Data/Ordmap/Ordset.lean index a51fcf32a77b9..fc25de56b3862 100644 --- a/Mathlib/Data/Ordmap/Ordset.lean +++ b/Mathlib/Data/Ordmap/Ordset.lean @@ -587,7 +587,7 @@ theorem balance_eq_balance' {l x r} (hl : Balanced l) (hr : Balanced r) (sl : Si · rfl · have : size rrl = 0 ∧ size rrr = 0 := by have := balancedSz_zero.1 hr.1.symm - rwa [size, sr.2.2.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero_iff] at this + rwa [size, sr.2.2.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero] at this cases sr.2.2.2.1.size_eq_zero.1 this.1 cases sr.2.2.2.2.size_eq_zero.1 this.2 obtain rfl : rrs = 1 := sr.2.2.1 @@ -595,7 +595,7 @@ theorem balance_eq_balance' {l x r} (hl : Balanced l) (hr : Balanced r) (sl : Si all_goals dsimp only [size]; decide · have : size rll = 0 ∧ size rlr = 0 := by have := balancedSz_zero.1 hr.1 - rwa [size, sr.2.1.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero_iff] at this + rwa [size, sr.2.1.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero] at this cases sr.2.1.2.1.size_eq_zero.1 this.1 cases sr.2.1.2.2.size_eq_zero.1 this.2 obtain rfl : rls = 1 := sr.2.1.1 @@ -614,7 +614,7 @@ theorem balance_eq_balance' {l x r} (hl : Balanced l) (hr : Balanced r) (sl : Si · rfl · have : size lrl = 0 ∧ size lrr = 0 := by have := balancedSz_zero.1 hl.1.symm - rwa [size, sl.2.2.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero_iff] at this + rwa [size, sl.2.2.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero] at this cases sl.2.2.2.1.size_eq_zero.1 this.1 cases sl.2.2.2.2.size_eq_zero.1 this.2 obtain rfl : lrs = 1 := sl.2.2.1 @@ -622,7 +622,7 @@ theorem balance_eq_balance' {l x r} (hl : Balanced l) (hr : Balanced r) (sl : Si all_goals dsimp only [size]; decide · have : size lll = 0 ∧ size llr = 0 := by have := balancedSz_zero.1 hl.1 - rwa [size, sl.2.1.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero_iff] at this + rwa [size, sl.2.1.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero] at this cases sl.2.1.2.1.size_eq_zero.1 this.1 cases sl.2.1.2.2.size_eq_zero.1 this.2 obtain rfl : lls = 1 := sl.2.1.1 @@ -670,7 +670,7 @@ theorem balanceL_eq_balance {l x r} (sl : Sized l) (sr : Sized r) (H1 : size l = · cases' l with ls ll lx lr · have : size rl = 0 ∧ size rr = 0 := by have := H1 rfl - rwa [size, sr.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero_iff] at this + rwa [size, sr.1, Nat.succ_le_succ_iff, Nat.le_zero, add_eq_zero] at this cases sr.2.1.size_eq_zero.1 this.1 cases sr.2.2.size_eq_zero.1 this.2 rw [sr.eq_node']; rfl @@ -1011,7 +1011,7 @@ theorem Valid'.node4L {l} {x : α} {m} {y : α} {r o₁ o₂} (hl : Valid' o₁ BalancedSz (size mr) (size r) ∧ BalancedSz (size l + size ml + 1) (size mr + size r + 1) from Valid'.node' (hl.node' hm.left this.1) (hm.right.node' hr this.2.1) this.2.2 rcases H with (⟨l0, m1, r0⟩ | ⟨l0, mr₁, lr₁, lr₂, mr₂⟩) - · rw [hm.2.size_eq, Nat.succ_inj', add_eq_zero_iff] at m1 + · rw [hm.2.size_eq, Nat.succ_inj', add_eq_zero] at m1 rw [l0, m1.1, m1.2]; revert r0; rcases size r with (_ | _ | _) <;> [decide; decide; (intro r0; unfold BalancedSz delta; omega)] · rcases Nat.eq_zero_or_pos (size r) with r0 | r0 diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/RightAngle.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/RightAngle.lean index 7dc41614cd346..a85b10c3cdb58 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/RightAngle.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/RightAngle.lean @@ -150,7 +150,7 @@ theorem cos_angle_add_mul_norm_of_inner_eq_zero {x y : V} (h : ⟪x, y⟫ = 0) : by_cases hxy : ‖x + y‖ = 0 · have h' := norm_add_sq_eq_norm_sq_add_norm_sq_real h rw [hxy, zero_mul, eq_comm, - add_eq_zero_iff' (mul_self_nonneg ‖x‖) (mul_self_nonneg ‖y‖), mul_self_eq_zero] at h' + add_eq_zero_iff_of_nonneg (mul_self_nonneg ‖x‖) (mul_self_nonneg ‖y‖), mul_self_eq_zero] at h' simp [h'.1] · exact div_mul_cancel₀ _ hxy diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean index 0f5414c2c8024..5ec5b7bcee160 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean @@ -167,7 +167,7 @@ theorem posDef_prod_iff [PartialOrder P] [CovariantClass P P (· + ·) (· ≤ · rintro ⟨⟨hle₁, ha₁⟩, ⟨hle₂, ha₂⟩⟩ refine ⟨⟨hle₁, hle₂⟩, ?_⟩ rintro ⟨x₁, x₂⟩ (hx : Q₁ x₁ + Q₂ x₂ = 0) - rw [add_eq_zero_iff' (hle₁ x₁) (hle₂ x₂), ha₁.eq_zero_iff, ha₂.eq_zero_iff] at hx + rw [add_eq_zero_iff_of_nonneg (hle₁ x₁) (hle₂ x₂), ha₁.eq_zero_iff, ha₂.eq_zero_iff] at hx rwa [Prod.mk_eq_zero] theorem PosDef.prod [PartialOrder P] [CovariantClass P P (· + ·) (· ≤ ·)] diff --git a/Mathlib/MeasureTheory/Covering/Besicovitch.lean b/Mathlib/MeasureTheory/Covering/Besicovitch.lean index e8ce34467ab0a..46af5641e5aa1 100644 --- a/Mathlib/MeasureTheory/Covering/Besicovitch.lean +++ b/Mathlib/MeasureTheory/Covering/Besicovitch.lean @@ -781,7 +781,7 @@ theorem exists_disjoint_closedBall_covering_ae_of_finiteMeasure_aux (μ : Measur rw [ENNReal.div_lt_iff, one_mul] · conv_lhs => rw [← add_zero (N : ℝ≥0∞)] exact ENNReal.add_lt_add_left (ENNReal.natCast_ne_top N) zero_lt_one - · simp only [true_or_iff, add_eq_zero_iff, Ne, not_false_iff, one_ne_zero, and_false_iff] + · simp only [true_or_iff, add_eq_zero, Ne, not_false_iff, one_ne_zero, and_false_iff] · simp only [ENNReal.natCast_ne_top, Ne, not_false_iff, or_true_iff] rw [zero_mul] at C apply le_bot_iff.1 diff --git a/Mathlib/MeasureTheory/Decomposition/Jordan.lean b/Mathlib/MeasureTheory/Decomposition/Jordan.lean index d7d28a7b8bd02..6ff1e74d9497e 100644 --- a/Mathlib/MeasureTheory/Decomposition/Jordan.lean +++ b/Mathlib/MeasureTheory/Decomposition/Jordan.lean @@ -454,7 +454,7 @@ theorem totalVariation_neg (s : SignedMeasure α) : (-s).totalVariation = s.tota theorem null_of_totalVariation_zero (s : SignedMeasure α) {i : Set α} (hs : s.totalVariation i = 0) : s i = 0 := by - rw [totalVariation, Measure.coe_add, Pi.add_apply, add_eq_zero_iff] at hs + rw [totalVariation, Measure.coe_add, Pi.add_apply, add_eq_zero] at hs rw [← toSignedMeasure_toJordanDecomposition s, toSignedMeasure, VectorMeasure.coe_sub, Pi.sub_apply, Measure.toSignedMeasure_apply, Measure.toSignedMeasure_apply] by_cases hi : MeasurableSet i @@ -484,7 +484,7 @@ theorem totalVariation_absolutelyContinuous_iff (s : SignedMeasure α) (μ : Mea all_goals refine Measure.AbsolutelyContinuous.mk fun S _ hS₂ => ?_ have := h hS₂ - rw [totalVariation, Measure.add_apply, add_eq_zero_iff] at this + rw [totalVariation, Measure.add_apply, add_eq_zero] at this exacts [this.1, this.2] · refine Measure.AbsolutelyContinuous.mk fun S _ hS₂ => ?_ rw [totalVariation, Measure.add_apply, h.1 hS₂, h.2 hS₂, add_zero] diff --git a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean index fc1891f03e347..a74a832aa9956 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean @@ -126,7 +126,7 @@ theorem singularPart_mutuallySingular (s : SignedMeasure α) (μ : Measure α) : · obtain ⟨i, hi, hpos, hneg⟩ := s.toJordanDecomposition.mutuallySingular rw [s.toJordanDecomposition.posPart.haveLebesgueDecomposition_add μ] at hpos rw [s.toJordanDecomposition.negPart.haveLebesgueDecomposition_add μ] at hneg - rw [add_apply, add_eq_zero_iff] at hpos hneg + rw [add_apply, add_eq_zero] at hpos hneg exact ⟨i, hi, hpos.1, hneg.1⟩ · rw [not_haveLebesgueDecomposition_iff] at hl cases' hl with hp hn diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index 09321be0ee129..f1612b1c061fe 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -592,7 +592,7 @@ theorem ae_smul_measure {p : α → Prop} [Monoid R] [DistribMulAction R ℝ≥0 theorem ae_add_measure_iff {p : α → Prop} {ν} : (∀ᵐ x ∂μ + ν, p x) ↔ (∀ᵐ x ∂μ, p x) ∧ ∀ᵐ x ∂ν, p x := - add_eq_zero_iff + add_eq_zero theorem ae_eq_comp' {ν : Measure β} {f : α → β} {g g' : β → δ} (hf : AEMeasurable f μ) (h : g =ᵐ[ν] g') (h2 : μ.map f ≪ ν) : g ∘ f =ᵐ[μ] g' ∘ f := diff --git a/Mathlib/NumberTheory/Bernoulli.lean b/Mathlib/NumberTheory/Bernoulli.lean index d50d5a91f1b05..e6c6d64dc34d4 100644 --- a/Mathlib/NumberTheory/Bernoulli.lean +++ b/Mathlib/NumberTheory/Bernoulli.lean @@ -256,7 +256,7 @@ theorem bernoulliPowerSeries_mul_exp_sub_one : bernoulliPowerSeries A * (exp A - · simp simp only [bernoulliPowerSeries, coeff_mul, coeff_X, sum_antidiagonal_succ', one_div, coeff_mk, coeff_one, coeff_exp, LinearMap.map_sub, factorial, if_pos, cast_succ, cast_one, cast_mul, - sub_zero, RingHom.map_one, add_eq_zero_iff, if_false, _root_.inv_one, zero_add, one_ne_zero, + sub_zero, RingHom.map_one, add_eq_zero, if_false, _root_.inv_one, zero_add, one_ne_zero, mul_zero, and_false_iff, sub_self, ← RingHom.map_mul, ← map_sum] cases' n with n · simp diff --git a/Mathlib/NumberTheory/Zsqrtd/Basic.lean b/Mathlib/NumberTheory/Zsqrtd/Basic.lean index 8afaf8159fc1c..1a153a411f763 100644 --- a/Mathlib/NumberTheory/Zsqrtd/Basic.lean +++ b/Mathlib/NumberTheory/Zsqrtd/Basic.lean @@ -521,7 +521,7 @@ theorem norm_eq_zero_iff {d : ℤ} (hd : d < 0) (z : ℤ√d) : z.norm = 0 ↔ z rw [norm_def, sub_eq_add_neg, mul_assoc] at h have left := mul_self_nonneg z.re have right := neg_nonneg.mpr (mul_nonpos_of_nonpos_of_nonneg hd.le (mul_self_nonneg z.im)) - obtain ⟨ha, hb⟩ := (add_eq_zero_iff' left right).mp h + obtain ⟨ha, hb⟩ := (add_eq_zero_iff_of_nonneg left right).mp h ext <;> apply eq_zero_of_mul_self_eq_zero · exact ha · rw [neg_eq_zero, mul_eq_zero] at hb diff --git a/Mathlib/RingTheory/Coprime/Basic.lean b/Mathlib/RingTheory/Coprime/Basic.lean index 4704beb924263..9e261696f011a 100644 --- a/Mathlib/RingTheory/Coprime/Basic.lean +++ b/Mathlib/RingTheory/Coprime/Basic.lean @@ -360,7 +360,7 @@ end CommRing theorem sq_add_sq_ne_zero {R : Type*} [LinearOrderedCommRing R] {a b : R} (h : IsCoprime a b) : a ^ 2 + b ^ 2 ≠ 0 := by intro h' - obtain ⟨ha, hb⟩ := (add_eq_zero_iff' (sq_nonneg _) (sq_nonneg _)).mp h' + obtain ⟨ha, hb⟩ := (add_eq_zero_iff_of_nonneg (sq_nonneg _) (sq_nonneg _)).mp h' obtain rfl := pow_eq_zero ha obtain rfl := pow_eq_zero hb exact not_isCoprime_zero_zero h diff --git a/Mathlib/RingTheory/GradedAlgebra/Basic.lean b/Mathlib/RingTheory/GradedAlgebra/Basic.lean index 8f91b495bb1d4..de86d0cc60388 100644 --- a/Mathlib/RingTheory/GradedAlgebra/Basic.lean +++ b/Mathlib/RingTheory/GradedAlgebra/Basic.lean @@ -262,10 +262,10 @@ def GradedRing.projZeroRingHom : A →+* A where by_cases h : i + j = 0 · rw [decompose_of_mem_same 𝒜 (show c * c' ∈ 𝒜 0 from h ▸ SetLike.GradedMul.mul_mem hc hc'), - decompose_of_mem_same 𝒜 (show c ∈ 𝒜 0 from (add_eq_zero_iff.mp h).1 ▸ hc), - decompose_of_mem_same 𝒜 (show c' ∈ 𝒜 0 from (add_eq_zero_iff.mp h).2 ▸ hc')] + decompose_of_mem_same 𝒜 (show c ∈ 𝒜 0 from (add_eq_zero.mp h).1 ▸ hc), + decompose_of_mem_same 𝒜 (show c' ∈ 𝒜 0 from (add_eq_zero.mp h).2 ▸ hc')] · rw [decompose_of_mem_ne 𝒜 (SetLike.GradedMul.mul_mem hc hc') h] - cases' show i ≠ 0 ∨ j ≠ 0 by rwa [add_eq_zero_iff, not_and_or] at h with h' h' + cases' show i ≠ 0 ∨ j ≠ 0 by rwa [add_eq_zero, not_and_or] at h with h' h' · simp only [decompose_of_mem_ne 𝒜 hc h', zero_mul] · simp only [decompose_of_mem_ne 𝒜 hc' h', mul_zero] · intro _ _ hd he diff --git a/Mathlib/RingTheory/PowerSeries/Order.lean b/Mathlib/RingTheory/PowerSeries/Order.lean index 0e7ceaefeb26b..b87f41f0ee1ca 100644 --- a/Mathlib/RingTheory/PowerSeries/Order.lean +++ b/Mathlib/RingTheory/PowerSeries/Order.lean @@ -296,7 +296,7 @@ theorem order_one : order (1 : R⟦X⟧) = 0 := by theorem order_zero_of_unit {f : PowerSeries R} : IsUnit f → f.order = 0 := by rintro ⟨⟨u, v, hu, hv⟩, hf⟩ apply And.left - rw [← add_eq_zero_iff, ← hf, ← nonpos_iff_eq_zero, ← @order_one R _ _, ← hu] + rw [← add_eq_zero, ← hf, ← nonpos_iff_eq_zero, ← @order_one R _ _, ← hu] exact order_mul_ge _ _ /-- The order of the formal power series `X` is `1`. -/ diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index 99569ecf6d0d0..fd3b1c5921f22 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -1762,7 +1762,7 @@ theorem eq_pow_of_mul_eq_pow {a b c : Associates α} (ha : a ≠ 0) (hb : b ≠ by_cases hk0 : k = 0 · use 1 rw [hk0, pow_zero] at h ⊢ - apply (mul_eq_one_iff.1 h).1 + apply (mul_eq_one.1 h).1 · refine is_pow_of_dvd_count ha fun p hp ↦ ?_ apply dvd_count_of_dvd_count_mul hb hp hab rw [h] diff --git a/Mathlib/SetTheory/Cardinal/Divisibility.lean b/Mathlib/SetTheory/Cardinal/Divisibility.lean index f586cb6fbc296..658f7943b383d 100644 --- a/Mathlib/SetTheory/Cardinal/Divisibility.lean +++ b/Mathlib/SetTheory/Cardinal/Divisibility.lean @@ -47,7 +47,7 @@ theorem isUnit_iff : IsUnit a ↔ a = 1 := by · exact (not_isUnit_zero h).elim rw [isUnit_iff_forall_dvd] at h cases' h 1 with t ht - rw [eq_comm, mul_eq_one_iff'] at ht + rw [eq_comm, mul_eq_one_iff_of_one_le] at ht · exact ht.1 · exact one_le_iff_ne_zero.mpr ha · apply one_le_iff_ne_zero.mpr From 7bb6e8b211eb21d9ca0f5f6f2eb4065692db4777 Mon Sep 17 00:00:00 2001 From: Patrick Massot Date: Sun, 11 Aug 2024 10:40:56 +0000 Subject: [PATCH 170/359] refactor: deprecate stdBasis (#15660) Deprecate stdBasis in favor of its duplicate LinearMap.single. Also changes the binder types in the later. Note the `simp` normal form of an applied (or even simply coerced to function) `LinearMap.single` is a `Pi.single`, as discussed during the community meeting. See the [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/LinearMap.2EstdBasis/near/457515331). Co-authored-by: Eric Wieser --- Mathlib/Analysis/Calculus/TangentCone.lean | 2 +- Mathlib/Analysis/Normed/Lp/PiLp.lean | 2 +- Mathlib/Data/Matrix/Basic.lean | 16 ++ Mathlib/Data/Matrix/Rank.lean | 2 +- .../LinearAlgebra/AffineSpace/AffineMap.lean | 4 +- Mathlib/LinearAlgebra/Basis/Defs.lean | 4 +- Mathlib/LinearAlgebra/FinsuppVectorSpace.lean | 13 +- Mathlib/LinearAlgebra/Matrix/Basis.lean | 2 +- .../LinearAlgebra/Matrix/BilinearForm.lean | 32 ++- Mathlib/LinearAlgebra/Matrix/Diagonal.lean | 19 +- Mathlib/LinearAlgebra/Matrix/DotProduct.lean | 20 +- .../Matrix/SesquilinearForm.lean | 47 +++-- Mathlib/LinearAlgebra/Matrix/ToLin.lean | 45 ++-- Mathlib/LinearAlgebra/Pi.lean | 97 ++++++++- .../LinearAlgebra/QuadraticForm/Complex.lean | 6 +- Mathlib/LinearAlgebra/QuadraticForm/Prod.lean | 2 +- Mathlib/LinearAlgebra/QuadraticForm/Real.lean | 2 +- Mathlib/LinearAlgebra/QuotientPi.lean | 2 +- Mathlib/LinearAlgebra/StdBasis.lean | 193 +++++++++--------- .../Measure/Lebesgue/EqHaar.lean | 2 +- Mathlib/RingTheory/AdjoinRoot.lean | 5 +- Mathlib/RingTheory/Finiteness.lean | 2 +- .../RingTheory/Flat/EquationalCriterion.lean | 2 +- Mathlib/RingTheory/MvPowerSeries/Basic.lean | 11 +- Mathlib/RingTheory/SimpleModule.lean | 2 +- 25 files changed, 321 insertions(+), 213 deletions(-) diff --git a/Mathlib/Analysis/Calculus/TangentCone.lean b/Mathlib/Analysis/Calculus/TangentCone.lean index af0dc8bc44f58..d75973e9bd895 100644 --- a/Mathlib/Analysis/Calculus/TangentCone.lean +++ b/Mathlib/Analysis/Calculus/TangentCone.lean @@ -169,7 +169,7 @@ theorem subset_tangentCone_prod_right {t : Set F} {y : F} (hs : x ∈ closure s) theorem mapsTo_tangentCone_pi {ι : Type*} [DecidableEq ι] {E : ι → Type*} [∀ i, NormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] {s : ∀ i, Set (E i)} {x : ∀ i, E i} {i : ι} (hi : ∀ j ≠ i, x j ∈ closure (s j)) : - MapsTo (LinearMap.single i : E i →ₗ[𝕜] ∀ j, E j) (tangentConeAt 𝕜 (s i) (x i)) + MapsTo (LinearMap.single 𝕜 E i) (tangentConeAt 𝕜 (s i) (x i)) (tangentConeAt 𝕜 (Set.pi univ s) x) := by rintro w ⟨c, d, hd, hc, hy⟩ have : ∀ n, ∀ j ≠ i, ∃ d', x j + d' ∈ s j ∧ ‖c n • d'‖ < (1 / 2 : ℝ) ^ n := fun n j hj ↦ by diff --git a/Mathlib/Analysis/Normed/Lp/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean index cabcd6dc2bf05..e1c8c3ba748a1 100644 --- a/Mathlib/Analysis/Normed/Lp/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -913,7 +913,7 @@ def basisFun : Basis ι 𝕜 (PiLp p fun _ : ι => 𝕜) := @[simp] theorem basisFun_apply [DecidableEq ι] (i) : basisFun p 𝕜 ι i = (WithLp.equiv p _).symm (Pi.single i 1) := by - simp_rw [basisFun, Basis.coe_ofEquivFun, WithLp.linearEquiv_symm_apply, Pi.single] + simp_rw [basisFun, Basis.coe_ofEquivFun, WithLp.linearEquiv_symm_apply] @[simp] theorem basisFun_repr (x : PiLp p fun _ : ι => 𝕜) (i : ι) : (basisFun p 𝕜 ι).repr x i = x i := diff --git a/Mathlib/Data/Matrix/Basic.lean b/Mathlib/Data/Matrix/Basic.lean index 29c4b9256a7d2..3d645f33d8402 100644 --- a/Mathlib/Data/Matrix/Basic.lean +++ b/Mathlib/Data/Matrix/Basic.lean @@ -783,6 +783,14 @@ variable [NonAssocSemiring α] theorem one_dotProduct_one : (1 : n → α) ⬝ᵥ 1 = Fintype.card n := by simp [dotProduct] +theorem dotProduct_single_one [DecidableEq n] (v : n → α) (i : n) : + dotProduct v (Pi.single i 1) = v i := by + rw [dotProduct_single, mul_one] + +theorem single_one_dotProduct [DecidableEq n] (i : n) (v : n → α) : + dotProduct (Pi.single i 1) v = v i := by + rw [single_dotProduct, one_mul] + end NonAssocSemiring section NonUnitalNonAssocRing @@ -1586,6 +1594,14 @@ theorem single_vecMul [Fintype m] [DecidableEq m] [NonUnitalNonAssocSemiring R] (i : m) (x : R) : Pi.single i x ᵥ* M = fun j => x * M i j := funext fun _ => single_dotProduct _ _ _ +theorem mulVec_single_one [Fintype n] [DecidableEq n] [NonAssocSemiring R] + (M : Matrix m n R) (j : n) : + M *ᵥ Pi.single j 1 = Mᵀ j := by ext; simp + +theorem single_one_vecMul [Fintype m] [DecidableEq m] [NonAssocSemiring R] + (i : m) (M : Matrix m n R) : + Pi.single i 1 ᵥ* M = M i := by simp + -- @[simp] -- Porting note: not in simpNF theorem diagonal_mulVec_single [Fintype n] [DecidableEq n] [NonUnitalNonAssocSemiring R] (v : n → R) (j : n) (x : R) : diagonal v *ᵥ Pi.single j x = Pi.single j (v j * x) := by diff --git a/Mathlib/Data/Matrix/Rank.lean b/Mathlib/Data/Matrix/Rank.lean index e62b41b4656d9..9c78955d6091f 100644 --- a/Mathlib/Data/Matrix/Rank.lean +++ b/Mathlib/Data/Matrix/Rank.lean @@ -141,7 +141,7 @@ theorem rank_eq_finrank_range_toLin [Finite m] [DecidableEq n] {M₁ M₂ : Type have aux₁ := toLin_self (Pi.basisFun R n) (Pi.basisFun R m) A i have aux₂ := Basis.equiv_apply (Pi.basisFun R n) i v₂ rw [toLin_eq_toLin', toLin'_apply'] at aux₁ - rw [Pi.basisFun_apply, LinearMap.coe_stdBasis] at aux₁ aux₂ + rw [Pi.basisFun_apply] at aux₁ aux₂ simp only [e₁, e₁, LinearMap.comp_apply, LinearEquiv.coe_coe, Equiv.refl_apply, aux₁, aux₂, LinearMap.coe_single, toLin_self, map_sum, LinearEquiv.map_smul, Basis.equiv_apply] diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean index 2bc4d04b8a70c..dfd2890057957 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean @@ -738,8 +738,8 @@ theorem pi_ext_nonempty [Nonempty ι] (h : ∀ i x, f (Pi.single i x) = g (Pi.si /-- This is used as the ext lemma instead of `AffineMap.pi_ext_nonempty` for reasons explained in note [partially-applied ext lemmas]. Analogous to `LinearMap.pi_ext'`-/ @[ext (iff := false)] -theorem pi_ext_nonempty' [Nonempty ι] (h : ∀ i, f.comp (LinearMap.single i).toAffineMap = - g.comp (LinearMap.single i).toAffineMap) : f = g := by +theorem pi_ext_nonempty' [Nonempty ι] (h : ∀ i, f.comp (LinearMap.single _ _ i).toAffineMap = + g.comp (LinearMap.single _ _ i).toAffineMap) : f = g := by refine pi_ext_nonempty fun i x => ?_ convert AffineMap.congr_fun (h i) x diff --git a/Mathlib/LinearAlgebra/Basis/Defs.lean b/Mathlib/LinearAlgebra/Basis/Defs.lean index 639fb724332cf..b77f16e72455a 100644 --- a/Mathlib/LinearAlgebra/Basis/Defs.lean +++ b/Mathlib/LinearAlgebra/Basis/Defs.lean @@ -740,11 +740,11 @@ theorem Basis.ofEquivFun_repr_apply [Finite ι] (e : M ≃ₗ[R] ι → R) (x : @[simp] theorem Basis.coe_ofEquivFun [Finite ι] [DecidableEq ι] (e : M ≃ₗ[R] ι → R) : - (Basis.ofEquivFun e : ι → M) = fun i => e.symm (Function.update 0 i 1) := + (Basis.ofEquivFun e : ι → M) = fun i => e.symm (Pi.single i 1) := funext fun i => e.injective <| funext fun j => by - simp [Basis.ofEquivFun, ← Finsupp.single_eq_pi_single, Finsupp.single_eq_update] + simp [Basis.ofEquivFun, ← Finsupp.single_eq_pi_single] @[simp] theorem Basis.ofEquivFun_equivFun [Finite ι] (v : Basis ι R M) : diff --git a/Mathlib/LinearAlgebra/FinsuppVectorSpace.lean b/Mathlib/LinearAlgebra/FinsuppVectorSpace.lean index 0599bc591d200..6c96e6b6ad6ae 100644 --- a/Mathlib/LinearAlgebra/FinsuppVectorSpace.lean +++ b/Mathlib/LinearAlgebra/FinsuppVectorSpace.lean @@ -154,9 +154,16 @@ theorem _root_.Finset.sum_single_ite [Fintype n] (a : R) (i : n) : simp only [apply_ite (Finsupp.single _), Finsupp.single_zero, Finset.sum_ite_eq, if_pos (Finset.mem_univ _)] -theorem equivFun_symm_stdBasis [Finite n] (b : Basis n R M) (i : n) : - b.equivFun.symm (LinearMap.stdBasis R (fun _ => R) i 1) = b i := by +@[simp] +theorem equivFun_symm_single [Finite n] (b : Basis n R M) (i : n) : + b.equivFun.symm (Pi.single i 1) = b i := by cases nonempty_fintype n - simp + simp [Pi.single_apply] + +set_option linter.deprecated false in +@[deprecated equivFun_symm_single (since := "2024-08-09")] +theorem equivFun_symm_stdBasis [Finite n] (b : Basis n R M) (i : n) : + b.equivFun.symm (LinearMap.stdBasis R (fun _ => R) i 1) = b i := + equivFun_symm_single .. end Basis diff --git a/Mathlib/LinearAlgebra/Matrix/Basis.lean b/Mathlib/LinearAlgebra/Matrix/Basis.lean index 6c0ce89d7f42e..3083418e9342d 100644 --- a/Mathlib/LinearAlgebra/Matrix/Basis.lean +++ b/Mathlib/LinearAlgebra/Matrix/Basis.lean @@ -220,7 +220,7 @@ theorem basis_toMatrix_basisFun_mul (b : Basis ι R (ι → R)) (A : Matrix ι simp only [basis_toMatrix_mul _ _ (Pi.basisFun R ι), Matrix.toLin_eq_toLin'] ext i j rw [LinearMap.toMatrix_apply, Matrix.toLin'_apply, Pi.basisFun_apply, - Matrix.mulVec_stdBasis_apply, Matrix.of_apply] + Matrix.mulVec_single_one, Matrix.of_apply] /-- See also `Basis.toMatrix_reindex` which gives the `simp` normal form of this result. -/ theorem Basis.toMatrix_reindex' [DecidableEq ι] [DecidableEq ι'] (b : Basis ι R M) (v : ι' → M) diff --git a/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean b/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean index c48d73e201987..1c8964c5c90ae 100644 --- a/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean +++ b/Mathlib/LinearAlgebra/Matrix/BilinearForm.lean @@ -60,10 +60,9 @@ This is an auxiliary definition for the equivalence `Matrix.toBilin'`. -/ def Matrix.toBilin'Aux [Fintype n] (M : Matrix n n R₂) : BilinForm R₂ (n → R₂) := Matrix.toLinearMap₂'Aux _ _ M -theorem Matrix.toBilin'Aux_stdBasis [Fintype n] [DecidableEq n] (M : Matrix n n R₂) (i j : n) : - M.toBilin'Aux (LinearMap.stdBasis R₂ (fun _ => R₂) i 1) - (LinearMap.stdBasis R₂ (fun _ => R₂) j 1) = M i j := - Matrix.toLinearMap₂'Aux_stdBasis _ _ _ _ _ +theorem Matrix.toBilin'Aux_single [Fintype n] [DecidableEq n] (M : Matrix n n R₂) (i j : n) : + M.toBilin'Aux (Pi.single i 1) (Pi.single j 1) = M i j := + Matrix.toLinearMap₂'Aux_single _ _ _ _ _ /-- The linear map from bilinear forms to `Matrix n n R` given an `n`-indexed basis. @@ -81,10 +80,8 @@ variable [Fintype n] [Fintype o] theorem toBilin'Aux_toMatrixAux [DecidableEq n] (B₂ : BilinForm R₂ (n → R₂)) : -- Porting note: had to hint the base ring even though it should be clear from context... - Matrix.toBilin'Aux (BilinForm.toMatrixAux (R₂ := R₂) - (fun j => stdBasis R₂ (fun _ => R₂) j 1) B₂) = B₂ := by - rw [BilinForm.toMatrixAux, Matrix.toBilin'Aux, - toLinearMap₂'Aux_toMatrix₂Aux] + Matrix.toBilin'Aux (BilinForm.toMatrixAux (R₂ := R₂) (fun j => Pi.single j 1) B₂) = B₂ := by + rw [BilinForm.toMatrixAux, Matrix.toBilin'Aux, toLinearMap₂'Aux_toMatrix₂Aux] section ToMatrix' @@ -100,13 +97,6 @@ variable [DecidableEq n] [DecidableEq o] def LinearMap.BilinForm.toMatrix' : BilinForm R₂ (n → R₂) ≃ₗ[R₂] Matrix n n R₂ := LinearMap.toMatrix₂' R₂ -@[simp] -theorem LinearMap.BilinForm.toMatrixAux_stdBasis (B : BilinForm R₂ (n → R₂)) : - -- Porting note: had to hint the base ring even though it should be clear from context... - BilinForm.toMatrixAux (R₂ := R₂) (fun j => stdBasis R₂ (fun _ => R₂) j 1) B = - BilinForm.toMatrix' B := - rfl - /-- The linear equivalence between `n × n` matrices and bilinear forms on `n → R` -/ def Matrix.toBilin' : Matrix n n R₂ ≃ₗ[R₂] BilinForm R₂ (n → R₂) := BilinForm.toMatrix'.symm @@ -124,10 +114,16 @@ theorem Matrix.toBilin'_apply' (M : Matrix n n R₂) (v w : n → R₂) : Matrix.toBilin' M v w = Matrix.dotProduct v (M *ᵥ w) := Matrix.toLinearMap₂'_apply' _ _ _ @[simp] +theorem Matrix.toBilin'_single (M : Matrix n n R₂) (i j : n) : + Matrix.toBilin' M (Pi.single i 1) (Pi.single j 1) = M i j := by + simp [Matrix.toBilin'_apply, Pi.single_apply] + +set_option linter.deprecated false in +@[simp, deprecated Matrix.toBilin'_single (since := "2024-08-09")] theorem Matrix.toBilin'_stdBasis (M : Matrix n n R₂) (i j : n) : Matrix.toBilin' M - (LinearMap.stdBasis R₂ (fun _ => R₂) i 1) - (LinearMap.stdBasis R₂ (fun _ => R₂) j 1) = M i j := Matrix.toLinearMap₂'_stdBasis _ _ _ + (LinearMap.stdBasis R₂ (fun _ ↦ R₂) i 1) + (LinearMap.stdBasis R₂ (fun _ ↦ R₂) j 1) = M i j := Matrix.toBilin'_single _ _ _ @[simp] theorem LinearMap.BilinForm.toMatrix'_symm : @@ -153,7 +149,7 @@ theorem BilinForm.toMatrix'_toBilin' (M : Matrix n n R₂) : @[simp] theorem BilinForm.toMatrix'_apply (B : BilinForm R₂ (n → R₂)) (i j : n) : - BilinForm.toMatrix' B i j = B (stdBasis R₂ (fun _ => R₂) i 1) (stdBasis R₂ (fun _ => R₂) j 1) := + BilinForm.toMatrix' B i j = B (Pi.single i 1) (Pi.single j 1) := LinearMap.toMatrix₂'_apply _ _ _ -- Porting note: dot notation for bundled maps doesn't work in the rest of this section diff --git a/Mathlib/LinearAlgebra/Matrix/Diagonal.lean b/Mathlib/LinearAlgebra/Matrix/Diagonal.lean index c1ad209b36602..a8b92d3cd2a99 100644 --- a/Mathlib/LinearAlgebra/Matrix/Diagonal.lean +++ b/Mathlib/LinearAlgebra/Matrix/Diagonal.lean @@ -32,6 +32,13 @@ variable {n : Type*} [Fintype n] [DecidableEq n] {R : Type v} [CommSemiring R] theorem proj_diagonal (i : n) (w : n → R) : (proj i).comp (toLin' (diagonal w)) = w i • proj i := LinearMap.ext fun _ => mulVec_diagonal _ _ _ +theorem diagonal_comp_single (w : n → R) (i : n) : + (diagonal w).toLin'.comp (LinearMap.single R (fun _ : n => R) i) = + w i • LinearMap.single R (fun _ : n => R) i := + LinearMap.ext fun x => (diagonal_mulVec_single w _ _).trans (Pi.single_smul' i (w i) x) + +set_option linter.deprecated false in +@[deprecated diagonal_comp_single (since := "2024-08-09")] theorem diagonal_comp_stdBasis (w : n → R) (i : n) : (diagonal w).toLin'.comp (LinearMap.stdBasis R (fun _ : n => R) i) = w i • LinearMap.stdBasis R (fun _ : n => R) i := @@ -50,21 +57,21 @@ variable {m n : Type*} [Fintype m] [Fintype n] {K : Type u} [Semifield K] -- maybe try to relax the universe constraint theorem ker_diagonal_toLin' [DecidableEq m] (w : m → K) : ker (toLin' (diagonal w)) = - ⨆ i ∈ { i | w i = 0 }, LinearMap.range (LinearMap.stdBasis K (fun _ => K) i) := by + ⨆ i ∈ { i | w i = 0 }, LinearMap.range (LinearMap.single K (fun _ => K) i) := by rw [← comap_bot, ← iInf_ker_proj, comap_iInf] have := fun i : m => ker_comp (toLin' (diagonal w)) (proj i) simp only [comap_iInf, ← this, proj_diagonal, ker_smul'] have : univ ⊆ { i : m | w i = 0 } ∪ { i : m | w i = 0 }ᶜ := by rw [Set.union_compl_self] - exact (iSup_range_stdBasis_eq_iInf_ker_proj K (fun _ : m => K) disjoint_compl_right this + exact (iSup_range_single_eq_iInf_ker_proj K (fun _ : m => K) disjoint_compl_right this (Set.toFinite _)).symm theorem range_diagonal [DecidableEq m] (w : m → K) : LinearMap.range (toLin' (diagonal w)) = - ⨆ i ∈ { i | w i ≠ 0 }, LinearMap.range (LinearMap.stdBasis K (fun _ => K) i) := by + ⨆ i ∈ { i | w i ≠ 0 }, LinearMap.range (LinearMap.single K (fun _ => K) i) := by dsimp only [mem_setOf_eq] - rw [← Submodule.map_top, ← iSup_range_stdBasis, Submodule.map_iSup] + rw [← Submodule.map_top, ← iSup_range_single, Submodule.map_iSup] congr; funext i - rw [← LinearMap.range_comp, diagonal_comp_stdBasis, ← range_smul'] + rw [← LinearMap.range_comp, diagonal_comp_single, ← range_smul'] end Semifield @@ -80,7 +87,7 @@ theorem rank_diagonal [DecidableEq m] [DecidableEq K] (w : m → K) : LinearMap.rank (toLin' (diagonal w)) = Fintype.card { i // w i ≠ 0 } := by have hu : univ ⊆ { i : m | w i = 0 }ᶜ ∪ { i : m | w i = 0 } := by rw [Set.compl_union_self] have hd : Disjoint { i : m | w i ≠ 0 } { i : m | w i = 0 } := disjoint_compl_left - have B₁ := iSup_range_stdBasis_eq_iInf_ker_proj K (fun _ : m => K) hd hu (Set.toFinite _) + have B₁ := iSup_range_single_eq_iInf_ker_proj K (fun _ : m => K) hd hu (Set.toFinite _) have B₂ := iInfKerProjEquiv K (fun _ ↦ K) hd hu rw [LinearMap.rank, range_diagonal, B₁, ← @rank_fun' K] apply LinearEquiv.rank_eq diff --git a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean index 2c95ece083bf7..4b648de325167 100644 --- a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean +++ b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean @@ -18,7 +18,7 @@ vectors `v w : n → R` to the sum of the entrywise products `v i * w i`. * `Matrix.dotProduct_stdBasis_one`: the dot product of `v` with the `i`th standard basis vector is `v i` -* `Matrix.dotProduct_eq_zero_iff`: if `v`'s' dot product with all `w` is zero, +* `Matrix.dotProduct_eq_zero_iff`: if `v`'s dot product with all `w` is zero, then `v` is zero ## Tags @@ -36,21 +36,21 @@ section Semiring variable [Semiring R] [Fintype n] -@[simp] +set_option linter.deprecated false in +@[simp, deprecated dotProduct_single (since := "2024-08-09")] theorem dotProduct_stdBasis_eq_mul [DecidableEq n] (v : n → R) (c : R) (i : n) : - dotProduct v (LinearMap.stdBasis R (fun _ => R) i c) = v i * c := by - rw [dotProduct, Finset.sum_eq_single i, LinearMap.stdBasis_same] - · exact fun _ _ hb => by rw [LinearMap.stdBasis_ne _ _ _ _ hb, mul_zero] - · exact fun hi => False.elim (hi <| Finset.mem_univ _) + dotProduct v (LinearMap.stdBasis R (fun _ => R) i c) = v i * c := + dotProduct_single .. --- @[simp] -- Porting note (#10618): simp can prove this +set_option linter.deprecated false in +@[deprecated dotProduct_single_one (since := "2024-08-09")] theorem dotProduct_stdBasis_one [DecidableEq n] (v : n → R) (i : n) : - dotProduct v (LinearMap.stdBasis R (fun _ => R) i 1) = v i := by - rw [dotProduct_stdBasis_eq_mul, mul_one] + dotProduct v (LinearMap.stdBasis R (fun _ => R) i 1) = v i := + dotProduct_single_one .. theorem dotProduct_eq (v w : n → R) (h : ∀ u, dotProduct v u = dotProduct w u) : v = w := by funext x - classical rw [← dotProduct_stdBasis_one v x, ← dotProduct_stdBasis_one w x, h] + classical rw [← dotProduct_single_one v x, ← dotProduct_single_one w x, h] theorem dotProduct_eq_iff {v w : n → R} : (∀ u, dotProduct v u = dotProduct w u) ↔ v = w := ⟨fun h => dotProduct_eq v w h, fun h _ => h ▸ rfl⟩ diff --git a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean index 53360f1ca3856..44fdd5263d583 100644 --- a/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/Matrix/SesquilinearForm.lean @@ -67,9 +67,8 @@ def Matrix.toLinearMap₂'Aux (f : Matrix n m N₂) : (n → R₁) →ₛₗ[σ variable [DecidableEq n] [DecidableEq m] -theorem Matrix.toLinearMap₂'Aux_stdBasis (f : Matrix n m N₂) (i : n) (j : m) : - f.toLinearMap₂'Aux σ₁ σ₂ (LinearMap.stdBasis R₁ (fun _ => R₁) i 1) - (LinearMap.stdBasis R₂ (fun _ => R₂) j 1) = f i j := by +theorem Matrix.toLinearMap₂'Aux_single (f : Matrix n m N₂) (i : n) (j : m) : + f.toLinearMap₂'Aux σ₁ σ₂ (Pi.single i 1) (Pi.single j 1) = f i j := by rw [Matrix.toLinearMap₂'Aux, mk₂'ₛₗ_apply] have : (∑ i', ∑ j', (if i = i' then (1 : S₁) else (0 : S₁)) • (if j = j' then (1 : S₂) else (0 : S₂)) • f i' j') = @@ -112,18 +111,17 @@ variable [DecidableEq n] [DecidableEq m] theorem LinearMap.toLinearMap₂'Aux_toMatrix₂Aux (f : (n → R₁) →ₛₗ[σ₁] (m → R₂) →ₛₗ[σ₂] N₂) : Matrix.toLinearMap₂'Aux σ₁ σ₂ - (LinearMap.toMatrix₂Aux R (fun i => stdBasis R₁ (fun _ => R₁) i 1) - (fun j => stdBasis R₂ (fun _ => R₂) j 1) f) = + (LinearMap.toMatrix₂Aux R (fun i => Pi.single i 1) (fun j => Pi.single j 1) f) = f := by refine ext_basis (Pi.basisFun R₁ n) (Pi.basisFun R₂ m) fun i j => ?_ - simp_rw [Pi.basisFun_apply, Matrix.toLinearMap₂'Aux_stdBasis, LinearMap.toMatrix₂Aux_apply] + simp_rw [Pi.basisFun_apply, Matrix.toLinearMap₂'Aux_single, LinearMap.toMatrix₂Aux_apply] theorem Matrix.toMatrix₂Aux_toLinearMap₂'Aux (f : Matrix n m N₂) : - LinearMap.toMatrix₂Aux R (fun i => LinearMap.stdBasis R₁ (fun _ => R₁) i 1) - (fun j => LinearMap.stdBasis R₂ (fun _ => R₂) j 1) (f.toLinearMap₂'Aux σ₁ σ₂) = + LinearMap.toMatrix₂Aux R (fun i => Pi.single i 1) + (fun j => Pi.single j 1) (f.toLinearMap₂'Aux σ₁ σ₂) = f := by ext i j - simp_rw [LinearMap.toMatrix₂Aux_apply, Matrix.toLinearMap₂'Aux_stdBasis] + simp_rw [LinearMap.toMatrix₂Aux_apply, Matrix.toLinearMap₂'Aux_single] end CommSemiring @@ -147,9 +145,7 @@ variable (R) /-- The linear equivalence between sesquilinear maps and `n × m` matrices -/ def LinearMap.toMatrixₛₗ₂' : ((n → R₁) →ₛₗ[σ₁] (m → R₂) →ₛₗ[σ₂] N₂) ≃ₗ[R] Matrix n m N₂ := - { LinearMap.toMatrix₂Aux R (fun i => stdBasis R₁ (fun _ => R₁) i 1) fun j => - stdBasis R₂ (fun _ => R₂) j - 1 with + { LinearMap.toMatrix₂Aux R (fun i => Pi.single i 1) (fun j => Pi.single j 1) with toFun := LinearMap.toMatrix₂Aux R _ _ invFun := Matrix.toLinearMap₂'Aux σ₁ σ₂ left_inv := LinearMap.toLinearMap₂'Aux_toMatrix₂Aux R @@ -197,16 +193,29 @@ theorem Matrix.toLinearMap₂'_apply' {T : Type*} [CommSemiring T] (M : Matrix n rw [smul_eq_mul, smul_eq_mul, mul_comm (w _), ← mul_assoc] @[simp] +theorem Matrix.toLinearMapₛₗ₂'_single (M : Matrix n m N₂) (i : n) (j : m) : + Matrix.toLinearMapₛₗ₂' R σ₁ σ₂ M (Pi.single i 1) (Pi.single j 1) = M i j := + Matrix.toLinearMap₂'Aux_single σ₁ σ₂ M i j + +set_option linter.deprecated false in +@[simp, deprecated Matrix.toLinearMapₛₗ₂'_single (since := "2024-08-09")] theorem Matrix.toLinearMapₛₗ₂'_stdBasis (M : Matrix n m N₂) (i : n) (j : m) : Matrix.toLinearMapₛₗ₂' R σ₁ σ₂ M (LinearMap.stdBasis R₁ (fun _ => R₁) i 1) (LinearMap.stdBasis R₂ (fun _ => R₂) j 1) = M i j := - Matrix.toLinearMap₂'Aux_stdBasis σ₁ σ₂ M i j + Matrix.toLinearMapₛₗ₂'_single .. @[simp] +theorem Matrix.toLinearMap₂'_single (M : Matrix n m N₂) (i : n) (j : m) : + Matrix.toLinearMap₂' R M (Pi.single i 1) (Pi.single j 1) = M i j := + Matrix.toLinearMap₂'Aux_single _ _ M i j + +set_option linter.deprecated false in +@[simp, deprecated Matrix.toLinearMap₂'_single (since := "2024-08-09")] theorem Matrix.toLinearMap₂'_stdBasis (M : Matrix n m N₂) (i : n) (j : m) : Matrix.toLinearMap₂' R M (LinearMap.stdBasis R (fun _ => R) i 1) (LinearMap.stdBasis R (fun _ => R) j 1) = M i j := - Matrix.toLinearMap₂'Aux_stdBasis _ _ M i j + show Matrix.toLinearMap₂' R M (Pi.single i 1) (Pi.single j 1) = M i j + from Matrix.toLinearMap₂'Aux_single _ _ M i j @[simp] theorem LinearMap.toMatrixₛₗ₂'_symm : @@ -240,14 +249,12 @@ theorem LinearMap.toMatrix'_toLinearMap₂' (M : Matrix n m N₂) : @[simp] theorem LinearMap.toMatrixₛₗ₂'_apply (B : (n → R₁) →ₛₗ[σ₁] (m → R₂) →ₛₗ[σ₂] N₂) (i : n) (j : m) : - LinearMap.toMatrixₛₗ₂' R B i j = - B (stdBasis R₁ (fun _ => R₁) i 1) (stdBasis R₂ (fun _ => R₂) j 1) := + LinearMap.toMatrixₛₗ₂' R B i j = B (Pi.single i 1) (Pi.single j 1) := rfl @[simp] theorem LinearMap.toMatrix₂'_apply (B : (n → S₁) →ₗ[S₁] (m → S₂) →ₗ[S₂] N₂) (i : n) (j : m) : - LinearMap.toMatrix₂' R B i j = - B (stdBasis S₁ (fun _ => S₁) i 1) (stdBasis S₂ (fun _ => S₂) j 1) := + LinearMap.toMatrix₂' R B i j = B (Pi.single i 1) (Pi.single j 1) := rfl end ToMatrix' @@ -351,8 +358,8 @@ noncomputable def Matrix.toLinearMap₂ : Matrix n m N₂ ≃ₗ[R] M₁ →ₗ[ theorem LinearMap.toMatrix₂_apply (B : M₁ →ₗ[R] M₂ →ₗ[R] N₂) (i : n) (j : m) : LinearMap.toMatrix₂ b₁ b₂ B i j = B (b₁ i) (b₂ j) := by simp only [toMatrix₂, LinearEquiv.trans_apply, toMatrix₂'_apply, LinearEquiv.arrowCongr_apply, - Basis.equivFun_symm_apply, stdBasis_apply', ite_smul, one_smul, zero_smul, sum_ite_eq, mem_univ, - ↓reduceIte, LinearEquiv.refl_apply] + Basis.equivFun_symm_apply, Pi.single_apply, ite_smul, one_smul, zero_smul, sum_ite_eq', + mem_univ, ↓reduceIte, LinearEquiv.refl_apply] @[simp] theorem Matrix.toLinearMap₂_apply (M : Matrix n m N₂) (x : M₁) (y : M₂) : diff --git a/Mathlib/LinearAlgebra/Matrix/ToLin.lean b/Mathlib/LinearAlgebra/Matrix/ToLin.lean index d0cc4f7aa9d4b..f7a4aba7c8c89 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLin.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLin.lean @@ -84,24 +84,19 @@ theorem Matrix.coe_vecMulLinear [Fintype m] (M : Matrix m n R) : variable [Fintype m] -@[simp] +set_option linter.deprecated false in +@[simp, deprecated Matrix.single_one_vecMul (since := "2024-08-09")] theorem Matrix.vecMul_stdBasis [DecidableEq m] (M : Matrix m n R) (i j) : - (LinearMap.stdBasis R (fun _ ↦ R) i 1 ᵥ* M) j = M i j := by - have : (∑ i', (if i = i' then 1 else 0) * M i' j) = M i j := by - simp_rw [boole_mul, Finset.sum_ite_eq, Finset.mem_univ, if_true] - simp only [vecMul, dotProduct] - convert this - split_ifs with h <;> simp only [stdBasis_apply] - · rw [h, Function.update_same] - · rw [Function.update_noteq (Ne.symm h), Pi.zero_apply] + (LinearMap.stdBasis R (fun _ ↦ R) i 1 ᵥ* M) j = M i j := + congr_fun (Matrix.single_one_vecMul ..) j theorem range_vecMulLinear (M : Matrix m n R) : LinearMap.range M.vecMulLinear = span R (range M) := by letI := Classical.decEq m - simp_rw [range_eq_map, ← iSup_range_stdBasis, Submodule.map_iSup, range_eq_map, ← + simp_rw [range_eq_map, ← iSup_range_single, Submodule.map_iSup, range_eq_map, ← Ideal.span_singleton_one, Ideal.span, Submodule.map_span, image_image, image_singleton, Matrix.vecMulLinear_apply, iSup_span, range_eq_iUnion, iUnion_singleton_eq_range, - LinearMap.stdBasis, coe_single] + LinearMap.single, LinearMap.coe_mk, AddHom.coe_mk] unfold vecMul simp_rw [single_dotProduct, one_mul] @@ -125,15 +120,15 @@ variable [DecidableEq m] by having matrices act by right multiplication. -/ def LinearMap.toMatrixRight' : ((m → R) →ₗ[R] n → R) ≃ₗ[Rᵐᵒᵖ] Matrix m n R where - toFun f i j := f (stdBasis R (fun _ ↦ R) i 1) j + toFun f i j := f (single R (fun _ ↦ R) i 1) j invFun := Matrix.vecMulLinear right_inv M := by ext i j - simp only [Matrix.vecMul_stdBasis, Matrix.vecMulLinear_apply] + simp left_inv f := by apply (Pi.basisFun R m).ext intro j; ext i - simp only [Pi.basisFun_apply, Matrix.vecMul_stdBasis, Matrix.vecMulLinear_apply] + simp map_add' f g := by ext i j simp only [Pi.add_apply, LinearMap.add_apply, Matrix.add_apply] @@ -167,7 +162,7 @@ theorem Matrix.toLinearMapRight'_mul_apply [Fintype l] [DecidableEq l] (M : Matr theorem Matrix.toLinearMapRight'_one : Matrix.toLinearMapRight' (1 : Matrix m m R) = LinearMap.id := by ext - simp [LinearMap.one_apply, stdBasis_apply] + simp [LinearMap.one_apply] /-- If `M` and `M'` are each other's inverse matrices, they provide an equivalence between `n → A` and `m → A` corresponding to `M.vecMul` and `M'.vecMul`. -/ @@ -258,14 +253,17 @@ theorem Matrix.ker_mulVecLin_eq_bot_iff {M : Matrix m n R} : (LinearMap.ker M.mulVecLin) = ⊥ ↔ ∀ v, M *ᵥ v = 0 → v = 0 := by simp only [Submodule.eq_bot_iff, LinearMap.mem_ker, Matrix.mulVecLin_apply] +set_option linter.deprecated false in +@[deprecated Matrix.mulVec_single_one (since := "2024-08-09")] theorem Matrix.mulVec_stdBasis [DecidableEq n] (M : Matrix m n R) (i j) : (M *ᵥ LinearMap.stdBasis R (fun _ ↦ R) j 1) i = M i j := - (congr_fun (Matrix.mulVec_single _ _ (1 : R)) i).trans <| mul_one _ + congr_fun (Matrix.mulVec_single_one ..) i -@[simp] +set_option linter.deprecated false in +@[simp, deprecated Matrix.mulVec_single_one (since := "2024-08-09")] theorem Matrix.mulVec_stdBasis_apply [DecidableEq n] (M : Matrix m n R) (j) : M *ᵥ LinearMap.stdBasis R (fun _ ↦ R) j 1 = Mᵀ j := - funext fun i ↦ Matrix.mulVec_stdBasis M i j + Matrix.mulVec_single_one .. theorem Matrix.range_mulVecLin (M : Matrix m n R) : LinearMap.range M.mulVecLin = span R (range Mᵀ) := by @@ -285,15 +283,16 @@ variable {k l m n : Type*} [DecidableEq n] [Fintype n] /-- Linear maps `(n → R) →ₗ[R] (m → R)` are linearly equivalent to `Matrix m n R`. -/ def LinearMap.toMatrix' : ((n → R) →ₗ[R] m → R) ≃ₗ[R] Matrix m n R where - toFun f := of fun i j ↦ f (stdBasis R (fun _ ↦ R) j 1) i + toFun f := of fun i j ↦ f (Pi.single j 1) i invFun := Matrix.mulVecLin right_inv M := by ext i j - simp only [Matrix.mulVec_stdBasis, Matrix.mulVecLin_apply, of_apply] + simp only [Matrix.mulVec_single_one, Matrix.mulVecLin_apply, of_apply, transpose_apply] left_inv f := by apply (Pi.basisFun R n).ext intro j; ext i - simp only [Pi.basisFun_apply, Matrix.mulVec_stdBasis, Matrix.mulVecLin_apply, of_apply] + simp only [Pi.basisFun_apply, Matrix.mulVec_single_one, + Matrix.mulVecLin_apply, of_apply, transpose_apply] map_add' f g := by ext i j simp only [Pi.add_apply, LinearMap.add_apply, of_apply, Matrix.add_apply] @@ -337,8 +336,8 @@ theorem LinearMap.toMatrix'_apply (f : (n → R) →ₗ[R] m → R) (i j) : congr ext j' split_ifs with h - · rw [h, stdBasis_same] - apply stdBasis_ne _ _ _ _ h + · rw [h, Pi.single_eq_same] + apply Pi.single_eq_of_ne h @[simp] theorem Matrix.toLin'_apply (M : Matrix m n R) (v : n → R) : Matrix.toLin' M v = M *ᵥ v := diff --git a/Mathlib/LinearAlgebra/Pi.lean b/Mathlib/LinearAlgebra/Pi.lean index 59907ee20f3e1..a6c4da9560237 100644 --- a/Mathlib/LinearAlgebra/Pi.lean +++ b/Mathlib/LinearAlgebra/Pi.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Group.Fin.Tuple import Mathlib.Algebra.BigOperators.Pi import Mathlib.Algebra.Module.Prod import Mathlib.Algebra.Module.Submodule.Ker +import Mathlib.Algebra.Module.Submodule.Range import Mathlib.Algebra.Module.Equiv.Basic import Mathlib.GroupTheory.GroupAction.BigOperators import Mathlib.Logic.Equiv.Fin @@ -113,17 +114,85 @@ theorem apply_single [AddCommMonoid M] [Module R M] [DecidableEq ι] (f : (i : (i j : ι) (x : φ i) : f j (Pi.single i x j) = (Pi.single i (f i x) : ι → M) j := Pi.apply_single (fun i => f i) (fun i => (f i).map_zero) _ _ _ +variable (R φ) + /-- The `LinearMap` version of `AddMonoidHom.single` and `Pi.single`. -/ def single [DecidableEq ι] (i : ι) : φ i →ₗ[R] (i : ι) → φ i := { AddMonoidHom.single φ i with toFun := Pi.single i map_smul' := Pi.single_smul i } +lemma single_apply [DecidableEq ι] {i : ι} (v : φ i) : + single R φ i v = Pi.single i v := + rfl + @[simp] -theorem coe_single [DecidableEq ι] (i : ι) : ⇑(single i : φ i →ₗ[R] (i : ι) → φ i) = Pi.single i := +theorem coe_single [DecidableEq ι] (i : ι) : + ⇑(single R φ i : φ i →ₗ[R] (i : ι) → φ i) = Pi.single i := rfl -variable (R φ) +variable [DecidableEq ι] + +theorem proj_comp_single_same (i : ι) : (proj i).comp (single R φ i) = id := + LinearMap.ext <| Pi.single_eq_same i + +theorem proj_comp_single_ne (i j : ι) (h : i ≠ j) : (proj i).comp (single R φ j) = 0 := + LinearMap.ext <| Pi.single_eq_of_ne h + +theorem iSup_range_single_le_iInf_ker_proj (I J : Set ι) (h : Disjoint I J) : + ⨆ i ∈ I, range (single R φ i) ≤ ⨅ i ∈ J, ker (proj i : (∀ i, φ i) →ₗ[R] φ i) := by + refine iSup_le fun i => iSup_le fun hi => range_le_iff_comap.2 ?_ + simp only [← ker_comp, eq_top_iff, SetLike.le_def, mem_ker, comap_iInf, mem_iInf] + rintro b - j hj + rw [proj_comp_single_ne R φ j i, zero_apply] + rintro rfl + exact h.le_bot ⟨hi, hj⟩ + +theorem iInf_ker_proj_le_iSup_range_single {I : Finset ι} {J : Set ι} (hu : Set.univ ⊆ ↑I ∪ J) : + ⨅ i ∈ J, ker (proj i : (∀ i, φ i) →ₗ[R] φ i) ≤ ⨆ i ∈ I, range (single R φ i) := + SetLike.le_def.2 + (by + intro b hb + simp only [mem_iInf, mem_ker, proj_apply] at hb + rw [← + show (∑ i ∈ I, Pi.single i (b i)) = b by + ext i + rw [Finset.sum_apply, ← Pi.single_eq_same i (b i)] + refine Finset.sum_eq_single i (fun j _ ne => Pi.single_eq_of_ne ne.symm _) ?_ + intro hiI + rw [Pi.single_eq_same] + exact hb _ ((hu trivial).resolve_left hiI)] + exact sum_mem_biSup fun i _ => mem_range_self (single R φ i) (b i)) + +theorem iSup_range_single_eq_iInf_ker_proj {I J : Set ι} (hd : Disjoint I J) + (hu : Set.univ ⊆ I ∪ J) (hI : Set.Finite I) : + ⨆ i ∈ I, range (single R φ i) = ⨅ i ∈ J, ker (proj i : (∀ i, φ i) →ₗ[R] φ i) := by + refine le_antisymm (iSup_range_single_le_iInf_ker_proj _ _ _ _ hd) ?_ + have : Set.univ ⊆ ↑hI.toFinset ∪ J := by rwa [hI.coe_toFinset] + refine le_trans (iInf_ker_proj_le_iSup_range_single R φ this) (iSup_mono fun i => ?_) + rw [Set.Finite.mem_toFinset] + +theorem iSup_range_single [Finite ι] : ⨆ i, range (single R φ i) = ⊤ := by + cases nonempty_fintype ι + convert top_unique (iInf_emptyset.ge.trans <| iInf_ker_proj_le_iSup_range_single R φ _) + · rename_i i + exact ((@iSup_pos _ _ _ fun _ => range <| single R φ i) <| Finset.mem_univ i).symm + · rw [Finset.coe_univ, Set.union_empty] + +theorem disjoint_single_single (I J : Set ι) (h : Disjoint I J) : + Disjoint (⨆ i ∈ I, range (single R φ i)) (⨆ i ∈ J, range (single R φ i)) := by + refine + Disjoint.mono (iSup_range_single_le_iInf_ker_proj _ _ _ _ <| disjoint_compl_right) + (iSup_range_single_le_iInf_ker_proj _ _ _ _ <| disjoint_compl_right) ?_ + simp only [disjoint_iff_inf_le, SetLike.le_def, mem_iInf, mem_inf, mem_ker, mem_bot, proj_apply, + funext_iff] + rintro b ⟨hI, hJ⟩ i + classical + by_cases hiI : i ∈ I + · by_cases hiJ : i ∈ J + · exact (h.le_bot ⟨hiI, hiJ⟩).elim + · exact hJ i hiJ + · exact hI i hiI /-- The linear equivalence between linear functions on a finite product of modules and families of functions on these modules. See note [bundled maps over different rings]. -/ @@ -131,7 +200,7 @@ families of functions on these modules. See note [bundled maps over different ri def lsum (S) [AddCommMonoid M] [Module R M] [Fintype ι] [DecidableEq ι] [Semiring S] [Module S M] [SMulCommClass R S M] : ((i : ι) → φ i →ₗ[R] M) ≃ₗ[S] ((i : ι) → φ i) →ₗ[R] M where toFun f := ∑ i : ι, (f i).comp (proj i) - invFun f i := f.comp (single i) + invFun f i := f.comp (single R φ i) map_add' f g := by simp only [Pi.add_apply, add_comp, Finset.sum_add_distrib] map_smul' c f := by simp only [Pi.smul_apply, smul_comp, Finset.smul_sum, RingHom.id_apply] left_inv f := by @@ -150,7 +219,7 @@ theorem lsum_apply (S) [AddCommMonoid M] [Module R M] [Fintype ι] [DecidableEq @[simp high] theorem lsum_single {ι R : Type*} [Fintype ι] [DecidableEq ι] [CommRing R] {M : ι → Type*} [(i : ι) → AddCommGroup (M i)] [(i : ι) → Module R (M i)] : - LinearMap.lsum R M R LinearMap.single = LinearMap.id := + LinearMap.lsum R M R (LinearMap.single R M) = LinearMap.id := LinearMap.ext fun x => by simp [Finset.univ_sum_single] variable {R φ} @@ -168,7 +237,7 @@ theorem pi_ext_iff : f = g ↔ ∀ i x, f (Pi.single i x) = g (Pi.single i x) := /-- This is used as the ext lemma instead of `LinearMap.pi_ext` for reasons explained in note [partially-applied ext lemmas]. -/ @[ext] -theorem pi_ext' (h : ∀ i, f.comp (single i) = g.comp (single i)) : f = g := by +theorem pi_ext' (h : ∀ i, f.comp (single R φ i) = g.comp (single R φ i)) : f = g := by refine pi_ext fun i x => ?_ convert LinearMap.congr_fun (h i) x @@ -224,6 +293,20 @@ theorem update_apply (f : (i : ι) → M₂ →ₗ[R] φ i) (c : M₂) (i j : ι · rw [h, update_same, update_same] · rw [update_noteq h, update_noteq h] +variable (R φ) + +theorem single_eq_pi_diag (i : ι) : single R φ i = pi (diag i) := by + ext x j + -- Porting note: made types explicit + convert (update_apply (R := R) (φ := φ) (ι := ι) 0 x i j _).symm + rfl + +theorem ker_single (i : ι) : ker (single R φ i) = ⊥ := + ker_eq_bot_of_injective <| Pi.single_injective _ _ + +theorem proj_comp_single (i j : ι) : (proj i).comp (single R φ j) = diag j i := by + rw [single_eq_pi_diag, proj_pi] + end /-- A linear map `f` applied to `x : ι → R` can be computed using the image under `f` of elements @@ -283,7 +366,7 @@ theorem iInf_comap_proj : simp theorem iSup_map_single [DecidableEq ι] [Finite ι] : - ⨆ i, map (LinearMap.single i : φ i →ₗ[R] (i : ι) → φ i) (p i) = pi Set.univ p := by + ⨆ i, map (LinearMap.single R φ i : φ i →ₗ[R] (i : ι) → φ i) (p i) = pi Set.univ p := by cases nonempty_fintype ι refine (iSup_le fun i => ?_).antisymm ?_ · rintro _ ⟨x, hx : x ∈ p i, rfl⟩ j - @@ -293,7 +376,7 @@ theorem iSup_map_single [DecidableEq ι] [Finite ι] : exact sum_mem_iSup fun i => mem_map_of_mem (hx i trivial) theorem le_comap_single_pi [DecidableEq ι] (p : (i : ι) → Submodule R (φ i)) {i} : - p i ≤ Submodule.comap (LinearMap.single i : φ i →ₗ[R] _) (Submodule.pi Set.univ p) := by + p i ≤ Submodule.comap (LinearMap.single R φ i : φ i →ₗ[R] _) (Submodule.pi Set.univ p) := by intro x hx rw [Submodule.mem_comap, Submodule.mem_pi] rintro j - diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean b/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean index 88d50e41cd269..980063ffab082 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean @@ -41,11 +41,11 @@ noncomputable def isometryEquivSumSquares (w' : ι → ℂ) : v j • w j ^ (-(1 / 2 : ℂ)) := by classical rw [Finset.sum_apply, sum_eq_single j, Pi.basisFun_apply, IsUnit.unit_spec, - LinearMap.stdBasis_apply, Pi.smul_apply, Pi.smul_apply, Function.update_same, smul_eq_mul, + Pi.smul_apply, Pi.smul_apply, Pi.single_eq_same, smul_eq_mul, smul_eq_mul, smul_eq_mul, mul_one] · intro i _ hij - rw [Pi.basisFun_apply, LinearMap.stdBasis_apply, Pi.smul_apply, Pi.smul_apply, - Function.update_noteq hij.symm, Pi.zero_apply, smul_eq_mul, smul_eq_mul, + rw [Pi.basisFun_apply, Pi.smul_apply, Pi.smul_apply, + Pi.single_eq_of_ne hij.symm, smul_eq_mul, smul_eq_mul, mul_zero, mul_zero] intro hj'; exact False.elim (hj' hj) simp_rw [Basis.unitsSMul_apply] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean index 5ec5b7bcee160..891097a9ae0b6 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Prod.lean @@ -268,7 +268,7 @@ def IsometryEquiv.pi [Fintype ι] @[simps!] def Isometry.single [Fintype ι] [DecidableEq ι] (Q : ∀ i, QuadraticMap R (Mᵢ i) P) (i : ι) : Q i →qᵢ pi Q where - toLinearMap := LinearMap.single i + toLinearMap := LinearMap.single _ _ i map_app' := pi_apply_single _ _ /-- `LinearMap.proj` as an isometry, when all but one quadratic form is zero. -/ diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Real.lean b/Mathlib/LinearAlgebra/QuadraticForm/Real.lean index c5985ce4da09f..7ec3a6f4947d8 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Real.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Real.lean @@ -46,7 +46,7 @@ noncomputable def isometryEquivSignWeightedSumSquares (w : ι → ℝ) : ext1 v classical suffices ∑ i, (w i / |(u i : ℝ)|) * v i ^ 2 = ∑ i, w i * (v i ^ 2 * |(u i : ℝ)|⁻¹) by - simpa [basisRepr_apply, Basis.unitsSMul_apply, ← _root_.sq, mul_pow, ← hwu] + simpa [basisRepr_apply, Basis.unitsSMul_apply, ← _root_.sq, mul_pow, ← hwu, Pi.single_apply] exact sum_congr rfl fun j _ ↦ by ring /-- **Sylvester's law of inertia**: A nondegenerate real quadratic form is equivalent to a weighted diff --git a/Mathlib/LinearAlgebra/QuotientPi.lean b/Mathlib/LinearAlgebra/QuotientPi.lean index b3f8f3b83c3c1..e4771d9be3c91 100644 --- a/Mathlib/LinearAlgebra/QuotientPi.lean +++ b/Mathlib/LinearAlgebra/QuotientPi.lean @@ -90,7 +90,7 @@ variable [Fintype ι] [DecidableEq ι] @[simp] def invFun : (∀ i, Ms i ⧸ p i) → (∀ i, Ms i) ⧸ pi Set.univ p := - piQuotientLift p (pi Set.univ p) single fun _ => le_comap_single_pi p + piQuotientLift p (pi Set.univ p) _ fun _ => le_comap_single_pi p theorem left_inv : Function.LeftInverse (invFun p) (toFun p) := fun x => Quotient.inductionOn' x fun x' => by diff --git a/Mathlib/LinearAlgebra/StdBasis.lean b/Mathlib/LinearAlgebra/StdBasis.lean index 97b36e6f8dbd0..d058dc9b63955 100644 --- a/Mathlib/LinearAlgebra/StdBasis.lean +++ b/Mathlib/LinearAlgebra/StdBasis.lean @@ -13,21 +13,19 @@ import Mathlib.LinearAlgebra.Basis.Defs This file defines the standard basis `Pi.basis (s : ∀ j, Basis (ι j) R (M j))`, which is the `Σ j, ι j`-indexed basis of `Π j, M j`. The basis vectors are given by -`Pi.basis s ⟨j, i⟩ j' = LinearMap.stdBasis R M j' (s j) i = if j = j' then s i else 0`. +`Pi.basis s ⟨j, i⟩ j' = Pi.single j' (s j) i = if j = j' then s i else 0`. The standard basis on `R^η`, i.e. `η → R` is called `Pi.basisFun`. -To give a concrete example, `LinearMap.stdBasis R (fun (i : Fin 3) ↦ R) i 1` +To give a concrete example, `Pi.single (i : Fin 3) (1 : R)` gives the `i`th unit basis vector in `R³`, and `Pi.basisFun R (Fin 3)` proves this is a basis over `Fin 3 → R`. ## Main definitions - - `LinearMap.stdBasis R M`: if `x` is a basis vector of `M i`, then - `LinearMap.stdBasis R M i x` is the `i`th standard basis vector of `Π i, M i`. - `Pi.basis s`: given a basis `s i` for each `M i`, the standard basis on `Π i, M i` - `Pi.basisFun R η`: the standard basis on `R^η`, i.e. `η → R`, given by - `Pi.basisFun R η i j = if i = j then 1 else 0`. + `Pi.basisFun R η i j = Pi.single i 1 j = if i = j then 1 else 0`. - `Matrix.stdBasis R n m`: the standard basis on `Matrix n m R`, given by `Matrix.stdBasis R n m (i, j) i' j' = if (i, j) = (i', j') then 1 else 0`. @@ -42,100 +40,92 @@ variable (R : Type*) {ι : Type*} [Semiring R] (φ : ι → Type*) [∀ i, AddCo [∀ i, Module R (φ i)] [DecidableEq ι] /-- The standard basis of the product of `φ`. -/ +@[deprecated LinearMap.single (since := "2024-08-09")] def stdBasis : ∀ i : ι, φ i →ₗ[R] ∀ i, φ i := - single + single R φ +set_option linter.deprecated false in +@[deprecated Pi.single (since := "2024-08-09")] theorem stdBasis_apply (i : ι) (b : φ i) : stdBasis R φ i b = update (0 : (a : ι) → φ a) i b := rfl -@[simp] +set_option linter.deprecated false in +@[simp, deprecated Pi.single_apply (since := "2024-08-09")] theorem stdBasis_apply' (i i' : ι) : (stdBasis R (fun _x : ι => R) i) 1 i' = ite (i = i') 1 0 := by - rw [LinearMap.stdBasis_apply, Function.update_apply, Pi.zero_apply] - congr 1; rw [eq_iff_iff, eq_comm] + simp_rw [stdBasis, single_apply, Pi.single_apply, eq_comm] +set_option linter.deprecated false in +@[deprecated LinearMap.coe_single (since := "2024-08-09")] theorem coe_stdBasis (i : ι) : ⇑(stdBasis R φ i) = Pi.single i := rfl -@[simp] +set_option linter.deprecated false in +@[simp, deprecated Pi.single_eq_same (since := "2024-08-09")] theorem stdBasis_same (i : ι) (b : φ i) : stdBasis R φ i b i = b := - Pi.single_eq_same i b + Pi.single_eq_same .. +set_option linter.deprecated false in +@[deprecated Pi.single_eq_of_ne (since := "2024-08-09")] theorem stdBasis_ne (i j : ι) (h : j ≠ i) (b : φ i) : stdBasis R φ i b j = 0 := Pi.single_eq_of_ne h b -theorem stdBasis_eq_pi_diag (i : ι) : stdBasis R φ i = pi (diag i) := by - ext x j - -- Porting note: made types explicit - convert (update_apply (R := R) (φ := φ) (ι := ι) 0 x i j _).symm - rfl +set_option linter.deprecated false in +@[deprecated single_eq_pi_diag (since := "2024-08-09")] +theorem stdBasis_eq_pi_diag (i : ι) : stdBasis R φ i = pi (diag i) := + single_eq_pi_diag .. +set_option linter.deprecated false in +@[deprecated ker_single (since := "2024-08-09")] theorem ker_stdBasis (i : ι) : ker (stdBasis R φ i) = ⊥ := - ker_eq_bot_of_injective <| Pi.single_injective _ _ + ker_single .. -theorem proj_comp_stdBasis (i j : ι) : (proj i).comp (stdBasis R φ j) = diag j i := by - rw [stdBasis_eq_pi_diag, proj_pi] +set_option linter.deprecated false in +@[deprecated proj_comp_single (since := "2024-08-09")] +theorem proj_comp_stdBasis (i j : ι) : (proj i).comp (stdBasis R φ j) = diag j i := + proj_comp_single .. +set_option linter.deprecated false in +@[deprecated proj_comp_single_same (since := "2024-08-09")] theorem proj_stdBasis_same (i : ι) : (proj i).comp (stdBasis R φ i) = id := - LinearMap.ext <| stdBasis_same R φ i + proj_comp_single_same .. +set_option linter.deprecated false in +@[deprecated proj_comp_single_ne (since := "2024-08-09")] theorem proj_stdBasis_ne (i j : ι) (h : i ≠ j) : (proj i).comp (stdBasis R φ j) = 0 := - LinearMap.ext <| stdBasis_ne R φ _ _ h + proj_comp_single_ne R φ i j h +set_option linter.deprecated false in +@[deprecated iSup_range_single_le_iInf_ker_proj (since := "2024-08-09")] theorem iSup_range_stdBasis_le_iInf_ker_proj (I J : Set ι) (h : Disjoint I J) : - ⨆ i ∈ I, range (stdBasis R φ i) ≤ ⨅ i ∈ J, ker (proj i : (∀ i, φ i) →ₗ[R] φ i) := by - refine iSup_le fun i => iSup_le fun hi => range_le_iff_comap.2 ?_ - simp only [← ker_comp, eq_top_iff, SetLike.le_def, mem_ker, comap_iInf, mem_iInf] - rintro b - j hj - rw [proj_stdBasis_ne R φ j i, zero_apply] - rintro rfl - exact h.le_bot ⟨hi, hj⟩ + ⨆ i ∈ I, range (stdBasis R φ i) ≤ ⨅ i ∈ J, ker (proj i : (∀ i, φ i) →ₗ[R] φ i) := + iSup_range_single_le_iInf_ker_proj R φ I J h +set_option linter.deprecated false in +@[deprecated iInf_ker_proj_le_iSup_range_single (since := "2024-08-09")] theorem iInf_ker_proj_le_iSup_range_stdBasis {I : Finset ι} {J : Set ι} (hu : Set.univ ⊆ ↑I ∪ J) : ⨅ i ∈ J, ker (proj i : (∀ i, φ i) →ₗ[R] φ i) ≤ ⨆ i ∈ I, range (stdBasis R φ i) := - SetLike.le_def.2 - (by - intro b hb - simp only [mem_iInf, mem_ker, proj_apply] at hb - rw [← - show (∑ i ∈ I, stdBasis R φ i (b i)) = b by - ext i - rw [Finset.sum_apply, ← stdBasis_same R φ i (b i)] - refine Finset.sum_eq_single i (fun j _ ne => stdBasis_ne _ _ _ _ ne.symm _) ?_ - intro hiI - rw [stdBasis_same] - exact hb _ ((hu trivial).resolve_left hiI)] - exact sum_mem_biSup fun i _ => mem_range_self (stdBasis R φ i) (b i)) + iInf_ker_proj_le_iSup_range_single R φ hu +set_option linter.deprecated false in +@[deprecated iSup_range_single_eq_iInf_ker_proj (since := "2024-08-09")] theorem iSup_range_stdBasis_eq_iInf_ker_proj {I J : Set ι} (hd : Disjoint I J) (hu : Set.univ ⊆ I ∪ J) (hI : Set.Finite I) : - ⨆ i ∈ I, range (stdBasis R φ i) = ⨅ i ∈ J, ker (proj i : (∀ i, φ i) →ₗ[R] φ i) := by - refine le_antisymm (iSup_range_stdBasis_le_iInf_ker_proj _ _ _ _ hd) ?_ - have : Set.univ ⊆ ↑hI.toFinset ∪ J := by rwa [hI.coe_toFinset] - refine le_trans (iInf_ker_proj_le_iSup_range_stdBasis R φ this) (iSup_mono fun i => ?_) - rw [Set.Finite.mem_toFinset] - -theorem iSup_range_stdBasis [Finite ι] : ⨆ i, range (stdBasis R φ i) = ⊤ := by - cases nonempty_fintype ι - convert top_unique (iInf_emptyset.ge.trans <| iInf_ker_proj_le_iSup_range_stdBasis R φ _) - · rename_i i - exact ((@iSup_pos _ _ _ fun _ => range <| stdBasis R φ i) <| Finset.mem_univ i).symm - · rw [Finset.coe_univ, Set.union_empty] + ⨆ i ∈ I, range (stdBasis R φ i) = ⨅ i ∈ J, ker (proj i : (∀ i, φ i) →ₗ[R] φ i) := + iSup_range_single_eq_iInf_ker_proj _ _ hd hu hI + +set_option linter.deprecated false in +@[deprecated iSup_range_single (since := "2024-08-09")] +theorem iSup_range_stdBasis [Finite ι] : ⨆ i, range (stdBasis R φ i) = ⊤ := + iSup_range_single .. +set_option linter.deprecated false in +@[deprecated disjoint_single_single (since := "2024-08-09")] theorem disjoint_stdBasis_stdBasis (I J : Set ι) (h : Disjoint I J) : - Disjoint (⨆ i ∈ I, range (stdBasis R φ i)) (⨆ i ∈ J, range (stdBasis R φ i)) := by - refine - Disjoint.mono (iSup_range_stdBasis_le_iInf_ker_proj _ _ _ _ <| disjoint_compl_right) - (iSup_range_stdBasis_le_iInf_ker_proj _ _ _ _ <| disjoint_compl_right) ?_ - simp only [disjoint_iff_inf_le, SetLike.le_def, mem_iInf, mem_inf, mem_ker, mem_bot, proj_apply, - funext_iff] - rintro b ⟨hI, hJ⟩ i - classical - by_cases hiI : i ∈ I - · by_cases hiJ : i ∈ J - · exact (h.le_bot ⟨hiI, hiJ⟩).elim - · exact hJ i hiJ - · exact hI i hiI + Disjoint (⨆ i ∈ I, range (stdBasis R φ i)) (⨆ i ∈ J, range (stdBasis R φ i)) := + disjoint_single_single R φ I J h +set_option linter.deprecated false in +@[deprecated "You can probably use Finsupp.single_eq_pi_single here" (since := "2024-08-09")] theorem stdBasis_eq_single {a : R} : (fun i : ι => (stdBasis R (fun _ : ι => R) i) a) = fun i : ι => ↑(Finsupp.single i a) := funext fun i => (Finsupp.single_eq_pi_single i a).symm @@ -154,32 +144,39 @@ section Module variable {η : Type*} {ιs : η → Type*} {Ms : η → Type*} -theorem linearIndependent_stdBasis [Ring R] [∀ i, AddCommGroup (Ms i)] [∀ i, Module R (Ms i)] +theorem linearIndependent_single [Ring R] [∀ i, AddCommGroup (Ms i)] [∀ i, Module R (Ms i)] [DecidableEq η] (v : ∀ j, ιs j → Ms j) (hs : ∀ i, LinearIndependent R (v i)) : - LinearIndependent R fun ji : Σj, ιs j => stdBasis R Ms ji.1 (v ji.1 ji.2) := by - have hs' : ∀ j : η, LinearIndependent R fun i : ιs j => stdBasis R Ms j (v j i) := by + LinearIndependent R fun ji : Σj, ιs j ↦ Pi.single ji.1 (v ji.1 ji.2) := by + have hs' : ∀ j : η, LinearIndependent R fun i : ιs j => LinearMap.single R Ms j (v j i) := by intro j - exact (hs j).map' _ (ker_stdBasis _ _ _) + exact (hs j).map' _ (LinearMap.ker_single _ _ _) apply linearIndependent_iUnion_finite hs' intro j J _ hiJ have h₀ : - ∀ j, span R (range fun i : ιs j => stdBasis R Ms j (v j i)) ≤ - LinearMap.range (stdBasis R Ms j) := by + ∀ j, span R (range fun i : ιs j => LinearMap.single R Ms j (v j i)) ≤ + LinearMap.range (LinearMap.single R Ms j) := by intro j rw [span_le, LinearMap.range_coe] apply range_comp_subset_range have h₁ : - span R (range fun i : ιs j => stdBasis R Ms j (v j i)) ≤ - ⨆ i ∈ ({j} : Set _), LinearMap.range (stdBasis R Ms i) := by - rw [@iSup_singleton _ _ _ fun i => LinearMap.range (stdBasis R (Ms) i)] + span R (range fun i : ιs j => LinearMap.single R Ms j (v j i)) ≤ + ⨆ i ∈ ({j} : Set _), LinearMap.range (LinearMap.single R Ms i) := by + rw [@iSup_singleton _ _ _ fun i => LinearMap.range (LinearMap.single R (Ms) i)] apply h₀ have h₂ : - ⨆ j ∈ J, span R (range fun i : ιs j => stdBasis R Ms j (v j i)) ≤ - ⨆ j ∈ J, LinearMap.range (stdBasis R (fun j : η => Ms j) j) := + ⨆ j ∈ J, span R (range fun i : ιs j => LinearMap.single R Ms j (v j i)) ≤ + ⨆ j ∈ J, LinearMap.range (LinearMap.single R (fun j : η => Ms j) j) := iSup₂_mono fun i _ => h₀ i have h₃ : Disjoint (fun i : η => i ∈ ({j} : Set _)) J := by convert Set.disjoint_singleton_left.2 hiJ using 0 - exact (disjoint_stdBasis_stdBasis _ _ _ _ h₃).mono h₁ h₂ + exact (disjoint_single_single _ _ _ _ h₃).mono h₁ h₂ + +set_option linter.deprecated false in +@[deprecated linearIndependent_single (since := "2024-08-09")] +theorem linearIndependent_stdBasis [Ring R] [∀ i, AddCommGroup (Ms i)] [∀ i, Module R (Ms i)] + [DecidableEq η] (v : ∀ j, ιs j → Ms j) (hs : ∀ i, LinearIndependent R (v i)) : + LinearIndependent R fun ji : Σj, ιs j => stdBasis R Ms ji.1 (v ji.1 ji.2) := + linearIndependent_single _ hs variable [Semiring R] [∀ i, AddCommMonoid (Ms i)] [∀ i, Module R (Ms i)] @@ -199,39 +196,37 @@ protected noncomputable def basis (s : ∀ j, Basis (ιs j) R (Ms j)) : Basis.ofRepr ((LinearEquiv.piCongrRight fun j => (s j).repr) ≪≫ₗ (Finsupp.sigmaFinsuppLEquivPiFinsupp R).symm) - -- Porting note: was - -- -- The `AddCommMonoid (Π j, Ms j)` instance was hard to find. - -- -- Defining this in tactic mode seems to shake up instance search enough - -- -- that it works by itself. - -- refine Basis.ofRepr (?_ ≪≫ₗ (Finsupp.sigmaFinsuppLEquivPiFinsupp R).symm) - -- exact LinearEquiv.piCongrRight fun j => (s j).repr @[simp] -theorem basis_repr_stdBasis [DecidableEq η] (s : ∀ j, Basis (ιs j) R (Ms j)) (j i) : - (Pi.basis s).repr (stdBasis R _ j (s j i)) = Finsupp.single ⟨j, i⟩ 1 := by +theorem basis_repr_single [DecidableEq η] (s : ∀ j, Basis (ιs j) R (Ms j)) (j i) : + (Pi.basis s).repr (Pi.single j (s j i)) = Finsupp.single ⟨j, i⟩ 1 := by + classical ext ⟨j', i'⟩ by_cases hj : j = j' · subst hj -- Porting note: needed to add more lemmas - simp only [Pi.basis, LinearEquiv.trans_apply, Basis.repr_self, stdBasis_same, + simp only [Pi.basis, LinearEquiv.trans_apply, Basis.repr_self, Pi.single_eq_same, LinearEquiv.piCongrRight, Finsupp.sigmaFinsuppLEquivPiFinsupp_symm_apply, Basis.repr_symm_apply, LinearEquiv.coe_mk, ne_eq, Sigma.mk.inj_iff, heq_eq_eq, true_and] symm - -- Porting note: `Sigma.mk.inj` not found in the following, replaced by `Sigma.mk.inj_iff.mp` - exact - Finsupp.single_apply_left - (fun i i' (h : (⟨j, i⟩ : Σj, ιs j) = ⟨j, i'⟩) => eq_of_heq (Sigma.mk.inj_iff.mp h).2) _ _ _ + simp [Finsupp.single_apply] simp only [Pi.basis, LinearEquiv.trans_apply, Finsupp.sigmaFinsuppLEquivPiFinsupp_symm_apply, - LinearEquiv.piCongrRight] + LinearEquiv.piCongrRight, coe_single] dsimp - rw [stdBasis_ne _ _ _ _ (Ne.symm hj), LinearEquiv.map_zero, Finsupp.zero_apply, + rw [Pi.single_eq_of_ne (Ne.symm hj), LinearEquiv.map_zero, Finsupp.zero_apply, Finsupp.single_eq_of_ne] rintro ⟨⟩ contradiction +set_option linter.deprecated false in +@[simp, deprecated basis_repr_single (since := "2024-08-09")] +theorem basis_repr_stdBasis [DecidableEq η] (s : ∀ j, Basis (ιs j) R (Ms j)) (j i) : + (Pi.basis s).repr (stdBasis R _ j (s j i)) = Finsupp.single ⟨j, i⟩ 1 := + basis_repr_single .. + @[simp] theorem basis_apply [DecidableEq η] (s : ∀ j, Basis (ιs j) R (Ms j)) (ji) : - Pi.basis s ji = stdBasis R _ ji.1 (s ji.1 ji.2) := + Pi.basis s ji = Pi.single ji.1 (s ji.1 ji.2) := Basis.apply_eq_iff.mpr (by simp) @[simp] @@ -251,9 +246,9 @@ noncomputable def basisFun : Basis η R (η → R) := Basis.ofEquivFun (LinearEquiv.refl _ _) @[simp] -theorem basisFun_apply [DecidableEq η] (i) : basisFun R η i = stdBasis R (fun _ : η => R) i 1 := by - simp only [basisFun, Basis.coe_ofEquivFun, LinearEquiv.refl_symm, LinearEquiv.refl_apply, - stdBasis_apply] +theorem basisFun_apply [DecidableEq η] (i) : + basisFun R η i = Pi.single i 1 := by + simp only [basisFun, Basis.coe_ofEquivFun, LinearEquiv.refl_symm, LinearEquiv.refl_apply] @[simp] theorem basisFun_repr (x : η → R) (i : η) : (Pi.basisFun R η).repr x i = x i := by simp [basisFun] @@ -318,10 +313,10 @@ theorem stdBasis_eq_stdBasisMatrix (i : m) (j : n) [DecidableEq m] [DecidableEq · simp only [stdBasis, hj, Basis.coe_reindex, comp_apply, Equiv.sigmaEquivProd_symm_apply, hi, and_true, not_false_iff, StdBasisMatrix.apply_of_ne] erw [Pi.basis_apply] - simp [hi, hj, Ne.symm hi, LinearMap.stdBasis_ne] + simp [hi, hj, Ne.symm hi, Pi.single_eq_of_ne] · simp only [stdBasis, Basis.coe_reindex, comp_apply, Equiv.sigmaEquivProd_symm_apply, hi, hj, and_self, not_false_iff, StdBasisMatrix.apply_of_ne] erw [Pi.basis_apply] - simp [hi, hj, Ne.symm hj, Ne.symm hi, LinearMap.stdBasis_ne] + simp [hi, hj, Ne.symm hj, Ne.symm hi, Pi.single_eq_of_ne] end Matrix diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 69bfc0ff36949..e40e2e449318e 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -83,7 +83,7 @@ theorem Basis.parallelepiped_eq_map {ι E : Type*} [Fintype ι] [NormedAddCommG classical rw [← Basis.parallelepiped_basisFun, ← Basis.parallelepiped_map] congr with x - simp + simp [Pi.single_apply] open MeasureTheory MeasureTheory.Measure diff --git a/Mathlib/RingTheory/AdjoinRoot.lean b/Mathlib/RingTheory/AdjoinRoot.lean index b0337aad5d373..75dc495a1ff01 100644 --- a/Mathlib/RingTheory/AdjoinRoot.lean +++ b/Mathlib/RingTheory/AdjoinRoot.lean @@ -465,10 +465,9 @@ def powerBasis' (hg : g.Monic) : PowerBasis R (AdjoinRoot g) where basis_eq_pow i := by simp only [powerBasisAux', Basis.coe_ofEquivFun, LinearEquiv.coe_symm_mk] rw [Finset.sum_eq_single i] - · rw [Function.update_same, monomial_one_right_eq_X_pow, (mk g).map_pow, mk_X] + · rw [Pi.single_eq_same, monomial_one_right_eq_X_pow, (mk g).map_pow, mk_X] · intro j _ hj - rw [← monomial_zero_right _] - convert congr_arg _ (Function.update_noteq hj _ _) + rw [← monomial_zero_right _, Pi.single_eq_of_ne hj] -- Fix `DecidableEq` mismatch · intros have := Finset.mem_univ i diff --git a/Mathlib/RingTheory/Finiteness.lean b/Mathlib/RingTheory/Finiteness.lean index fae5966bb346c..208f3f7f870cb 100644 --- a/Mathlib/RingTheory/Finiteness.lean +++ b/Mathlib/RingTheory/Finiteness.lean @@ -227,7 +227,7 @@ theorem fg_pi {ι : Type*} {M : ι → Type*} [Finite ι] [∀ i, AddCommMonoid simp_rw [fg_def] at hsb ⊢ choose t htf hts using hsb refine - ⟨⋃ i, (LinearMap.single i : _ →ₗ[R] _) '' t i, Set.finite_iUnion fun i => (htf i).image _, ?_⟩ + ⟨⋃ i, (LinearMap.single R _ i) '' t i, Set.finite_iUnion fun i => (htf i).image _, ?_⟩ -- Note: #8386 changed `span_image` into `span_image _` simp_rw [span_iUnion, span_image _, hts, Submodule.iSup_map_single] diff --git a/Mathlib/RingTheory/Flat/EquationalCriterion.lean b/Mathlib/RingTheory/Flat/EquationalCriterion.lean index 8740523c9b048..42f4b8bf28bd1 100644 --- a/Mathlib/RingTheory/Flat/EquationalCriterion.lean +++ b/Mathlib/RingTheory/Flat/EquationalCriterion.lean @@ -149,7 +149,7 @@ theorem tfae_equational_criterion : List.TFAE [ constructor · apply Finsupp.basisSingleOne.ext intro i - simpa [total_apply, sum_fintype, single_apply] using ha'y' i + simpa [total_apply, sum_fintype, Finsupp.single_apply] using ha'y' i · ext j simp only [total_apply, zero_smul, implies_true, sum_fintype, finset_sum_apply] exact ha' j diff --git a/Mathlib/RingTheory/MvPowerSeries/Basic.lean b/Mathlib/RingTheory/MvPowerSeries/Basic.lean index f05450e9a3df8..8e2b9f0fb7e79 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Basic.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Basic.lean @@ -8,7 +8,6 @@ import Mathlib.Algebra.MvPolynomial.Basic import Mathlib.Algebra.Order.Antidiag.Finsupp import Mathlib.Data.Finsupp.Antidiagonal import Mathlib.Data.Finsupp.Weight -import Mathlib.LinearAlgebra.StdBasis import Mathlib.Tactic.Linarith /-! @@ -129,7 +128,7 @@ variable (R) [Semiring R] and sending all other `x : σ →₀ ℕ` different from `n` to `0`. -/ def monomial (n : σ →₀ ℕ) : R →ₗ[R] MvPowerSeries σ R := letI := Classical.decEq σ - LinearMap.stdBasis R (fun _ ↦ R) n + LinearMap.single R (fun _ ↦ R) n /-- The `n`th coefficient of a multivariate formal power series. -/ def coeff (n : σ →₀ ℕ) : MvPowerSeries σ R →ₗ[R] R := @@ -147,7 +146,7 @@ if and only if all their coefficients are equal. -/ add_decl_doc MvPowerSeries.ext_iff theorem monomial_def [DecidableEq σ] (n : σ →₀ ℕ) : - (monomial R n) = LinearMap.stdBasis R (fun _ ↦ R) n := by + (monomial R n) = LinearMap.single R (fun _ ↦ R) n := by rw [monomial] -- unify the `Decidable` arguments convert rfl @@ -157,18 +156,18 @@ theorem coeff_monomial [DecidableEq σ] (m n : σ →₀ ℕ) (a : R) : -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [coeff, monomial_def, LinearMap.proj_apply (i := m)] -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - erw [LinearMap.stdBasis_apply, Function.update_apply, Pi.zero_apply] + erw [LinearMap.single_apply, Pi.single_apply] @[simp] theorem coeff_monomial_same (n : σ →₀ ℕ) (a : R) : coeff R n (monomial R n a) = a := by classical rw [monomial_def] - exact LinearMap.stdBasis_same R (fun _ ↦ R) n a + exact Pi.single_eq_same _ _ theorem coeff_monomial_ne {m n : σ →₀ ℕ} (h : m ≠ n) (a : R) : coeff R m (monomial R n a) = 0 := by classical rw [monomial_def] - exact LinearMap.stdBasis_ne R (fun _ ↦ R) _ _ h a + exact Pi.single_eq_of_ne h _ theorem eq_of_coeff_monomial_ne_zero {m n : σ →₀ ℕ} {a : R} (h : coeff R m (monomial R n a) ≠ 0) : m = n := diff --git a/Mathlib/RingTheory/SimpleModule.lean b/Mathlib/RingTheory/SimpleModule.lean index 6bd08e9df362c..263f1455ae1d0 100644 --- a/Mathlib/RingTheory/SimpleModule.lean +++ b/Mathlib/RingTheory/SimpleModule.lean @@ -300,7 +300,7 @@ instance {ι} [Finite ι] (R : ι → Type*) [∀ i, Ring (R i)] [∀ i, IsSemis have (i) : IsSemisimpleModule (∀ i, R i) (R i) := ((e i).isSemisimpleModule_iff_of_bijective Function.bijective_id).mpr inferInstance classical - exact isSemisimpleModule_of_isSemisimpleModule_submodule' (p := (range <| single ·)) + exact isSemisimpleModule_of_isSemisimpleModule_submodule' (p := (range <| single _ _ ·)) (fun i ↦ .range _) (by simp_rw [range_eq_map, Submodule.iSup_map_single, Submodule.pi_top]) /-- A binary product of semisimple rings is semisimple. -/ From 0b0394f530ed451a7431d78fcf3832effcf080e8 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 11 Aug 2024 11:38:08 +0000 Subject: [PATCH 171/359] chore(Filter/Pi): add missing `*` in `Type*` (#15684) --- Mathlib/Order/Filter/Pi.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Order/Filter/Pi.lean b/Mathlib/Order/Filter/Pi.lean index a62858e00d968..375092c5e7acd 100644 --- a/Mathlib/Order/Filter/Pi.lean +++ b/Mathlib/Order/Filter/Pi.lean @@ -100,7 +100,7 @@ theorem Eventually.eval_pi {i : ι} (hf : ∀ᶠ x : α i in f i, p i x) : theorem eventually_pi [Finite ι] (hf : ∀ i, ∀ᶠ x in f i, p i x) : ∀ᶠ x : ∀ i, α i in pi f, ∀ i, p i (x i) := eventually_all.2 fun _i => (hf _).eval_pi -theorem hasBasis_pi {ι' : ι → Type} {s : ∀ i, ι' i → Set (α i)} {p : ∀ i, ι' i → Prop} +theorem hasBasis_pi {ι' : ι → Type*} {s : ∀ i, ι' i → Set (α i)} {p : ∀ i, ι' i → Prop} (h : ∀ i, (f i).HasBasis (p i) (s i)) : (pi f).HasBasis (fun If : Set ι × ∀ i, ι' i => If.1.Finite ∧ ∀ i ∈ If.1, p i (If.2 i)) fun If : Set ι × ∀ i, ι' i => If.1.pi fun i => s i <| If.2 i := by From f76c72af16b3425d2b8916d396468d76c3ab4a4a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 11 Aug 2024 16:02:34 +0000 Subject: [PATCH 172/359] chore: Frattini no longer needs Fintype.ofFinite (#15687) --- Mathlib/GroupTheory/Frattini.lean | 4 ---- 1 file changed, 4 deletions(-) diff --git a/Mathlib/GroupTheory/Frattini.lean b/Mathlib/GroupTheory/Frattini.lean index 131e0c34f2ae1..24428121357c5 100644 --- a/Mathlib/GroupTheory/Frattini.lean +++ b/Mathlib/GroupTheory/Frattini.lean @@ -50,10 +50,6 @@ theorem frattini_nongenerating [IsCoatomic (Subgroup G)] {K : Subgroup G} (h : K ⊔ frattini G = ⊤) : K = ⊤ := Order.radical_nongenerating h --- The Sylow files unnecessarily use `Fintype` (computable) where often `Finite` would suffice, --- so we need this: -attribute [local instance] Fintype.ofFinite - /-- When `G` is finite, the Frattini subgroup is nilpotent. -/ theorem frattini_nilpotent [Finite G] : Group.IsNilpotent (frattini G) := by -- We use the characterisation of nilpotency in terms of all Sylow subgroups being normal. From c78e6881c0732682e4bd15eb06ba55a5e1d7e810 Mon Sep 17 00:00:00 2001 From: damiano Date: Sun, 11 Aug 2024 16:11:42 +0000 Subject: [PATCH 173/359] chore: remove unused `set_option in` (#15653) Found by the linter #13653. --- Mathlib/Analysis/PSeries.lean | 1 - Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean | 4 ---- Mathlib/Analysis/SpecialFunctions/Stirling.lean | 4 ---- Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean | 1 - Mathlib/Data/Complex/ExponentialBounds.lean | 1 - Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean | 4 ---- Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean | 1 + Mathlib/NumberTheory/Cyclotomic/Discriminant.lean | 3 --- Mathlib/NumberTheory/ModularForms/SlashActions.lean | 4 ---- Mathlib/RingTheory/RootsOfUnity/Complex.lean | 1 - 10 files changed, 1 insertion(+), 23 deletions(-) diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index aff7e120ffa68..88436cd3124d9 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -374,7 +374,6 @@ open Finset variable {α : Type*} [LinearOrderedField α] -set_option tactic.skipAssignedInstances false in theorem sum_Ioc_inv_sq_le_sub {k n : ℕ} (hk : k ≠ 0) (h : k ≤ n) : (∑ i ∈ Ioc k n, ((i : α) ^ 2)⁻¹) ≤ (k : α)⁻¹ - (n : α)⁻¹ := by refine Nat.le_induction ?_ ?_ n h diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean index 4e974d98df1d6..52f3f9747cae0 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean @@ -292,10 +292,6 @@ theorem hasSum_log_sub_log_of_abs_lt_one {x : ℝ} (h : |x| < 1) : @[deprecated (since := "2024-01-31")] alias hasSum_log_sub_log_of_abs_lt_1 := hasSum_log_sub_log_of_abs_lt_one -#adaptation_note /-- after v4.7.0-rc1, there is a performance problem in `field_simp`. -(Part of the code was ignoring the `maxDischargeDepth` setting: - now that we have to increase it, other paths becomes slow.) -/ -set_option maxHeartbeats 400000 in /-- Expansion of `log (1 + a⁻¹)` as a series in powers of `1 / (2 * a + 1)`. -/ theorem hasSum_log_one_add_inv {a : ℝ} (h : 0 < a) : HasSum (fun k : ℕ => (2 : ℝ) * (1 / (2 * k + 1)) * (1 / (2 * a + 1)) ^ (2 * k + 1)) diff --git a/Mathlib/Analysis/SpecialFunctions/Stirling.lean b/Mathlib/Analysis/SpecialFunctions/Stirling.lean index 1d5205947fa4d..0ff3d40698098 100644 --- a/Mathlib/Analysis/SpecialFunctions/Stirling.lean +++ b/Mathlib/Analysis/SpecialFunctions/Stirling.lean @@ -110,10 +110,6 @@ theorem log_stirlingSeq_diff_le_geo_sum (n : ℕ) : exact inv_le_one (le_add_of_nonneg_left <| by positivity) exact hasSum_le hab (log_stirlingSeq_diff_hasSum n) g -#adaptation_note /-- after v4.7.0-rc1, there is a performance problem in `field_simp`. -(Part of the code was ignoring the `maxDischargeDepth` setting: - now that we have to increase it, other paths become slow.) -/ -set_option maxHeartbeats 400000 in /-- We have the bound `log (stirlingSeq n) - log (stirlingSeq (n+1))` ≤ 1/(4 n^2) -/ theorem log_stirlingSeq_sub_log_stirlingSeq_succ (n : ℕ) : diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean index ae2375694f43c..1086cfd9ef8b8 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean @@ -456,7 +456,6 @@ private theorem edgeDensity_star_not_uniform [Nonempty α] · left; linarith · right; linarith -set_option tactic.skipAssignedInstances false in /-- Lower bound on the edge densities between non-uniform parts of `SzemerediRegularity.increment`. -/ theorem edgeDensity_chunk_not_uniform [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) diff --git a/Mathlib/Data/Complex/ExponentialBounds.lean b/Mathlib/Data/Complex/ExponentialBounds.lean index dcd9a472bfd4e..cb4022b3c3fe8 100644 --- a/Mathlib/Data/Complex/ExponentialBounds.lean +++ b/Mathlib/Data/Complex/ExponentialBounds.lean @@ -47,7 +47,6 @@ theorem exp_neg_one_lt_d9 : exp (-1) < 0.3678794412 := by norm_num · norm_num -set_option tactic.skipAssignedInstances false in theorem log_two_near_10 : |log 2 - 287209 / 414355| ≤ 1 / 10 ^ 10 := by suffices |log 2 - 287209 / 414355| ≤ 1 / 17179869184 + (1 / 10 ^ 10 - 1 / 2 ^ 34) by norm_num1 at * diff --git a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean index 8a716d2787e2c..61e01c28bc811 100644 --- a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean +++ b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean @@ -334,10 +334,6 @@ theorem exists_normalized_aux1 {N : ℕ} {τ : ℝ} (a : SatelliteConfig E N τ) variable [NormedSpace ℝ E] -#adaptation_note /-- after v4.7.0-rc1, there is a performance problem in `field_simp`. -(Part of the code was ignoring the `maxDischargeDepth` setting: - now that we have to increase it, other paths becomes slow.) -/ -set_option maxHeartbeats 400000 in theorem exists_normalized_aux2 {N : ℕ} {τ : ℝ} (a : SatelliteConfig E N τ) (lastc : a.c (last N) = 0) (lastr : a.r (last N) = 1) (hτ : 1 ≤ τ) (δ : ℝ) (hδ1 : τ ≤ 1 + δ / 4) (hδ2 : δ ≤ 1) (i j : Fin N.succ) (inej : i ≠ j) (hi : ‖a.c i‖ ≤ 2) (hj : 2 < ‖a.c j‖) : diff --git a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean index bbbf341b4fcef..5c2649473b452 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean @@ -213,6 +213,7 @@ theorem Basis.parallelepiped_map (b : Basis ι ℝ E) (e : E ≃ₗ[ℝ] F) : LinearMap.isOpenMap_of_finiteDimensional _ e.surjective) := PositiveCompacts.ext (image_parallelepiped e.toLinearMap _).symm +-- removing this option makes elaboration approximately 1 second slower set_option tactic.skipAssignedInstances false in theorem Basis.prod_parallelepiped (v : Basis ι ℝ E) (w : Basis ι' ℝ F) : (v.prod w).parallelepiped = v.parallelepiped.prod w.parallelepiped := by diff --git a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean index 7f0f0d7205f37..3ca7d4838404a 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean @@ -52,7 +52,6 @@ namespace IsCyclotomicExtension variable {p : ℕ+} {k : ℕ} {K : Type u} {L : Type v} {ζ : L} [Field K] [Field L] variable [Algebra K L] -set_option tactic.skipAssignedInstances false in /-- If `p` is a prime and `IsCyclotomicExtension {p ^ (k + 1)} K L`, then the discriminant of `hζ.powerBasis K` is `(-1) ^ ((p ^ (k + 1).totient) / 2) * p ^ (p ^ k * ((p - 1) * (k + 1) - 1))` if `Irreducible (cyclotomic (p ^ (k + 1)) K))`, and `p ^ (k + 1) ≠ 2`. -/ @@ -127,7 +126,6 @@ theorem discr_prime_pow_ne_two' [IsCyclotomicExtension {p ^ (k + 1)} K L] [hp : (-1) ^ ((p : ℕ) ^ k * (p - 1) / 2) * p ^ ((p : ℕ) ^ k * ((p - 1) * (k + 1) - 1)) := by simpa [totient_prime_pow hp.out (succ_pos k)] using discr_prime_pow_ne_two hζ hirr hk -set_option tactic.skipAssignedInstances false in /-- If `p` is a prime and `IsCyclotomicExtension {p ^ k} K L`, then the discriminant of `hζ.powerBasis K` is `(-1) ^ ((p ^ k).totient / 2) * p ^ (p ^ (k - 1) * ((p - 1) * k - 1))` if `Irreducible (cyclotomic (p ^ k) K))`. Beware that in the cases `p ^ k = 1` and `p ^ k = 2` @@ -180,7 +178,6 @@ theorem discr_prime_pow [hcycl : IsCyclotomicExtension {p ^ k} K L] [hp : Fact ( rw [← (algebraMap K L).map_one, trace_algebraMap, finrank _ hirr]; norm_num · exact discr_prime_pow_ne_two hζ hirr hk -set_option tactic.skipAssignedInstances false in /-- If `p` is a prime and `IsCyclotomicExtension {p ^ k} K L`, then there are `u : ℤˣ` and `n : ℕ` such that the discriminant of `hζ.powerBasis K` is `u * p ^ n`. Often this is enough and less cumbersome to use than `IsCyclotomicExtension.discr_prime_pow`. -/ diff --git a/Mathlib/NumberTheory/ModularForms/SlashActions.lean b/Mathlib/NumberTheory/ModularForms/SlashActions.lean index 74b41fc97eb79..4a597199dca6a 100644 --- a/Mathlib/NumberTheory/ModularForms/SlashActions.lean +++ b/Mathlib/NumberTheory/ModularForms/SlashActions.lean @@ -93,10 +93,6 @@ section -- temporary notation until the instance is built local notation:100 f " ∣[" k "]" γ:100 => ModularForm.slash k γ f -#adaptation_note /-- after v4.7.0-rc1, there is a performance problem in `field_simp`. -(Part of the code was ignoring the `maxDischargeDepth` setting: - now that we have to increase it, other paths become slow.) -/ -set_option maxHeartbeats 400000 in private theorem slash_mul (k : ℤ) (A B : GL(2, ℝ)⁺) (f : ℍ → ℂ) : f ∣[k](A * B) = (f ∣[k]A) ∣[k]B := by ext1 x diff --git a/Mathlib/RingTheory/RootsOfUnity/Complex.lean b/Mathlib/RingTheory/RootsOfUnity/Complex.lean index fc45f3a23d202..7f034a01f4772 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Complex.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Complex.lean @@ -117,7 +117,6 @@ theorem IsPrimitiveRoot.arg_eq_pi_iff {n : ℕ} {ζ : ℂ} (hζ : IsPrimitiveRoo (h.trans Complex.arg_neg_one.symm), fun h => h.symm ▸ Complex.arg_neg_one⟩ -set_option tactic.skipAssignedInstances false in theorem IsPrimitiveRoot.arg {n : ℕ} {ζ : ℂ} (h : IsPrimitiveRoot ζ n) (hn : n ≠ 0) : ∃ i : ℤ, ζ.arg = i / n * (2 * Real.pi) ∧ IsCoprime i n ∧ i.natAbs < n := by rw [Complex.isPrimitiveRoot_iff _ _ hn] at h From 22e0c3ba3eb7a646840302a7a9af1b71eaa34b55 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 11 Aug 2024 18:43:44 +0000 Subject: [PATCH 174/359] chore: reduce Cardinal imports (#15579) --- Mathlib.lean | 1 + Mathlib/Algebra/MvPolynomial/Cardinal.lean | 2 +- Mathlib/Algebra/Polynomial/Cardinal.lean | 2 +- Mathlib/LinearAlgebra/Dimension/Free.lean | 1 + Mathlib/SetTheory/Cardinal/Basic.lean | 10 --- Mathlib/SetTheory/Cardinal/Finsupp.lean | 72 ++++++++++++++++++++++ Mathlib/SetTheory/Cardinal/Ordinal.lean | 51 +-------------- 7 files changed, 78 insertions(+), 61 deletions(-) create mode 100644 Mathlib/SetTheory/Cardinal/Finsupp.lean diff --git a/Mathlib.lean b/Mathlib.lean index 6b53163fcfd2d..3290f9d7ca702 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3973,6 +3973,7 @@ import Mathlib.SetTheory.Cardinal.CountableCover import Mathlib.SetTheory.Cardinal.Divisibility import Mathlib.SetTheory.Cardinal.ENat import Mathlib.SetTheory.Cardinal.Finite +import Mathlib.SetTheory.Cardinal.Finsupp import Mathlib.SetTheory.Cardinal.Ordinal import Mathlib.SetTheory.Cardinal.PartENat import Mathlib.SetTheory.Cardinal.SchroederBernstein diff --git a/Mathlib/Algebra/MvPolynomial/Cardinal.lean b/Mathlib/Algebra/MvPolynomial/Cardinal.lean index f18398f31a910..b86b3f22fca31 100644 --- a/Mathlib/Algebra/MvPolynomial/Cardinal.lean +++ b/Mathlib/Algebra/MvPolynomial/Cardinal.lean @@ -5,7 +5,7 @@ Authors: Chris Hughes, Junyan Xu -/ import Mathlib.Algebra.MvPolynomial.Equiv import Mathlib.Data.Finsupp.Fintype -import Mathlib.SetTheory.Cardinal.Ordinal +import Mathlib.SetTheory.Cardinal.Finsupp /-! # Cardinality of Multivariate Polynomial Ring diff --git a/Mathlib/Algebra/Polynomial/Cardinal.lean b/Mathlib/Algebra/Polynomial/Cardinal.lean index 6475cb427828c..b46c1a997b8a6 100644 --- a/Mathlib/Algebra/Polynomial/Cardinal.lean +++ b/Mathlib/Algebra/Polynomial/Cardinal.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Junyan Xu -/ import Mathlib.Algebra.Polynomial.Basic -import Mathlib.SetTheory.Cardinal.Ordinal +import Mathlib.SetTheory.Cardinal.Finsupp /-! # Cardinality of Polynomial Ring diff --git a/Mathlib/LinearAlgebra/Dimension/Free.lean b/Mathlib/LinearAlgebra/Dimension/Free.lean index 5234a91b3b9d2..3b06f6f965e75 100644 --- a/Mathlib/LinearAlgebra/Dimension/Free.lean +++ b/Mathlib/LinearAlgebra/Dimension/Free.lean @@ -6,6 +6,7 @@ Authors: Riccardo Brasca import Mathlib.LinearAlgebra.Dimension.StrongRankCondition import Mathlib.LinearAlgebra.FreeModule.Basic import Mathlib.LinearAlgebra.FreeModule.Finite.Basic +import Mathlib.SetTheory.Cardinal.Finsupp /-! # Rank of free modules diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index bc9c1b0e031a5..302fa05874ae6 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn -/ import Mathlib.Data.Fintype.BigOperators -import Mathlib.Data.Finsupp.Defs import Mathlib.Data.Set.Countable import Mathlib.Logic.Small.Set import Mathlib.Order.SuccPred.CompleteLinearOrder @@ -78,7 +77,6 @@ Cantor's theorem, König's theorem, Konig's theorem -/ assert_not_exists Field -assert_not_exists Module open Mathlib (Vector) open Function Set Order @@ -1217,14 +1215,6 @@ theorem mk_coe_finset {α : Type u} {s : Finset α} : #s = ↑(Finset.card s) := theorem mk_finset_of_fintype [Fintype α] : #(Finset α) = 2 ^ Fintype.card α := by simp [Pow.pow] -@[simp] -theorem mk_finsupp_lift_of_fintype (α : Type u) (β : Type v) [Fintype α] [Zero β] : - #(α →₀ β) = lift.{u} #β ^ Fintype.card α := by - simpa using (@Finsupp.equivFunOnFinite α β _ _).cardinal_eq - -theorem mk_finsupp_of_fintype (α β : Type u) [Fintype α] [Zero β] : - #(α →₀ β) = #β ^ Fintype.card α := by simp - theorem card_le_of_finset {α} (s : Finset α) : (s.card : Cardinal) ≤ #α := @mk_coe_finset _ s ▸ mk_set_le _ diff --git a/Mathlib/SetTheory/Cardinal/Finsupp.lean b/Mathlib/SetTheory/Cardinal/Finsupp.lean new file mode 100644 index 0000000000000..fc2818b357c56 --- /dev/null +++ b/Mathlib/SetTheory/Cardinal/Finsupp.lean @@ -0,0 +1,72 @@ +/- +Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Violeta Hernández Palacios, Junyan Xu +-/ +import Mathlib.SetTheory.Cardinal.Ordinal +import Mathlib.Data.Finsupp.Basic +import Mathlib.Data.Finsupp.Multiset + +/-! # Results on the cardinality of finitely supported functions and multisets. -/ + +universe u v + +namespace Cardinal + +@[simp] +theorem mk_finsupp_lift_of_fintype (α : Type u) (β : Type v) [Fintype α] [Zero β] : + #(α →₀ β) = lift.{u} #β ^ Fintype.card α := by + simpa using (@Finsupp.equivFunOnFinite α β _ _).cardinal_eq + +theorem mk_finsupp_of_fintype (α β : Type u) [Fintype α] [Zero β] : + #(α →₀ β) = #β ^ Fintype.card α := by simp + +@[simp] +theorem mk_finsupp_lift_of_infinite (α : Type u) (β : Type v) [Infinite α] [Zero β] [Nontrivial β] : + #(α →₀ β) = max (lift.{v} #α) (lift.{u} #β) := by + apply le_antisymm + · calc + #(α →₀ β) ≤ #(Finset (α × β)) := mk_le_of_injective (Finsupp.graph_injective α β) + _ = #(α × β) := mk_finset_of_infinite _ + _ = max (lift.{v} #α) (lift.{u} #β) := by + rw [mk_prod, mul_eq_max_of_aleph0_le_left] <;> simp + + · apply max_le <;> rw [← lift_id #(α →₀ β), ← lift_umax] + · cases' exists_ne (0 : β) with b hb + exact lift_mk_le.{v}.2 ⟨⟨_, Finsupp.single_left_injective hb⟩⟩ + · inhabit α + exact lift_mk_le.{u}.2 ⟨⟨_, Finsupp.single_injective default⟩⟩ + +theorem mk_finsupp_of_infinite (α β : Type u) [Infinite α] [Zero β] [Nontrivial β] : + #(α →₀ β) = max #α #β := by simp + +@[simp] +theorem mk_finsupp_lift_of_infinite' (α : Type u) (β : Type v) [Nonempty α] [Zero β] [Infinite β] : + #(α →₀ β) = max (lift.{v} #α) (lift.{u} #β) := by + cases fintypeOrInfinite α + · rw [mk_finsupp_lift_of_fintype] + have : ℵ₀ ≤ (#β).lift := aleph0_le_lift.2 (aleph0_le_mk β) + rw [max_eq_right (le_trans _ this), power_nat_eq this] + exacts [Fintype.card_pos, lift_le_aleph0.2 (lt_aleph0_of_finite _).le] + · apply mk_finsupp_lift_of_infinite + +theorem mk_finsupp_of_infinite' (α β : Type u) [Nonempty α] [Zero β] [Infinite β] : + #(α →₀ β) = max #α #β := by simp + +theorem mk_finsupp_nat (α : Type u) [Nonempty α] : #(α →₀ ℕ) = max #α ℵ₀ := by simp + +theorem mk_multiset_of_isEmpty (α : Type u) [IsEmpty α] : #(Multiset α) = 1 := + Multiset.toFinsupp.toEquiv.cardinal_eq.trans (by simp) + +open scoped Classical + +@[simp] +theorem mk_multiset_of_nonempty (α : Type u) [Nonempty α] : #(Multiset α) = max #α ℵ₀ := + Multiset.toFinsupp.toEquiv.cardinal_eq.trans (mk_finsupp_nat α) + +theorem mk_multiset_of_infinite (α : Type u) [Infinite α] : #(Multiset α) = #α := by simp + +theorem mk_multiset_of_countable (α : Type u) [Countable α] [Nonempty α] : #(Multiset α) = ℵ₀ := + Multiset.toFinsupp.toEquiv.cardinal_eq.trans (by simp) + +end Cardinal diff --git a/Mathlib/SetTheory/Cardinal/Ordinal.lean b/Mathlib/SetTheory/Cardinal/Ordinal.lean index cccfc9d1719b0..98fb9b0b0ef76 100644 --- a/Mathlib/SetTheory/Cardinal/Ordinal.lean +++ b/Mathlib/SetTheory/Cardinal/Ordinal.lean @@ -3,7 +3,6 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn -/ -import Mathlib.Data.Finsupp.Multiset import Mathlib.Order.Bounded import Mathlib.SetTheory.Cardinal.PartENat import Mathlib.SetTheory.Ordinal.Principal @@ -43,6 +42,8 @@ using ordinals. cardinal arithmetic (for infinite cardinals) -/ +assert_not_exists Module +assert_not_exists Finsupp noncomputable section @@ -1073,54 +1074,6 @@ theorem mk_finset_of_infinite (α : Type u) [Infinite α] : #(Finset α) = #α : #(Finset α) ≤ #(List α) := mk_le_of_surjective List.toFinset_surjective _ = #α := mk_list_eq_mk α -@[simp] -theorem mk_finsupp_lift_of_infinite (α : Type u) (β : Type v) [Infinite α] [Zero β] [Nontrivial β] : - #(α →₀ β) = max (lift.{v} #α) (lift.{u} #β) := by - apply le_antisymm - · calc - #(α →₀ β) ≤ #(Finset (α × β)) := mk_le_of_injective (Finsupp.graph_injective α β) - _ = #(α × β) := mk_finset_of_infinite _ - _ = max (lift.{v} #α) (lift.{u} #β) := by - rw [mk_prod, mul_eq_max_of_aleph0_le_left] <;> simp - - · apply max_le <;> rw [← lift_id #(α →₀ β), ← lift_umax] - · cases' exists_ne (0 : β) with b hb - exact lift_mk_le.{v}.2 ⟨⟨_, Finsupp.single_left_injective hb⟩⟩ - · inhabit α - exact lift_mk_le.{u}.2 ⟨⟨_, Finsupp.single_injective default⟩⟩ - -theorem mk_finsupp_of_infinite (α β : Type u) [Infinite α] [Zero β] [Nontrivial β] : - #(α →₀ β) = max #α #β := by simp - -@[simp] -theorem mk_finsupp_lift_of_infinite' (α : Type u) (β : Type v) [Nonempty α] [Zero β] [Infinite β] : - #(α →₀ β) = max (lift.{v} #α) (lift.{u} #β) := by - cases fintypeOrInfinite α - · rw [mk_finsupp_lift_of_fintype] - have : ℵ₀ ≤ (#β).lift := aleph0_le_lift.2 (aleph0_le_mk β) - rw [max_eq_right (le_trans _ this), power_nat_eq this] - exacts [Fintype.card_pos, lift_le_aleph0.2 (lt_aleph0_of_finite _).le] - · apply mk_finsupp_lift_of_infinite - -theorem mk_finsupp_of_infinite' (α β : Type u) [Nonempty α] [Zero β] [Infinite β] : - #(α →₀ β) = max #α #β := by simp - -theorem mk_finsupp_nat (α : Type u) [Nonempty α] : #(α →₀ ℕ) = max #α ℵ₀ := by simp - -@[simp] -theorem mk_multiset_of_nonempty (α : Type u) [Nonempty α] : #(Multiset α) = max #α ℵ₀ := by - classical - exact Multiset.toFinsupp.toEquiv.cardinal_eq.trans (mk_finsupp_nat α) - -theorem mk_multiset_of_infinite (α : Type u) [Infinite α] : #(Multiset α) = #α := by simp - -theorem mk_multiset_of_isEmpty (α : Type u) [IsEmpty α] : #(Multiset α) = 1 := - Multiset.toFinsupp.toEquiv.cardinal_eq.trans (by simp) - -theorem mk_multiset_of_countable (α : Type u) [Countable α] [Nonempty α] : #(Multiset α) = ℵ₀ := by - classical - exact Multiset.toFinsupp.toEquiv.cardinal_eq.trans (by simp) - theorem mk_bounded_set_le_of_infinite (α : Type u) [Infinite α] (c : Cardinal) : #{ t : Set α // #t ≤ c } ≤ #α ^ c := by refine le_trans ?_ (by rw [← add_one_eq (aleph0_le_mk α)]) From d59c78e4718ea11e5cbf74da01933b876e702783 Mon Sep 17 00:00:00 2001 From: damiano Date: Sun, 11 Aug 2024 19:24:47 +0000 Subject: [PATCH 175/359] chore: uniformize some author names (#15256) I simply went for similarities: maybe some of these merges are actually different authors with very similar names! Suggested by the list in [this Zulip message](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/copyright.20header.20line.20lint/near/454640469). --- Archive/Hairer.lean | 4 ++-- Archive/Sensitivity.lean | 7 ++++--- Mathlib/Algebra/Bounds.lean | 4 ++-- Mathlib/Algebra/Field/Basic.lean | 4 ++-- Mathlib/Algebra/Field/Defs.lean | 4 ++-- Mathlib/Algebra/Field/IsField.lean | 4 ++-- Mathlib/Algebra/Field/Power.lean | 4 ++-- Mathlib/Algebra/Group/Fin/Basic.lean | 4 ++-- Mathlib/Algebra/MonoidAlgebra/Basic.lean | 2 +- Mathlib/Algebra/Order/Field/Basic.lean | 4 ++-- Mathlib/Algebra/Order/Field/Canonical/Basic.lean | 4 ++-- Mathlib/Algebra/Order/Field/Canonical/Defs.lean | 4 ++-- Mathlib/Algebra/Order/Field/Defs.lean | 4 ++-- Mathlib/Algebra/Order/Field/InjSurj.lean | 4 ++-- Mathlib/Algebra/Order/Field/Power.lean | 4 ++-- Mathlib/Algebra/Order/Field/Unbundled/Basic.lean | 4 ++-- Mathlib/Algebra/Order/Interval/Set/Monoid.lean | 4 ++-- Mathlib/Algebra/Order/Invertible.lean | 4 ++-- Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean | 2 +- Mathlib/Algebra/Polynomial/Derivation.lean | 2 +- Mathlib/Algebra/Star/Rat.lean | 4 ++-- Mathlib/Analysis/Analytic/Linear.lean | 4 ++-- Mathlib/Analysis/Asymptotics/Theta.lean | 4 ++-- Mathlib/Analysis/Calculus/DiffContOnCl.lean | 4 ++-- .../Analysis/Calculus/InverseFunctionTheorem/Deriv.lean | 4 ++-- Mathlib/Analysis/Complex/AbsMax.lean | 4 ++-- Mathlib/Analysis/Complex/Convex.lean | 4 ++-- Mathlib/Analysis/Complex/Liouville.lean | 4 ++-- Mathlib/Analysis/Complex/Schwarz.lean | 4 ++-- Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean | 4 ++-- Mathlib/Analysis/Convex/Basic.lean | 2 +- Mathlib/Analysis/Convex/Combination.lean | 4 ++-- Mathlib/Analysis/Convex/Complex.lean | 4 ++-- Mathlib/Analysis/Convex/Hull.lean | 4 ++-- Mathlib/Analysis/Convex/Integral.lean | 4 ++-- Mathlib/Analysis/Convex/Jensen.lean | 2 +- Mathlib/Analysis/Convex/Slope.lean | 4 ++-- Mathlib/Analysis/Normed/Module/Completion.lean | 4 ++-- Mathlib/Analysis/NormedSpace/ENorm.lean | 4 ++-- Mathlib/Analysis/ODE/PicardLindelof.lean | 4 ++-- Mathlib/Analysis/PSeries.lean | 4 ++-- Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean | 4 ++-- Mathlib/Analysis/SpecialFunctions/Sqrt.lean | 4 ++-- Mathlib/Analysis/SpecificLimits/Basic.lean | 2 +- Mathlib/Analysis/SpecificLimits/Normed.lean | 2 +- Mathlib/CategoryTheory/Endofunctor/Algebra.lean | 2 +- .../CategoryTheory/Sites/Coherent/CoherentTopology.lean | 2 +- Mathlib/CategoryTheory/Sites/Grothendieck.lean | 4 ++-- Mathlib/CategoryTheory/Sites/Sieves.lean | 2 +- Mathlib/Condensed/Equivalence.lean | 2 +- Mathlib/Control/Monad/Writer.lean | 2 +- Mathlib/Data/Bool/Set.lean | 4 ++-- Mathlib/Data/Finsupp/AList.lean | 4 ++-- Mathlib/Data/Finsupp/Weight.lean | 4 ++-- Mathlib/Data/List/NodupEquivFin.lean | 4 ++-- Mathlib/Data/Set/BoolIndicator.lean | 4 ++-- Mathlib/Data/Set/Pointwise/BoundedMul.lean | 4 ++-- Mathlib/Data/Set/Pointwise/Interval.lean | 4 ++-- Mathlib/Data/Sum/Basic.lean | 2 +- .../Dynamics/Circle/RotationNumber/TranslationNumber.lean | 4 ++-- Mathlib/Dynamics/Minimal.lean | 4 ++-- Mathlib/Dynamics/PeriodicPts.lean | 4 ++-- Mathlib/Geometry/Euclidean/Inversion/Basic.lean | 4 ++-- Mathlib/Geometry/Manifold/PartitionOfUnity.lean | 4 ++-- Mathlib/Geometry/Manifold/WhitneyEmbedding.lean | 4 ++-- Mathlib/Lean/Exception.lean | 4 ++-- Mathlib/Lean/Expr/Basic.lean | 2 +- Mathlib/Lean/Expr/ReplaceRec.lean | 2 +- Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean | 4 ++-- Mathlib/LinearAlgebra/AffineSpace/Ordered.lean | 4 ++-- Mathlib/LinearAlgebra/AffineSpace/Slope.lean | 4 ++-- Mathlib/MeasureTheory/Function/Floor.lean | 4 ++-- Mathlib/MeasureTheory/Group/Action.lean | 4 ++-- Mathlib/MeasureTheory/Group/FundamentalDomain.lean | 4 ++-- Mathlib/MeasureTheory/Group/MeasurableEquiv.lean | 4 ++-- Mathlib/MeasureTheory/Group/Pointwise.lean | 4 ++-- Mathlib/MeasureTheory/Integral/Average.lean | 4 ++-- Mathlib/MeasureTheory/Integral/FundThmCalculus.lean | 4 ++-- Mathlib/MeasureTheory/Integral/IntervalIntegral.lean | 4 ++-- .../MeasureTheory/MeasurableSpace/CountablyGenerated.lean | 2 +- Mathlib/MeasureTheory/Measure/Regular.lean | 2 +- Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean | 2 +- Mathlib/NumberTheory/Liouville/LiouvilleWith.lean | 4 ++-- Mathlib/NumberTheory/Liouville/Measure.lean | 4 ++-- Mathlib/NumberTheory/Ostrowski.lean | 2 +- Mathlib/Order/Birkhoff.lean | 2 +- Mathlib/Order/Filter/CountableInter.lean | 4 ++-- Mathlib/Order/Filter/Germ/Basic.lean | 4 ++-- Mathlib/Order/Filter/Germ/OrderedMonoid.lean | 4 ++-- Mathlib/Order/Filter/Interval.lean | 4 ++-- Mathlib/Order/Filter/Pi.lean | 4 ++-- Mathlib/Order/Filter/Prod.lean | 2 +- Mathlib/Order/Interval/Set/OrdConnected.lean | 4 ++-- Mathlib/Order/Interval/Set/ProjIcc.lean | 4 ++-- Mathlib/Order/Interval/Set/WithBotTop.lean | 4 ++-- Mathlib/Order/Iterate.lean | 4 ++-- Mathlib/Order/OrdContinuous.lean | 4 ++-- Mathlib/Order/RelClasses.lean | 2 +- Mathlib/Order/SemiconjSup.lean | 4 ++-- Mathlib/RingTheory/PowerSeries/WellKnown.lean | 4 ++-- Mathlib/SetTheory/Cardinal/ENat.lean | 4 ++-- Mathlib/Tactic/FunProp/Attr.lean | 4 ++-- Mathlib/Tactic/FunProp/Core.lean | 4 ++-- Mathlib/Tactic/FunProp/Decl.lean | 4 ++-- Mathlib/Tactic/FunProp/Elab.lean | 4 ++-- Mathlib/Tactic/FunProp/FunctionData.lean | 4 ++-- Mathlib/Tactic/FunProp/Mor.lean | 4 ++-- Mathlib/Tactic/FunProp/Theorems.lean | 4 ++-- Mathlib/Tactic/FunProp/ToBatteries.lean | 4 ++-- Mathlib/Tactic/FunProp/Types.lean | 4 ++-- Mathlib/Tactic/Polyrith.lean | 2 +- Mathlib/Tactic/TypeStar.lean | 4 ++-- Mathlib/Topology/Algebra/GroupWithZero.lean | 4 ++-- Mathlib/Topology/Algebra/Order/Archimedean.lean | 4 ++-- Mathlib/Topology/Algebra/UniformMulAction.lean | 4 ++-- Mathlib/Topology/Bornology/Constructions.lean | 4 ++-- Mathlib/Topology/EMetricSpace/Paracompact.lean | 4 ++-- Mathlib/Topology/Inseparable.lean | 2 +- Mathlib/Topology/Instances/Irrational.lean | 4 ++-- Mathlib/Topology/MetricSpace/Holder.lean | 4 ++-- Mathlib/Topology/MetricSpace/ShrinkingLemma.lean | 4 ++-- Mathlib/Topology/Order/Category/FrameAdjunction.lean | 4 ++-- Mathlib/Topology/Order/ExtrClosure.lean | 4 ++-- Mathlib/Topology/Order/IntermediateValue.lean | 4 ++-- Mathlib/Topology/Order/MonotoneContinuity.lean | 4 ++-- Mathlib/Topology/TietzeExtension.lean | 4 ++-- Mathlib/Topology/UrysohnsLemma.lean | 4 ++-- Mathlib/Topology/VectorBundle/Basic.lean | 2 +- Mathlib/Util/MemoFix.lean | 2 +- test/Set.lean | 2 +- 130 files changed, 235 insertions(+), 234 deletions(-) diff --git a/Archive/Hairer.lean b/Archive/Hairer.lean index 74e24e2876566..acd326c0b13bf 100644 --- a/Archive/Hairer.lean +++ b/Archive/Hairer.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2023 Floris Van Doorn. All rights reserved. +Copyright (c) 2023 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johan Commelin, Sébastien Gouëzel, Patrick Massot, Ruben Van de Velde, Floris Van Doorn, +Authors: Johan Commelin, Sébastien Gouëzel, Patrick Massot, Ruben Van de Velde, Floris van Doorn, Junyan Xu -/ import Mathlib.Algebra.MvPolynomial.Funext diff --git a/Archive/Sensitivity.lean b/Archive/Sensitivity.lean index 08871a933bd8c..ae88d8cc7fc27 100644 --- a/Archive/Sensitivity.lean +++ b/Archive/Sensitivity.lean @@ -1,8 +1,9 @@ /- -Copyright (c) 2019 Reid Barton, Johan Commelin, Jesse Han, Chris Hughes, Robert Y. Lewis, +Copyright (c) 2019 Reid Barton, Johan Commelin, Jesse Michael Han, Chris Hughes, Robert Y. Lewis, Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Reid Barton, Johan Commelin, Jesse Han, Chris Hughes, Robert Y. Lewis, Patrick Massot +Authors: Reid Barton, Johan Commelin, Jesse Michael Han, Chris Hughes, Robert Y. Lewis, + Patrick Massot -/ import Mathlib.Tactic.FinCases import Mathlib.Tactic.ApplyFun @@ -19,7 +20,7 @@ dimension n ≥ 1, if one colors more than half the vertices then at least one vertex has at least √n colored neighbors. A fun summer collaboration by -Reid Barton, Johan Commelin, Jesse Han, Chris Hughes, Robert Y. Lewis, and Patrick Massot, +Reid Barton, Johan Commelin, Jesse Michael Han, Chris Hughes, Robert Y. Lewis, and Patrick Massot, based on Don Knuth's account of the story (https://www.cs.stanford.edu/~knuth/papers/huang.pdf), using the Lean theorem prover (https://leanprover.github.io/), diff --git a/Mathlib/Algebra/Bounds.lean b/Mathlib/Algebra/Bounds.lean index 4cd78689220c4..95219c22b4fa9 100644 --- a/Mathlib/Algebra/Bounds.lean +++ b/Mathlib/Algebra/Bounds.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Algebra.Order.Group.OrderIso import Mathlib.Data.Set.Pointwise.Basic diff --git a/Mathlib/Algebra/Field/Basic.lean b/Mathlib/Algebra/Field/Basic.lean index 30ccec87b7083..b0dabde1c5792 100644 --- a/Mathlib/Algebra/Field/Basic.lean +++ b/Mathlib/Algebra/Field/Basic.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Robert Lewis. All rights reserved. +Copyright (c) 2014 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro +Authors: Robert Y. Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Field.Defs import Mathlib.Algebra.GroupWithZero.Units.Lemmas diff --git a/Mathlib/Algebra/Field/Defs.lean b/Mathlib/Algebra/Field/Defs.lean index 4888fed53964a..f7a7742c83c9d 100644 --- a/Mathlib/Algebra/Field/Defs.lean +++ b/Mathlib/Algebra/Field/Defs.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Robert Lewis. All rights reserved. +Copyright (c) 2014 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro, Yaël Dillies +Authors: Robert Y. Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro, Yaël Dillies -/ import Mathlib.Algebra.Ring.Defs import Mathlib.Data.Rat.Init diff --git a/Mathlib/Algebra/Field/IsField.lean b/Mathlib/Algebra/Field/IsField.lean index 4890d9fed4f35..509721d425184 100644 --- a/Mathlib/Algebra/Field/IsField.lean +++ b/Mathlib/Algebra/Field/IsField.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Robert Lewis. All rights reserved. +Copyright (c) 2014 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro +Authors: Robert Y. Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Field.Defs import Mathlib.Tactic.Common diff --git a/Mathlib/Algebra/Field/Power.lean b/Mathlib/Algebra/Field/Power.lean index 2e6d21899a6a5..8f85d4bef98a9 100644 --- a/Mathlib/Algebra/Field/Power.lean +++ b/Mathlib/Algebra/Field/Power.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Robert Lewis. All rights reserved. +Copyright (c) 2014 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro +Authors: Robert Y. Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro -/ import Mathlib.Algebra.Field.Defs import Mathlib.Algebra.Ring.Int diff --git a/Mathlib/Algebra/Group/Fin/Basic.lean b/Mathlib/Algebra/Group/Fin/Basic.lean index 2a21fd4c5aa71..ec862c8479a6a 100644 --- a/Mathlib/Algebra/Group/Fin/Basic.lean +++ b/Mathlib/Algebra/Group/Fin/Basic.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yakov Peckersky. All rights reserved. +Copyright (c) 2021 Yakov Pechersky. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yakov Peckersky +Authors: Yakov Pechersky -/ import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.NeZero diff --git a/Mathlib/Algebra/MonoidAlgebra/Basic.lean b/Mathlib/Algebra/MonoidAlgebra/Basic.lean index 41dceb8c077b0..4022f37abb36e 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Basic.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johannes Hölzl, Yury G. Kudryashov, Scott Morrison +Authors: Johannes Hölzl, Yury Kudryashov, Scott Morrison -/ import Mathlib.Algebra.Algebra.Equiv import Mathlib.Algebra.Algebra.NonUnitalHom diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index edf647b081fba..61dd3d24f00a7 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Robert Lewis. All rights reserved. +Copyright (c) 2014 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn +Authors: Robert Y. Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.CharZero.Lemmas import Mathlib.Algebra.Order.Field.Defs diff --git a/Mathlib/Algebra/Order/Field/Canonical/Basic.lean b/Mathlib/Algebra/Order/Field/Canonical/Basic.lean index 860c00b1294ed..ac3b9f3e55d5d 100644 --- a/Mathlib/Algebra/Order/Field/Canonical/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Canonical/Basic.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Robert Lewis. All rights reserved. +Copyright (c) 2014 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn +Authors: Robert Y. Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.Order.Field.Canonical.Defs diff --git a/Mathlib/Algebra/Order/Field/Canonical/Defs.lean b/Mathlib/Algebra/Order/Field/Canonical/Defs.lean index 496ce58ce4222..f09b170c98fab 100644 --- a/Mathlib/Algebra/Order/Field/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Field/Canonical/Defs.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Robert Lewis. All rights reserved. +Copyright (c) 2014 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn +Authors: Robert Y. Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.Order.Field.Defs import Mathlib.Algebra.Order.Ring.Canonical diff --git a/Mathlib/Algebra/Order/Field/Defs.lean b/Mathlib/Algebra/Order/Field/Defs.lean index 83ae7171dd273..5e36efa7b7f81 100644 --- a/Mathlib/Algebra/Order/Field/Defs.lean +++ b/Mathlib/Algebra/Order/Field/Defs.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Robert Lewis. All rights reserved. +Copyright (c) 2014 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn +Authors: Robert Y. Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Algebra.Field.Defs diff --git a/Mathlib/Algebra/Order/Field/InjSurj.lean b/Mathlib/Algebra/Order/Field/InjSurj.lean index ce215d1ed616d..73fb2c67d8eac 100644 --- a/Mathlib/Algebra/Order/Field/InjSurj.lean +++ b/Mathlib/Algebra/Order/Field/InjSurj.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Robert Lewis. All rights reserved. +Copyright (c) 2014 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn +Authors: Robert Y. Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.Field.Basic import Mathlib.Algebra.Order.Field.Defs diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index aefad3506d9bd..8cb0dc8de386a 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Robert Lewis. All rights reserved. +Copyright (c) 2014 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn +Authors: Robert Y. Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.CharZero.Lemmas import Mathlib.Algebra.GroupWithZero.Commute diff --git a/Mathlib/Algebra/Order/Field/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Field/Unbundled/Basic.lean index 35d309de2f0a3..19f180e9e1699 100644 --- a/Mathlib/Algebra/Order/Field/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Unbundled/Basic.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Robert Lewis. All rights reserved. +Copyright (c) 2014 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn +Authors: Robert Y. Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn -/ import Mathlib.Algebra.Field.Defs import Mathlib.Algebra.GroupWithZero.Basic diff --git a/Mathlib/Algebra/Order/Interval/Set/Monoid.lean b/Mathlib/Algebra/Order/Interval/Set/Monoid.lean index e963b968a5756..c61a84f70e84a 100644 --- a/Mathlib/Algebra/Order/Interval/Set/Monoid.lean +++ b/Mathlib/Algebra/Order/Interval/Set/Monoid.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Patrick Massot +Authors: Yury Kudryashov, Patrick Massot -/ import Mathlib.Algebra.Group.Basic import Mathlib.Data.Set.Function diff --git a/Mathlib/Algebra/Order/Invertible.lean b/Mathlib/Algebra/Order/Invertible.lean index 2f03e065abdc8..ebf2c06de0127 100644 --- a/Mathlib/Algebra/Order/Invertible.lean +++ b/Mathlib/Algebra/Order/Invertible.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Algebra.Ring.Invertible diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean index 08d3cc925d652..ca33dfd14cedd 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean @@ -1,7 +1,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, Yury G. Kudryashov +Authors: Jeremy Avigad, Robert Y. Lewis, Yury Kudryashov -/ import Mathlib.Algebra.Order.Monoid.Unbundled.Basic import Mathlib.Algebra.Order.Monoid.Unbundled.OrderDual diff --git a/Mathlib/Algebra/Polynomial/Derivation.lean b/Mathlib/Algebra/Polynomial/Derivation.lean index 6e2cb709b04c3..5a14afb6cc51f 100644 --- a/Mathlib/Algebra/Polynomial/Derivation.lean +++ b/Mathlib/Algebra/Polynomial/Derivation.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Kevin Buzzard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kevin Buzzard, Richard Hill +Authors: Kevin Buzzard, Richard M. Hill -/ import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.Algebra.Polynomial.Derivative diff --git a/Mathlib/Algebra/Star/Rat.lean b/Mathlib/Algebra/Star/Rat.lean index 65ac65ceee0be..927ff6812ec1d 100644 --- a/Mathlib/Algebra/Star/Rat.lean +++ b/Mathlib/Algebra/Star/Rat.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Yael Dillies. All rights reserved. +Copyright (c) 2024 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yael Dillies +Authors: Yaël Dillies -/ import Mathlib.Algebra.Field.Opposite import Mathlib.Algebra.Star.Basic diff --git a/Mathlib/Analysis/Analytic/Linear.lean b/Mathlib/Analysis/Analytic/Linear.lean index 6addbfac86dae..f7afb9b160fc2 100644 --- a/Mathlib/Analysis/Analytic/Linear.lean +++ b/Mathlib/Analysis/Analytic/Linear.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Analytic.Basic diff --git a/Mathlib/Analysis/Asymptotics/Theta.lean b/Mathlib/Analysis/Asymptotics/Theta.lean index e352e6b93cbfc..ab9b0e2af090e 100644 --- a/Mathlib/Analysis/Asymptotics/Theta.lean +++ b/Mathlib/Analysis/Asymptotics/Theta.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Asymptotics.Asymptotics import Mathlib.Analysis.Normed.Module.Basic diff --git a/Mathlib/Analysis/Calculus/DiffContOnCl.lean b/Mathlib/Analysis/Calculus/DiffContOnCl.lean index da4a7d25f2c2b..10a6fd3b0a289 100644 --- a/Mathlib/Analysis/Calculus/DiffContOnCl.lean +++ b/Mathlib/Analysis/Calculus/DiffContOnCl.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Calculus.Deriv.Inv import Mathlib.Analysis.NormedSpace.Real diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean index 638229d72659c..f4db529395d41 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Calculus.Deriv.Inverse import Mathlib.Analysis.Calculus.InverseFunctionTheorem.FDeriv diff --git a/Mathlib/Analysis/Complex/AbsMax.lean b/Mathlib/Analysis/Complex/AbsMax.lean index 931e271b760d3..9bb3354667e7c 100644 --- a/Mathlib/Analysis/Complex/AbsMax.lean +++ b/Mathlib/Analysis/Complex/AbsMax.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Complex.CauchyIntegral import Mathlib.Analysis.Normed.Module.Completion diff --git a/Mathlib/Analysis/Complex/Convex.lean b/Mathlib/Analysis/Complex/Convex.lean index 211bd245b6638..2bac30b8eb149 100644 --- a/Mathlib/Analysis/Complex/Convex.lean +++ b/Mathlib/Analysis/Complex/Convex.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2023 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2023 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Convex.Combination import Mathlib.Analysis.Complex.Basic diff --git a/Mathlib/Analysis/Complex/Liouville.lean b/Mathlib/Analysis/Complex/Liouville.lean index 13ec934118ccc..1075a8ade0521 100644 --- a/Mathlib/Analysis/Complex/Liouville.lean +++ b/Mathlib/Analysis/Complex/Liouville.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Complex.CauchyIntegral import Mathlib.Analysis.Calculus.FDeriv.Analytic diff --git a/Mathlib/Analysis/Complex/Schwarz.lean b/Mathlib/Analysis/Complex/Schwarz.lean index 76d7b547b0fe4..cf851529fbf6f 100644 --- a/Mathlib/Analysis/Complex/Schwarz.lean +++ b/Mathlib/Analysis/Complex/Schwarz.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Complex.AbsMax import Mathlib.Analysis.Complex.RemovableSingularity diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean index 7ea463a66a4b7..28f5dd099e13f 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Complex.UpperHalfPlane.Basic import Mathlib.Analysis.Convex.Contractible diff --git a/Mathlib/Analysis/Convex/Basic.lean b/Mathlib/Analysis/Convex/Basic.lean index 89b70137948b7..24f8e44ef54c4 100644 --- a/Mathlib/Analysis/Convex/Basic.lean +++ b/Mathlib/Analysis/Convex/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Alexander Bentkamp. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Alexander Bentkamp, Yury Kudriashov, Yaël Dillies +Authors: Alexander Bentkamp, Yury Kudryashov, Yaël Dillies -/ import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Algebra.Order.Module.OrderedSMul diff --git a/Mathlib/Analysis/Convex/Combination.lean b/Mathlib/Analysis/Convex/Combination.lean index a14b272ac8de7..58926a2e62815 100644 --- a/Mathlib/Analysis/Convex/Combination.lean +++ b/Mathlib/Analysis/Convex/Combination.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2019 Yury Kudriashov. All rights reserved. +Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury Kudriashov +Authors: Yury Kudryashov -/ import Mathlib.Algebra.Order.BigOperators.Ring.Finset import Mathlib.Analysis.Convex.Hull diff --git a/Mathlib/Analysis/Convex/Complex.lean b/Mathlib/Analysis/Convex/Complex.lean index 11bfb59e83a93..489c384d7e4f1 100644 --- a/Mathlib/Analysis/Convex/Complex.lean +++ b/Mathlib/Analysis/Convex/Complex.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2019 Yury Kudriashov. All rights reserved. +Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury Kudriashov, Yaël Dillies +Authors: Yury Kudryashov, Yaël Dillies -/ import Mathlib.Analysis.Convex.Basic import Mathlib.Data.Complex.Module diff --git a/Mathlib/Analysis/Convex/Hull.lean b/Mathlib/Analysis/Convex/Hull.lean index 93049c95f690f..996ebbb415ab1 100644 --- a/Mathlib/Analysis/Convex/Hull.lean +++ b/Mathlib/Analysis/Convex/Hull.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury Kudriashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury Kudriashov, Yaël Dillies +Authors: Yury Kudryashov, Yaël Dillies -/ import Mathlib.Analysis.Convex.Basic import Mathlib.Order.Closure diff --git a/Mathlib/Analysis/Convex/Integral.lean b/Mathlib/Analysis/Convex/Integral.lean index 668b85bf22562..435c9f7ec4a19 100644 --- a/Mathlib/Analysis/Convex/Integral.lean +++ b/Mathlib/Analysis/Convex/Integral.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Convex.Function import Mathlib.Analysis.Convex.StrictConvexSpace diff --git a/Mathlib/Analysis/Convex/Jensen.lean b/Mathlib/Analysis/Convex/Jensen.lean index 9de725de0be6c..cbfbf7d8bf46b 100644 --- a/Mathlib/Analysis/Convex/Jensen.lean +++ b/Mathlib/Analysis/Convex/Jensen.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Alexander Bentkamp. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Alexander Bentkamp, Yury Kudriashov +Authors: Alexander Bentkamp, Yury Kudryashov -/ import Mathlib.Analysis.Convex.Combination import Mathlib.Analysis.Convex.Function diff --git a/Mathlib/Analysis/Convex/Slope.lean b/Mathlib/Analysis/Convex/Slope.lean index 0be23d8a46fcc..4bfa3fe7e4643 100644 --- a/Mathlib/Analysis/Convex/Slope.lean +++ b/Mathlib/Analysis/Convex/Slope.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury Kudriashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury Kudriashov, Malo Jaffré +Authors: Yury Kudryashov, Malo Jaffré -/ import Mathlib.Analysis.Convex.Function import Mathlib.Tactic.AdaptationNote diff --git a/Mathlib/Analysis/Normed/Module/Completion.lean b/Mathlib/Analysis/Normed/Module/Completion.lean index 4d52281665885..4e50f3b1130b0 100644 --- a/Mathlib/Analysis/Normed/Module/Completion.lean +++ b/Mathlib/Analysis/Normed/Module/Completion.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Normed.Group.Completion import Mathlib.Analysis.NormedSpace.OperatorNorm.NormedSpace diff --git a/Mathlib/Analysis/NormedSpace/ENorm.lean b/Mathlib/Analysis/NormedSpace/ENorm.lean index fccadf993c631..fbbb8c726b3ad 100644 --- a/Mathlib/Analysis/NormedSpace/ENorm.lean +++ b/Mathlib/Analysis/NormedSpace/ENorm.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Normed.Module.Basic diff --git a/Mathlib/Analysis/ODE/PicardLindelof.lean b/Mathlib/Analysis/ODE/PicardLindelof.lean index 945ba6e8e3483..ca5ac0ac8a0a4 100644 --- a/Mathlib/Analysis/ODE/PicardLindelof.lean +++ b/Mathlib/Analysis/ODE/PicardLindelof.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Winston Yin +Authors: Yury Kudryashov, Winston Yin -/ import Mathlib.Analysis.SpecialFunctions.Integrals import Mathlib.Topology.MetricSpace.Contracting diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 88436cd3124d9..8a19619191174 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.SpecialFunctions.Pow.NNReal import Mathlib.Analysis.SpecialFunctions.Pow.Continuity diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean index 72cfc152f87d1..1dde0cd951592 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Complex.Circle import Mathlib.Analysis.SpecialFunctions.Complex.Log diff --git a/Mathlib/Analysis/SpecialFunctions/Sqrt.lean b/Mathlib/Analysis/SpecialFunctions/Sqrt.lean index 9af0fe3aa8414..4a26cd60c9cbc 100644 --- a/Mathlib/Analysis/SpecialFunctions/Sqrt.lean +++ b/Mathlib/Analysis/SpecialFunctions/Sqrt.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Calculus.ContDiff.Basic import Mathlib.Analysis.Calculus.Deriv.Pow diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 314ebd5e16ce3..cc06f030ef173 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel, Johannes Hölzl, Yury G. Kudryashov, Patrick Massot +Authors: Sébastien Gouëzel, Johannes Hölzl, Yury Kudryashov, Patrick Massot -/ import Mathlib.Algebra.GeomSum import Mathlib.Order.Filter.Archimedean diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index 22e2390134e42..21ee99951f714 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Anatole Dedecker, Sébastien Gouëzel, Yury G. Kudryashov, Dylan MacKenzie, Patrick Massot +Authors: Anatole Dedecker, Sébastien Gouëzel, Yury Kudryashov, Dylan MacKenzie, Patrick Massot -/ import Mathlib.Algebra.BigOperators.Module import Mathlib.Algebra.Order.Field.Basic diff --git a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean index 26791079e378f..0ba3c1ced4604 100644 --- a/Mathlib/CategoryTheory/Endofunctor/Algebra.lean +++ b/Mathlib/CategoryTheory/Endofunctor/Algebra.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Joseph Hua. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison, Bhavik Mehta, Johan Commelin, Reid Barton, Rob Lewis, Joseph Hua +Authors: Scott Morrison, Bhavik Mehta, Johan Commelin, Reid Barton, Robert Y. Lewis, Joseph Hua -/ import Mathlib.CategoryTheory.Limits.Shapes.Terminal diff --git a/Mathlib/CategoryTheory/Sites/Coherent/CoherentTopology.lean b/Mathlib/CategoryTheory/Sites/Coherent/CoherentTopology.lean index 366489770146d..bd53932b0ed07 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/CoherentTopology.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/CoherentTopology.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Adam Topaz. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Adam Topaz, Nick Kuhn +Authors: Adam Topaz, Nikolas Kuhn -/ import Mathlib.CategoryTheory.Sites.Coherent.CoherentSheaves /-! diff --git a/Mathlib/CategoryTheory/Sites/Grothendieck.lean b/Mathlib/CategoryTheory/Sites/Grothendieck.lean index f7773b4fc0fb2..faa04b5947877 100644 --- a/Mathlib/CategoryTheory/Sites/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Sites/Grothendieck.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Bhavik Mehta, E. W. Ayers. All rights reserved. +Copyright (c) 2020 Bhavik Mehta, Edward Ayers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Bhavik Mehta, E. W. Ayers +Authors: Bhavik Mehta, Edward Ayers -/ import Mathlib.CategoryTheory.Sites.Sieves import Mathlib.CategoryTheory.Limits.Shapes.Multiequalizer diff --git a/Mathlib/CategoryTheory/Sites/Sieves.lean b/Mathlib/CategoryTheory/Sites/Sieves.lean index 16cee3b83af67..d745ad87d9aed 100644 --- a/Mathlib/CategoryTheory/Sites/Sieves.lean +++ b/Mathlib/CategoryTheory/Sites/Sieves.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Bhavik Mehta, E. W. Ayers +Authors: Bhavik Mehta, Edward Ayers -/ import Mathlib.CategoryTheory.Comma.Over import Mathlib.CategoryTheory.Limits.Shapes.Pullback.HasPullback diff --git a/Mathlib/Condensed/Equivalence.lean b/Mathlib/Condensed/Equivalence.lean index 41a393ec6337c..579875bfae18a 100644 --- a/Mathlib/Condensed/Equivalence.lean +++ b/Mathlib/Condensed/Equivalence.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Adam Topaz. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Adam Topaz, Nick Kuhn, Dagur Asgeirsson +Authors: Adam Topaz, Nikolas Kuhn, Dagur Asgeirsson -/ import Mathlib.Topology.Category.Profinite.EffectiveEpi import Mathlib.Topology.Category.Stonean.EffectiveEpi diff --git a/Mathlib/Control/Monad/Writer.lean b/Mathlib/Control/Monad/Writer.lean index bd76158292c0c..fcbf590189c9f 100644 --- a/Mathlib/Control/Monad/Writer.lean +++ b/Mathlib/Control/Monad/Writer.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Simon Hudon, E.W.Ayers +Authors: Simon Hudon, Edward Ayers -/ import Mathlib.Algebra.Group.Defs import Mathlib.Logic.Equiv.Defs diff --git a/Mathlib/Data/Bool/Set.lean b/Mathlib/Data/Bool/Set.lean index 4b9691909a876..4c63a9edfac52 100644 --- a/Mathlib/Data/Bool/Set.lean +++ b/Mathlib/Data/Bool/Set.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Data.Set.Image diff --git a/Mathlib/Data/Finsupp/AList.lean b/Mathlib/Data/Finsupp/AList.lean index 10d6f833c569c..d7f3418427f67 100644 --- a/Mathlib/Data/Finsupp/AList.lean +++ b/Mathlib/Data/Finsupp/AList.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Violeta Hernández. All rights reserved. +Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Violeta Hernández +Authors: Violeta Hernández Palacios -/ import Mathlib.Data.Finsupp.Basic import Mathlib.Data.List.AList diff --git a/Mathlib/Data/Finsupp/Weight.lean b/Mathlib/Data/Finsupp/Weight.lean index 74ecc12d9a715..8c24117e1132d 100644 --- a/Mathlib/Data/Finsupp/Weight.lean +++ b/Mathlib/Data/Finsupp/Weight.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Antoine Chambert-Loir, María Inés de Frutos Fernández. All rights reserved. +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 +Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández -/ import Mathlib.Algebra.Order.BigOperators.Group.Finset diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index b65f0f158ae71..36b450b8c49a8 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Data.List.Duplicate import Mathlib.Data.List.Sort diff --git a/Mathlib/Data/Set/BoolIndicator.lean b/Mathlib/Data/Set/BoolIndicator.lean index fb3eb84a90b3c..09d93329c0505 100644 --- a/Mathlib/Data/Set/BoolIndicator.lean +++ b/Mathlib/Data/Set/BoolIndicator.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Dagur Tómas Ásgeirsson. All rights reserved. +Copyright (c) 2022 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Dagur Tómas Ásgeirsson, Leonardo de Moura +Authors: Dagur Asgeirsson, Leonardo de Moura -/ import Mathlib.Data.Set.Basic diff --git a/Mathlib/Data/Set/Pointwise/BoundedMul.lean b/Mathlib/Data/Set/Pointwise/BoundedMul.lean index 90d8349d3bf11..b98cbbde6ca66 100644 --- a/Mathlib/Data/Set/Pointwise/BoundedMul.lean +++ b/Mathlib/Data/Set/Pointwise/BoundedMul.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. KudryashovJ +Authors: Yury KudryashovJ -/ import Mathlib.Algebra.Order.Monoid.Defs import Mathlib.Data.Set.Pointwise.Basic diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Data/Set/Pointwise/Interval.lean index ad4c9218ade0c..06375bbd88b1d 100644 --- a/Mathlib/Data/Set/Pointwise/Interval.lean +++ b/Mathlib/Data/Set/Pointwise/Interval.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Patrick Massot +Authors: Yury Kudryashov, Patrick Massot -/ import Mathlib.Order.Interval.Set.UnorderedInterval import Mathlib.Algebra.Order.Interval.Set.Monoid diff --git a/Mathlib/Data/Sum/Basic.lean b/Mathlib/Data/Sum/Basic.lean index e099eb1892281..9265b7b297c2d 100644 --- a/Mathlib/Data/Sum/Basic.lean +++ b/Mathlib/Data/Sum/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro, Yury G. Kudryashov +Authors: Mario Carneiro, Yury Kudryashov -/ import Mathlib.Logic.Function.Basic import Mathlib.Tactic.MkIffOfInductiveProp diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index 7e51244947644..644d722bfefa4 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Analysis.SpecificLimits.Basic diff --git a/Mathlib/Dynamics/Minimal.lean b/Mathlib/Dynamics/Minimal.lean index cd3549dff4f79..394bd9cb13203 100644 --- a/Mathlib/Dynamics/Minimal.lean +++ b/Mathlib/Dynamics/Minimal.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.GroupTheory.GroupAction.Basic import Mathlib.Topology.Algebra.ConstMulAction diff --git a/Mathlib/Dynamics/PeriodicPts.lean b/Mathlib/Dynamics/PeriodicPts.lean index 95604ab07bdae..658cf28c6f47d 100644 --- a/Mathlib/Dynamics/PeriodicPts.lean +++ b/Mathlib/Dynamics/PeriodicPts.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Algebra.GroupPower.IterateHom import Mathlib.Algebra.Ring.Divisibility.Basic diff --git a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean index 744c70d6ccb26..5d868b2a052ee 100644 --- a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.Normed.Group.AddTorsor import Mathlib.Analysis.InnerProductSpace.Basic diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index 70c9c57e64c52..e4fa08c639677 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Geometry.Manifold.Algebra.Structures import Mathlib.Geometry.Manifold.BumpFunction diff --git a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean index f45a88c138f56..e40360732b4b1 100644 --- a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean +++ b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Geometry.Manifold.Diffeomorph import Mathlib.Geometry.Manifold.Instances.Real diff --git a/Mathlib/Lean/Exception.lean b/Mathlib/Lean/Exception.lean index b03baf1351cd1..9eb2a70deb8a5 100644 --- a/Mathlib/Lean/Exception.lean +++ b/Mathlib/Lean/Exception.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 E.W.Ayers. All rights reserved. +Copyright (c) 2022 Edward Ayers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: E.W.Ayers +Authors: Edward Ayers -/ import Lean.Exception diff --git a/Mathlib/Lean/Expr/Basic.lean b/Mathlib/Lean/Expr/Basic.lean index a21111425b25f..b7a76c6339391 100644 --- a/Mathlib/Lean/Expr/Basic.lean +++ b/Mathlib/Lean/Expr/Basic.lean @@ -2,7 +2,7 @@ Copyright (c) 2019 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Simon Hudon, Scott Morrison, Keeley Hoek, Robert Y. Lewis, -Floris van Doorn, E.W.Ayers, Arthur Paulino +Floris van Doorn, Edward Ayers, Arthur Paulino -/ import Lean.Meta.Tactic.Rewrite import Batteries.Lean.Expr diff --git a/Mathlib/Lean/Expr/ReplaceRec.lean b/Mathlib/Lean/Expr/ReplaceRec.lean index 938190bca1e77..80f1f08324276 100644 --- a/Mathlib/Lean/Expr/ReplaceRec.lean +++ b/Mathlib/Lean/Expr/ReplaceRec.lean @@ -2,7 +2,7 @@ Copyright (c) 2019 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Simon Hudon, Scott Morrison, Keeley Hoek, Robert Y. Lewis, -Floris van Doorn, E.W.Ayers +Floris van Doorn, Edward Ayers -/ import Lean.Expr import Mathlib.Util.MemoFix diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean index 110da1dca9c07..c8fd0d0b5ebb8 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.LinearAlgebra.AffineSpace.AffineMap import Mathlib.LinearAlgebra.GeneralLinearGroup diff --git a/Mathlib/LinearAlgebra/AffineSpace/Ordered.lean b/Mathlib/LinearAlgebra/AffineSpace/Ordered.lean index 0793bb07f1de3..7835a5c734a01 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Ordered.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Ordered.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Algebra.CharP.Invertible import Mathlib.Algebra.Order.Invertible diff --git a/Mathlib/LinearAlgebra/AffineSpace/Slope.lean b/Mathlib/LinearAlgebra/AffineSpace/Slope.lean index 2713830981c3c..d5601f89df9af 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Slope.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Slope.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.LinearAlgebra.AffineSpace.AffineMap import Mathlib.Tactic.FieldSimp diff --git a/Mathlib/MeasureTheory/Function/Floor.lean b/Mathlib/MeasureTheory/Function/Floor.lean index d22f8acb0f7c2..56148ef42f928 100644 --- a/Mathlib/MeasureTheory/Function/Floor.lean +++ b/Mathlib/MeasureTheory/Function/Floor.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.MeasureTheory.Constructions.BorelSpace.Order diff --git a/Mathlib/MeasureTheory/Group/Action.lean b/Mathlib/MeasureTheory/Group/Action.lean index 51654f035648f..cd1b2c959c282 100644 --- a/Mathlib/MeasureTheory/Group/Action.lean +++ b/Mathlib/MeasureTheory/Group/Action.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Dynamics.Ergodic.MeasurePreserving import Mathlib.Dynamics.Minimal diff --git a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean index ded2b790c5d06..181c922b7b868 100644 --- a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean +++ b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Alex Kontorovich, Heather Macbeth +Authors: Yury Kudryashov, Alex Kontorovich, Heather Macbeth -/ import Mathlib.MeasureTheory.Group.Action import Mathlib.MeasureTheory.Integral.SetIntegral diff --git a/Mathlib/MeasureTheory/Group/MeasurableEquiv.lean b/Mathlib/MeasureTheory/Group/MeasurableEquiv.lean index 992075b2fcbd2..95bd76b8d5572 100644 --- a/Mathlib/MeasureTheory/Group/MeasurableEquiv.lean +++ b/Mathlib/MeasureTheory/Group/MeasurableEquiv.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.MeasureTheory.Group.Arithmetic diff --git a/Mathlib/MeasureTheory/Group/Pointwise.lean b/Mathlib/MeasureTheory/Group/Pointwise.lean index 108a2490b6595..84db89e4019cb 100644 --- a/Mathlib/MeasureTheory/Group/Pointwise.lean +++ b/Mathlib/MeasureTheory/Group/Pointwise.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Alex J. Best +Authors: Yury Kudryashov, Alex J. Best -/ import Mathlib.MeasureTheory.Group.Arithmetic diff --git a/Mathlib/MeasureTheory/Integral/Average.lean b/Mathlib/MeasureTheory/Integral/Average.lean index a779b9ccf468c..8ec1ba5cb9b64 100644 --- a/Mathlib/MeasureTheory/Integral/Average.lean +++ b/Mathlib/MeasureTheory/Integral/Average.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Yaël Dillies +Authors: Yury Kudryashov, Yaël Dillies -/ import Mathlib.MeasureTheory.Integral.SetIntegral diff --git a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean index 8de4e345f0f17..0ab09a7e17c02 100644 --- a/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean +++ b/Mathlib/MeasureTheory/Integral/FundThmCalculus.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Patrick Massot, Sébastien Gouëzel +Authors: Yury Kudryashov, Patrick Massot, Sébastien Gouëzel -/ import Mathlib.Analysis.Calculus.FDeriv.Measurable import Mathlib.Analysis.Calculus.Deriv.Comp diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index ec7fbc0202d17..5ca5b383e51ad 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Patrick Massot, Sébastien Gouëzel +Authors: Yury Kudryashov, Patrick Massot, Sébastien Gouëzel -/ import Mathlib.Order.Interval.Set.Disjoint import Mathlib.MeasureTheory.Integral.SetIntegral diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 67325d6b4ba98..3469ca0689e07 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Felix Weilacher. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Felix Weilacher, Yury G. Kudryashov, Rémy Degenne +Authors: Felix Weilacher, Yury Kudryashov, Rémy Degenne -/ import Mathlib.MeasureTheory.MeasurableSpace.Embedding import Mathlib.Data.Set.MemPartition diff --git a/Mathlib/MeasureTheory/Measure/Regular.lean b/Mathlib/MeasureTheory/Measure/Regular.lean index 72a01a4c40abd..3ba9dadd512e0 100644 --- a/Mathlib/MeasureTheory/Measure/Regular.lean +++ b/Mathlib/MeasureTheory/Measure/Regular.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel, Floris Van Doorn, Yury Kudryashov +Authors: Sébastien Gouëzel, Floris van Doorn, Yury Kudryashov -/ import Mathlib.Topology.MetricSpace.HausdorffDistance import Mathlib.MeasureTheory.Constructions.BorelSpace.Order diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index 2075c97a58718..7dec6765c845c 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Alex Best, Riccardo Brasca, Eric Rodriguez +Authors: Alex J. Best, Riccardo Brasca, Eric Rodriguez -/ import Mathlib.Data.PNat.Prime import Mathlib.Algebra.IsPrimePow diff --git a/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean b/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean index 1343850e11960..77b9e83e7512d 100644 --- a/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean +++ b/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.SpecialFunctions.Pow.Asymptotics import Mathlib.NumberTheory.Liouville.Basic diff --git a/Mathlib/NumberTheory/Liouville/Measure.lean b/Mathlib/NumberTheory/Liouville/Measure.lean index c571be5f3871c..c6bd2d0a641c4 100644 --- a/Mathlib/NumberTheory/Liouville/Measure.lean +++ b/Mathlib/NumberTheory/Liouville/Measure.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.MeasureTheory.Measure.Lebesgue.Basic import Mathlib.NumberTheory.Liouville.Residual diff --git a/Mathlib/NumberTheory/Ostrowski.lean b/Mathlib/NumberTheory/Ostrowski.lean index 602814e476419..bfdd99846ba8c 100644 --- a/Mathlib/NumberTheory/Ostrowski.lean +++ b/Mathlib/NumberTheory/Ostrowski.lean @@ -2,7 +2,7 @@ Copyright (c) 2024 María Inés de Frutos-Fernández. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata, Fabrizio Barroero, Laura Capuano, Nirvana Coppola, -María Inés de Frutos Fernández, Sam van Gool, Silvain Rideau-Kikuchi, Amos Turchet, +María Inés de Frutos-Fernández, Sam van Gool, Silvain Rideau-Kikuchi, Amos Turchet, Francesco Veneziano -/ diff --git a/Mathlib/Order/Birkhoff.lean b/Mathlib/Order/Birkhoff.lean index 35622b5b1f465..03f815acdee3c 100644 --- a/Mathlib/Order/Birkhoff.lean +++ b/Mathlib/Order/Birkhoff.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies, Filipo A. E. Nuccio, Sam van Gool +Authors: Yaël Dillies, Filippo A. E. Nuccio, Sam van Gool -/ import Mathlib.Order.Interval.Finset.Basic import Mathlib.Data.Fintype.Order diff --git a/Mathlib/Order/Filter/CountableInter.lean b/Mathlib/Order/Filter/CountableInter.lean index 133d492fa3182..ea7172ae2ea14 100644 --- a/Mathlib/Order/Filter/CountableInter.lean +++ b/Mathlib/Order/Filter/CountableInter.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Order.Filter.Curry import Mathlib.Data.Set.Countable diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index 8b35547927b1e..ba315579aa20a 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Abhimanyu Pallavi Sudhir +Authors: Yury Kudryashov, Abhimanyu Pallavi Sudhir -/ import Mathlib.Order.Filter.Basic import Mathlib.Algebra.Module.Pi diff --git a/Mathlib/Order/Filter/Germ/OrderedMonoid.lean b/Mathlib/Order/Filter/Germ/OrderedMonoid.lean index 4c235e9d1e4ba..f2d7d62077b7e 100644 --- a/Mathlib/Order/Filter/Germ/OrderedMonoid.lean +++ b/Mathlib/Order/Filter/Germ/OrderedMonoid.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Abhimanyu Pallavi Sudhir +Authors: Yury Kudryashov, Abhimanyu Pallavi Sudhir -/ import Mathlib.Algebra.Module.Pi import Mathlib.Algebra.Order.Monoid.Canonical.Defs diff --git a/Mathlib/Order/Filter/Interval.lean b/Mathlib/Order/Filter/Interval.lean index fc61850b49472..0bb96d5502cad 100644 --- a/Mathlib/Order/Filter/Interval.lean +++ b/Mathlib/Order/Filter/Interval.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Order.Interval.Set.OrdConnected import Mathlib.Order.Filter.SmallSets diff --git a/Mathlib/Order/Filter/Pi.lean b/Mathlib/Order/Filter/Pi.lean index 375092c5e7acd..147840d30d026 100644 --- a/Mathlib/Order/Filter/Pi.lean +++ b/Mathlib/Order/Filter/Pi.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Alex Kontorovich +Authors: Yury Kudryashov, Alex Kontorovich -/ import Mathlib.Order.Filter.Bases diff --git a/Mathlib/Order/Filter/Prod.lean b/Mathlib/Order/Filter/Prod.lean index e8b6061a038ad..0202f4d85b27e 100644 --- a/Mathlib/Order/Filter/Prod.lean +++ b/Mathlib/Order/Filter/Prod.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johanes Hölzl, Patrick Massot, Yury Kudryashov, Kevin Wilson, Heather Macbeth +Authors: Johannes Hölzl, Patrick Massot, Yury Kudryashov, Kevin H. Wilson, Heather Macbeth -/ import Mathlib.Order.Filter.Basic diff --git a/Mathlib/Order/Interval/Set/OrdConnected.lean b/Mathlib/Order/Interval/Set/OrdConnected.lean index 42ed4e013b795..85255d6b9f744 100644 --- a/Mathlib/Order/Interval/Set/OrdConnected.lean +++ b/Mathlib/Order/Interval/Set/OrdConnected.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Order.Interval.Set.OrderEmbedding import Mathlib.Order.Antichain diff --git a/Mathlib/Order/Interval/Set/ProjIcc.lean b/Mathlib/Order/Interval/Set/ProjIcc.lean index 66e0520f27cf0..c6956311951c5 100644 --- a/Mathlib/Order/Interval/Set/ProjIcc.lean +++ b/Mathlib/Order/Interval/Set/ProjIcc.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Patrick Massot +Authors: Yury Kudryashov, Patrick Massot -/ import Mathlib.Data.Set.Function import Mathlib.Order.Interval.Set.OrdConnected diff --git a/Mathlib/Order/Interval/Set/WithBotTop.lean b/Mathlib/Order/Interval/Set/WithBotTop.lean index 67f2d8093f357..d5f2e62b5a1b9 100644 --- a/Mathlib/Order/Interval/Set/WithBotTop.lean +++ b/Mathlib/Order/Interval/Set/WithBotTop.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Data.Set.Image import Mathlib.Order.Interval.Set.Basic diff --git a/Mathlib/Order/Iterate.lean b/Mathlib/Order/Iterate.lean index a6bd7612424ec..6e475b270ddfb 100644 --- a/Mathlib/Order/Iterate.lean +++ b/Mathlib/Order/Iterate.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Logic.Function.Iterate import Mathlib.Order.Monotone.Basic diff --git a/Mathlib/Order/OrdContinuous.lean b/Mathlib/Order/OrdContinuous.lean index 73f54a8267dba..1af016fa24c89 100644 --- a/Mathlib/Order/OrdContinuous.lean +++ b/Mathlib/Order/OrdContinuous.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Johannes Hölzl +Authors: Yury Kudryashov, Johannes Hölzl -/ import Mathlib.Order.ConditionallyCompleteLattice.Basic import Mathlib.Order.RelIso.Basic diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index b2fd3766945b1..1a1ad02cbedff 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad, Mario Carneiro, Yury G. Kudryashov +Authors: Jeremy Avigad, Mario Carneiro, Yury Kudryashov -/ import Mathlib.Data.Nat.Defs import Mathlib.Logic.IsEmpty diff --git a/Mathlib/Order/SemiconjSup.lean b/Mathlib/Order/SemiconjSup.lean index 264dada072538..1a7a4802f052c 100644 --- a/Mathlib/Order/SemiconjSup.lean +++ b/Mathlib/Order/SemiconjSup.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Logic.Function.Conjugate diff --git a/Mathlib/RingTheory/PowerSeries/WellKnown.lean b/Mathlib/RingTheory/PowerSeries/WellKnown.lean index 6ef5b6366a493..8b3ce3c34a0ce 100644 --- a/Mathlib/RingTheory/PowerSeries/WellKnown.lean +++ b/Mathlib/RingTheory/PowerSeries/WellKnown.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Algebra.Algebra.Rat import Mathlib.Algebra.BigOperators.NatAntidiagonal diff --git a/Mathlib/SetTheory/Cardinal/ENat.lean b/Mathlib/SetTheory/Cardinal/ENat.lean index a8b7b9ddf45c9..07957687d320d 100644 --- a/Mathlib/SetTheory/Cardinal/ENat.lean +++ b/Mathlib/SetTheory/Cardinal/ENat.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2024 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Algebra.Order.Hom.Ring import Mathlib.Data.ENat.Basic diff --git a/Mathlib/Tactic/FunProp/Attr.lean b/Mathlib/Tactic/FunProp/Attr.lean index 28ca1408f16dd..e3fa55a39e7af 100644 --- a/Mathlib/Tactic/FunProp/Attr.lean +++ b/Mathlib/Tactic/FunProp/Attr.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Tomas Skrivan. All rights reserved. +Copyright (c) 2024 Tomáš Skřivan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Tomas Skrivan +Authors: Tomáš Skřivan -/ import Lean diff --git a/Mathlib/Tactic/FunProp/Core.lean b/Mathlib/Tactic/FunProp/Core.lean index 0d1cc6fed6612..74e9c8a628be3 100644 --- a/Mathlib/Tactic/FunProp/Core.lean +++ b/Mathlib/Tactic/FunProp/Core.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Tomas Skrivan. All rights reserved. +Copyright (c) 2024 Tomáš Skřivan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Tomas Skrivan +Authors: Tomáš Skřivan -/ import Lean import Mathlib.Tactic.FunProp.Theorems diff --git a/Mathlib/Tactic/FunProp/Decl.lean b/Mathlib/Tactic/FunProp/Decl.lean index f1e5aef0744f2..690077cb3a1b2 100644 --- a/Mathlib/Tactic/FunProp/Decl.lean +++ b/Mathlib/Tactic/FunProp/Decl.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Tomas Skrivan. All rights reserved. +Copyright (c) 2024 Tomáš Skřivan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Tomas Skrivan +Authors: Tomáš Skřivan -/ import Lean diff --git a/Mathlib/Tactic/FunProp/Elab.lean b/Mathlib/Tactic/FunProp/Elab.lean index fd6a7f5a01b6b..1500bc79931fa 100644 --- a/Mathlib/Tactic/FunProp/Elab.lean +++ b/Mathlib/Tactic/FunProp/Elab.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Tomas Skrivan. All rights reserved. +Copyright (c) 2024 Tomáš Skřivan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Tomas Skrivan +Authors: Tomáš Skřivan -/ import Lean import Mathlib.Tactic.FunProp.Core diff --git a/Mathlib/Tactic/FunProp/FunctionData.lean b/Mathlib/Tactic/FunProp/FunctionData.lean index 9c4ef56b7ec80..26a9f42315b07 100644 --- a/Mathlib/Tactic/FunProp/FunctionData.lean +++ b/Mathlib/Tactic/FunProp/FunctionData.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Tomas Skrivan. All rights reserved. +Copyright (c) 2024 Tomáš Skřivan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Tomas Skrivan +Authors: Tomáš Skřivan -/ import Lean import Qq diff --git a/Mathlib/Tactic/FunProp/Mor.lean b/Mathlib/Tactic/FunProp/Mor.lean index 212b344ed3c5b..afe269094f0de 100644 --- a/Mathlib/Tactic/FunProp/Mor.lean +++ b/Mathlib/Tactic/FunProp/Mor.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Tomas Skrivan. All rights reserved. +Copyright (c) 2024 Tomáš Skřivan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Tomas Skrivan +Authors: Tomáš Skřivan -/ import Lean import Mathlib.Data.FunLike.Basic diff --git a/Mathlib/Tactic/FunProp/Theorems.lean b/Mathlib/Tactic/FunProp/Theorems.lean index e5cdb0566e323..9993eef51a215 100644 --- a/Mathlib/Tactic/FunProp/Theorems.lean +++ b/Mathlib/Tactic/FunProp/Theorems.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Tomas Skrivan. All rights reserved. +Copyright (c) 2024 Tomáš Skřivan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Tomas Skrivan +Authors: Tomáš Skřivan -/ import Mathlib.Tactic.FunProp.Decl import Mathlib.Tactic.FunProp.Types diff --git a/Mathlib/Tactic/FunProp/ToBatteries.lean b/Mathlib/Tactic/FunProp/ToBatteries.lean index d0b3799d41870..5cf39afb38db6 100644 --- a/Mathlib/Tactic/FunProp/ToBatteries.lean +++ b/Mathlib/Tactic/FunProp/ToBatteries.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Tomas Skrivan. All rights reserved. +Copyright (c) 2024 Tomáš Skřivan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Tomas Skrivan +Authors: Tomáš Skřivan -/ import Lean diff --git a/Mathlib/Tactic/FunProp/Types.lean b/Mathlib/Tactic/FunProp/Types.lean index 4ed8dea3ee855..00daaab6c8b4d 100644 --- a/Mathlib/Tactic/FunProp/Types.lean +++ b/Mathlib/Tactic/FunProp/Types.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2024 Tomas Skrivan. All rights reserved. +Copyright (c) 2024 Tomáš Skřivan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Tomas Skrivan +Authors: Tomáš Skřivan -/ import Mathlib.Tactic.FunProp.FunctionData import Batteries.Data.RBMap.Basic diff --git a/Mathlib/Tactic/Polyrith.lean b/Mathlib/Tactic/Polyrith.lean index 130d9752819a9..c2a538579b186 100644 --- a/Mathlib/Tactic/Polyrith.lean +++ b/Mathlib/Tactic/Polyrith.lean @@ -55,7 +55,7 @@ remember to force recompilation of any files that call `polyrith`. * See the book [*Ideals, Varieties, and Algorithms*][coxlittleOshea1997] by David Cox, John Little, and Donal O'Shea for the background theory on Groebner bases * This code was heavily inspired by the code for the tactic `linarith`, which was written by - Robert Lewis, who advised me on this project as part of a Computer Science independent study + Robert Y. Lewis, who advised me on this project as part of a Computer Science independent study at Brown University. -/ diff --git a/Mathlib/Tactic/TypeStar.lean b/Mathlib/Tactic/TypeStar.lean index 5decbc6e5e4ff..d020ece4a2492 100644 --- a/Mathlib/Tactic/TypeStar.lean +++ b/Mathlib/Tactic/TypeStar.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2023 Matthew Ballard. All rights reserved. +Copyright (c) 2023 Matthew Robert Ballard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Matthew Ballard +Authors: Kyle Miller -/ import Lean.Elab.Term diff --git a/Mathlib/Topology/Algebra/GroupWithZero.lean b/Mathlib/Topology/Algebra/GroupWithZero.lean index f67dd259425c7..a363dc902e11b 100644 --- a/Mathlib/Topology/Algebra/GroupWithZero.lean +++ b/Mathlib/Topology/Algebra/GroupWithZero.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2020 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Pi.Lemmas import Mathlib.Topology.Algebra.Monoid diff --git a/Mathlib/Topology/Algebra/Order/Archimedean.lean b/Mathlib/Topology/Algebra/Order/Archimedean.lean index a7395faf8d693..5d3d42abf04f6 100644 --- a/Mathlib/Topology/Algebra/Order/Archimedean.lean +++ b/Mathlib/Topology/Algebra/Order/Archimedean.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.GroupTheory.Archimedean import Mathlib.Topology.Order.Basic diff --git a/Mathlib/Topology/Algebra/UniformMulAction.lean b/Mathlib/Topology/Algebra/UniformMulAction.lean index b729b368c569b..4c23eda8c9359 100644 --- a/Mathlib/Topology/Algebra/UniformMulAction.lean +++ b/Mathlib/Topology/Algebra/UniformMulAction.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Topology.Algebra.UniformGroup import Mathlib.Topology.UniformSpace.Completion diff --git a/Mathlib/Topology/Bornology/Constructions.lean b/Mathlib/Topology/Bornology/Constructions.lean index bf35360d4243a..e53a16415f67c 100644 --- a/Mathlib/Topology/Bornology/Constructions.lean +++ b/Mathlib/Topology/Bornology/Constructions.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Topology.Bornology.Basic diff --git a/Mathlib/Topology/EMetricSpace/Paracompact.lean b/Mathlib/Topology/EMetricSpace/Paracompact.lean index cedc03fd28486..15a258752ae20 100644 --- a/Mathlib/Topology/EMetricSpace/Paracompact.lean +++ b/Mathlib/Topology/EMetricSpace/Paracompact.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 202 Yury G. Kudryashov. All rights reserved. +Copyright (c) 202 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.SetTheory.Ordinal.Basic import Mathlib.Tactic.GCongr diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index 8f614e06f86e5..d2089d77aed05 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Andrew Yang, Yury G. Kudryashov +Authors: Andrew Yang, Yury Kudryashov -/ import Mathlib.Tactic.TFAE import Mathlib.Topology.ContinuousOn diff --git a/Mathlib/Topology/Instances/Irrational.lean b/Mathlib/Topology/Instances/Irrational.lean index c354b27701bc5..fc7c1f606e691 100644 --- a/Mathlib/Topology/Instances/Irrational.lean +++ b/Mathlib/Topology/Instances/Irrational.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Data.Real.Irrational import Mathlib.Data.Rat.Encodable diff --git a/Mathlib/Topology/MetricSpace/Holder.lean b/Mathlib/Topology/MetricSpace/Holder.lean index 55c2b340269f9..c60629e1cc045 100644 --- a/Mathlib/Topology/MetricSpace/Holder.lean +++ b/Mathlib/Topology/MetricSpace/Holder.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Topology.MetricSpace.Lipschitz import Mathlib.Analysis.SpecialFunctions.Pow.Continuity diff --git a/Mathlib/Topology/MetricSpace/ShrinkingLemma.lean b/Mathlib/Topology/MetricSpace/ShrinkingLemma.lean index b562744a70f69..058f2115e3541 100644 --- a/Mathlib/Topology/MetricSpace/ShrinkingLemma.lean +++ b/Mathlib/Topology/MetricSpace/ShrinkingLemma.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Topology.EMetricSpace.Paracompact import Mathlib.Topology.MetricSpace.Basic diff --git a/Mathlib/Topology/Order/Category/FrameAdjunction.lean b/Mathlib/Topology/Order/Category/FrameAdjunction.lean index 9d5db86970bb8..66dd40b808326 100644 --- a/Mathlib/Topology/Order/Category/FrameAdjunction.lean +++ b/Mathlib/Topology/Order/Category/FrameAdjunction.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2023 Anne Baanen, Sam v. Gool, Leo Mayer, Brendan S. Murphy. All rights reserved. +Copyright (c) 2023 Anne Baanen, Sam van Gool, Leo Mayer, Brendan Murphy. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Anne Baanen, Sam v. Gool, Leo Mayer, Brendan S. Murphy +Authors: Anne Baanen, Sam van Gool, Leo Mayer, Brendan Murphy -/ import Mathlib.Topology.Category.Locale diff --git a/Mathlib/Topology/Order/ExtrClosure.lean b/Mathlib/Topology/Order/ExtrClosure.lean index 9104c5e4d0bac..d1b2049789057 100644 --- a/Mathlib/Topology/Order/ExtrClosure.lean +++ b/Mathlib/Topology/Order/ExtrClosure.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Topology.Order.OrderClosed import Mathlib.Topology.Order.LocalExtr diff --git a/Mathlib/Topology/Order/IntermediateValue.lean b/Mathlib/Topology/Order/IntermediateValue.lean index 5ba397119d8e2..00431baf1a323 100644 --- a/Mathlib/Topology/Order/IntermediateValue.lean +++ b/Mathlib/Topology/Order/IntermediateValue.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Alistair Tucker, Wen Yang +Authors: Yury Kudryashov, Alistair Tucker, Wen Yang -/ import Mathlib.Order.Interval.Set.Image import Mathlib.Order.CompleteLatticeIntervals diff --git a/Mathlib/Topology/Order/MonotoneContinuity.lean b/Mathlib/Topology/Order/MonotoneContinuity.lean index f05c7f61f828b..9c1a04ed46bec 100644 --- a/Mathlib/Topology/Order/MonotoneContinuity.lean +++ b/Mathlib/Topology/Order/MonotoneContinuity.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov, Heather Macbeth +Authors: Yury Kudryashov, Heather Macbeth -/ import Mathlib.Topology.Homeomorph import Mathlib.Topology.Order.LeftRightNhds diff --git a/Mathlib/Topology/TietzeExtension.lean b/Mathlib/Topology/TietzeExtension.lean index ebbe95a8f4538..55ec301b782ca 100644 --- a/Mathlib/Topology/TietzeExtension.lean +++ b/Mathlib/Topology/TietzeExtension.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.SpecificLimits.Basic import Mathlib.Order.Interval.Set.IsoIoo diff --git a/Mathlib/Topology/UrysohnsLemma.lean b/Mathlib/Topology/UrysohnsLemma.lean index 1dff9ac7af4ab..5226cd1fb3438 100644 --- a/Mathlib/Topology/UrysohnsLemma.lean +++ b/Mathlib/Topology/UrysohnsLemma.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury G. Kudryashov +Authors: Yury Kudryashov -/ import Mathlib.Analysis.NormedSpace.AddTorsor import Mathlib.LinearAlgebra.AffineSpace.Ordered diff --git a/Mathlib/Topology/VectorBundle/Basic.lean b/Mathlib/Topology/VectorBundle/Basic.lean index d5f3015592180..c6c5091999c82 100644 --- a/Mathlib/Topology/VectorBundle/Basic.lean +++ b/Mathlib/Topology/VectorBundle/Basic.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Nicolò Cavalleri. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Nicolò Cavalleri, Sebastien Gouezel, Heather Macbeth, Patrick Massot, Floris van Doorn +Authors: Nicolò Cavalleri, Sébastien Gouëzel, Heather Macbeth, Patrick Massot, Floris van Doorn -/ import Mathlib.Analysis.Normed.Operator.BoundedLinearMaps import Mathlib.Topology.FiberBundle.Basic diff --git a/Mathlib/Util/MemoFix.lean b/Mathlib/Util/MemoFix.lean index 9a5a501ae00c4..ab8344cbc6c6c 100644 --- a/Mathlib/Util/MemoFix.lean +++ b/Mathlib/Util/MemoFix.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Gabriel Ebner, E.W.Ayers +Authors: Gabriel Ebner, Edward Ayers -/ import Lean.Data.HashMap diff --git a/test/Set.lean b/test/Set.lean index 27b4dd3eed024..8a053f5e287ac 100644 --- a/test/Set.lean +++ b/test/Set.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Ian Benway. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Ian Benway. +Authors: Ian Benway -/ import Mathlib.Tactic.Set From 5bb47cdc9ef1d4807c3f148a620f71924c650086 Mon Sep 17 00:00:00 2001 From: Rida Hamadani Date: Sun, 11 Aug 2024 20:16:27 +0000 Subject: [PATCH 176/359] feat(Combinatorics/SimpleGraph): bot is not connected (#15675) --- Mathlib/Combinatorics/SimpleGraph/Path.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index 0f5aaeef7169e..1731440ff45a7 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -719,6 +719,17 @@ theorem Preconnected.map {G : SimpleGraph V} {H : SimpleGraph V'} (f : G →g H) protected lemma Preconnected.mono {G G' : SimpleGraph V} (h : G ≤ G') (hG : G.Preconnected) : G'.Preconnected := fun u v => (hG u v).mono h +lemma bot_preconnected_iff_subsingleton : (⊥ : SimpleGraph V).Preconnected ↔ Subsingleton V := by + refine ⟨fun h ↦ ?_, fun h ↦ by simpa [subsingleton_iff, ← reachable_bot] using h⟩ + contrapose h + simp [nontrivial_iff.mp <| not_subsingleton_iff_nontrivial.mp h, Preconnected, reachable_bot, h] + +lemma bot_preconnected [Subsingleton V] : (⊥ : SimpleGraph V).Preconnected := + bot_preconnected_iff_subsingleton.mpr ‹_› + +lemma bot_not_preconnected [Nontrivial V] : ¬(⊥ : SimpleGraph V).Preconnected := + bot_preconnected_iff_subsingleton.not.mpr <| not_subsingleton_iff_nontrivial.mpr ‹_› + lemma top_preconnected : (⊤ : SimpleGraph V).Preconnected := fun x y => by if h : x = y then rw [h] else exact Adj.reachable h @@ -758,6 +769,9 @@ protected lemma Connected.mono {G G' : SimpleGraph V} (h : G ≤ G') preconnected := hG.preconnected.mono h nonempty := hG.nonempty +lemma bot_not_connected [Nontrivial V] : ¬(⊥ : SimpleGraph V).Connected := by + simp [bot_not_preconnected, connected_iff, ‹_›] + lemma top_connected [Nonempty V] : (⊤ : SimpleGraph V).Connected where preconnected := top_preconnected From 85e0f4e9f798d0c0ff4e9a96430403182d4bab55 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 11 Aug 2024 21:00:25 +0000 Subject: [PATCH 177/359] chore: final cleanup of Mathport (#15630) --- Mathlib.lean | 3 +- Mathlib/Data/PSigma/Order.lean | 2 +- Mathlib/Data/Quot.lean | 2 +- Mathlib/Data/Set/Notation.lean | 2 +- Mathlib/Data/Sigma/Order.lean | 2 +- Mathlib/Data/Vector3.lean | 2 +- Mathlib/Mathport/README.md | 19 +- Mathlib/Mathport/Syntax.lean | 281 ------------------ Mathlib/Order/SetNotation.lean | 2 +- .../Notation.lean => Util/Notation3.lean} | 0 test/notation3.lean | 2 +- 11 files changed, 16 insertions(+), 301 deletions(-) delete mode 100644 Mathlib/Mathport/Syntax.lean rename Mathlib/{Mathport/Notation.lean => Util/Notation3.lean} (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 3290f9d7ca702..7b99ae598c9d8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3077,8 +3077,6 @@ import Mathlib.Logic.Small.Ring import Mathlib.Logic.Small.Set import Mathlib.Logic.Unique import Mathlib.Logic.UnivLE -import Mathlib.Mathport.Notation -import Mathlib.Mathport.Syntax import Mathlib.MeasureTheory.Category.MeasCat import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic import Mathlib.MeasureTheory.Constructions.BorelSpace.Complex @@ -4643,6 +4641,7 @@ import Mathlib.Util.GetAllModules import Mathlib.Util.IncludeStr import Mathlib.Util.LongNames import Mathlib.Util.MemoFix +import Mathlib.Util.Notation3 import Mathlib.Util.Qq import Mathlib.Util.SleepHeartbeats import Mathlib.Util.Superscript diff --git a/Mathlib/Data/PSigma/Order.lean b/Mathlib/Data/PSigma/Order.lean index 89c2d715e11a6..097dd92225da5 100644 --- a/Mathlib/Data/PSigma/Order.lean +++ b/Mathlib/Data/PSigma/Order.lean @@ -5,7 +5,7 @@ Authors: Scott Morrison, Minchao Wu -/ import Mathlib.Data.Sigma.Lex import Mathlib.Order.BoundedOrder -import Mathlib.Mathport.Notation +import Mathlib.Util.Notation3 import Init.NotationExtra import Mathlib.Data.Sigma.Basic diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index ed7af7bc3672e..c78bf64552825 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -6,7 +6,7 @@ Authors: Johannes Hölzl import Mathlib.Init.Data.Quot import Mathlib.Logic.Relator import Mathlib.Logic.Unique -import Mathlib.Mathport.Notation +import Mathlib.Util.Notation3 /-! # Quotient types diff --git a/Mathlib/Data/Set/Notation.lean b/Mathlib/Data/Set/Notation.lean index cf31a4e683cff..fd312c88c90a2 100644 --- a/Mathlib/Data/Set/Notation.lean +++ b/Mathlib/Data/Set/Notation.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Peter Nelson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Peter Nelson -/ -import Mathlib.Mathport.Notation +import Mathlib.Util.Notation3 import Mathlib.Lean.Expr.ExtraRecognizers /-! diff --git a/Mathlib/Data/Sigma/Order.lean b/Mathlib/Data/Sigma/Order.lean index f5acd3a0b8adc..f7d007a46e857 100644 --- a/Mathlib/Data/Sigma/Order.lean +++ b/Mathlib/Data/Sigma/Order.lean @@ -5,7 +5,7 @@ Authors: Yaël Dillies -/ import Mathlib.Data.Sigma.Lex import Mathlib.Order.BoundedOrder -import Mathlib.Mathport.Notation +import Mathlib.Util.Notation3 import Mathlib.Data.Sigma.Basic /-! diff --git a/Mathlib/Data/Vector3.lean b/Mathlib/Data/Vector3.lean index 26d0da1c9b371..4e0562d15422e 100644 --- a/Mathlib/Data/Vector3.lean +++ b/Mathlib/Data/Vector3.lean @@ -5,7 +5,7 @@ Authors: Mario Carneiro -/ import Mathlib.Data.Fin.Fin2 import Mathlib.Init.Logic -import Mathlib.Mathport.Notation +import Mathlib.Util.Notation3 import Mathlib.Tactic.TypeStar /-! diff --git a/Mathlib/Mathport/README.md b/Mathlib/Mathport/README.md index 4a25af5bc8789..59db89b9c6665 100644 --- a/Mathlib/Mathport/README.md +++ b/Mathlib/Mathport/README.md @@ -1,16 +1,13 @@ Mathport prelude === -This directory contains instructions for `mathport`, +This directory used to contain instructions for `mathport`, helping it to translate `mathlib` and align declarations and tactics with `mathlib4`. -(These files were formerly part of `mathport`, in the directory `Mathport/Prelude/`.) -`Syntax.lean` -: Contains unimplemented stubs of tactics which need to be migrated from Lean3 to Lean4. - When porting tactics, you can move the relevant stubs to a new file and - use them as a starting point. - Please make sure this file stays in sync with new tactic implementations - (and in particular that the syntax is not defined twice). - Please preserve the syntax of existing mathlib tactics, - so that there are no unnecessary parse errors in the source files generated by `synport`. - It is fine to fail with an error message for unimplemented features for now. +If you still need to use mathport, you will need to move back to the `v3-eol` tag of mathlib, +which still contains mathport infrastructure, and migrate your project to that tag, +and then upgrade the rest of the way. + +(That tag also contains the final form of `Syntax.lean`, +which implicitly contained a TODO list of unported tactics from Lean 3, +which may still be of some interest.) diff --git a/Mathlib/Mathport/Syntax.lean b/Mathlib/Mathport/Syntax.lean deleted file mode 100644 index 247469d916e96..0000000000000 --- a/Mathlib/Mathport/Syntax.lean +++ /dev/null @@ -1,281 +0,0 @@ -/- -Copyright (c) 2021 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro --/ -import Lean.Elab.Command -import Lean.Elab.Quotation -import Batteries.Tactic.Where -import Mathlib.Data.Matrix.Notation -import Mathlib.Logic.Equiv.PartialEquiv -import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic -import Mathlib.Order.Filter.Basic -import Mathlib.RingTheory.WittVector.Basic -import Mathlib.RingTheory.WittVector.IsPoly -import Mathlib.SetTheory.Game.PGame -import Mathlib.Tactic.Abel -import Mathlib.Tactic.ApplyCongr -import Mathlib.Tactic.ApplyFun -import Mathlib.Tactic.ApplyWith -import Mathlib.Tactic.ByContra -import Mathlib.Tactic.CC -import Mathlib.Tactic.CancelDenoms -import Mathlib.Tactic.Cases -import Mathlib.Tactic.CasesM -import Mathlib.Tactic.CategoryTheory.Coherence -import Mathlib.Tactic.CategoryTheory.Elementwise -import Mathlib.Tactic.CategoryTheory.Slice -import Mathlib.Tactic.Choose -import Mathlib.Tactic.Clean -import Mathlib.Tactic.Clear_ -import Mathlib.Tactic.ClearExclamation -import Mathlib.Tactic.ClearExcept -import Mathlib.Tactic.Constructor -import Mathlib.Tactic.CongrM -import Mathlib.Tactic.Continuity -import Mathlib.Tactic.Contrapose -import Mathlib.Tactic.Conv -import Mathlib.Tactic.Convert -import Mathlib.Tactic.Core -import Mathlib.Tactic.DefEqTransformations -import Mathlib.Tactic.ExtractGoal -import Mathlib.Tactic.ExistsI -import Mathlib.Tactic.FieldSimp -import Mathlib.Tactic.FinCases -import Mathlib.Tactic.Find -import Mathlib.Tactic.GeneralizeProofs -import Mathlib.Tactic.Group -import Mathlib.Tactic.GuardHypNums -import Mathlib.Tactic.Hint -import Mathlib.Tactic.ITauto -import Mathlib.Tactic.InferParam -import Mathlib.Tactic.IntervalCases -import Mathlib.Tactic.Inhabit -import Mathlib.Tactic.IrreducibleDef -import Mathlib.Tactic.Lift -import Mathlib.Tactic.Linarith -import Mathlib.Tactic.LinearCombination -import Mathlib.Tactic.Measurability -import Mathlib.Tactic.MkIffOfInductiveProp -import Mathlib.Tactic.ModCases -import Mathlib.Tactic.Monotonicity -import Mathlib.Tactic.NoncommRing -import Mathlib.Tactic.Nontriviality -import Mathlib.Tactic.NormNum -import Mathlib.Tactic.NthRewrite -import Mathlib.Tactic.Polyrith -import Mathlib.Tactic.Positivity -import Mathlib.Tactic.PushNeg -import Mathlib.Tactic.Qify -import Mathlib.Tactic.Recover -import Mathlib.Tactic.Relation.Trans -import Mathlib.Tactic.Rename -import Mathlib.Tactic.RenameBVar -import Mathlib.Tactic.Replace -import Mathlib.Tactic.Ring -import Mathlib.Tactic.RSuffices -import Mathlib.Tactic.ScopedNS -import Mathlib.Tactic.Set -import Mathlib.Tactic.SimpIntro -import Mathlib.Tactic.SimpRw -import Mathlib.Tactic.Simps.Basic -import Mathlib.Tactic.SplitIfs -import Mathlib.Tactic.Substs -import Mathlib.Tactic.SwapVar -import Mathlib.Tactic.Tauto -import Mathlib.Tactic.TFAE -import Mathlib.Tactic.Trace -import Mathlib.Tactic.TypeCheck -import Mathlib.Tactic.Use -import Mathlib.Tactic.WLOG -import Mathlib.Tactic.Zify -import Mathlib.Util.WithWeakNamespace -import Mathlib.Mathport.Notation - --- To fix upstream: --- * bracketedExplicitBinders doesn't support optional types - -/-! -This file defines all the tactics that are required by mathport. Most of the `syntax` declarations -in this file (as opposed to the imported files) are not defined anywhere and effectively form the -TODO list before we can port mathlib to lean 4 for real. - -For tactic writers: I (Mario) have put a comment before each syntax declaration to represent the -estimated difficulty of writing the tactic. The key is as follows: - -* `E`: Easy. It's a simple macro in terms of existing things, - or an elab tactic for which we have many similar examples. Example: `left` -* `M`: Medium. An elab tactic, not too hard, perhaps a 100-200 lines file. Example: `have` -* `N`: Possibly requires new mechanisms in lean 4, some investigation required -* `B`: Hard, because it is a big and complicated tactic -* `S`: Possibly easy, because we can just stub it out or replace with something else -* `?`: uncategorized --/ -namespace Mathlib.Tactic -open Lean Parser.Tactic - -/- N -/ elab (name := omit) "omit" (ppSpace ident)+ : command => pure () -/- N -/ syntax (name := parameter) "parameter" (ppSpace bracketedBinder)+ : command - -/- S -/ syntax (name := propagateTags) "propagate_tags " tacticSeq : tactic -/- S -/ syntax (name := mapply) "mapply " term : tactic -/- S -/ syntax "destruct " term : tactic -/- N -/ syntax (name := abstract) "abstract" (ppSpace ident)? ppSpace tacticSeq : tactic - -/- S -/ syntax (name := rsimp) "rsimp" : tactic -/- S -/ syntax (name := compVal) "comp_val" : tactic -/- S -/ syntax (name := async) "async " tacticSeq : tactic - -/- M -/ syntax (name := injectionsAndClear) "injections_and_clear" : tactic - -/- E -/ syntax (name := tryFor) "try_for " term:max tacticSeq : tactic -/- E -/ syntax (name := unfoldCoes) "unfold_coes" (location)? : tactic -/- E -/ syntax (name := unfoldWf) "unfold_wf" : tactic -/- M -/ syntax (name := unfoldAux) "unfold_aux" : tactic -/- S -/ syntax (name := «continue») "continue " tacticSeq : tactic -/- B -/ syntax (name := refineStruct) "refine_struct " term : tactic -/- M -/ syntax (name := matchHyp) "match_hyp " ("(" &"m" " := " term ") ")? ident " : " term : - tactic -/- S -/ syntax (name := guardTags) "guard_tags" (ppSpace ident)* : tactic -/- S -/ syntax (name := guardProofTerm) "guard_proof_term " tactic:51 " => " term : tactic -/- S -/ syntax (name := failIfSuccess?) "fail_if_success? " str ppSpace tacticSeq : tactic -/- N -/ syntax (name := field) "field " ident " => " tacticSeq : tactic -/- S -/ syntax (name := haveField) "have_field" : tactic -/- S -/ syntax (name := applyField) "apply_field" : tactic -/- S -/ syntax (name := hGeneralize) "h_generalize " atomic(binderIdent " : ")? term:51 " = " ident - (" with " binderIdent)? : tactic -/- S -/ syntax (name := hGeneralize!) "h_generalize! " atomic(binderIdent " : ")? - term:51 " = " ident (" with " binderIdent)? : tactic -/- S -/ syntax (name := extractGoal!) "extract_goal!" (ppSpace ident)? - (" with" (ppSpace colGt ident)*)? : tactic -/- S -/ syntax (name := revertDeps) "revert_deps" (ppSpace colGt ident)* : tactic -/- S -/ syntax (name := revertAfter) "revert_after " ident : tactic -/- S -/ syntax (name := revertTargetDeps) "revert_target_deps" : tactic - -/- S -/ syntax (name := rcases?) "rcases?" casesTarget,* (" : " num)? : tactic -/- S -/ syntax (name := rintro?) "rintro?" (" : " num)? : tactic - -/- M -/ syntax (name := decide!) "decide!" : tactic - -/- M -/ syntax (name := deltaInstance) "delta_instance" (ppSpace ident)* : tactic - -/- S -/ syntax (name := elide) "elide " num (location)? : tactic -/- S -/ syntax (name := unelide) "unelide" (location)? : tactic - -/-- `ext1? pat*` is like `ext1 pat*` but gives a suggestion on what pattern to use -/ -/- M -/ syntax (name := ext1?) "ext1?" (colGt ppSpace rintroPat)* : tactic -/-- `ext? pat*` is like `ext pat*` but gives a suggestion on what pattern to use -/ -/- M -/ syntax (name := ext?) "ext?" (colGt ppSpace rintroPat)* (" : " num)? : tactic - -/- S -/ syntax (name := clarify) "clarify" (config)? - (Parser.Tactic.simpArgs)? (" using " term,+)? : tactic -/- S -/ syntax (name := safe) "safe" (config)? - (Parser.Tactic.simpArgs)? (" using " term,+)? : tactic -/- S -/ syntax (name := finish) "finish" (config)? - (Parser.Tactic.simpArgs)? (" using " term,+)? : tactic - -syntax generalizesArg := (ident " : ")? term:51 " = " ident -/- M -/ syntax (name := generalizes) "generalizes " "[" generalizesArg,* "]" : tactic - -syntax withPattern := "-" <|> "_" <|> ident -/- S -/ syntax (name := cases'') "cases''" casesTarget - (" with" (ppSpace colGt withPattern)+)? : tactic -syntax fixingClause := " fixing" (" *" <|> (ppSpace ident)+) -syntax generalizingClause := " generalizing" (ppSpace ident)+ -/- S -/ syntax (name := induction'') "induction''" casesTarget - (fixingClause <|> generalizingClause)? (" with" (ppSpace colGt withPattern)+)? : tactic - -/- B -/ syntax (name := obviously) "obviously" : tactic - -/- S -/ syntax (name := prettyCases) "pretty_cases" : tactic - -/- M -/ syntax (name := assocRw) "assoc_rw " rwRuleSeq (location)? : tactic - -/- N -/ syntax (name := dsimpResult) "dsimp_result" - (&" only")? (dsimpArgs)? " => " tacticSeq : tactic -/- N -/ syntax (name := simpResult) "simp_result" - (&" only")? (simpArgs)? " => " tacticSeq : tactic - -/- S -/ syntax (name := suggest) "suggest" (config)? (ppSpace num)? - (simpArgs)? (" using" (ppSpace colGt binderIdent)+)? : tactic - -/- M -/ syntax (name := truncCases) "trunc_cases " term - (" with" (ppSpace colGt binderIdent)+)? : tactic - -/- E -/ syntax (name := applyNormed) "apply_normed " term : tactic - -/- B -/ syntax (name := acMono) "ac_mono" ("*" <|> ("^" num))? - (config)? ((" : " term) <|> (" := " term))? : tactic - -/- M -/ syntax (name := reassoc) "reassoc" (ppSpace colGt ident)* : tactic -/- M -/ syntax (name := reassoc!) "reassoc!" (ppSpace colGt ident)* : tactic -/- M -/ syntax (name := deriveReassocProof) "derive_reassoc_proof" : tactic - -/- S -/ syntax (name := subtypeInstance) "subtype_instance" : tactic - -/- S -/ syntax (name := transport) "transport" (ppSpace term)? " using " term : tactic - -/- M -/ syntax (name := unfoldCases) "unfold_cases " tacticSeq : tactic - -/- B -/ syntax (name := equivRw) "equiv_rw" (config)? - ((" [" term,* "]") <|> (ppSpace term)) (location)? : tactic -/- B -/ syntax (name := equivRwType) "equiv_rw_type" (config)? ppSpace term : tactic - -/- E -/ syntax (name := nthRwLHS) "nth_rw_lhs " num rwRuleSeq (location)? : tactic -/- E -/ syntax (name := nthRwRHS) "nth_rw_rhs " num rwRuleSeq (location)? : tactic - -/- M -/ syntax (name := piInstanceDeriveField) "pi_instance_derive_field" : tactic -/- M -/ syntax (name := piInstance) "pi_instance" : tactic - -/- B -/ syntax (name := tidy) "tidy" (config)? : tactic -/- B -/ syntax (name := tidy?) "tidy?" (config)? : tactic - -/- M -/ syntax (name := deriveElementwiseProof) "derive_elementwise_proof" : tactic - -/- M -/ syntax (name := computeDegreeLE) "compute_degree_le" : tactic - -/- S -/ syntax (name := mkDecorations) "mk_decorations" : tactic - -/- E -/ syntax (name := isBounded_default) "isBounded_default" : tactic - -/- S -/ syntax (name := mvBisim) "mv_bisim" (ppSpace colGt term)? - (" with" (ppSpace binderIdent)+)? : tactic - -/- E -/ syntax (name := unitInterval) "unit_interval" : tactic - -/- M -/ syntax (name := padicIndexSimp) "padic_index_simp" " [" term,* "]" (location)? : tactic - -/- E -/ syntax (name := uniqueDiffWithinAt_Ici_Iic_univ) "uniqueDiffWithinAt_Ici_Iic_univ" : tactic - -/- E -/ syntax (name := wittTruncateFunTac) "witt_truncate_fun_tac" : tactic - -/- S -/ syntax (name := intro) "intro" : attr -/- S -/ syntax (name := intro!) "intro!" : attr - -/- S -/ syntax (name := interactive) "interactive" : attr - -/- M -/ syntax (name := expandExists) "expand_exists" (ppSpace ident)+ : attr - --- TODO: this should be handled in mathport -/- S -/ syntax (name := protectProj) "protect_proj" (&" without" (ppSpace ident)+)? : attr - -/- M -/ syntax (name := notationClass) "notation_class" "*"? (ppSpace ident)? : attr - -/- N -/ syntax (name := addTacticDoc) (docComment)? "add_tactic_doc " term : command - -/- S -/ syntax (name := listUnusedDecls) "#list_unused_decls" : command - -/- N -/ syntax (name := defReplacer) "def_replacer " ident Parser.Term.optType : command - -/- M -/ syntax (name := reassocAxiom) "reassoc_axiom " ident : command - -/- S -/ syntax (name := sample) "#sample " term : command - -/- S -/ syntax (name := printSorryIn) "#print_sorry_in " ident : command - -/- E -/ syntax (name := assertInstance) "assert_instance " term : command -/- E -/ syntax (name := assertNoInstance) "assert_no_instance " term : command - -end Tactic - -end Mathlib diff --git a/Mathlib/Order/SetNotation.lean b/Mathlib/Order/SetNotation.lean index 18ea3a90e75a7..dc586e637b640 100644 --- a/Mathlib/Order/SetNotation.lean +++ b/Mathlib/Order/SetNotation.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Yury Kudryashov -/ import Mathlib.Data.Set.Defs -import Mathlib.Mathport.Notation +import Mathlib.Util.Notation3 /-! # Notation classes for set supremum and infimum diff --git a/Mathlib/Mathport/Notation.lean b/Mathlib/Util/Notation3.lean similarity index 100% rename from Mathlib/Mathport/Notation.lean rename to Mathlib/Util/Notation3.lean diff --git a/test/notation3.lean b/test/notation3.lean index 0aa79ee858955..44116cca2ce50 100644 --- a/test/notation3.lean +++ b/test/notation3.lean @@ -1,4 +1,4 @@ -import Mathlib.Mathport.Notation +import Mathlib.Util.Notation3 import Mathlib.Init.Data.Nat.Lemmas import Mathlib.Data.Nat.Defs From 1c0bc97cf3db27c7843717ddb46f908eb84d987f Mon Sep 17 00:00:00 2001 From: damiano Date: Sun, 11 Aug 2024 22:04:38 +0000 Subject: [PATCH 178/359] fix: add year (#15703) Ruben mentioned this year in #15256 since he found https://github.com/leanprover-community/mathlib3/pull/6395. --- Mathlib/Topology/EMetricSpace/Paracompact.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Topology/EMetricSpace/Paracompact.lean b/Mathlib/Topology/EMetricSpace/Paracompact.lean index 15a258752ae20..db4fdd6346d82 100644 --- a/Mathlib/Topology/EMetricSpace/Paracompact.lean +++ b/Mathlib/Topology/EMetricSpace/Paracompact.lean @@ -1,5 +1,5 @@ /- -Copyright (c) 202 Yury Kudryashov. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ From 8b3112eb81a2941c5e99d7ede2683b012beaf995 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Sun, 11 Aug 2024 22:38:58 +0000 Subject: [PATCH 179/359] chore: turn `induction'` into `induction` (#15697) This PR substitutes the (soon to be deprecated) `induction'` with `induction`. --- Mathlib/Algebra/Group/Support.lean | 7 ++++--- Mathlib/Algebra/Polynomial/Basic.lean | 6 +++--- Mathlib/Algebra/Polynomial/Div.lean | 7 ++++--- Mathlib/Algebra/Polynomial/Laurent.lean | 7 +++---- Mathlib/Algebra/Polynomial/Monic.lean | 6 +++--- Mathlib/Algebra/Regular/SMul.lean | 7 ++++--- .../Asymptotics/SuperpolynomialDecay.lean | 13 +++++++------ .../SpecialFunctions/Gamma/Basic.lean | 10 ++++++---- .../SpecialFunctions/Gamma/BohrMollerup.lean | 7 ++++--- .../Triangulated/Pretriangulated.lean | 12 ++++++------ .../Combinatorics/Derangements/Finite.lean | 7 ++++--- Mathlib/FieldTheory/Finite/Basic.lean | 19 ++++++++++--------- Mathlib/GroupTheory/Perm/Basic.lean | 12 ++++++------ Mathlib/GroupTheory/Perm/Cycle/Factors.lean | 7 ++++--- Mathlib/GroupTheory/Perm/List.lean | 14 ++++++++------ Mathlib/GroupTheory/Solvable.lean | 6 +++--- Mathlib/NumberTheory/Padics/PadicVal.lean | 7 ++++--- Mathlib/RingTheory/Binomial.lean | 7 ++++--- 18 files changed, 87 insertions(+), 74 deletions(-) diff --git a/Mathlib/Algebra/Group/Support.lean b/Mathlib/Algebra/Group/Support.lean index 0a81d3998d168..51a1ada635050 100644 --- a/Mathlib/Algebra/Group/Support.lean +++ b/Mathlib/Algebra/Group/Support.lean @@ -209,9 +209,10 @@ theorem mulSupport_mul [MulOneClass M] (f g : α → M) : @[to_additive] theorem mulSupport_pow [Monoid M] (f : α → M) (n : ℕ) : (mulSupport fun x => f x ^ n) ⊆ mulSupport f := by - induction' n with n hfn - · simp [pow_zero, mulSupport_one] - · simpa only [pow_succ'] using (mulSupport_mul f _).trans (union_subset Subset.rfl hfn) + induction n with + | zero => simp [pow_zero, mulSupport_one] + | succ n hfn => + simpa only [pow_succ'] using (mulSupport_mul f _).trans (union_subset Subset.rfl hfn) section DivisionMonoid diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index 095de48696935..fa151eea0e863 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -772,9 +772,9 @@ theorem support_trinomial' (k m n : ℕ) (x y z : R) : end Fewnomials theorem X_pow_eq_monomial (n) : X ^ n = monomial n (1 : R) := by - induction' n with n hn - · rw [pow_zero, monomial_zero_one] - · rw [pow_succ, hn, X, monomial_mul_monomial, one_mul] + induction n with + | zero => rw [pow_zero, monomial_zero_one] + | succ n hn => rw [pow_succ, hn, X, monomial_mul_monomial, one_mul] @[simp high] theorem toFinsupp_X_pow (n : ℕ) : (X ^ n).toFinsupp = Finsupp.single n (1 : R) := by diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index dd585a94cd62a..f50623b79b73a 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -41,9 +41,10 @@ theorem X_pow_dvd_iff {f : R[X]} {n : ℕ} : X ^ n ∣ f ↔ ∀ d < n, f.coeff ⟨fun ⟨g, hgf⟩ d hd => by simp only [hgf, coeff_X_pow_mul', ite_eq_right_iff, not_le_of_lt hd, IsEmpty.forall_iff], fun hd => by - induction' n with n hn - · simp [pow_zero, one_dvd] - · obtain ⟨g, hgf⟩ := hn fun d : ℕ => fun H : d < n => hd _ (Nat.lt_succ_of_lt H) + induction n with + | zero => simp [pow_zero, one_dvd] + | succ n hn => + obtain ⟨g, hgf⟩ := hn fun d : ℕ => fun H : d < n => hd _ (Nat.lt_succ_of_lt H) have := coeff_X_pow_mul g n 0 rw [zero_add, ← hgf, hd n (Nat.lt_succ_self n)] at this obtain ⟨k, hgk⟩ := Polynomial.X_dvd_iff.mpr this.symm diff --git a/Mathlib/Algebra/Polynomial/Laurent.lean b/Mathlib/Algebra/Polynomial/Laurent.lean index f5842de474972..ca06a13294c58 100644 --- a/Mathlib/Algebra/Polynomial/Laurent.lean +++ b/Mathlib/Algebra/Polynomial/Laurent.lean @@ -365,10 +365,9 @@ it follow that `Q` is true on all Laurent polynomials. -/ theorem reduce_to_polynomial_of_mul_T (f : R[T;T⁻¹]) {Q : R[T;T⁻¹] → Prop} (Qf : ∀ f : R[X], Q (toLaurent f)) (QT : ∀ f, Q (f * T 1) → Q f) : Q f := by induction' f using LaurentPolynomial.induction_on_mul_T with f n - induction' n with n hn - · simpa only [Nat.zero_eq, Nat.cast_zero, neg_zero, T_zero, mul_one] using Qf _ - · convert QT _ _ - simpa using hn + induction n with + | zero => simpa only [Nat.zero_eq, Nat.cast_zero, neg_zero, T_zero, mul_one] using Qf _ + | succ n hn => convert QT _ _; simpa using hn section Support diff --git a/Mathlib/Algebra/Polynomial/Monic.lean b/Mathlib/Algebra/Polynomial/Monic.lean index 9783d0e918278..60c51bbad7bce 100644 --- a/Mathlib/Algebra/Polynomial/Monic.lean +++ b/Mathlib/Algebra/Polynomial/Monic.lean @@ -206,9 +206,9 @@ theorem eq_one_of_map_eq_one {S : Type*} [Semiring S] [Nontrivial S] (f : R →+ rw [← hndeg, ← Polynomial.leadingCoeff, hp.leadingCoeff, C.map_one] theorem natDegree_pow (hp : p.Monic) (n : ℕ) : (p ^ n).natDegree = n * p.natDegree := by - induction' n with n hn - · simp - · rw [pow_succ, (hp.pow n).natDegree_mul hp, hn, Nat.succ_mul, add_comm] + induction n with + | zero => simp + | succ n hn => rw [pow_succ, (hp.pow n).natDegree_mul hp, hn, Nat.succ_mul, add_comm] end Monic diff --git a/Mathlib/Algebra/Regular/SMul.lean b/Mathlib/Algebra/Regular/SMul.lean index fca8434b15e2c..5895185e0774c 100644 --- a/Mathlib/Algebra/Regular/SMul.lean +++ b/Mathlib/Algebra/Regular/SMul.lean @@ -133,9 +133,10 @@ theorem of_mul_eq_one (h : a * b = 1) : IsSMulRegular M b := /-- Any power of an `M`-regular element is `M`-regular. -/ theorem pow (n : ℕ) (ra : IsSMulRegular M a) : IsSMulRegular M (a ^ n) := by - induction' n with n hn - · rw [pow_zero]; simp only [one] - · rw [pow_succ'] + induction n with + | zero => rw [pow_zero]; simp only [one] + | succ n hn => + rw [pow_succ'] exact (ra.smul_iff (a ^ n)).mpr hn /-- An element `a` is `M`-regular if and only if a positive power of `a` is `M`-regular. -/ diff --git a/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean b/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean index 3dfa44a453e03..29a5817582655 100644 --- a/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean +++ b/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean @@ -103,9 +103,9 @@ theorem SuperpolynomialDecay.mul_param (hf : SuperpolynomialDecay l k f) : theorem SuperpolynomialDecay.param_pow_mul (hf : SuperpolynomialDecay l k f) (n : ℕ) : SuperpolynomialDecay l k (k ^ n * f) := by - induction' n with n hn - · simpa only [Nat.zero_eq, one_mul, pow_zero] using hf - · simpa only [pow_succ', mul_assoc] using hn.param_mul + induction n with + | zero => simpa only [Nat.zero_eq, one_mul, pow_zero] using hf + | succ n hn => simpa only [pow_succ', mul_assoc] using hn.param_mul theorem SuperpolynomialDecay.mul_param_pow (hf : SuperpolynomialDecay l k f) (n : ℕ) : SuperpolynomialDecay l k (f * k ^ n) := @@ -260,9 +260,10 @@ theorem superpolynomialDecay_mul_param_iff (hk : Tendsto k l atTop) : theorem superpolynomialDecay_param_pow_mul_iff (hk : Tendsto k l atTop) (n : ℕ) : SuperpolynomialDecay l k (k ^ n * f) ↔ SuperpolynomialDecay l k f := by - induction' n with n hn - · simp - · simpa [pow_succ, ← mul_comm k, mul_assoc, + induction n with + | zero => simp + | succ n hn => + simpa [pow_succ, ← mul_comm k, mul_assoc, superpolynomialDecay_param_mul_iff (k ^ n * f) hk] using hn theorem superpolynomialDecay_mul_param_pow_iff (hk : Tendsto k l atTop) (n : ℕ) : diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index 49bdffb59a123..4ba4ca7492589 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -324,10 +324,12 @@ theorem Gamma_eq_integral {s : ℂ} (hs : 0 < s.re) : Gamma s = GammaIntegral s theorem Gamma_one : Gamma 1 = 1 := by rw [Gamma_eq_integral] <;> simp theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n + 1) = n ! := by - induction' n with n hn - · simp - · rw [Gamma_add_one n.succ <| Nat.cast_ne_zero.mpr <| Nat.succ_ne_zero n] - simp only [Nat.cast_succ, Nat.factorial_succ, Nat.cast_mul]; congr + induction n with + | zero => simp + | succ n hn => + rw [Gamma_add_one n.succ <| Nat.cast_ne_zero.mpr <| Nat.succ_ne_zero n] + simp only [Nat.cast_succ, Nat.factorial_succ, Nat.cast_mul] + congr @[simp] theorem Gamma_ofNat_eq_factorial (n : ℕ) [(n + 1).AtLeastTwo] : diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean index 69dbed7be166e..62fae42a39804 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean @@ -152,9 +152,10 @@ theorem f_nat_eq (hf_feq : ∀ {y : ℝ}, 0 < y → f (y + 1) = f y + log y) (hn theorem f_add_nat_eq (hf_feq : ∀ {y : ℝ}, 0 < y → f (y + 1) = f y + log y) (hx : 0 < x) (n : ℕ) : f (x + n) = f x + ∑ m ∈ Finset.range n, log (x + m) := by - induction' n with n hn - · simp - · have : x + n.succ = x + n + 1 := by push_cast; ring + induction n with + | zero => simp + | succ n hn => + have : x + n.succ = x + n + 1 := by push_cast; ring rw [this, hf_feq, hn] · rw [Finset.range_succ, Finset.sum_insert Finset.not_mem_range_self] abel diff --git a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean index 2e80a2f1b2e21..077faaf31d4b1 100644 --- a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean +++ b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean @@ -399,12 +399,12 @@ lemma shift_distinguished (n : ℤ) : isomorphic_distinguished _ (hb _ (ha _ hT)) _ ((Triangle.shiftFunctorAdd' C _ _ _ hc).app T) obtain (n|n) := n - · induction' n with n hn - · exact H_zero - · exact H_add hn H_one rfl - · induction' n with n hn - · exact H_neg_one - · exact H_add hn H_neg_one rfl + · induction n with + | zero => exact H_zero + | succ n hn => exact H_add hn H_one rfl + · induction n with + | zero => exact H_neg_one + | succ n hn => exact H_add hn H_neg_one rfl end Triangle diff --git a/Mathlib/Combinatorics/Derangements/Finite.lean b/Mathlib/Combinatorics/Derangements/Finite.lean index cdc356dc495b9..31ea036e767de 100644 --- a/Mathlib/Combinatorics/Derangements/Finite.lean +++ b/Mathlib/Combinatorics/Derangements/Finite.lean @@ -78,9 +78,10 @@ theorem numDerangements_add_two (n : ℕ) : theorem numDerangements_succ (n : ℕ) : (numDerangements (n + 1) : ℤ) = (n + 1) * (numDerangements n : ℤ) - (-1) ^ n := by - induction' n with n hn - · rfl - · simp only [numDerangements_add_two, hn, pow_succ, Int.ofNat_mul, Int.ofNat_add, Int.ofNat_succ] + induction n with + | zero => rfl + | succ n hn => + simp only [numDerangements_add_two, hn, pow_succ, Int.ofNat_mul, Int.ofNat_add] ring theorem card_derangements_fin_eq_numDerangements {n : ℕ} : diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index 36e485c411639..a99d07385ebeb 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -221,9 +221,9 @@ theorem pow_card (a : K) : a ^ q = a := by pow_card_sub_one_eq_one a h, one_mul] theorem pow_card_pow (n : ℕ) (a : K) : a ^ q ^ n = a := by - induction' n with n ih - · simp - · simp [pow_succ, pow_mul, ih, pow_card] + induction n with + | zero => simp + | succ n ih => simp [pow_succ, pow_mul, ih, pow_card] end @@ -351,9 +351,10 @@ theorem frobenius_pow {p : ℕ} [Fact p.Prime] [CharP K p] {n : ℕ} (hcard : q frobenius K p ^ n = 1 := by ext x; conv_rhs => rw [RingHom.one_def, RingHom.id_apply, ← pow_card x, hcard] clear hcard - induction' n with n hn - · simp - · rw [pow_succ', pow_succ, pow_mul, RingHom.mul_def, RingHom.comp_apply, frobenius_def, hn] + induction n with + | zero => simp + | succ n hn => + rw [pow_succ', pow_succ, pow_mul, RingHom.mul_def, RingHom.comp_apply, frobenius_def, hn] open Polynomial @@ -471,9 +472,9 @@ theorem pow_card {p : ℕ} [Fact p.Prime] (x : ZMod p) : x ^ p = x := by @[simp] theorem pow_card_pow {n p : ℕ} [Fact p.Prime] (x : ZMod p) : x ^ p ^ n = x := by - induction' n with n ih - · simp - · simp [pow_succ, pow_mul, ih, pow_card] + induction n with + | zero => simp + | succ n ih => simp [pow_succ, pow_mul, ih, pow_card] @[simp] theorem frobenius_zmod (p : ℕ) [Fact p.Prime] : frobenius (ZMod p) p = RingHom.id _ := by diff --git a/Mathlib/GroupTheory/Perm/Basic.lean b/Mathlib/GroupTheory/Perm/Basic.lean index d6267ca0eb3cc..94a39ade46477 100644 --- a/Mathlib/GroupTheory/Perm/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Basic.lean @@ -346,9 +346,9 @@ private theorem pow_aux (hf : ∀ x, p x ↔ p (f x)) : ∀ {n : ℕ} (x), p x @[simp] theorem subtypePerm_pow (f : Perm α) (n : ℕ) (hf) : (f.subtypePerm hf : Perm { x // p x }) ^ n = (f ^ n).subtypePerm (pow_aux hf) := by - induction' n with n ih - · simp - · simp_rw [pow_succ', ih, subtypePerm_mul] + induction n with + | zero => simp + | succ n ih => simp_rw [pow_succ', ih, subtypePerm_mul] private theorem zpow_aux (hf : ∀ x, p x ↔ p (f x)) : ∀ {n : ℤ} (x), p x ↔ p ((f ^ n) x) | Int.ofNat n => pow_aux hf @@ -359,9 +359,9 @@ private theorem zpow_aux (hf : ∀ x, p x ↔ p (f x)) : ∀ {n : ℤ} (x), p x @[simp] theorem subtypePerm_zpow (f : Perm α) (n : ℤ) (hf) : (f.subtypePerm hf ^ n : Perm { x // p x }) = (f ^ n).subtypePerm (zpow_aux hf) := by - induction' n with n ih - · exact subtypePerm_pow _ _ _ - · simp only [zpow_negSucc, subtypePerm_pow, subtypePerm_inv] + induction n with + | ofNat n => exact subtypePerm_pow _ _ _ + | negSucc n => simp only [zpow_negSucc, subtypePerm_pow, subtypePerm_inv] variable [DecidablePred p] {a : α} diff --git a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean index 06e8e9cdedfb5..a718ff65870f1 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean @@ -61,9 +61,10 @@ theorem cycleOf_inv (f : Perm α) [DecidableRel f.SameCycle] (x : α) : theorem cycleOf_pow_apply_self (f : Perm α) [DecidableRel f.SameCycle] (x : α) : ∀ n : ℕ, (cycleOf f x ^ n) x = (f ^ n) x := by intro n - induction' n with n hn - · rfl - · rw [pow_succ', mul_apply, cycleOf_apply, hn, if_pos, pow_succ', mul_apply] + induction n with + | zero => rfl + | succ n hn => + rw [pow_succ', mul_apply, cycleOf_apply, hn, if_pos, pow_succ', mul_apply] exact ⟨n, rfl⟩ @[simp] diff --git a/Mathlib/GroupTheory/Perm/List.lean b/Mathlib/GroupTheory/Perm/List.lean index b19007fd40f9a..092595ac44ce4 100644 --- a/Mathlib/GroupTheory/Perm/List.lean +++ b/Mathlib/GroupTheory/Perm/List.lean @@ -267,9 +267,10 @@ theorem formPerm_rotate_one (l : List α) (h : Nodup l) : formPerm (l.rotate 1) theorem formPerm_rotate (l : List α) (h : Nodup l) (n : ℕ) : formPerm (l.rotate n) = formPerm l := by - induction' n with n hn - · simp - · rw [← rotate_rotate, formPerm_rotate_one, hn] + induction n with + | zero => simp + | succ n hn => + rw [← rotate_rotate, formPerm_rotate_one, hn] rwa [IsRotated.nodup_iff] exact IsRotated.forall l n @@ -294,9 +295,10 @@ theorem formPerm_reverse : ∀ l : List α, formPerm l.reverse = (formPerm l)⁻ theorem formPerm_pow_apply_getElem (l : List α) (w : Nodup l) (n : ℕ) (i : ℕ) (h : i < l.length) : (formPerm l ^ n) l[i] = l[(i + n) % l.length]'(Nat.mod_lt _ (i.zero_le.trans_lt h)) := by - induction' n with n hn - · simp [Nat.mod_eq_of_lt h] - · simp [pow_succ', mul_apply, hn, formPerm_apply_getElem _ w, Nat.succ_eq_add_one, + induction n with + | zero => simp [Nat.mod_eq_of_lt h] + | succ n hn => + simp [pow_succ', mul_apply, hn, formPerm_apply_getElem _ w, Nat.succ_eq_add_one, ← Nat.add_assoc] @[deprecated formPerm_pow_apply_getElem (since := "2024-08-03")] diff --git a/Mathlib/GroupTheory/Solvable.lean b/Mathlib/GroupTheory/Solvable.lean index 687ad0c74e7a2..9082598ea518c 100644 --- a/Mathlib/GroupTheory/Solvable.lean +++ b/Mathlib/GroupTheory/Solvable.lean @@ -47,9 +47,9 @@ theorem derivedSeries_succ (n : ℕ) : -- Porting note: had to provide inductive hypothesis explicitly theorem derivedSeries_normal (n : ℕ) : (derivedSeries G n).Normal := by - induction' n with n ih - · exact (⊤ : Subgroup G).normal_of_characteristic - · exact @Subgroup.commutator_normal G _ (derivedSeries G n) (derivedSeries G n) ih ih + induction n with + | zero => exact (⊤ : Subgroup G).normal_of_characteristic + | succ n ih => exact @Subgroup.commutator_normal G _ (derivedSeries G n) (derivedSeries G n) ih ih -- Porting note: higher simp priority to restore Lean 3 behavior @[simp 1100] diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 90bc427eeab82..15cddbfe318d9 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -636,9 +636,10 @@ theorem padicValNat_eq_zero_of_mem_Ioo {m k : ℕ} theorem padicValNat_factorial_mul_add {n : ℕ} (m : ℕ) [hp : Fact p.Prime] (h : n < p) : padicValNat p (p * m + n) ! = padicValNat p (p * m) ! := by - induction' n with n hn - · rw [add_zero] - · rw [add_succ, factorial_succ, + induction n with + | zero => rw [add_zero] + | succ n hn => + rw [add_succ, factorial_succ, padicValNat.mul (succ_ne_zero (p * m + n)) <| factorial_ne_zero (p * m + _), hn <| lt_of_succ_lt h, ← add_succ, padicValNat_eq_zero_of_mem_Ioo ⟨(Nat.lt_add_of_pos_right <| succ_pos n), diff --git a/Mathlib/RingTheory/Binomial.lean b/Mathlib/RingTheory/Binomial.lean index 17b424f495698..82a59a2f62044 100644 --- a/Mathlib/RingTheory/Binomial.lean +++ b/Mathlib/RingTheory/Binomial.lean @@ -168,9 +168,10 @@ namespace Polynomial theorem ascPochhammer_smeval_cast (R : Type*) [Semiring R] {S : Type*} [NonAssocSemiring S] [Pow S ℕ] [Module R S] [IsScalarTower R S S] [NatPowAssoc S] (x : S) (n : ℕ) : (ascPochhammer R n).smeval x = (ascPochhammer ℕ n).smeval x := by - induction' n with n hn - · simp only [Nat.zero_eq, ascPochhammer_zero, smeval_one, one_smul] - · simp only [ascPochhammer_succ_right, mul_add, smeval_add, smeval_mul_X, ← Nat.cast_comm] + induction n with + | zero => simp only [Nat.zero_eq, ascPochhammer_zero, smeval_one, one_smul] + | succ n hn => + simp only [ascPochhammer_succ_right, mul_add, smeval_add, smeval_mul_X, ← Nat.cast_comm] simp only [← C_eq_natCast, smeval_C_mul, hn, Nat.cast_smul_eq_nsmul R n] simp only [nsmul_eq_mul, Nat.cast_id] From b51014c15450d43f140584e7f29df5d1d2759a0c Mon Sep 17 00:00:00 2001 From: Salvatore Mercuri Date: Sun, 11 Aug 2024 23:33:25 +0000 Subject: [PATCH 180/359] feat(Algebra/Ring): add `RingEquiv` versions of splitting and product equivalences (#15075) Add `RingEquiv` versions of `Equiv.piCongrLeft`, `Equiv.piEquivPiSubtypeProd`, and `Equiv.prodCongr`. This has been split from the larger PR #13577. --- Mathlib/Algebra/Ring/Equiv.lean | 57 +++++++++++++++++++++++++++++++++ 1 file changed, 57 insertions(+) diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index 92eabfb25d020..33303fe81c670 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ +import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Algebra.GroupWithZero.InjSurj @@ -440,6 +441,62 @@ theorem piCongrRight_trans {ι : Type*} {R S T : ι → Type*} (piCongrRight e).trans (piCongrRight f) = piCongrRight fun i => (e i).trans (f i) := rfl +/-- Transport dependent functions through an equivalence of the base space. + +This is `Equiv.piCongrLeft'` as a `RingEquiv`. -/ +@[simps!] +def piCongrLeft' {ι ι' : Type*} (R : ι → Type*) (e : ι ≃ ι') + [∀ i, NonUnitalNonAssocSemiring (R i)] : + ((i : ι) → R i) ≃+* ((i : ι') → R (e.symm i)) where + toEquiv := Equiv.piCongrLeft' R e + map_mul' _ _ := rfl + map_add' _ _ := rfl + +@[simp] +theorem piCongrLeft'_symm {R : Type*} [NonUnitalNonAssocSemiring R] (e : α ≃ β) : + (RingEquiv.piCongrLeft' (fun _ => R) e).symm = RingEquiv.piCongrLeft' _ e.symm := by + simp only [piCongrLeft', RingEquiv.symm, MulEquiv.symm, Equiv.piCongrLeft'_symm] + +/-- Transport dependent functions through an equivalence of the base space. + +This is `Equiv.piCongrLeft` as a `RingEquiv`. -/ +@[simps!] +def piCongrLeft {ι ι' : Type*} (S : ι' → Type*) (e : ι ≃ ι') + [∀ i, NonUnitalNonAssocSemiring (S i)] : + ((i : ι) → S (e i)) ≃+* ((i : ι') → S i) := + (RingEquiv.piCongrLeft' S e.symm).symm + +/-- Splits the indices of ring `∀ (i : ι), Y i` along the predicate `p`. This is +`Equiv.piEquivPiSubtypeProd` as a `RingEquiv`. -/ +@[simps!] +def piEquivPiSubtypeProd {ι : Type*} (p : ι → Prop) [DecidablePred p] (Y : ι → Type*) + [∀ i, NonUnitalNonAssocSemiring (Y i)] : + ((i : ι) → Y i) ≃+* ((i : { x : ι // p x }) → Y i) × ((i : { x : ι // ¬p x }) → Y i) where + toEquiv := Equiv.piEquivPiSubtypeProd p Y + map_mul' _ _ := rfl + map_add' _ _ := rfl + +/-- Product of ring equivalences. This is `Equiv.prodCongr` as a `RingEquiv`. -/ +@[simps!] +def prodCongr {R R' S S' : Type*} [NonUnitalNonAssocSemiring R] [NonUnitalNonAssocSemiring R'] + [NonUnitalNonAssocSemiring S] [NonUnitalNonAssocSemiring S'] + (f : R ≃+* R') (g : S ≃+* S') : + R × S ≃+* R' × S' where + toEquiv := Equiv.prodCongr f g + map_mul' _ _ := by + simp only [Equiv.toFun_as_coe, Equiv.prodCongr_apply, EquivLike.coe_coe, + Prod.map, Prod.fst_mul, map_mul, Prod.snd_mul, Prod.mk_mul_mk] + map_add' _ _ := by + simp only [Equiv.toFun_as_coe, Equiv.prodCongr_apply, EquivLike.coe_coe, + Prod.map, Prod.fst_add, map_add, Prod.snd_add, Prod.mk_add_mk] + +@[simp] +theorem coe_prodCongr {R R' S S' : Type*} [NonUnitalNonAssocSemiring R] + [NonUnitalNonAssocSemiring R'] [NonUnitalNonAssocSemiring S] [NonUnitalNonAssocSemiring S'] + (f : R ≃+* R') (g : S ≃+* S') : + ⇑(RingEquiv.prodCongr f g) = Prod.map f g := + rfl + end NonUnitalSemiring section Semiring From ddc3dc13bd98b4b192bf184f487ad0301f21b9e7 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Sun, 11 Aug 2024 23:33:26 +0000 Subject: [PATCH 181/359] feat(Analysis/Complex/ReImTopology): Uniform continuity of re im (#15411) We show that re and im are uniformly continuous and then use this to show that re/im parts of a uniform limit converge uniformly. Co-authored-by: Chris Birkbeck --- Mathlib/Analysis/Complex/Basic.lean | 6 +++++ Mathlib/Analysis/Complex/ReImTopology.lean | 26 ++++++++++++++++++++++ 2 files changed, 32 insertions(+) diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index 74cabdacaa529..5cce5c3e7eeb1 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -241,6 +241,9 @@ def reCLM : ℂ →L[ℝ] ℝ := theorem continuous_re : Continuous re := reCLM.continuous +lemma uniformlyContinous_re : UniformContinuous re := + reCLM.uniformContinuous + @[simp] theorem reCLM_coe : (reCLM : ℂ →ₗ[ℝ] ℝ) = reLm := rfl @@ -257,6 +260,9 @@ def imCLM : ℂ →L[ℝ] ℝ := theorem continuous_im : Continuous im := imCLM.continuous +lemma uniformlyContinous_im : UniformContinuous im := + imCLM.uniformContinuous + @[simp] theorem imCLM_coe : (imCLM : ℂ →ₗ[ℝ] ℝ) = imLm := rfl diff --git a/Mathlib/Analysis/Complex/ReImTopology.lean b/Mathlib/Analysis/Complex/ReImTopology.lean index 5fff8e442adaf..b9eac29a030c2 100644 --- a/Mathlib/Analysis/Complex/ReImTopology.lean +++ b/Mathlib/Analysis/Complex/ReImTopology.lean @@ -175,3 +175,29 @@ theorem IsClosed.reProdIm (hs : IsClosed s) (ht : IsClosed t) : IsClosed (s × theorem Bornology.IsBounded.reProdIm (hs : IsBounded s) (ht : IsBounded t) : IsBounded (s ×ℂ t) := antilipschitz_equivRealProd.isBounded_preimage (hs.prod ht) + +section continuity + +variable {α ι : Type*} + +protected lemma TendstoUniformlyOn.re {f : ι → α → ℂ} {p : Filter ι} {g : α → ℂ} {K : Set α} + (hf : TendstoUniformlyOn f g p K) : + TendstoUniformlyOn (fun n x => (f n x).re) (fun y => (g y).re) p K := by + apply UniformContinuous.comp_tendstoUniformlyOn uniformlyContinous_re hf + +protected lemma TendstoUniformly.re {f : ι → α → ℂ} {p : Filter ι} {g : α → ℂ} + (hf : TendstoUniformly f g p) : + TendstoUniformly (fun n x => (f n x).re) (fun y => (g y).re) p := by + apply UniformContinuous.comp_tendstoUniformly uniformlyContinous_re hf + +protected lemma TendstoUniformlyOn.im {f : ι → α → ℂ} {p : Filter ι} {g : α → ℂ} {K : Set α} + (hf : TendstoUniformlyOn f g p K) : + TendstoUniformlyOn (fun n x => (f n x).im) (fun y => (g y).im) p K := by + apply UniformContinuous.comp_tendstoUniformlyOn uniformlyContinous_im hf + +protected lemma TendstoUniformly.im {f : ι → α → ℂ} {p : Filter ι} {g : α → ℂ} + (hf : TendstoUniformly f g p) : + TendstoUniformly (fun n x => (f n x).im) (fun y => (g y).im) p := by + apply UniformContinuous.comp_tendstoUniformly uniformlyContinous_im hf + +end continuity From 3486eabde9d145d64eb198be7d5c4714e1b0c9c4 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 11 Aug 2024 23:33:27 +0000 Subject: [PATCH 182/359] chore: Missing `Ring` and `CommRing` instances for graded rings (#15492) The analogous `Semiring` versions already exist. --- Mathlib/Algebra/DirectSum/Ring.lean | 37 ++++++++++++++++++++--------- 1 file changed, 26 insertions(+), 11 deletions(-) diff --git a/Mathlib/Algebra/DirectSum/Ring.lean b/Mathlib/Algebra/DirectSum/Ring.lean index c292fb743421f..05a579900593c 100644 --- a/Mathlib/Algebra/DirectSum/Ring.lean +++ b/Mathlib/Algebra/DirectSum/Ring.lean @@ -607,7 +607,8 @@ section Uniform variable (ι) -/-- A direct sum of copies of a `Semiring` inherits the multiplication structure. -/ +/-- A direct sum of copies of a `NonUnitalNonAssocSemiring` inherits the multiplication structure. +-/ instance NonUnitalNonAssocSemiring.directSumGNonUnitalNonAssocSemiring {R : Type*} [AddMonoid ι] [NonUnitalNonAssocSemiring R] : DirectSum.GNonUnitalNonAssocSemiring fun _ : ι => R := { -- Porting note: removed Mul.gMul ι with and we seem ok @@ -618,12 +619,20 @@ instance NonUnitalNonAssocSemiring.directSumGNonUnitalNonAssocSemiring {R : Type /-- A direct sum of copies of a `Semiring` inherits the multiplication structure. -/ instance Semiring.directSumGSemiring {R : Type*} [AddMonoid ι] [Semiring R] : - DirectSum.GSemiring fun _ : ι => R := - { NonUnitalNonAssocSemiring.directSumGNonUnitalNonAssocSemiring ι, - Monoid.gMonoid ι with - natCast := fun n => n - natCast_zero := Nat.cast_zero - natCast_succ := Nat.cast_succ } + DirectSum.GSemiring fun _ : ι => R where + __ := NonUnitalNonAssocSemiring.directSumGNonUnitalNonAssocSemiring ι + __ := Monoid.gMonoid ι + natCast n := n + natCast_zero := Nat.cast_zero + natCast_succ := Nat.cast_succ + +/-- A direct sum of copies of a `Ring` inherits the multiplication structure. -/ +instance Ring.directSumGRing {R : Type*} [AddMonoid ι] [Ring R] : + DirectSum.GRing fun _ : ι => R where + __ := Semiring.directSumGSemiring ι + intCast z := z + intCast_ofNat := Int.cast_natCast + intCast_negSucc_ofNat := Int.cast_negSucc open DirectSum @@ -632,10 +641,16 @@ example {R : Type*} [AddMonoid ι] [Semiring R] (i j : ι) (a b : R) : (DirectSum.of _ i a * DirectSum.of _ j b : ⨁ _, R) = DirectSum.of _ (i + j) (a * b) := by rw [DirectSum.of_mul_of, Mul.gMul_mul] -/-- A direct sum of copies of a `CommSemiring` inherits the commutative multiplication structure. --/ +/-- A direct sum of copies of a `CommSemiring` inherits the commutative multiplication structure. -/ instance CommSemiring.directSumGCommSemiring {R : Type*} [AddCommMonoid ι] [CommSemiring R] : - DirectSum.GCommSemiring fun _ : ι => R := - { CommMonoid.gCommMonoid ι, Semiring.directSumGSemiring ι with } + DirectSum.GCommSemiring fun _ : ι => R where + __ := Semiring.directSumGSemiring ι + __ := CommMonoid.gCommMonoid ι + +/-- A direct sum of copies of a `CommRing` inherits the commutative multiplication structure. -/ +instance CommmRing.directSumGCommRing {R : Type*} [AddCommMonoid ι] [CommRing R] : + DirectSum.GCommRing fun _ : ι => R where + __ := Ring.directSumGRing ι + __ := CommMonoid.gCommMonoid ι end Uniform From 14de54768eeaf150332491c16e49d9bb8f405ce4 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 11 Aug 2024 23:33:28 +0000 Subject: [PATCH 183/359] chore(Integral/Lebesgue): reorganize `lintegral_sum_measure` (#15502) Previously, we proved `lintegral_sum_measure`, then deduced special cases. The proof of `lintegral_sum_measure` essentially contained reduction to `lintegral_add_measure`, so I prove it first, then deduce `lintegral_finset_sum_measure`, then `lintegral_sum_measure`. --- .../MeasureTheory/Function/SimpleFunc.lean | 4 ++ Mathlib/MeasureTheory/Integral/Lebesgue.lean | 58 +++++++++---------- 2 files changed, 32 insertions(+), 30 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 60c0bc193f8dc..ab03e2640207b 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -868,6 +868,10 @@ theorem lintegral_smul (f : α →ₛ ℝ≥0∞) (c : ℝ≥0∞) : f.lintegral theorem lintegral_zero [MeasurableSpace α] (f : α →ₛ ℝ≥0∞) : f.lintegral 0 = 0 := (lintegralₗ f).map_zero +theorem lintegral_finset_sum {ι} (f : α →ₛ ℝ≥0∞) (μ : ι → Measure α) (s : Finset ι) : + f.lintegral (∑ i ∈ s, μ i) = ∑ i ∈ s, f.lintegral (μ i) := + map_sum (lintegralₗ f) .. + theorem lintegral_sum {m : MeasurableSpace α} {ι} (f : α →ₛ ℝ≥0∞) (μ : ι → Measure α) : f.lintegral (Measure.sum μ) = ∑' i, f.lintegral (μ i) := by simp only [lintegral, Measure.sum_apply, f.measurableSet_preimage, ← Finset.tsum_subtype, ← diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index a7db0fcab3509..78b990a4b3679 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -246,12 +246,12 @@ theorem lintegral_mono_ae {f g : α → ℝ≥0∞} (h : ∀ᵐ a ∂μ, f a ≤ rcases exists_measurable_superset_of_null h with ⟨t, hts, ht, ht0⟩ have : ∀ᵐ x ∂μ, x ∉ t := measure_zero_iff_ae_nmem.1 ht0 rw [lintegral, lintegral] - refine iSup_le fun s => iSup_le fun hfs => le_iSup_of_le (s.restrict tᶜ) <| le_iSup_of_le ?_ ?_ + refine iSup₂_le fun s hfs ↦ le_iSup₂_of_le (s.restrict tᶜ) ?_ ?_ · intro a by_cases h : a ∈ t <;> simp only [restrict_apply s ht.compl, mem_compl_iff, h, not_true, not_false_eq_true, indicator_of_not_mem, zero_le, not_false_eq_true, indicator_of_mem] - exact le_trans (hfs a) (_root_.by_contradiction fun hnfg => h (hts hnfg)) + exact le_trans (hfs a) (by_contradiction fun hnfg => h (hts hnfg)) · refine le_of_eq (SimpleFunc.lintegral_congr <| this.mono fun a hnt => ?_) by_cases hat : a ∈ t <;> simp only [restrict_apply s ht.compl, mem_compl_iff, hat, not_true, not_false_eq_true, indicator_of_not_mem, not_false_eq_true, indicator_of_mem] @@ -586,41 +586,39 @@ lemma setLIntegral_smul_measure (c : ℝ≥0∞) (f : α → ℝ≥0∞) (s : Se alias set_lintegral_smul_measure := setLIntegral_smul_measure @[simp] -theorem lintegral_sum_measure {m : MeasurableSpace α} {ι} (f : α → ℝ≥0∞) (μ : ι → Measure α) : - ∫⁻ a, f a ∂Measure.sum μ = ∑' i, ∫⁻ a, f a ∂μ i := by - classical - simp only [lintegral, iSup_subtype', SimpleFunc.lintegral_sum, ENNReal.tsum_eq_iSup_sum] - rw [iSup_comm] - congr; funext s - induction' s using Finset.induction_on with i s hi hs - · simp - simp only [Finset.sum_insert hi, ← hs] - refine (ENNReal.iSup_add_iSup ?_).symm - intro φ ψ - exact - ⟨⟨φ ⊔ ψ, fun x => sup_le (φ.2 x) (ψ.2 x)⟩, - add_le_add (SimpleFunc.lintegral_mono le_sup_left le_rfl) - (Finset.sum_le_sum fun j _ => SimpleFunc.lintegral_mono le_sup_right le_rfl)⟩ - -theorem hasSum_lintegral_measure {ι} {_ : MeasurableSpace α} (f : α → ℝ≥0∞) (μ : ι → Measure α) : - HasSum (fun i => ∫⁻ a, f a ∂μ i) (∫⁻ a, f a ∂Measure.sum μ) := - (lintegral_sum_measure f μ).symm ▸ ENNReal.summable.hasSum +theorem lintegral_zero_measure {m : MeasurableSpace α} (f : α → ℝ≥0∞) : + ∫⁻ a, f a ∂(0 : Measure α) = 0 := by + simp [lintegral] @[simp] -theorem lintegral_add_measure {m : MeasurableSpace α} (f : α → ℝ≥0∞) (μ ν : Measure α) : +theorem lintegral_add_measure (f : α → ℝ≥0∞) (μ ν : Measure α) : ∫⁻ a, f a ∂(μ + ν) = ∫⁻ a, f a ∂μ + ∫⁻ a, f a ∂ν := by - simpa [tsum_fintype] using lintegral_sum_measure f fun b => cond b μ ν + simp only [lintegral, SimpleFunc.lintegral_add, iSup_subtype'] + refine (ENNReal.iSup_add_iSup ?_).symm + rintro ⟨φ, hφ⟩ ⟨ψ, hψ⟩ + refine ⟨⟨φ ⊔ ψ, sup_le hφ hψ⟩, ?_⟩ + apply_rules [add_le_add, SimpleFunc.lintegral_mono, le_rfl] -- TODO: use `gcongr` + exacts [le_sup_left, le_sup_right] @[simp] -theorem lintegral_finset_sum_measure {ι} {m : MeasurableSpace α} (s : Finset ι) (f : α → ℝ≥0∞) - (μ : ι → Measure α) : ∫⁻ a, f a ∂(∑ i ∈ s, μ i) = ∑ i ∈ s, ∫⁻ a, f a ∂μ i := by - rw [← Measure.sum_coe_finset, lintegral_sum_measure, ← Finset.tsum_subtype'] - simp only [Finset.coe_sort_coe] +theorem lintegral_finset_sum_measure {ι} (s : Finset ι) (f : α → ℝ≥0∞) (μ : ι → Measure α) : + ∫⁻ a, f a ∂(∑ i ∈ s, μ i) = ∑ i ∈ s, ∫⁻ a, f a ∂μ i := + let F : Measure α →+ ℝ≥0∞ := + { toFun := (lintegral · f), + map_zero' := lintegral_zero_measure f, + map_add' := lintegral_add_measure f } + map_sum F μ s @[simp] -theorem lintegral_zero_measure {m : MeasurableSpace α} (f : α → ℝ≥0∞) : - ∫⁻ a, f a ∂(0 : Measure α) = 0 := by - simp [lintegral] +theorem lintegral_sum_measure {m : MeasurableSpace α} {ι} (f : α → ℝ≥0∞) (μ : ι → Measure α) : + ∫⁻ a, f a ∂Measure.sum μ = ∑' i, ∫⁻ a, f a ∂μ i := by + simp_rw [ENNReal.tsum_eq_iSup_sum, ← lintegral_finset_sum_measure, + lintegral, SimpleFunc.lintegral_sum, ENNReal.tsum_eq_iSup_sum, + SimpleFunc.lintegral_finset_sum, iSup_comm (ι := Finset ι)] + +theorem hasSum_lintegral_measure {ι} {_ : MeasurableSpace α} (f : α → ℝ≥0∞) (μ : ι → Measure α) : + HasSum (fun i => ∫⁻ a, f a ∂μ i) (∫⁻ a, f a ∂Measure.sum μ) := + (lintegral_sum_measure f μ).symm ▸ ENNReal.summable.hasSum @[simp] theorem lintegral_of_isEmpty {α} [MeasurableSpace α] [IsEmpty α] (μ : Measure α) (f : α → ℝ≥0∞) : From e848aeef83386bbba5e26c848810c108c997b5e3 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 11 Aug 2024 23:33:29 +0000 Subject: [PATCH 184/359] refactor(Cyclic): redefine, move definition (#15544) Redefine `IsCyclic` without `Subgroup` dependency, move the definition to `Algebra/Group/Defs`. --- Mathlib/Algebra/Group/Defs.lean | 17 ++++++++- .../GroupTheory/SpecificGroups/Cyclic.lean | 38 ++++++------------- Mathlib/Logic/Function/Defs.lean | 2 +- 3 files changed, 29 insertions(+), 28 deletions(-) diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 74d035dc670c9..5e8035c15123a 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -5,6 +5,7 @@ Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro -/ import Mathlib.Data.Int.Notation import Mathlib.Algebra.Group.ZeroOne +import Mathlib.Logic.Function.Defs import Mathlib.Tactic.Lemma import Mathlib.Tactic.TypeStar import Mathlib.Tactic.Simps.Basic @@ -861,6 +862,20 @@ instance SubNegMonoid.SMulInt {M} [SubNegMonoid M] : SMul ℤ M := attribute [to_additive existing SubNegMonoid.SMulInt] DivInvMonoid.Pow +/-- A group is called *cyclic* if it is generated by a single element. -/ +class IsAddCyclic (G : Type u) [SMul ℤ G] : Prop where + protected exists_zsmul_surjective : ∃ g : G, Function.Surjective (· • g : ℤ → G) + +/-- A group is called *cyclic* if it is generated by a single element. -/ +@[to_additive] +class IsCyclic (G : Type u) [Pow G ℤ] : Prop where + protected exists_zpow_surjective : ∃ g : G, Function.Surjective (g ^ · : ℤ → G) + +@[to_additive] +theorem exists_zpow_surjective (G : Type*) [Pow G ℤ] [IsCyclic G] : + ∃ g : G, Function.Surjective (g ^ · : ℤ → G) := + IsCyclic.exists_zpow_surjective + section DivInvMonoid variable [DivInvMonoid G] {a b : G} @@ -1164,5 +1179,5 @@ initialize_simps_projections AddGroup initialize_simps_projections CommGroup initialize_simps_projections AddCommGroup -assert_not_exists Function.Injective +assert_not_exists Function.Injective.eq_iff assert_not_exists IsCommutative diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 749712cd51ac2..71c6a80871b74 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -49,20 +49,13 @@ attribute [local instance] setFintype open Subgroup -/-- A group is called *cyclic* if it is generated by a single element. -/ -class IsAddCyclic (α : Type u) [AddGroup α] : Prop where - exists_generator : ∃ g : α, ∀ x, x ∈ AddSubgroup.zmultiples g - -/-- A group is called *cyclic* if it is generated by a single element. -/ @[to_additive] -class IsCyclic (α : Type u) [Group α] : Prop where - exists_generator : ∃ g : α, ∀ x, x ∈ zpowers g +theorem IsCyclic.exists_generator [Group α] [IsCyclic α] : ∃ g : α, ∀ x, x ∈ Subgroup.zpowers g := + exists_zpow_surjective α @[to_additive] instance (priority := 100) isCyclic_of_subsingleton [Group α] [Subsingleton α] : IsCyclic α := - ⟨⟨1, fun x => by - rw [Subsingleton.elim x 1] - exact mem_zpowers 1⟩⟩ + ⟨⟨1, fun _ => ⟨0, Subsingleton.elim _ _⟩⟩⟩ @[simp] theorem isCyclic_multiplicative_iff [AddGroup α] : IsCyclic (Multiplicative α) ↔ IsAddCyclic α := @@ -113,11 +106,10 @@ MonoidAddHom.map_add_cyclic := AddMonoidHom.map_addCyclic @[to_additive] theorem isCyclic_of_orderOf_eq_card [Fintype α] (x : α) (hx : orderOf x = Fintype.card α) : IsCyclic α := by - classical - use x - simp_rw [← SetLike.mem_coe, ← Set.eq_univ_iff_forall] - rw [← Fintype.card_congr (Equiv.Set.univ α), ← Fintype.card_zpowers] at hx - exact Set.eq_of_subset_of_card_le (Set.subset_univ _) (ge_of_eq hx) + use x + rw [← Set.range_iff_surjective, ← coe_zpowers] + rw [← Fintype.card_congr (Equiv.Set.univ α), ← Fintype.card_zpowers] at hx + exact Set.eq_of_subset_of_card_le (Set.subset_univ _) (ge_of_eq hx) @[deprecated (since := "2024-02-21")] alias isAddCyclic_of_orderOf_eq_card := isAddCyclic_of_addOrderOf_eq_card @@ -560,17 +552,11 @@ variable [CommGroup α] [IsSimpleGroup α] @[to_additive] instance (priority := 100) isCyclic : IsCyclic α := by - cases' subsingleton_or_nontrivial α with hi hi <;> haveI := hi - · apply isCyclic_of_subsingleton - · obtain ⟨g, hg⟩ := exists_ne (1 : α) - refine ⟨⟨g, fun x => ?_⟩⟩ - cases' IsSimpleOrder.eq_bot_or_eq_top (Subgroup.zpowers g) with hb ht - · exfalso - apply hg - rw [← Subgroup.mem_bot, ← hb] - apply Subgroup.mem_zpowers - · rw [ht] - apply Subgroup.mem_top + nontriviality α + obtain ⟨g, hg⟩ := exists_ne (1 : α) + have : Subgroup.zpowers g = ⊤ := + (eq_bot_or_eq_top (Subgroup.zpowers g)).resolve_left (Subgroup.zpowers_ne_bot.2 hg) + exact ⟨⟨g, (Subgroup.eq_top_iff' _).1 this⟩⟩ @[to_additive] theorem prime_card [Fintype α] : (Fintype.card α).Prime := by diff --git a/Mathlib/Logic/Function/Defs.lean b/Mathlib/Logic/Function/Defs.lean index 97aa6b6452579..57c703c66bfff 100644 --- a/Mathlib/Logic/Function/Defs.lean +++ b/Mathlib/Logic/Function/Defs.lean @@ -5,7 +5,7 @@ Authors: Leonardo de Moura, Jeremy Avigad, Haitao Zhang -/ import Mathlib.Tactic.AdaptationNote import Mathlib.Tactic.Attr.Register -import Mathlib.Tactic.Basic +import Mathlib.Tactic.Lemma import Mathlib.Tactic.Eqns import Mathlib.Tactic.TypeStar import Batteries.Logic From 48d7a9502b8c17c42528a9c071bd7216e27d70f0 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 11 Aug 2024 23:33:31 +0000 Subject: [PATCH 185/359] feat(QuotientGroup): add `QuotientAddGroup.mk_{nat,int}_mul` (#15596) --- Mathlib/GroupTheory/QuotientGroup.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/GroupTheory/QuotientGroup.lean b/Mathlib/GroupTheory/QuotientGroup.lean index 06dad0c62e756..649d28c78dc0d 100644 --- a/Mathlib/GroupTheory/QuotientGroup.lean +++ b/Mathlib/GroupTheory/QuotientGroup.lean @@ -598,6 +598,20 @@ theorem comap_comap_center {H₁ : Subgroup G} [H₁.Normal] {H₂ : Subgroup (G end QuotientGroup +namespace QuotientAddGroup + +variable {R : Type*} [NonAssocRing R] (N : AddSubgroup R) [N.Normal] + +@[simp] +theorem mk_nat_mul (n : ℕ) (a : R) : ((n * a : R) : R ⧸ N) = n • ↑a := by + rw [← nsmul_eq_mul, mk_nsmul N a n] + +@[simp] +theorem mk_int_mul (n : ℤ) (a : R) : ((n * a : R) : R ⧸ N) = n • ↑a := by + rw [← zsmul_eq_mul, mk_zsmul N a n] + +end QuotientAddGroup + namespace Group open scoped Classical From e5667e513f06ff96a97ffc08d335cfc21aecdc3a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Dupuis?= Date: Sun, 11 Aug 2024 23:33:32 +0000 Subject: [PATCH 186/359] =?UTF-8?q?feat(Integral):=20lemma=20about=20integ?= =?UTF-8?q?rating=20a=20function=20`X=20=E2=86=92=20C(Y,=20E)`=20(#15627)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a function `f : X → C(Y, E)`, we have that `(∫ x, f x ∂μ) y = ∫ x, f x y ∂μ`. --- Mathlib/MeasureTheory/Integral/SetIntegral.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index ff79440e0ab5c..b8cfddf9ab91b 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -1261,6 +1261,18 @@ theorem integral_comp_comm (L : E ≃L[𝕜] F) (φ : X → E) : ∫ x, L (φ x) end ContinuousLinearEquiv +namespace ContinuousMap + +lemma integral_apply [TopologicalSpace Y] [CompactSpace Y] [NormedAddCommGroup E] [NormedSpace ℝ E] + [CompleteSpace E] {f : X → C(Y, E)} (hf : Integrable f μ) (y : Y) : + (∫ x, f x ∂μ) y = ∫ x, f x y ∂μ := by + calc (∫ x, f x ∂μ) y = ContinuousMap.evalCLM ℝ y (∫ x, f x ∂μ) := rfl + _ = ∫ x, ContinuousMap.evalCLM ℝ y (f x) ∂μ := + (ContinuousLinearMap.integral_comp_comm _ hf).symm + _ = _ := rfl + +end ContinuousMap + @[norm_cast] theorem integral_ofReal {f : X → ℝ} : ∫ x, (f x : 𝕜) ∂μ = ↑(∫ x, f x ∂μ) := (@RCLike.ofRealLI 𝕜 _).integral_comp_comm f From 52db037ec27f860489f7aa09462cfaf06a563f66 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 11 Aug 2024 23:33:33 +0000 Subject: [PATCH 187/359] chore(AEDisjoint): generalize 2 lemmas to `Sort*` (#15662) --- Mathlib/MeasureTheory/Measure/AEDisjoint.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/AEDisjoint.lean b/Mathlib/MeasureTheory/Measure/AEDisjoint.lean index d582abce1f308..bbf8c6bbd9839 100644 --- a/Mathlib/MeasureTheory/Measure/AEDisjoint.lean +++ b/Mathlib/MeasureTheory/Measure/AEDisjoint.lean @@ -77,12 +77,12 @@ protected theorem congr (h : AEDisjoint μ s t) (hu : u =ᵐ[μ] s) (hv : v =ᵐ mono_ae h (Filter.EventuallyEq.le hu) (Filter.EventuallyEq.le hv) @[simp] -theorem iUnion_left_iff [Countable ι] {s : ι → Set α} : +theorem iUnion_left_iff {ι : Sort*} [Countable ι] {s : ι → Set α} : AEDisjoint μ (⋃ i, s i) t ↔ ∀ i, AEDisjoint μ (s i) t := by simp only [AEDisjoint, iUnion_inter, measure_iUnion_null_iff] @[simp] -theorem iUnion_right_iff [Countable ι] {t : ι → Set α} : +theorem iUnion_right_iff {ι : Sort*} [Countable ι] {t : ι → Set α} : AEDisjoint μ s (⋃ i, t i) ↔ ∀ i, AEDisjoint μ s (t i) := by simp only [AEDisjoint, inter_iUnion, measure_iUnion_null_iff] From 8410437117b89f8d0e0ca0f8a29e680b17105c44 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 11 Aug 2024 23:33:34 +0000 Subject: [PATCH 188/359] chore(MeasureTheory/Group): move defs to `Defs` (#15673) --- Mathlib.lean | 1 + Mathlib/MeasureTheory/Group/Action.lean | 18 +---- Mathlib/MeasureTheory/Group/Defs.lean | 99 ++++++++++++++++++++++++ Mathlib/MeasureTheory/Group/Measure.lean | 38 --------- 4 files changed, 101 insertions(+), 55 deletions(-) create mode 100644 Mathlib/MeasureTheory/Group/Defs.lean diff --git a/Mathlib.lean b/Mathlib.lean index 7b99ae598c9d8..5c6cfabfa151f 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3159,6 +3159,7 @@ import Mathlib.MeasureTheory.Group.Action import Mathlib.MeasureTheory.Group.AddCircle import Mathlib.MeasureTheory.Group.Arithmetic import Mathlib.MeasureTheory.Group.Convolution +import Mathlib.MeasureTheory.Group.Defs import Mathlib.MeasureTheory.Group.FundamentalDomain import Mathlib.MeasureTheory.Group.GeometryOfNumbers import Mathlib.MeasureTheory.Group.Integral diff --git a/Mathlib/MeasureTheory/Group/Action.lean b/Mathlib/MeasureTheory/Group/Action.lean index cd1b2c959c282..3238971b09711 100644 --- a/Mathlib/MeasureTheory/Group/Action.lean +++ b/Mathlib/MeasureTheory/Group/Action.lean @@ -8,6 +8,7 @@ import Mathlib.Dynamics.Minimal import Mathlib.GroupTheory.GroupAction.Hom import Mathlib.MeasureTheory.Group.MeasurableEquiv import Mathlib.MeasureTheory.Measure.Regular +import Mathlib.MeasureTheory.Group.Defs import Mathlib.Order.Filter.EventuallyConst /-! @@ -29,23 +30,6 @@ universe u v w variable {G : Type u} {M : Type v} {α : Type w} {s : Set α} -/-- A measure `μ : Measure α` is invariant under an additive action of `M` on `α` if for any -measurable set `s : Set α` and `c : M`, the measure of its preimage under `fun x => c +ᵥ x` is equal -to the measure of `s`. -/ -class VAddInvariantMeasure (M α : Type*) [VAdd M α] {_ : MeasurableSpace α} (μ : Measure α) : - Prop where - measure_preimage_vadd : ∀ (c : M) ⦃s : Set α⦄, MeasurableSet s → μ ((fun x => c +ᵥ x) ⁻¹' s) = μ s - -/-- A measure `μ : Measure α` is invariant under a multiplicative action of `M` on `α` if for any -measurable set `s : Set α` and `c : M`, the measure of its preimage under `fun x => c • x` is equal -to the measure of `s`. -/ -@[to_additive, mk_iff smulInvariantMeasure_iff] -class SMulInvariantMeasure (M α : Type*) [SMul M α] {_ : MeasurableSpace α} (μ : Measure α) : - Prop where - measure_preimage_smul : ∀ (c : M) ⦃s : Set α⦄, MeasurableSet s → μ ((fun x => c • x) ⁻¹' s) = μ s - -attribute [to_additive] smulInvariantMeasure_iff - namespace SMulInvariantMeasure @[to_additive] diff --git a/Mathlib/MeasureTheory/Group/Defs.lean b/Mathlib/MeasureTheory/Group/Defs.lean new file mode 100644 index 0000000000000..35caa44a94cae --- /dev/null +++ b/Mathlib/MeasureTheory/Group/Defs.lean @@ -0,0 +1,99 @@ +/- +Copyright (c) 2020 Floris van Doorn. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn, Yury Kudryashov +-/ +import Mathlib.MeasureTheory.Measure.MeasureSpace + +/-! +# Definitions about invariant measures + +In this file we define typeclasses for measures invariant under (scalar) multiplication. + +- `MeasureTheory.SMulInvariantMeasure M α μ` + says that the measure `μ` is invariant under scalar multiplication by `c : M`; +- `MeasureTheory.VAddInvariantMeasure M α μ` is the additive version of this typeclass; +- `MeasureTheory.Measure.IsMulLeftInvariant μ`, `MeasureTheory.Measure.IsMulRightInvariant μ` + say that the measure `μ` is invariant under multiplication on the left and on the right, + respectively. +- `MeasureTheory.Measure.IsAddLeftInvariant μ`, `MeasureTheory.Measure.IsAddRightInvariant μ` + are the additive versions of these typeclasses. + +For basic facts about the first two typeclasses, see `Mathlib/MeasureTheory/Group/Action`. +For facts about left/right-invariant measures, see `Mathlib/MeasureTheory/Group/Measure`. + +## Implementation Notes + +The `smul`/`vadd` typeclasses and the left/right multiplication typeclasses +were defined by different people with different tastes, +so the former explicitly use measures of the preimages, +while the latter use `MeasureTheory.Measure.map`. + +If the left/right multiplication is measurable +(which is the case in most if not all interesting examples), +these definitions are equivalent. + +The definitions that use `MeasureTheory.Measure.map` +imply that the left (resp., right) multiplication is `AEMeasurable`. +-/ + +namespace MeasureTheory + +/-- A measure `μ : Measure α` is invariant under an additive action of `M` on `α` if for any +measurable set `s : Set α` and `c : M`, the measure of its preimage under `fun x => c +ᵥ x` is equal +to the measure of `s`. -/ +class VAddInvariantMeasure (M α : Type*) [VAdd M α] {_ : MeasurableSpace α} (μ : Measure α) : + Prop where + measure_preimage_vadd : ∀ (c : M) ⦃s : Set α⦄, MeasurableSet s → μ ((fun x => c +ᵥ x) ⁻¹' s) = μ s + +/-- A measure `μ : Measure α` is invariant under a multiplicative action of `M` on `α` if for any +measurable set `s : Set α` and `c : M`, the measure of its preimage under `fun x => c • x` is equal +to the measure of `s`. -/ +@[to_additive, mk_iff smulInvariantMeasure_iff] +class SMulInvariantMeasure (M α : Type*) [SMul M α] {_ : MeasurableSpace α} (μ : Measure α) : + Prop where + measure_preimage_smul : ∀ (c : M) ⦃s : Set α⦄, MeasurableSet s → μ ((fun x => c • x) ⁻¹' s) = μ s + +attribute [to_additive] smulInvariantMeasure_iff + +namespace Measure + +variable {G : Type*} [MeasurableSpace G] + +/-- A measure `μ` on a measurable additive group is left invariant + if the measure of left translations of a set are equal to the measure of the set itself. -/ +class IsAddLeftInvariant [Add G] (μ : Measure G) : Prop where + map_add_left_eq_self : ∀ g : G, map (g + ·) μ = μ + +/-- A measure `μ` on a measurable group is left invariant + if the measure of left translations of a set are equal to the measure of the set itself. -/ +@[to_additive existing] +class IsMulLeftInvariant [Mul G] (μ : Measure G) : Prop where + map_mul_left_eq_self : ∀ g : G, map (g * ·) μ = μ + +/-- A measure `μ` on a measurable additive group is right invariant + if the measure of right translations of a set are equal to the measure of the set itself. -/ +class IsAddRightInvariant [Add G] (μ : Measure G) : Prop where + map_add_right_eq_self : ∀ g : G, map (· + g) μ = μ + +/-- A measure `μ` on a measurable group is right invariant + if the measure of right translations of a set are equal to the measure of the set itself. -/ +@[to_additive existing] +class IsMulRightInvariant [Mul G] (μ : Measure G) : Prop where + map_mul_right_eq_self : ∀ g : G, map (· * g) μ = μ + +variable [Mul G] {μ : Measure G} + +@[to_additive] +instance IsMulLeftInvariant.smulInvariantMeasure [IsMulLeftInvariant μ] : + SMulInvariantMeasure G G μ := + ⟨fun _x _s hs => measure_preimage_of_map_eq_self (map_mul_left_eq_self _) hs.nullMeasurableSet⟩ + +@[to_additive] +instance IsMulRightInvariant.toSMulInvariantMeasure_op [μ.IsMulRightInvariant] : + SMulInvariantMeasure Gᵐᵒᵖ G μ := + ⟨fun _x _s hs => measure_preimage_of_map_eq_self (map_mul_right_eq_self _) hs.nullMeasurableSet⟩ + +end Measure + +end MeasureTheory diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 8172325ce37ac..f6a2650672d2e 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -38,34 +38,6 @@ variable {𝕜 G H : Type*} [MeasurableSpace G] [MeasurableSpace H] namespace MeasureTheory -namespace Measure - -/-- A measure `μ` on a measurable additive group is left invariant - if the measure of left translations of a set are equal to the measure of the set itself. -/ -class IsAddLeftInvariant [Add G] (μ : Measure G) : Prop where - map_add_left_eq_self : ∀ g : G, map (g + ·) μ = μ - -/-- A measure `μ` on a measurable group is left invariant - if the measure of left translations of a set are equal to the measure of the set itself. -/ -@[to_additive existing] -class IsMulLeftInvariant [Mul G] (μ : Measure G) : Prop where - map_mul_left_eq_self : ∀ g : G, map (g * ·) μ = μ - -/-- A measure `μ` on a measurable additive group is right invariant - if the measure of right translations of a set are equal to the measure of the set itself. -/ -class IsAddRightInvariant [Add G] (μ : Measure G) : Prop where - map_add_right_eq_self : ∀ g : G, map (· + g) μ = μ - -/-- A measure `μ` on a measurable group is right invariant - if the measure of right translations of a set are equal to the measure of the set itself. -/ -@[to_additive existing] -class IsMulRightInvariant [Mul G] (μ : Measure G) : Prop where - map_mul_right_eq_self : ∀ g : G, map (· * g) μ = μ - -end Measure - -open Measure - section Mul variable [Mul G] {μ : Measure G} @@ -98,16 +70,6 @@ instance isMulRightInvariant_smul_nnreal [IsMulRightInvariant μ] (c : ℝ≥0) IsMulRightInvariant (c • μ) := MeasureTheory.isMulRightInvariant_smul (c : ℝ≥0∞) -@[to_additive] -instance IsMulLeftInvariant.smulInvariantMeasure [IsMulLeftInvariant μ] : - SMulInvariantMeasure G G μ := - ⟨fun _x _s hs => measure_preimage_of_map_eq_self (map_mul_left_eq_self _ _) hs.nullMeasurableSet⟩ - -@[to_additive] -instance IsMulRightInvariant.toSMulInvariantMeasure_op [μ.IsMulRightInvariant] : - SMulInvariantMeasure Gᵐᵒᵖ G μ := - ⟨fun _x _s hs => measure_preimage_of_map_eq_self (map_mul_right_eq_self _ _) hs.nullMeasurableSet⟩ - section MeasurableMul variable [MeasurableMul G] From 444c41c24fe06d70141b61b109ba46b3cda1d7d5 Mon Sep 17 00:00:00 2001 From: Julian Berman Date: Sun, 11 Aug 2024 23:33:35 +0000 Subject: [PATCH 189/359] fix: typo in the cdot linter typing suggestion (#15677) (This was previously showing cdot itself as how to type it). --- Mathlib/Tactic/Linter/Lint.lean | 2 +- test/Lint.lean | 8 ++++---- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index 8c049afef1034..bdd5e856f065d 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -212,7 +212,7 @@ def cdotLinter : Linter where run := withSetOptionIn fun stx => do if (← MonadState.get).messages.hasErrors then return for s in unwanted_cdot stx do - Linter.logLint linter.cdot s m!"Please, use '·' (typed as `\\·`) instead of '{s}' as 'cdot'." + Linter.logLint linter.cdot s m!"Please, use '·' (typed as `\\.`) instead of '{s}' as 'cdot'." initialize addLinter cdotLinter diff --git a/test/Lint.lean b/test/Lint.lean index dce066f67ad12..423b30a42e594 100644 --- a/test/Lint.lean +++ b/test/Lint.lean @@ -60,13 +60,13 @@ end add set_option linter.cdot false in /-- -warning: Please, use '·' (typed as `\·`) instead of '.' as 'cdot'. +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. note: this linter can be disabled with `set_option linter.cdot false` --- -warning: Please, use '·' (typed as `\·`) instead of '.' as 'cdot'. +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. note: this linter can be disabled with `set_option linter.cdot false` --- -warning: Please, use '·' (typed as `\·`) instead of '.' as 'cdot'. +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. note: this linter can be disabled with `set_option linter.cdot false` -/ #guard_msgs in @@ -80,7 +80,7 @@ instance : Inhabited Nat where set_option linter.cdot false in /-- -warning: Please, use '·' (typed as `\·`) instead of '.' as 'cdot'. +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. note: this linter can be disabled with `set_option linter.cdot false` -/ #guard_msgs in From 13fd24632acc75769ac3858a9a8a65aa70108ab2 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sun, 11 Aug 2024 23:33:36 +0000 Subject: [PATCH 190/359] feat(NumberTheory/GaussSum): add two results (#15680) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds ```lean lemma gaussSum_mul {R : Type u} [CommRing R] [Fintype R] {R' : Type v} [CommRing R'] (χ φ : MulChar R R') (ψ : AddChar R R') : gaussSum χ ψ * gaussSum φ ψ = ∑ t : R, ∑ x : R, χ x * φ (t - x) * ψ t ``` and ```lean lemma gaussSum_ne_zero_of_nontrivial {R : Type u} [Field R] [Fintype R] {R' : Type v} [CommRing R'] [IsDomain R'] (h : (Fintype.card R : R') ≠ 0) {χ : MulChar R R'} (hχ : χ ≠ 1) {ψ : AddChar R R'} (hψ : ψ.IsPrimitive) : gaussSum χ ψ ≠ 0 ``` which will be needed in results on Jacobi sums (coming soon). --- Mathlib/NumberTheory/GaussSum.lean | 27 ++++++++++++++++++++++----- 1 file changed, 22 insertions(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/GaussSum.lean b/Mathlib/NumberTheory/GaussSum.lean index 8b6c441cec12a..8dcfe9dbeb4e8 100644 --- a/Mathlib/NumberTheory/GaussSum.lean +++ b/Mathlib/NumberTheory/GaussSum.lean @@ -62,7 +62,6 @@ variable {R' : Type v} [CommRing R'] ### Definition and first properties -/ - /-- Definition of the Gauss sum associated to a multiplicative and an additive character. -/ def gaussSum (χ : MulChar R R') (ψ : AddChar R R') : R' := ∑ a, χ a * ψ a @@ -80,9 +79,24 @@ end GaussSumDef ### The product of two Gauss sums -/ - section GaussSumProd +open Finset in +/-- A formula for the product of two Gauss sums with the same additive character. -/ +lemma gaussSum_mul {R : Type u} [CommRing R] [Fintype R] {R' : Type v} [CommRing R'] + (χ φ : MulChar R R') (ψ : AddChar R R') : + gaussSum χ ψ * gaussSum φ ψ = ∑ t : R, ∑ x : R, χ x * φ (t - x) * ψ t := by + rw [gaussSum, gaussSum, sum_mul_sum] + conv => enter [1, 2, x, 2, x_1]; rw [mul_mul_mul_comm] + simp only [← ψ.map_add_eq_mul] + have sum_eq x : ∑ y : R, χ x * φ y * ψ (x + y) = ∑ y : R, χ x * φ (y - x) * ψ y := by + rw [sum_bij (fun a _ ↦ a + x)] + · simp only [mem_univ, forall_true_left, forall_const] + · simp only [mem_univ, add_left_inj, imp_self, forall_const] + · exact fun b _ ↦ ⟨b - x, mem_univ _, by rw [sub_add_cancel]⟩ + · exact fun a _ ↦ by rw [add_sub_cancel_right, add_comm] + rw [sum_congr rfl fun x _ ↦ sum_eq x, sum_comm] + -- In the following, we need `R` to be a finite field and `R'` to be a domain. variable {R : Type u} [Field R] [Fintype R] {R' : Type v} [CommRing R'] [IsDomain R'] @@ -116,6 +130,12 @@ theorem gaussSum_mul_gaussSum_eq_card {χ : MulChar R R'} (hχ : χ ≠ 1) {ψ : rw [Finset.sum_ite_eq' Finset.univ (1 : R)] simp only [Finset.mem_univ, map_one, one_mul, if_true] +/-- The Gauss sum of a nontrivial character on a finite field does not vanish. -/ +lemma gaussSum_ne_zero_of_nontrivial (h : (Fintype.card R : R') ≠ 0) {χ : MulChar R R'} + (hχ : χ ≠ 1) {ψ : AddChar R R'} (hψ : ψ.IsPrimitive) : + gaussSum χ ψ ≠ 0 := + fun H ↦ h.symm <| zero_mul (gaussSum χ⁻¹ _) ▸ H ▸ gaussSum_mul_gaussSum_eq_card hχ hψ + /-- When `χ` is a nontrivial quadratic character, then the square of `gaussSum χ ψ` is `χ(-1)` times the cardinality of `R`. -/ theorem gaussSum_sq {χ : MulChar R R'} (hχ₁ : χ ≠ 1) (hχ₂ : IsQuadratic χ) @@ -132,7 +152,6 @@ end GaussSumProd ### Gauss sums and Frobenius -/ - section gaussSum_frob variable {R : Type u} [CommRing R] [Fintype R] {R' : Type v} [CommRing R'] @@ -175,7 +194,6 @@ end gaussSum_frob ### Values of quadratic characters -/ - section GaussSumValues variable {R : Type u} [CommRing R] [Fintype R] {R' : Type v} [CommRing R'] [IsDomain R'] @@ -235,7 +253,6 @@ The proof uses the Gauss sum of `χ₈` and a primitive additive character on ` in this way, the result is reduced to `card_pow_char_pow`. -/ - open ZMod /-- For every finite field `F` of odd characteristic, we have `2^(#F/2) = χ₈ #F` in `F`. -/ From f121f99cdd7fea19ac1b0b5f20f7b4e8d1236cc1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 11 Aug 2024 23:33:38 +0000 Subject: [PATCH 191/359] chore(SetTheory/Surreal/Multiplication): fix typos (#15685) Fix some small typos and re-wrap a few lines. --- Mathlib/SetTheory/Surreal/Multiplication.lean | 19 +++++++++---------- 1 file changed, 9 insertions(+), 10 deletions(-) diff --git a/Mathlib/SetTheory/Surreal/Multiplication.lean b/Mathlib/SetTheory/Surreal/Multiplication.lean index 362a8f555a1af..3f931a7db4fd7 100644 --- a/Mathlib/SetTheory/Surreal/Multiplication.lean +++ b/Mathlib/SetTheory/Surreal/Multiplication.lean @@ -30,7 +30,7 @@ In the argument, P3 is stated with four variables `x₁`, `x₂`, `y₁`, `y₂` `y₁ < y₂`, and says that `x₁ * y₂ + x₂ * x₁ < x₁ * y₁ + x₂ * y₂`, which is equivalent to `0 < x₂ - x₁ → 0 < y₂ - y₁ → 0 < (x₂ - x₁) * (y₂ - y₁)`, i.e. `@mul_pos PGame _ (x₂ - x₁) (y₂ - y₁)`. It has to be stated in this form and not in terms of -`mul_pos` because we need to show show P1, P2 and (a specialized form of) P3 simultaneously, and +`mul_pos` because we need to show P1, P2 and (a specialized form of) P3 simultaneously, and for example `P1 x y` will be deduced from P3 with variables taking values simpler than `x` or `y` (among other induction hypotheses), but if you subtract two pregames simpler than `x` or `y`, the result may no longer be simpler. @@ -40,19 +40,19 @@ requires that `y₂ = y` or `-y` and that `y₁` is a left option of `y₂`. Aft shown, a further inductive argument (this time using the `GameAdd` relation) proves P3 in full. Implementation strategy of the inductive argument: we -* extract specialized versions (`IH1`, `IH2`, `IH3`, `IH4` and `IH24`) of the - induction hypothesis that are easier to apply (takes `IsOption` arguments directly), and -* show they are invariant under certain symmetries (permutation and negation of arguments) - and that the induction hypothesis indeed implies the specialized versions. +* extract specialized versions (`IH1`, `IH2`, `IH3`, `IH4` and `IH24`) of the induction hypothesis + that are easier to apply (taking `IsOption` arguments directly), and +* show they are invariant under certain symmetries (permutation and negation of arguments) and that + the induction hypothesis indeed implies the specialized versions. * utilize the symmetries to minimize calculation. The whole proof features a clear separation into lemmas of different roles: * verification of symmetry properties of P and IH (`P3_comm`, `ih1_neg_left`, etc.), -* calculations that connects P1, P2, P3, and inequalities between the product of +* calculations that connect P1, P2, P3, and inequalities between the product of two surreals and its options (`mulOption_lt_iff_P1`, etc.), * specializations of the induction hypothesis (`numeric_option_mul`, `ih1`, `ih1_swap`, `ih₁₂`, `ih4`, etc.), -* application of specialized indution hypothesis +* application of specialized induction hypothesis (`P1_of_ih`, `mul_right_le_of_equiv`, `P3_of_lt`, etc.). ## References @@ -68,8 +68,8 @@ open SetTheory Game PGame WellFounded namespace Surreal.Multiplication -/-- The nontrivial part of P1 in [SchleicherStoll] says that the left options of `x * y` - are less than the right options, and this is the general form of these statements. -/ +/-- The nontrivial part of P1 in [SchleicherStoll] says that the left options of `x * y` are less + than the right options, and this is the general form of these statements. -/ def P1 (x₁ x₂ x₃ y₁ y₂ y₃ : PGame) := ⟦x₁ * y₁⟧ + ⟦x₂ * y₂⟧ - ⟦x₁ * y₂⟧ < ⟦x₃ * y₁⟧ + ⟦x₂ * y₃⟧ - (⟦x₃ * y₃⟧ : Game) @@ -206,7 +206,6 @@ lemma ih1_neg_right : IH1 x y → IH1 x (-y) := /-! #### Specialize `ih` to obtain specialized induction hypotheses for P1 -/ - lemma numeric_option_mul (ih : ∀ a, ArgsRel a (Args.P1 x y) → P124 a) (h : IsOption x' x) : (x' * y).Numeric := ih (Args.P1 x' y) (TransGen.single <| cutExpand_pair_left h) From d14d0e5bc3b0c2489d362aaa498bfcf43934afa6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 11 Aug 2024 23:33:39 +0000 Subject: [PATCH 192/359] chore: clean up some naming around `Matrix.stdBasisMatrix` (#15693) `std_basis` is a lean3-ism. In future we might want to rename this to `Matrix.single`. --- Mathlib/Data/Matrix/Basis.lean | 19 +++++++++++++------ Mathlib/RingTheory/MatrixAlgebra.lean | 9 ++++++--- Mathlib/RingTheory/PolynomialAlgebra.lean | 2 +- 3 files changed, 20 insertions(+), 10 deletions(-) diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index 104482a97f690..1204f7bb9f157 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -56,7 +56,7 @@ theorem mulVec_stdBasisMatrix [Fintype m] (i : n) (j : m) (c : α) (x : m → α · simp simp [h, h.symm] -theorem matrix_eq_sum_std_basis [Fintype m] [Fintype n] (x : Matrix m n α) : +theorem matrix_eq_sum_stdBasisMatrix [Fintype m] [Fintype n] (x : Matrix m n α) : x = ∑ i : m, ∑ j : n, stdBasisMatrix i j (x i j) := by ext i j; symm iterate 2 rw [Finset.sum_apply] @@ -67,15 +67,22 @@ theorem matrix_eq_sum_std_basis [Fintype m] [Fintype n] (x : Matrix m n α) : -- Porting note(#12717): `simp` seems unwilling to apply `Fintype.sum_apply` simp (config := { unfoldPartialApp := true }) [stdBasisMatrix, (Fintype.sum_apply), hj'] +@[deprecated (since := "2024-08-11")] alias matrix_eq_sum_std_basis := matrix_eq_sum_stdBasisMatrix + +theorem stdBasisMatrix_eq_single_vecMulVec_single (i : m) (j : n) : + stdBasisMatrix i j (1 : α) = vecMulVec (Pi.single i 1) (Pi.single j 1) := by + ext i' j' + -- Porting note: lean3 didn't apply `mul_ite`. + simp [-mul_ite, stdBasisMatrix, vecMulVec, ite_and, Pi.single_apply, eq_comm] + -- TODO: tie this up with the `Basis` machinery of linear algebra -- this is not completely trivial because we are indexing by two types, instead of one --- TODO: add `std_basis_vec` +@[deprecated stdBasisMatrix_eq_single_vecMulVec_single (since := "2024-08-11")] theorem std_basis_eq_basis_mul_basis (i : m) (j : n) : stdBasisMatrix i j (1 : α) = vecMulVec (fun i' => ite (i = i') 1 0) fun j' => ite (j = j') 1 0 := by - ext i' j' - -- Porting note: lean3 didn't apply `mul_ite`. - simp [-mul_ite, stdBasisMatrix, vecMulVec, ite_and] + rw [stdBasisMatrix_eq_single_vecMulVec_single] + congr! with i <;> simp only [Pi.single_apply, eq_comm] -- todo: the old proof used fintypes, I don't know `Finsupp` but this feels generalizable @[elab_as_elim] @@ -83,7 +90,7 @@ protected theorem induction_on' [Finite m] [Finite n] {P : Matrix m n α → Pro (h_zero : P 0) (h_add : ∀ p q, P p → P q → P (p + q)) (h_std_basis : ∀ (i : m) (j : n) (x : α), P (stdBasisMatrix i j x)) : P M := by cases nonempty_fintype m; cases nonempty_fintype n - rw [matrix_eq_sum_std_basis M, ← Finset.sum_product'] + rw [matrix_eq_sum_stdBasisMatrix M, ← Finset.sum_product'] apply Finset.sum_induction _ _ h_add h_zero · intros apply h_std_basis diff --git a/Mathlib/RingTheory/MatrixAlgebra.lean b/Mathlib/RingTheory/MatrixAlgebra.lean index 873a460abb979..8525a367e4fcf 100644 --- a/Mathlib/RingTheory/MatrixAlgebra.lean +++ b/Mathlib/RingTheory/MatrixAlgebra.lean @@ -95,7 +95,7 @@ theorem invFun_algebraMap (M : Matrix n n R) : invFun R A n (M.map (algebraMap R dsimp [invFun] simp only [Algebra.algebraMap_eq_smul_one, smul_tmul, ← tmul_sum, mul_boole] congr - conv_rhs => rw [matrix_eq_sum_std_basis M] + conv_rhs => rw [matrix_eq_sum_stdBasisMatrix M] convert Finset.sum_product (β := Matrix n n R); simp theorem right_inv (M : Matrix n n A) : (toFunAlgHom R A n) (invFun R A n M) = M := by @@ -103,7 +103,7 @@ theorem right_inv (M : Matrix n n A) : (toFunAlgHom R A n) (invFun R A n M) = M mul_boole, toFunAlgHom_apply, RingHom.map_zero, RingHom.map_one, Matrix.map_apply, Pi.smul_def] convert Finset.sum_product (β := Matrix n n A) - conv_lhs => rw [matrix_eq_sum_std_basis M] + conv_lhs => rw [matrix_eq_sum_stdBasisMatrix M] refine Finset.sum_congr rfl fun i _ => Finset.sum_congr rfl fun j _ => Matrix.ext fun a b => ?_ simp only [stdBasisMatrix, smul_apply, Matrix.map_apply] split_ifs <;> aesop @@ -144,11 +144,14 @@ theorem matrixEquivTensor_apply (M : Matrix n n A) : -- Porting note: short circuiting simplifier from simplifying left hand side @[simp (high)] -theorem matrixEquivTensor_apply_std_basis (i j : n) (x : A) : +theorem matrixEquivTensor_apply_stdBasisMatrix (i j : n) (x : A) : matrixEquivTensor R A n (stdBasisMatrix i j x) = x ⊗ₜ stdBasisMatrix i j 1 := by have t : ∀ p : n × n, i = p.1 ∧ j = p.2 ↔ p = (i, j) := by aesop simp [ite_tmul, t, stdBasisMatrix] +@[deprecated (since := "2024-08-11")] alias matrixEquivTensor_apply_std_basis := + matrixEquivTensor_apply_stdBasisMatrix + @[simp] theorem matrixEquivTensor_apply_symm (a : A) (M : Matrix n n R) : (matrixEquivTensor R A n).symm (a ⊗ₜ M) = M.map fun x => a * algebraMap R A x := diff --git a/Mathlib/RingTheory/PolynomialAlgebra.lean b/Mathlib/RingTheory/PolynomialAlgebra.lean index b8f5983c86a8e..d21058ee4247c 100644 --- a/Mathlib/RingTheory/PolynomialAlgebra.lean +++ b/Mathlib/RingTheory/PolynomialAlgebra.lean @@ -222,7 +222,7 @@ open Finset unseal Algebra.TensorProduct.mul in theorem matPolyEquiv_coeff_apply_aux_1 (i j : n) (k : ℕ) (x : R) : matPolyEquiv (stdBasisMatrix i j <| monomial k x) = monomial k (stdBasisMatrix i j x) := by - simp only [matPolyEquiv, AlgEquiv.trans_apply, matrixEquivTensor_apply_std_basis] + simp only [matPolyEquiv, AlgEquiv.trans_apply, matrixEquivTensor_apply_stdBasisMatrix] apply (polyEquivTensor R (Matrix n n R)).injective simp only [AlgEquiv.apply_symm_apply,Algebra.TensorProduct.comm_tmul, polyEquivTensor_apply, eval₂_monomial] From 0b82895500c544ded438039bab8faa59e808918b Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Sun, 11 Aug 2024 23:33:40 +0000 Subject: [PATCH 193/359] feat(Algebra/Group/Subgroup/Basic): `inclusion_inj` (#15699) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add the lemma ```lean @[to_additive (attr := simp)] lemma inclusion_inj {H K : Subgroup G} (h : H ≤ K) {x y : H} : inclusion h x = inclusion h y ↔ x = y := (inclusion_injective h).eq_iff ``` following usual practice for injectivity lemmas. From AperiodicMonotilesLean. --- Mathlib/Algebra/Group/Subgroup/Basic.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index f60917ae388d5..e7386db452760 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -623,6 +623,11 @@ theorem coe_inclusion {H K : Subgroup G} {h : H ≤ K} (a : H) : (inclusion h a theorem inclusion_injective {H K : Subgroup G} (h : H ≤ K) : Function.Injective <| inclusion h := Set.inclusion_injective h +@[to_additive (attr := simp)] +lemma inclusion_inj {H K : Subgroup G} (h : H ≤ K) {x y : H} : + inclusion h x = inclusion h y ↔ x = y := + (inclusion_injective h).eq_iff + @[to_additive (attr := simp)] theorem subtype_comp_inclusion {H K : Subgroup G} (hH : H ≤ K) : K.subtype.comp (inclusion hH) = H.subtype := From c13b15d0da693837a0221e65cfcb9ba3cf4f949f Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Sun, 11 Aug 2024 23:33:41 +0000 Subject: [PATCH 194/359] feat(GroupTheory/Index): `exists_pow_mem_of_index_ne_zero` (#15701) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add the lemma ```lean @[to_additive] lemma exists_pow_mem_of_index_ne_zero (h : H.index ≠ 0) (a : G) : ∃ n, 0 < n ∧ n ≤ H.index ∧ a ^ n ∈ H := by ``` and three similar lemmas. These are essentially more general versions of `sq_mem_of_index_two` for arbitrary finite index. From AperiodicMonotilesLean. --- Mathlib/GroupTheory/Index.lean | 50 ++++++++++++++++++++++++++++++++++ 1 file changed, 50 insertions(+) diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 20a4102ae16ee..ff288a79a6b5f 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -397,6 +397,56 @@ lemma finite_quotient_of_pretransitive_of_index_ne_zero {X : Type*} [MulAction G have := (MulAction.pretransitive_iff_subsingleton_quotient G X).1 inferInstance exact finite_quotient_of_finite_quotient_of_index_ne_zero hi +@[to_additive] +lemma exists_pow_mem_of_index_ne_zero (h : H.index ≠ 0) (a : G) : + ∃ n, 0 < n ∧ n ≤ H.index ∧ a ^ n ∈ H := by + suffices ∃ n₁ n₂, n₁ < n₂ ∧ n₂ ≤ H.index ∧ ((a ^ n₂ : G) : G ⧸ H) = ((a ^ n₁ : G) : G ⧸ H) by + rcases this with ⟨n₁, n₂, hlt, hle, he⟩ + refine ⟨n₂ - n₁, by omega, by omega, ?_⟩ + rw [eq_comm, QuotientGroup.eq, ← zpow_natCast, ← zpow_natCast, ← zpow_neg, ← zpow_add, + add_comm] at he + rw [← zpow_natCast] + convert he + omega + suffices ∃ n₁ n₂, n₁ ≠ n₂ ∧ n₁ ≤ H.index ∧ n₂ ≤ H.index ∧ + ((a ^ n₂ : G) : G ⧸ H) = ((a ^ n₁ : G) : G ⧸ H) by + rcases this with ⟨n₁, n₂, hne, hle₁, hle₂, he⟩ + rcases hne.lt_or_lt with hlt | hlt + · exact ⟨n₁, n₂, hlt, hle₂, he⟩ + · exact ⟨n₂, n₁, hlt, hle₁, he.symm⟩ + by_contra hc + simp_rw [not_exists] at hc + let f : (Set.Icc 0 H.index) → G ⧸ H := fun n ↦ (a ^ (n : ℕ) : G) + have hf : Function.Injective f := by + rintro ⟨n₁, h₁, hle₁⟩ ⟨n₂, h₂, hle₂⟩ he + have hc' := hc n₁ n₂ + dsimp only [f] at he + simpa [hle₁, hle₂, he] using hc' + have := (fintypeOfIndexNeZero h).finite + have hcard := Finite.card_le_of_injective f hf + simp [← index_eq_card] at hcard + +@[to_additive] +lemma exists_pow_mem_of_relindex_ne_zero (h : H.relindex K ≠ 0) {a : G} (ha : a ∈ K) : + ∃ n, 0 < n ∧ n ≤ H.relindex K ∧ a ^ n ∈ H ⊓ K := by + rcases exists_pow_mem_of_index_ne_zero h ⟨a, ha⟩ with ⟨n, hlt, hle, he⟩ + refine ⟨n, hlt, hle, ?_⟩ + simpa [pow_mem ha, mem_subgroupOf] using he + +@[to_additive] +lemma pow_mem_of_index_ne_zero_of_dvd (h : H.index ≠ 0) (a : G) {n : ℕ} + (hn : ∀ m, 0 < m → m ≤ H.index → m ∣ n) : a ^ n ∈ H := by + rcases exists_pow_mem_of_index_ne_zero h a with ⟨m, hlt, hle, he⟩ + rcases hn m hlt hle with ⟨k, rfl⟩ + rw [pow_mul] + exact pow_mem he _ + +@[to_additive] +lemma pow_mem_of_relindex_ne_zero_of_dvd (h : H.relindex K ≠ 0) {a : G} (ha : a ∈ K) {n : ℕ} + (hn : ∀ m, 0 < m → m ≤ H.relindex K → m ∣ n) : a ^ n ∈ H ⊓ K := by + convert pow_mem_of_index_ne_zero_of_dvd h ⟨a, ha⟩ hn + simp [pow_mem ha, mem_subgroupOf] + section FiniteIndex variable (H K) From 52853729e61ad3c60e989ccd9977dc64c8c12c5e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 11 Aug 2024 23:33:42 +0000 Subject: [PATCH 195/359] chore(SetTheory/Game/Basic): fix Game.PGame namespace (#15708) We move the `Fuzzy` definition so that the `PGame` namespace doesn't need to be opened twice. We also move the infix notation definition a bit later down the file. --- Mathlib/SetTheory/Game/Basic.lean | 41 ++++++++++++++++----------- Mathlib/SetTheory/Game/Impartial.lean | 8 +++--- 2 files changed, 28 insertions(+), 21 deletions(-) diff --git a/Mathlib/SetTheory/Game/Basic.lean b/Mathlib/SetTheory/Game/Basic.lean index 20a32c27731a5..175b69c1a18db 100644 --- a/Mathlib/SetTheory/Game/Basic.lean +++ b/Mathlib/SetTheory/Game/Basic.lean @@ -101,20 +101,24 @@ If `0 ⧏ x` (less or fuzzy with), then Left can win `x` as the first player. -/ def LF : Game → Game → Prop := Quotient.lift₂ PGame.LF fun _ _ _ _ hx hy => propext (lf_congr hx hy) -local infixl:50 " ⧏ " => LF - /-- On `Game`, simp-normal inequalities should use as few negations as possible. -/ @[simp] -theorem not_le : ∀ {x y : Game}, ¬x ≤ y ↔ y ⧏ x := by +theorem not_le : ∀ {x y : Game}, ¬x ≤ y ↔ Game.LF y x := by rintro ⟨x⟩ ⟨y⟩ exact PGame.not_le /-- On `Game`, simp-normal inequalities should use as few negations as possible. -/ @[simp] -theorem not_lf : ∀ {x y : Game}, ¬x ⧏ y ↔ y ≤ x := by +theorem not_lf : ∀ {x y : Game}, ¬Game.LF x y ↔ y ≤ x := by rintro ⟨x⟩ ⟨y⟩ exact PGame.not_lf +/-- The fuzzy, confused, or incomparable relation on games. + +If `x ‖ 0`, then the first player can always win `x`. -/ +def Fuzzy : Game → Game → Prop := + Quotient.lift₂ PGame.Fuzzy fun _ _ _ _ hx hy => propext (fuzzy_congr hx hy) + -- Porting note: had to replace ⧏ with LF, otherwise cannot differentiate with the operator on PGame instance : IsTrichotomous Game LF := ⟨by @@ -126,31 +130,34 @@ instance : IsTrichotomous Game LF := /-! It can be useful to use these lemmas to turn `PGame` inequalities into `Game` inequalities, as the `AddCommGroup` structure on `Game` often simplifies many proofs. -/ +end Game + +namespace PGame + -- Porting note: In a lot of places, I had to add explicitely that the quotient element was a Game. -- In Lean4, quotients don't have the setoid as an instance argument, -- but as an explicit argument, see https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/confusion.20between.20equivalence.20and.20instance.20setoid/near/360822354 -theorem PGame.le_iff_game_le {x y : PGame} : x ≤ y ↔ (⟦x⟧ : Game) ≤ ⟦y⟧ := +theorem le_iff_game_le {x y : PGame} : x ≤ y ↔ (⟦x⟧ : Game) ≤ ⟦y⟧ := Iff.rfl -theorem PGame.lf_iff_game_lf {x y : PGame} : PGame.LF x y ↔ ⟦x⟧ ⧏ ⟦y⟧ := +theorem lf_iff_game_lf {x y : PGame} : x ⧏ y ↔ Game.LF ⟦x⟧ ⟦y⟧ := Iff.rfl -theorem PGame.lt_iff_game_lt {x y : PGame} : x < y ↔ (⟦x⟧ : Game) < ⟦y⟧ := +theorem lt_iff_game_lt {x y : PGame} : x < y ↔ (⟦x⟧ : Game) < ⟦y⟧ := Iff.rfl -theorem PGame.equiv_iff_game_eq {x y : PGame} : x ≈ y ↔ (⟦x⟧ : Game) = ⟦y⟧ := +theorem equiv_iff_game_eq {x y : PGame} : x ≈ y ↔ (⟦x⟧ : Game) = ⟦y⟧ := (@Quotient.eq' _ _ x y).symm -/-- The fuzzy, confused, or incomparable relation on games. +theorem fuzzy_iff_game_fuzzy {x y : PGame} : x ‖ y ↔ Game.Fuzzy ⟦x⟧ ⟦y⟧ := + Iff.rfl -If `x ‖ 0`, then the first player can always win `x`. -/ -def Fuzzy : Game → Game → Prop := - Quotient.lift₂ PGame.Fuzzy fun _ _ _ _ hx hy => propext (fuzzy_congr hx hy) +end PGame -local infixl:50 " ‖ " => Fuzzy +namespace Game -theorem PGame.fuzzy_iff_game_fuzzy {x y : PGame} : PGame.Fuzzy x y ↔ ⟦x⟧ ‖ ⟦y⟧ := - Iff.rfl +local infixl:50 " ⧏ " => LF +local infixl:50 " ‖ " => Fuzzy instance covariantClass_add_le : CovariantClass Game Game (· + ·) (· ≤ ·) := ⟨by @@ -225,8 +232,8 @@ theorem quot_sub (a b : PGame) : ⟦a - b⟧ = (⟦a⟧ : Game) - ⟦b⟧ := theorem quot_eq_of_mk'_quot_eq {x y : PGame} (L : x.LeftMoves ≃ y.LeftMoves) (R : x.RightMoves ≃ y.RightMoves) (hl : ∀ i, (⟦x.moveLeft i⟧ : Game) = ⟦y.moveLeft (L i)⟧) (hr : ∀ j, (⟦x.moveRight j⟧ : Game) = ⟦y.moveRight (R j)⟧) : (⟦x⟧ : Game) = ⟦y⟧ := by - exact Quot.sound (equiv_of_mk_equiv L R (fun _ => Game.PGame.equiv_iff_game_eq.2 (hl _)) - (fun _ => Game.PGame.equiv_iff_game_eq.2 (hr _))) + exact Quot.sound (equiv_of_mk_equiv L R (fun _ => equiv_iff_game_eq.2 (hl _)) + (fun _ => equiv_iff_game_eq.2 (hr _))) /-! Multiplicative operations can be defined at the level of pre-games, but to prove their properties we need to use the abelian group structure of games. diff --git a/Mathlib/SetTheory/Game/Impartial.lean b/Mathlib/SetTheory/Game/Impartial.lean index bd034de0b9af6..815f1fae0b646 100644 --- a/Mathlib/SetTheory/Game/Impartial.lean +++ b/Mathlib/SetTheory/Game/Impartial.lean @@ -141,14 +141,14 @@ theorem mk'_add_self : (⟦G⟧ : Quotient setoid) + ⟦G⟧ = 0 := /-- This lemma doesn't require `H` to be impartial. -/ theorem equiv_iff_add_equiv_zero (H : PGame) : (H ≈ G) ↔ (H + G ≈ 0) := by - rw [Game.PGame.equiv_iff_game_eq, ← @add_right_cancel_iff _ _ _ ⟦G⟧, mk'_add_self, ← quot_add, - Game.PGame.equiv_iff_game_eq] + rw [equiv_iff_game_eq, ← @add_right_cancel_iff _ _ _ ⟦G⟧, mk'_add_self, ← quot_add, + equiv_iff_game_eq] rfl /-- This lemma doesn't require `H` to be impartial. -/ theorem equiv_iff_add_equiv_zero' (H : PGame) : (G ≈ H) ↔ (G + H ≈ 0) := by - rw [Game.PGame.equiv_iff_game_eq, ← @add_left_cancel_iff _ _ _ ⟦G⟧, mk'_add_self, ← quot_add, - Game.PGame.equiv_iff_game_eq] + rw [equiv_iff_game_eq, ← @add_left_cancel_iff _ _ _ ⟦G⟧, mk'_add_self, ← quot_add, + equiv_iff_game_eq] exact ⟨Eq.symm, Eq.symm⟩ theorem le_zero_iff {G : PGame} [G.Impartial] : G ≤ 0 ↔ 0 ≤ G := by From 055d766c68bd10e70d7adf8fe87f44add3d806fa Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 12 Aug 2024 00:35:48 +0000 Subject: [PATCH 196/359] feat(GroupTheory/GroupAction/Basic): `orbitRel_subgroupOf` (#15505) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add the lemma ```lean @[to_additive] lemma orbitRel_subgroupOf (H K : Subgroup G) : orbitRel (H.subgroupOf K) α = orbitRel (H ⊓ K : Subgroup G) α := by ``` From AperiodicMonotilesLean. --- Mathlib/GroupTheory/GroupAction/Basic.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index 02cadaba567a8..71bc41b29ddaf 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -378,6 +378,25 @@ lemma orbitRel_r_apply {a b : α} : (orbitRel G _).r a b ↔ a ∈ orbit G b := lemma orbitRel_subgroup_le (H : Subgroup G) : orbitRel H α ≤ orbitRel G α := Setoid.le_def.2 mem_orbit_of_mem_orbit_subgroup +@[to_additive] +lemma orbitRel_subgroupOf (H K : Subgroup G) : + orbitRel (H.subgroupOf K) α = orbitRel (H ⊓ K : Subgroup G) α := by + rw [← Subgroup.subgroupOf_map_subtype] + ext x + simp_rw [orbitRel_apply] + refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ + · rcases h with ⟨⟨gv, gp⟩, rfl⟩ + simp only [Submonoid.mk_smul] + refine mem_orbit _ (⟨gv, ?_⟩ : Subgroup.map K.subtype (H.subgroupOf K)) + simpa using gp + · rcases h with ⟨⟨gv, gp⟩, rfl⟩ + simp only [Submonoid.mk_smul] + simp only [Subgroup.subgroupOf_map_subtype, Subgroup.mem_inf] at gp + refine mem_orbit _ (⟨⟨gv, ?_⟩, ?_⟩ : H.subgroupOf K) + · exact gp.2 + · simp only [Subgroup.mem_subgroupOf] + exact gp.1 + /-- 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 From e5d297f01115fdd418d45ee9a0f4ea5df7639835 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 12 Aug 2024 02:13:44 +0000 Subject: [PATCH 197/359] chore(*): use `getElem` instead of `List.get`/`List.nthLe` (#15455) Deletions: - List.toFinsupp_apply_lt' --- Mathlib/Algebra/BigOperators/Fin.lean | 8 ++-- Mathlib/Algebra/BigOperators/Group/List.lean | 4 +- Mathlib/Analysis/Analytic/Composition.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Mul.lean | 38 +++++++++---------- Mathlib/Combinatorics/Young/YoungDiagram.lean | 14 +++---- Mathlib/Data/Finset/Sort.lean | 17 ++++----- Mathlib/Data/List/Basic.lean | 22 +++++++---- Mathlib/Data/List/GetD.lean | 22 +++++++---- Mathlib/Data/List/ToFinsupp.lean | 13 ++----- Mathlib/Data/Nat/Nth.lean | 2 +- Mathlib/GroupTheory/Coxeter/Inversion.lean | 16 ++++---- .../LinearAlgebra/Matrix/Transvection.lean | 22 +++++++++-- Mathlib/Order/RelSeries.lean | 4 +- .../RingTheory/Regular/RegularSequence.lean | 2 +- Mathlib/SetTheory/Game/PGame.lean | 8 ++-- 15 files changed, 106 insertions(+), 88 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Fin.lean b/Mathlib/Algebra/BigOperators/Fin.lean index e8acaf6380069..328d5d841314d 100644 --- a/Mathlib/Algebra/BigOperators/Fin.lean +++ b/Mathlib/Algebra/BigOperators/Fin.lean @@ -421,18 +421,18 @@ end CommMonoid @[to_additive] theorem alternatingProd_eq_finset_prod {G : Type*} [CommGroup G] : - ∀ (L : List G), alternatingProd L = ∏ i : Fin L.length, L.get i ^ (-1 : ℤ) ^ (i : ℕ) + ∀ (L : List G), alternatingProd L = ∏ i : Fin L.length, L[i] ^ (-1 : ℤ) ^ (i : ℕ) | [] => by rw [alternatingProd, Finset.prod_eq_one] rintro ⟨i, ⟨⟩⟩ | g::[] => by - show g = ∏ i : Fin 1, [g].get i ^ (-1 : ℤ) ^ (i : ℕ) + show g = ∏ i : Fin 1, [g][i] ^ (-1 : ℤ) ^ (i : ℕ) rw [Fin.prod_univ_succ]; simp | g::h::L => calc g * h⁻¹ * L.alternatingProd - = g * h⁻¹ * ∏ i : Fin L.length, L.get i ^ (-1 : ℤ) ^ (i : ℕ) := + = g * h⁻¹ * ∏ i : Fin L.length, L[i] ^ (-1 : ℤ) ^ (i : ℕ) := congr_arg _ (alternatingProd_eq_finset_prod _) - _ = ∏ i : Fin (L.length + 2), List.get (g::h::L) i ^ (-1 : ℤ) ^ (i : ℕ) := by + _ = ∏ i : Fin (L.length + 2), (g::h::L)[i] ^ (-1 : ℤ) ^ (i : ℕ) := by { rw [Fin.prod_univ_succ, Fin.prod_univ_succ, mul_assoc] simp [pow_add]} diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 9c6cfd3293ff3..6b3d964510bb8 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -414,7 +414,7 @@ theorem prod_reverse_noncomm : ∀ L : List G, L.reverse.prod = (L.map fun x => @[to_additive (attr := simp) "Counterpart to `List.sum_take_succ` when we have a negation operation"] theorem prod_drop_succ : - ∀ (L : List G) (i : ℕ) (p), (L.drop (i + 1)).prod = (L.get ⟨i, p⟩)⁻¹ * (L.drop i).prod + ∀ (L : List G) (i : ℕ) (p : i < L.length), (L.drop (i + 1)).prod = L[i]⁻¹ * (L.drop i).prod | [], i, p => False.elim (Nat.not_lt_zero _ p) | x :: xs, 0, _ => by simp | x :: xs, i + 1, p => prod_drop_succ xs i _ @@ -459,7 +459,7 @@ theorem prod_range_div (n : ℕ) (f : ℕ → G) : /-- Alternative version of `List.prod_set` when the list is over a group -/ @[to_additive "Alternative version of `List.sum_set` when the list is over a group"] theorem prod_set' (L : List G) (n : ℕ) (a : G) : - (L.set n a).prod = L.prod * if hn : n < L.length then (L.get ⟨n, hn⟩)⁻¹ * a else 1 := by + (L.set n a).prod = L.prod * if hn : n < L.length then L[n]⁻¹ * a else 1 := by refine (prod_set L n a).trans ?_ split_ifs with hn · rw [mul_comm _ a, mul_assoc a, prod_drop_succ L n hn, mul_comm _ (drop n L).prod, ← diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index d71cf4e190305..5cf9bd1117d0b 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -379,7 +379,7 @@ theorem comp_id (p : FormalMultilinearSeries 𝕜 E F) : p.comp (id 𝕜 E) = p intro b _ hb obtain ⟨k, hk, lt_k⟩ : ∃ (k : ℕ), k ∈ Composition.blocks b ∧ 1 < k := Composition.ne_ones_iff.1 hb - obtain ⟨i, hi⟩ : ∃ (i : Fin b.blocks.length), b.blocks.get i = k := + obtain ⟨i, hi⟩ : ∃ (i : Fin b.blocks.length), b.blocks[i] = k := List.get_of_mem hk let j : Fin b.length := ⟨i.val, b.blocks_length ▸ i.prop⟩ diff --git a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean index f25414a4c94c0..9c79b4bf2279c 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean @@ -550,7 +550,7 @@ variable {ι : Type*} {𝔸 𝔸' : Type*} [NormedRing 𝔸] [NormedCommRing theorem hasStrictFDerivAt_list_prod' [Fintype ι] {l : List ι} {x : ι → 𝔸} : HasStrictFDerivAt (𝕜 := 𝕜) (fun x ↦ (l.map x).prod) (∑ i : Fin l.length, ((l.take i).map x).prod • - smulRight (proj (l.get i)) ((l.drop (.succ i)).map x).prod) x := by + smulRight (proj l[i]) ((l.drop (.succ i)).map x).prod) x := by induction l with | nil => simp [hasStrictFDerivAt_const] | cons a l IH => @@ -570,7 +570,7 @@ theorem hasStrictFDerivAt_list_prod_finRange' {n : ℕ} {x : Fin n → 𝔸} : theorem hasStrictFDerivAt_list_prod_attach' [DecidableEq ι] {l : List ι} {x : {i // i ∈ l} → 𝔸} : HasStrictFDerivAt (𝕜 := 𝕜) (fun x ↦ (l.attach.map x).prod) (∑ i : Fin l.length, ((l.attach.take i).map x).prod • - smulRight (proj (l.attach.get (i.cast l.length_attach.symm))) + smulRight (proj l.attach[i.cast l.length_attach.symm]) ((l.attach.drop (.succ i)).map x).prod) x := hasStrictFDerivAt_list_prod'.congr_fderiv <| Eq.symm <| Finset.sum_equiv (finCongr l.length_attach.symm) (by simp) (by simp) @@ -579,7 +579,7 @@ theorem hasStrictFDerivAt_list_prod_attach' [DecidableEq ι] {l : List ι} {x : theorem hasFDerivAt_list_prod' [Fintype ι] {l : List ι} {x : ι → 𝔸'} : HasFDerivAt (𝕜 := 𝕜) (fun x ↦ (l.map x).prod) (∑ i : Fin l.length, ((l.take i).map x).prod • - smulRight (proj (l.get i)) ((l.drop (.succ i)).map x).prod) x := + smulRight (proj l[i]) ((l.drop (.succ i)).map x).prod) x := hasStrictFDerivAt_list_prod'.hasFDerivAt @[fun_prop] @@ -593,7 +593,7 @@ theorem hasFDerivAt_list_prod_finRange' {n : ℕ} {x : Fin n → 𝔸} : theorem hasFDerivAt_list_prod_attach' [DecidableEq ι] {l : List ι} {x : {i // i ∈ l} → 𝔸} : HasFDerivAt (𝕜 := 𝕜) (fun x ↦ (l.attach.map x).prod) (∑ i : Fin l.length, ((l.attach.take i).map x).prod • - smulRight (proj (l.attach.get (i.cast l.length_attach.symm))) + smulRight (proj l.attach[i.cast l.length_attach.symm]) ((l.attach.drop (.succ i)).map x).prod) x := hasStrictFDerivAt_list_prod_attach'.hasFDerivAt @@ -606,11 +606,11 @@ For `NormedCommRing 𝔸'`, can rewrite as `Multiset` using `Multiset.prod_coe`. theorem hasStrictFDerivAt_list_prod [DecidableEq ι] [Fintype ι] {l : List ι} {x : ι → 𝔸'} : HasStrictFDerivAt (𝕜 := 𝕜) (fun x ↦ (l.map x).prod) (l.map fun i ↦ ((l.erase i).map x).prod • proj i).sum x := by - refine .congr_fderiv hasStrictFDerivAt_list_prod' ?_ + refine hasStrictFDerivAt_list_prod'.congr_fderiv ?_ conv_rhs => arg 1; arg 2; rw [← List.finRange_map_get l] simp only [List.map_map, ← List.sum_toFinset _ (List.nodup_finRange _), List.toFinset_finRange, - Function.comp_def, ((List.erase_get _).map _).prod_eq, List.eraseIdx_eq_take_drop_succ, - List.map_append, List.prod_append] + Function.comp_def, ((List.erase_getElem _).map _).prod_eq, List.eraseIdx_eq_take_drop_succ, + List.map_append, List.prod_append, List.get_eq_getElem, Fin.getElem_fin, Nat.succ_eq_add_one] exact Finset.sum_congr rfl fun i _ ↦ by ext; simp only [smul_apply, smulRight_apply, smul_eq_mul]; ring @@ -642,12 +642,12 @@ theorem HasStrictFDerivAt.list_prod' {l : List ι} {x : E} (h : ∀ i ∈ l, HasStrictFDerivAt (f i ·) (f' i) x) : HasStrictFDerivAt (fun x ↦ (l.map (f · x)).prod) (∑ i : Fin l.length, ((l.take i).map (f · x)).prod • - smulRight (f' (l.get i)) ((l.drop (.succ i)).map (f · x)).prod) x := by + 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 (hasStrictFDerivAt_list_prod_finRange'.comp x - (hasStrictFDerivAt_pi.mpr fun i ↦ h (l.get i) (l.get_mem i i.isLt))) ?_ + (hasStrictFDerivAt_pi.mpr fun i ↦ h l[i] (l.getElem_mem ..))) ?_ ext m - simp [← Function.comp_def (f · x) (l.get ·), ← List.map_map, List.map_take, List.map_drop] + simp [← List.map_map] /-- Unlike `HasFDerivAt.finset_prod`, supports non-commutative multiply and duplicate elements. @@ -657,39 +657,37 @@ theorem HasFDerivAt.list_prod' {l : List ι} {x : E} (h : ∀ i ∈ l, HasFDerivAt (f i ·) (f' i) x) : HasFDerivAt (fun x ↦ (l.map (f · x)).prod) (∑ i : Fin l.length, ((l.take i).map (f · x)).prod • - smulRight (f' (l.get i)) ((l.drop (.succ i)).map (f · x)).prod) x := by + 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.get i) (l.get_mem i i.isLt))) ?_ + (hasFDerivAt_pi.mpr fun i ↦ h l[i] (l.getElem_mem i i.isLt))) ?_ ext m - simp [← Function.comp_def (f · x) (l.get ·), ← List.map_map, List.map_take, List.map_drop] + simp [← List.map_map] @[fun_prop] theorem HasFDerivWithinAt.list_prod' {l : List ι} {x : E} (h : ∀ i ∈ l, HasFDerivWithinAt (f i ·) (f' i) s x) : HasFDerivWithinAt (fun x ↦ (l.map (f · x)).prod) (∑ i : Fin l.length, ((l.take i).map (f · x)).prod • - smulRight (f' (l.get i)) ((l.drop (.succ i)).map (f · x)).prod) s x := by + 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.get i) (l.get_mem i i.isLt))) ?_ + (hasFDerivWithinAt_pi.mpr fun i ↦ h l[i] (l.get_mem i i.isLt))) ?_ ext m - simp [← Function.comp_def (f · x) (l.get ·), ← List.map_map, List.map_take, List.map_drop] + simp [← List.map_map] theorem fderiv_list_prod' {l : List ι} {x : E} (h : ∀ i ∈ l, DifferentiableAt 𝕜 (f i ·) x) : fderiv 𝕜 (fun x ↦ (l.map (f · x)).prod) x = ∑ i : Fin l.length, ((l.take i).map (f · x)).prod • - smulRight (fderiv 𝕜 (fun x ↦ f (l.get i) x) x) - ((l.drop (.succ i)).map (f · x)).prod := + smulRight (fderiv 𝕜 (fun x ↦ f l[i] x) x) ((l.drop (.succ i)).map (f · x)).prod := (HasFDerivAt.list_prod' fun i hi ↦ (h i hi).hasFDerivAt).fderiv theorem fderivWithin_list_prod' {l : List ι} {x : E} (hxs : UniqueDiffWithinAt 𝕜 s x) (h : ∀ i ∈ l, DifferentiableWithinAt 𝕜 (f i ·) s x) : fderivWithin 𝕜 (fun x ↦ (l.map (f · x)).prod) s x = ∑ i : Fin l.length, ((l.take i).map (f · x)).prod • - smulRight (fderivWithin 𝕜 (fun x ↦ f (l.get i) x) s x) - ((l.drop (.succ i)).map (f · x)).prod := + smulRight (fderivWithin 𝕜 (fun x ↦ f l[i] x) s x) ((l.drop (.succ i)).map (f · x)).prod := (HasFDerivWithinAt.list_prod' fun i hi ↦ (h i hi).hasFDerivWithinAt).fderivWithin hxs @[fun_prop] diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index 7e6c9dafd5dc9..79b3e2892c891 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -403,7 +403,7 @@ protected def cellsOfRowLens : List ℕ → Finset (ℕ × ℕ) (Embedding.prodMap ⟨_, Nat.succ_injective⟩ (Embedding.refl ℕ)) protected theorem mem_cellsOfRowLens {w : List ℕ} {c : ℕ × ℕ} : - c ∈ YoungDiagram.cellsOfRowLens w ↔ ∃ h : c.fst < w.length, c.snd < w.get ⟨c.fst, h⟩ := by + c ∈ YoungDiagram.cellsOfRowLens w ↔ ∃ h : c.fst < w.length, c.snd < w[c.fst] := by induction' w with w_hd w_tl w_ih generalizing c <;> rw [YoungDiagram.cellsOfRowLens] · simp [YoungDiagram.cellsOfRowLens] · rcases c with ⟨⟨_, _⟩, _⟩ @@ -421,14 +421,14 @@ def ofRowLens (w : List ℕ) (hw : w.Sorted (· ≥ ·)) : YoungDiagram where refine ⟨hi.trans_lt h1, ?_⟩ calc j1 ≤ j2 := hj - _ < w.get ⟨i2, _⟩ := h2 - _ ≤ w.get ⟨i1, _⟩ := by + _ < w[i2] := h2 + _ ≤ w[i1] := by obtain rfl | h := eq_or_lt_of_le hi - · convert le_refl (w.get ⟨i1, h1⟩) + · rfl · exact List.pairwise_iff_get.mp hw _ _ h theorem mem_ofRowLens {w : List ℕ} {hw : w.Sorted (· ≥ ·)} {c : ℕ × ℕ} : - c ∈ ofRowLens w hw ↔ ∃ h : c.fst < w.length, c.snd < w.get ⟨c.fst, h⟩ := + c ∈ ofRowLens w hw ↔ ∃ h : c.fst < w.length, c.snd < w[c.fst] := YoungDiagram.mem_cellsOfRowLens /-- The number of rows in `ofRowLens w hw` is the length of `w` -/ @@ -436,11 +436,11 @@ theorem rowLens_length_ofRowLens {w : List ℕ} {hw : w.Sorted (· ≥ ·)} (hpo (ofRowLens w hw).rowLens.length = w.length := by simp only [length_rowLens, colLen, Nat.find_eq_iff, mem_cells, mem_ofRowLens, lt_self_iff_false, IsEmpty.exists_iff, Classical.not_not] - exact ⟨not_false, fun n hn => ⟨hn, hpos _ (List.get_mem _ _ hn)⟩⟩ + exact ⟨not_false, fun n hn => ⟨hn, hpos _ (List.getElem_mem _ _ hn)⟩⟩ /-- The length of the `i`th row in `ofRowLens w hw` is the `i`th entry of `w` -/ theorem rowLen_ofRowLens {w : List ℕ} {hw : w.Sorted (· ≥ ·)} (i : Fin w.length) : - (ofRowLens w hw).rowLen i = w.get i := by + (ofRowLens w hw).rowLen i = w[i] := by simp [rowLen, Nat.find_eq_iff, mem_ofRowLens] /-- The left_inv direction of the equivalence -/ diff --git a/Mathlib/Data/Finset/Sort.lean b/Mathlib/Data/Finset/Sort.lean index f2963984cc6cc..418f537711a77 100644 --- a/Mathlib/Data/Finset/Sort.lean +++ b/Mathlib/Data/Finset/Sort.lean @@ -93,16 +93,16 @@ theorem sorted_zero_eq_min'_aux (s : Finset α) (h : 0 < (s.sort (· ≤ ·)).le exact s.min'_le _ this theorem sorted_zero_eq_min' {s : Finset α} {h : 0 < (s.sort (· ≤ ·)).length} : - (s.sort (· ≤ ·)).get ⟨0, h⟩ = s.min' (card_pos.1 <| by rwa [length_sort] at h) := + (s.sort (· ≤ ·))[0] = s.min' (card_pos.1 <| by rwa [length_sort] at h) := sorted_zero_eq_min'_aux _ _ _ theorem min'_eq_sorted_zero {s : Finset α} {h : s.Nonempty} : - s.min' h = (s.sort (· ≤ ·)).get ⟨0, (by rw [length_sort]; exact card_pos.2 h)⟩ := + s.min' h = (s.sort (· ≤ ·))[0]'(by rw [length_sort]; exact card_pos.2 h) := (sorted_zero_eq_min'_aux _ _ _).symm theorem sorted_last_eq_max'_aux (s : Finset α) (h : (s.sort (· ≤ ·)).length - 1 < (s.sort (· ≤ ·)).length) (H : s.Nonempty) : - (s.sort (· ≤ ·)).get ⟨(s.sort (· ≤ ·)).length - 1, h⟩ = s.max' H := by + (s.sort (· ≤ ·))[(s.sort (· ≤ ·)).length - 1] = s.max' H := by let l := s.sort (· ≤ ·) apply le_antisymm · have : l.get ⟨(s.sort (· ≤ ·)).length - 1, h⟩ ∈ s := @@ -115,14 +115,14 @@ theorem sorted_last_eq_max'_aux (s : Finset α) theorem sorted_last_eq_max' {s : Finset α} {h : (s.sort (· ≤ ·)).length - 1 < (s.sort (· ≤ ·)).length} : - (s.sort (· ≤ ·)).get ⟨(s.sort (· ≤ ·)).length - 1, h⟩ = + (s.sort (· ≤ ·))[(s.sort (· ≤ ·)).length - 1] = s.max' (by rw [length_sort] at h; exact card_pos.1 (lt_of_le_of_lt bot_le h)) := sorted_last_eq_max'_aux _ _ _ theorem max'_eq_sorted_last {s : Finset α} {h : s.Nonempty} : s.max' h = - (s.sort (· ≤ ·)).get ⟨(s.sort (· ≤ ·)).length - 1, - by simpa using Nat.sub_lt (card_pos.mpr h) Nat.zero_lt_one⟩ := + (s.sort (· ≤ ·))[(s.sort (· ≤ ·)).length - 1]' + (by simpa using Nat.sub_lt (card_pos.mpr h) Nat.zero_lt_one) := (sorted_last_eq_max'_aux _ _ _).symm /-- Given a finset `s` of cardinality `k` in a linear order `α`, the map `orderIsoOfFin s h` @@ -150,8 +150,7 @@ theorem orderIsoOfFin_symm_apply (s : Finset α) {k : ℕ} (h : s.card = k) (x : rfl theorem orderEmbOfFin_apply (s : Finset α) {k : ℕ} (h : s.card = k) (i : Fin k) : - s.orderEmbOfFin h i = - (s.sort (· ≤ ·)).get ⟨i, by rw [length_sort, h]; exact i.2⟩ := + s.orderEmbOfFin h i = (s.sort (· ≤ ·))[i]'(by rw [length_sort, h]; exact i.2) := rfl @[simp] @@ -170,7 +169,7 @@ theorem range_orderEmbOfFin (s : Finset α) {k : ℕ} (h : s.card = k) : /-- The bijection `orderEmbOfFin s h` sends `0` to the minimum of `s`. -/ theorem orderEmbOfFin_zero {s : Finset α} {k : ℕ} (h : s.card = k) (hz : 0 < k) : orderEmbOfFin s h ⟨0, hz⟩ = s.min' (card_pos.mp (h.symm ▸ hz)) := by - simp only [orderEmbOfFin_apply, Fin.val_mk, sorted_zero_eq_min'] + simp only [orderEmbOfFin_apply, Fin.getElem_fin, sorted_zero_eq_min'] /-- The bijection `orderEmbOfFin s h` sends `k-1` to the maximum of `s`. -/ theorem orderEmbOfFin_last {s : Finset α} {k : ℕ} (h : s.card = k) (hz : 0 < k) : diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 2bbfd9513f043..f0e9fc75a0292 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -2170,18 +2170,24 @@ theorem map_foldl_erase [DecidableEq β] {f : α → β} (finj : Injective f) {l map f (foldl List.erase l₁ l₂) = foldl (fun l a => l.erase (f a)) (map f l₁) l₂ := by induction l₂ generalizing l₁ <;> [rfl; simp only [foldl_cons, map_erase finj, *]] -theorem erase_get [DecidableEq ι] {l : List ι} (i : Fin l.length) : - Perm (l.erase (l.get i)) (l.eraseIdx ↑i) := by - induction l with +theorem erase_getElem [DecidableEq ι] {l : List ι} {i : ℕ} (hi : i < l.length) : + Perm (l.erase l[i]) (l.eraseIdx i) := by + induction l generalizing i with | nil => simp | cons a l IH => - cases i using Fin.cases with + cases i with | zero => simp | succ i => - by_cases ha : a = l.get i - · simpa [ha] using .trans (perm_cons_erase (l.get_mem i i.isLt)) (.cons _ (IH i)) - · simp only [get_eq_getElem] at IH ha ⊢ - simpa [ha] using IH i + have hi' : i < l.length := by simpa using hi + if ha : a = l[i] then + simpa [ha] using .trans (perm_cons_erase (l.getElem_mem i _)) (.cons _ (IH hi')) + else + simpa [ha] using IH hi' + +@[deprecated erase_getElem (since := "2024-08-03")] +theorem erase_get [DecidableEq ι] {l : List ι} (i : Fin l.length) : + Perm (l.erase (l.get i)) (l.eraseIdx ↑i) := + erase_getElem i.isLt theorem length_eraseIdx_add_one {l : List ι} {i : ℕ} (h : i < l.length) : (l.eraseIdx i).length + 1 = l.length := calc diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index 07ad3fab49f97..9b888110c53b9 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -30,7 +30,7 @@ section getD variable (d : α) -theorem getD_eq_get {n : ℕ} (hn : n < l.length) : l.getD n d = l.get ⟨n, hn⟩ := by +theorem getD_eq_getElem {n : ℕ} (hn : n < l.length) : l.getD n d = l[n] := by induction l generalizing n with | nil => simp at hn | cons head tail ih => @@ -38,6 +38,10 @@ theorem getD_eq_get {n : ℕ} (hn : n < l.length) : l.getD n d = l.get ⟨n, hn · exact getD_cons_zero · exact ih _ +@[deprecated getD_eq_getElem (since := "2024-08-02")] +theorem getD_eq_get {n : ℕ} (hn : n < l.length) : l.getD n d = l.get ⟨n, hn⟩ := + getD_eq_getElem l d hn + theorem getD_map {n : ℕ} (f : α → β) : (map f l).getD n (f d) = f (l.getD n d) := by simp theorem getD_eq_default {n : ℕ} (hn : l.length ≤ n) : l.getD n d = d := by @@ -68,18 +72,16 @@ theorem getElem?_getD_replicate_default_eq (r n : ℕ) : (replicate r d)[n]?.get @[deprecated (since := "2024-06-12")] alias getD_replicate_default_eq := getElem?_getD_replicate_default_eq -set_option linter.deprecated false in theorem getD_append (l l' : List α) (d : α) (n : ℕ) (h : n < l.length) : (l ++ l').getD n d = l.getD n d := by - rw [getD_eq_get _ _ (Nat.lt_of_lt_of_le h (length_append _ _ ▸ Nat.le_add_right _ _)), - get_append _ h, getD_eq_get] + rw [getD_eq_getElem _ _ (Nat.lt_of_lt_of_le h (length_append _ _ ▸ Nat.le_add_right _ _)), + getElem_append _ h, getD_eq_getElem] theorem getD_append_right (l l' : List α) (d : α) (n : ℕ) (h : l.length ≤ n) : (l ++ l').getD n d = l'.getD (n - l.length) d := by cases Nat.lt_or_ge n (l ++ l').length with | inl h' => - rw [getD_eq_get (l ++ l') d h', get_eq_getElem, getElem_append_right, getD_eq_get, - get_eq_getElem] + rw [getD_eq_getElem (l ++ l') d h', getElem_append_right, getD_eq_getElem] · rw [length_append] at h' exact Nat.sub_lt_left_of_lt_add h h' · exact Nat.not_lt_of_le h @@ -89,7 +91,7 @@ theorem getD_append_right (l l' : List α) (d : α) (n : ℕ) (h : l.length ≤ theorem getD_eq_getD_get? (n : ℕ) : l.getD n d = (l.get? n).getD d := by cases Nat.lt_or_ge n l.length with - | inl h => rw [getD_eq_get _ _ h, get?_eq_get h, Option.getD_some] + | inl h => rw [getD_eq_getElem _ _ h, get?_eq_get h, get_eq_getElem, Option.getD_some] | inr h => rw [getD_eq_default _ _ h, get?_eq_none.mpr h, Option.getD_none] end getD @@ -110,8 +112,12 @@ theorem getI_cons_zero : getI (x :: xs) 0 = x := theorem getI_cons_succ : getI (x :: xs) (n + 1) = getI xs n := rfl +theorem getI_eq_getElem {n : ℕ} (hn : n < l.length) : l.getI n = l[n] := + getD_eq_getElem .. + +@[deprecated getI_eq_getElem (since := "2024-08-02")] theorem getI_eq_get {n : ℕ} (hn : n < l.length) : l.getI n = l.get ⟨n, hn⟩ := - getD_eq_get .. + getD_eq_getElem .. theorem getI_eq_default {n : ℕ} (hn : l.length ≤ n) : l.getI n = default := getD_eq_default _ _ hn diff --git a/Mathlib/Data/List/ToFinsupp.lean b/Mathlib/Data/List/ToFinsupp.lean index afa5e1b2d5d4c..b8fbb880f704e 100644 --- a/Mathlib/Data/List/ToFinsupp.lean +++ b/Mathlib/Data/List/ToFinsupp.lean @@ -61,16 +61,11 @@ theorem toFinsupp_support : l.toFinsupp.support = (Finset.range l.length).filter (getD l · 0 ≠ 0) := rfl -theorem toFinsupp_apply_lt (hn : n < l.length) : l.toFinsupp n = l.get ⟨n, hn⟩ := - getD_eq_get _ _ _ +theorem toFinsupp_apply_lt (hn : n < l.length) : l.toFinsupp n = l[n] := + getD_eq_getElem _ _ _ -theorem toFinsupp_apply_fin (n : Fin l.length) : l.toFinsupp n = l.get n := - getD_eq_get _ _ _ - -set_option linter.deprecated false in -@[deprecated (since := "2023-04-10")] -theorem toFinsupp_apply_lt' (hn : n < l.length) : l.toFinsupp n = l.nthLe n hn := - getD_eq_get _ _ _ +theorem toFinsupp_apply_fin (n : Fin l.length) : l.toFinsupp n = l[n] := + getD_eq_getElem _ _ _ theorem toFinsupp_apply_le (hn : l.length ≤ n) : l.toFinsupp n = 0 := getD_eq_default _ _ hn diff --git a/Mathlib/Data/Nat/Nth.lean b/Mathlib/Data/Nat/Nth.lean index 0ad148511a5d1..83719860919c9 100644 --- a/Mathlib/Data/Nat/Nth.lean +++ b/Mathlib/Data/Nat/Nth.lean @@ -66,7 +66,7 @@ theorem nth_eq_getD_sort (h : (setOf p).Finite) (n : ℕ) : theorem nth_eq_orderEmbOfFin (hf : (setOf p).Finite) {n : ℕ} (hn : n < hf.toFinset.card) : nth p n = hf.toFinset.orderEmbOfFin rfl ⟨n, hn⟩ := by - rw [nth_eq_getD_sort hf, Finset.orderEmbOfFin_apply, List.getD_eq_get] + rw [nth_eq_getD_sort hf, Finset.orderEmbOfFin_apply, List.getD_eq_getElem, Fin.getElem_fin] theorem nth_strictMonoOn (hf : (setOf p).Finite) : StrictMonoOn (nth p) (Set.Iio hf.toFinset.card) := by diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index cdc9cb64b7dab..8ca8f8ad2b96a 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -347,8 +347,8 @@ theorem isRightInversion_of_mem_rightInvSeq {ω : List B} (hω : cs.IsReduced ω (ht : t ∈ ris ω) : cs.IsRightInversion (π ω) t := by constructor · exact cs.isReflection_of_mem_rightInvSeq ω ht - · obtain ⟨⟨j, hj⟩, rfl⟩ := List.mem_iff_get.mp ht - rw [← List.getD_eq_get _ 1 hj, wordProd_mul_getD_rightInvSeq] + · obtain ⟨j, hj, rfl⟩ := List.mem_iff_getElem.mp ht + rw [← List.getD_eq_getElem _ 1 hj, wordProd_mul_getD_rightInvSeq] rw [cs.length_rightInvSeq] at hj calc ℓ (π (ω.eraseIdx j)) @@ -360,8 +360,8 @@ theorem isLeftInversion_of_mem_leftInvSeq {ω : List B} (hω : cs.IsReduced ω) (ht : t ∈ lis ω) : cs.IsLeftInversion (π ω) t := by constructor · exact cs.isReflection_of_mem_leftInvSeq ω ht - · obtain ⟨⟨j, hj⟩, rfl⟩ := List.mem_iff_get.mp ht - rw [← List.getD_eq_get _ 1 hj, getD_leftInvSeq_mul_wordProd] + · obtain ⟨j, hj, rfl⟩ := List.mem_iff_getElem.mp ht + rw [← List.getD_eq_getElem _ 1 hj, getD_leftInvSeq_mul_wordProd] rw [cs.length_leftInvSeq] at hj calc ℓ (π (ω.eraseIdx j)) @@ -389,13 +389,13 @@ theorem prod_leftInvSeq (ω : List B) : prod (lis ω) = (π ω)⁻¹ := by exact cs.prod_rightInvSeq _ theorem IsReduced.nodup_rightInvSeq {ω : List B} (rω : cs.IsReduced ω) : List.Nodup (ris ω) := by - apply List.nodup_iff_get?_ne_get?.mpr - intro j j' j_lt_j' j'_lt_length (dup : get? (rightInvSeq cs ω) j = get? (rightInvSeq cs ω) j') + apply List.nodup_iff_getElem?_ne_getElem?.mpr + intro j j' j_lt_j' j'_lt_length (dup : (rightInvSeq cs ω)[j]? = (rightInvSeq cs ω)[j']?) show False replace j'_lt_length : j' < List.length ω := by simpa using j'_lt_length - rw [get?_eq_get (by simp; omega), get?_eq_get (by simp; omega)] at dup + rw [getElem?_eq_getElem (by simp; omega), getElem?_eq_getElem (by simp; omega)] at dup apply Option.some_injective at dup - rw [← getD_eq_get _ 1, ← getD_eq_get _ 1] at dup + rw [← getD_eq_getElem _ 1, ← getD_eq_getElem _ 1] at dup set! t := (ris ω).getD j 1 with h₁ set! t' := (ris (ω.eraseIdx j)).getD (j' - 1) 1 with h₂ have h₃ : t' = (ris ω).getD j' 1 := by diff --git a/Mathlib/LinearAlgebra/Matrix/Transvection.lean b/Mathlib/LinearAlgebra/Matrix/Transvection.lean index 59a213005ef4a..1e06f78ec9172 100644 --- a/Mathlib/LinearAlgebra/Matrix/Transvection.lean +++ b/Mathlib/LinearAlgebra/Matrix/Transvection.lean @@ -340,20 +340,34 @@ def listTransvecRow : List (Matrix (Fin r ⊕ Unit) (Fin r ⊕ Unit) 𝕜) := @[simp] theorem length_listTransvecCol : (listTransvecCol M).length = r := by simp [listTransvecCol] +theorem listTransvecCol_getElem {i : ℕ} (h : i < (listTransvecCol M).length) : + (listTransvecCol M)[i] = + letI i' : Fin r := ⟨i, length_listTransvecCol M ▸ h⟩ + transvection (inl i') (inr unit) <| -M (inl i') (inr unit) / M (inr unit) (inr unit) := by + simp [listTransvecCol] + +@[deprecated listTransvecCol_getElem (since := "2024-08-03")] theorem listTransvecCol_get (i : Fin (listTransvecCol M).length) : (listTransvecCol M).get i = letI i' := Fin.cast (length_listTransvecCol M) i - transvection (inl i') (inr unit) <| -M (inl i') (inr unit) / M (inr unit) (inr unit) := by - simp [listTransvecCol, Fin.cast] + transvection (inl i') (inr unit) <| -M (inl i') (inr unit) / M (inr unit) (inr unit) := + listTransvecCol_getElem .. @[simp] theorem length_listTransvecRow : (listTransvecRow M).length = r := by simp [listTransvecRow] +theorem listTransvecRow_getElem {i : ℕ} (h : i < (listTransvecRow M).length) : + (listTransvecRow M)[i] = + letI i' : Fin r := ⟨i, length_listTransvecRow M ▸ h⟩ + transvection (inr unit) (inl i') <| -M (inr unit) (inl i') / M (inr unit) (inr unit) := by + simp [listTransvecRow, Fin.cast] + +@[deprecated listTransvecRow_getElem (since := "2024-08-03")] theorem listTransvecRow_get (i : Fin (listTransvecRow M).length) : (listTransvecRow M).get i = letI i' := Fin.cast (length_listTransvecRow M) i - transvection (inr unit) (inl i') <| -M (inr unit) (inl i') / M (inr unit) (inr unit) := by - simp [listTransvecRow, Fin.cast] + transvection (inr unit) (inl i') <| -M (inr unit) (inl i') / M (inr unit) (inr unit) := + listTransvecRow_getElem .. /-- Multiplying by some of the matrices in `listTransvecCol M` does not change the last row. -/ theorem listTransvecCol_mul_last_row_drop (i : Fin r ⊕ Unit) {k : ℕ} (hk : k ≤ r) : diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index 36c5009b8d87c..6b905496d8037 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -105,8 +105,8 @@ lemma toList_ne_nil (x : RelSeries r) : x.toList ≠ [] := fun m => /-- Every nonempty list satisfying the chain condition gives a relation series-/ @[simps] def fromListChain' (x : List α) (x_ne_nil : x ≠ []) (hx : x.Chain' r) : RelSeries r where - length := x.length.pred - toFun := x.get ∘ Fin.cast (Nat.succ_pred_eq_of_pos <| List.length_pos.mpr x_ne_nil) + length := x.length - 1 + toFun i := x[Fin.cast (Nat.succ_pred_eq_of_pos <| List.length_pos.mpr x_ne_nil) i] step i := List.chain'_iff_get.mp hx i i.2 /-- Relation series of `r` and nonempty list of `α` satisfying `r`-chain condition bijectively diff --git a/Mathlib/RingTheory/Regular/RegularSequence.lean b/Mathlib/RingTheory/Regular/RegularSequence.lean index 27b18cacdb514..6da2991f4d4ca 100644 --- a/Mathlib/RingTheory/Regular/RegularSequence.lean +++ b/Mathlib/RingTheory/Regular/RegularSequence.lean @@ -140,7 +140,7 @@ structure IsWeaklyRegular (rs : List R) : Prop where lemma isWeaklyRegular_iff_Fin (rs : List R) : IsWeaklyRegular M rs ↔ ∀ (i : Fin rs.length), - IsSMulRegular (M ⧸ (ofList (rs.take i) • ⊤ : Submodule R M)) (rs.get i) := + IsSMulRegular (M ⧸ (ofList (rs.take i) • ⊤ : Submodule R M)) rs[i] := Iff.trans (isWeaklyRegular_iff M rs) (Iff.symm Fin.forall_iff) /-- A weakly regular sequence `rs` on `M` is regular if also `M/rsM ≠ 0`. -/ diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index 42a29db62e8be..3f2affe6986e1 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -169,21 +169,21 @@ def toOfListsRightMoves {L R : List PGame} : Fin R.length ≃ (ofLists L R).Righ ((Equiv.cast (rightMoves_ofLists L R).symm).trans Equiv.ulift).symm theorem ofLists_moveLeft {L R : List PGame} (i : Fin L.length) : - (ofLists L R).moveLeft (toOfListsLeftMoves i) = L.get i := + (ofLists L R).moveLeft (toOfListsLeftMoves i) = L[i] := rfl @[simp] theorem ofLists_moveLeft' {L R : List PGame} (i : (ofLists L R).LeftMoves) : - (ofLists L R).moveLeft i = L.get (toOfListsLeftMoves.symm i) := + (ofLists L R).moveLeft i = L[toOfListsLeftMoves.symm i] := rfl theorem ofLists_moveRight {L R : List PGame} (i : Fin R.length) : - (ofLists L R).moveRight (toOfListsRightMoves i) = R.get i := + (ofLists L R).moveRight (toOfListsRightMoves i) = R[i] := rfl @[simp] theorem ofLists_moveRight' {L R : List PGame} (i : (ofLists L R).RightMoves) : - (ofLists L R).moveRight i = R.get (toOfListsRightMoves.symm i) := + (ofLists L R).moveRight i = R[toOfListsRightMoves.symm i] := rfl /-- A variant of `PGame.recOn` expressed in terms of `PGame.moveLeft` and `PGame.moveRight`. From 267c64bef5534e7011cb2b75a66c22b758d0e2a9 Mon Sep 17 00:00:00 2001 From: Lorenzo Luccioli Date: Mon, 12 Aug 2024 04:09:17 +0000 Subject: [PATCH 198/359] feat (MeasureTheory-Probabilty): Congruence lemma for double integrals. (#15460) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit - Add `ProbabilityTheory.Kernel.integral_congr_ae₂` and `MeasureTheory.integral_congr_ae₂`. --- Mathlib/MeasureTheory/Integral/Bochner.lean | 8 ++++++++ Mathlib/Probability/Kernel/Basic.lean | 8 ++++++++ 2 files changed, 16 insertions(+) diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 3490d60b63ba6..74bbe759404b4 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -839,6 +839,14 @@ theorem integral_congr_ae {f g : α → G} (h : f =ᵐ[μ] g) : ∫ a, f a ∂μ exact setToFun_congr_ae (dominatedFinMeasAdditive_weightedSMul μ) h · simp [integral, hG] +lemma integral_congr_ae₂ {β : Type*} {_ : MeasurableSpace β} {ν : Measure β} {f g : α → β → G} + (h : ∀ᵐ a ∂μ, f a =ᵐ[ν] g a) : + ∫ a, ∫ b, f a b ∂ν ∂μ = ∫ a, ∫ b, g a b ∂ν ∂μ := by + apply integral_congr_ae + filter_upwards [h] with _ ha + apply integral_congr_ae + filter_upwards [ha] with _ hb using hb + -- Porting note: `nolint simpNF` added because simplify fails on left-hand side @[simp, nolint simpNF] theorem L1.integral_of_fun_eq_integral {f : α → G} (hf : Integrable f μ) : diff --git a/Mathlib/Probability/Kernel/Basic.lean b/Mathlib/Probability/Kernel/Basic.lean index 87c0e64295061..5a8caec90e730 100644 --- a/Mathlib/Probability/Kernel/Basic.lean +++ b/Mathlib/Probability/Kernel/Basic.lean @@ -234,6 +234,14 @@ lemma IsMarkovKernel.integrable (μ : Measure α) [IsFiniteMeasure μ] Integrable (fun x => (κ x s).toReal) μ := IsFiniteKernel.integrable μ κ hs +lemma integral_congr_ae₂ {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {f g : α → β → E} + {μ : Measure α} (h : ∀ᵐ a ∂μ, f a =ᵐ[κ a] g a) : + ∫ a, ∫ b, f a b ∂(κ a) ∂μ = ∫ a, ∫ b, g a b ∂(κ a) ∂μ := by + apply integral_congr_ae + filter_upwards [h] with _ ha + apply integral_congr_ae + filter_upwards [ha] with _ hb using hb + section Sum /-- Sum of an indexed family of kernels. -/ From 2c4b300a09c200cb3d5bfca7e032f9c8a1379d86 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 12 Aug 2024 04:26:19 +0000 Subject: [PATCH 199/359] chore: move files into new `Analysis.Normed.Affine` (#15184) --- Mathlib.lean | 10 +++++----- Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean | 2 +- Mathlib/Analysis/Calculus/AddTorsor/Coord.lean | 2 +- Mathlib/Analysis/Convex/Intrinsic.lean | 2 +- Mathlib/Analysis/Convex/Measure.lean | 2 +- Mathlib/Analysis/Convex/Normed.lean | 2 +- Mathlib/Analysis/Convex/StrictConvexBetween.lean | 2 +- Mathlib/Analysis/Convex/StrictConvexSpace.lean | 2 +- .../{NormedSpace => Normed/Affine}/AddTorsor.lean | 0 .../{NormedSpace => Normed/Affine}/AddTorsorBases.lean | 0 .../Affine}/ContinuousAffineMap.lean | 2 +- .../Affine/Isometry.lean} | 0 .../{NormedSpace => Normed/Affine}/MazurUlam.lean | 2 +- Mathlib/Analysis/Normed/Module/FiniteDimension.lean | 4 ++-- Mathlib/Analysis/Normed/Operator/Banach.lean | 2 +- .../Geometry/Euclidean/Angle/Unoriented/Affine.lean | 2 +- Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean | 2 +- Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean | 3 ++- Mathlib/Topology/UrysohnsLemma.lean | 2 +- 19 files changed, 22 insertions(+), 21 deletions(-) rename Mathlib/Analysis/{NormedSpace => Normed/Affine}/AddTorsor.lean (100%) rename Mathlib/Analysis/{NormedSpace => Normed/Affine}/AddTorsorBases.lean (100%) rename Mathlib/Analysis/{NormedSpace => Normed/Affine}/ContinuousAffineMap.lean (99%) rename Mathlib/Analysis/{NormedSpace/AffineIsometry.lean => Normed/Affine/Isometry.lean} (100%) rename Mathlib/Analysis/{NormedSpace => Normed/Affine}/MazurUlam.lean (99%) diff --git a/Mathlib.lean b/Mathlib.lean index 5c6cfabfa151f..eb1c20596def4 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1137,6 +1137,11 @@ import Mathlib.Analysis.MeanInequalities import Mathlib.Analysis.MeanInequalitiesPow import Mathlib.Analysis.MellinInversion import Mathlib.Analysis.MellinTransform +import Mathlib.Analysis.Normed.Affine.AddTorsor +import Mathlib.Analysis.Normed.Affine.AddTorsorBases +import Mathlib.Analysis.Normed.Affine.ContinuousAffineMap +import Mathlib.Analysis.Normed.Affine.Isometry +import Mathlib.Analysis.Normed.Affine.MazurUlam import Mathlib.Analysis.Normed.Algebra.Basic import Mathlib.Analysis.Normed.Algebra.Exponential import Mathlib.Analysis.Normed.Algebra.MatrixExponential @@ -1202,13 +1207,9 @@ import Mathlib.Analysis.Normed.Ring.Seminorm import Mathlib.Analysis.Normed.Ring.SeminormFromBounded import Mathlib.Analysis.Normed.Ring.SeminormFromConst import Mathlib.Analysis.Normed.Ring.Units -import Mathlib.Analysis.NormedSpace.AddTorsor -import Mathlib.Analysis.NormedSpace.AddTorsorBases -import Mathlib.Analysis.NormedSpace.AffineIsometry import Mathlib.Analysis.NormedSpace.BallAction import Mathlib.Analysis.NormedSpace.ConformalLinearMap import Mathlib.Analysis.NormedSpace.Connected -import Mathlib.Analysis.NormedSpace.ContinuousAffineMap import Mathlib.Analysis.NormedSpace.DualNumber import Mathlib.Analysis.NormedSpace.ENorm import Mathlib.Analysis.NormedSpace.Extend @@ -1221,7 +1222,6 @@ import Mathlib.Analysis.NormedSpace.HomeomorphBall import Mathlib.Analysis.NormedSpace.IndicatorFunction import Mathlib.Analysis.NormedSpace.Int import Mathlib.Analysis.NormedSpace.MStructure -import Mathlib.Analysis.NormedSpace.MazurUlam import Mathlib.Analysis.NormedSpace.Multilinear.Basic import Mathlib.Analysis.NormedSpace.Multilinear.Curry import Mathlib.Analysis.NormedSpace.OperatorNorm.Asymptotics diff --git a/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean b/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean index 72b3d3c4c8d5a..5eb7534bb2ff4 100644 --- a/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean +++ b/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ -import Mathlib.Analysis.NormedSpace.ContinuousAffineMap +import Mathlib.Analysis.Normed.Affine.ContinuousAffineMap import Mathlib.Analysis.Calculus.ContDiff.Basic /-! diff --git a/Mathlib/Analysis/Calculus/AddTorsor/Coord.lean b/Mathlib/Analysis/Calculus/AddTorsor/Coord.lean index c6ef3c74856e0..e35d693084775 100644 --- a/Mathlib/Analysis/Calculus/AddTorsor/Coord.lean +++ b/Mathlib/Analysis/Calculus/AddTorsor/Coord.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ import Mathlib.Analysis.Calculus.AddTorsor.AffineMap -import Mathlib.Analysis.NormedSpace.AddTorsorBases +import Mathlib.Analysis.Normed.Affine.AddTorsorBases /-! # Barycentric coordinates are smooth diff --git a/Mathlib/Analysis/Convex/Intrinsic.lean b/Mathlib/Analysis/Convex/Intrinsic.lean index bcc5dacfc8ebc..7e5b5e4ee2e13 100644 --- a/Mathlib/Analysis/Convex/Intrinsic.lean +++ b/Mathlib/Analysis/Convex/Intrinsic.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Paul Reichert. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Paul Reichert, Yaël Dillies -/ -import Mathlib.Analysis.NormedSpace.AddTorsorBases +import Mathlib.Analysis.Normed.Affine.AddTorsorBases /-! # Intrinsic frontier and interior diff --git a/Mathlib/Analysis/Convex/Measure.lean b/Mathlib/Analysis/Convex/Measure.lean index 09fbbdf95ba4d..6dec1b9d305d3 100644 --- a/Mathlib/Analysis/Convex/Measure.lean +++ b/Mathlib/Analysis/Convex/Measure.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Analysis.Convex.Topology -import Mathlib.Analysis.NormedSpace.AddTorsorBases +import Mathlib.Analysis.Normed.Affine.AddTorsorBases import Mathlib.MeasureTheory.Measure.Lebesgue.EqHaar /-! diff --git a/Mathlib/Analysis/Convex/Normed.lean b/Mathlib/Analysis/Convex/Normed.lean index 7756e3de54a43..2359391cb0ece 100644 --- a/Mathlib/Analysis/Convex/Normed.lean +++ b/Mathlib/Analysis/Convex/Normed.lean @@ -7,7 +7,7 @@ import Mathlib.Analysis.Convex.Between import Mathlib.Analysis.Convex.Jensen import Mathlib.Analysis.Convex.Topology import Mathlib.Analysis.Normed.Group.Pointwise -import Mathlib.Analysis.NormedSpace.AddTorsor +import Mathlib.Analysis.Normed.Affine.AddTorsor /-! # Topological and metric properties of convex sets in normed spaces diff --git a/Mathlib/Analysis/Convex/StrictConvexBetween.lean b/Mathlib/Analysis/Convex/StrictConvexBetween.lean index c067d9e67638d..c3bb34bd7fb30 100644 --- a/Mathlib/Analysis/Convex/StrictConvexBetween.lean +++ b/Mathlib/Analysis/Convex/StrictConvexBetween.lean @@ -5,7 +5,7 @@ Authors: Joseph Myers -/ import Mathlib.Analysis.Convex.Between import Mathlib.Analysis.Convex.StrictConvexSpace -import Mathlib.Analysis.NormedSpace.AffineIsometry +import Mathlib.Analysis.Normed.Affine.Isometry /-! # Betweenness in affine spaces for strictly convex spaces diff --git a/Mathlib/Analysis/Convex/StrictConvexSpace.lean b/Mathlib/Analysis/Convex/StrictConvexSpace.lean index ea048d3b233ce..80a4039dba35a 100644 --- a/Mathlib/Analysis/Convex/StrictConvexSpace.lean +++ b/Mathlib/Analysis/Convex/StrictConvexSpace.lean @@ -6,7 +6,7 @@ Authors: Yaël Dillies, Yury Kudryashov import Mathlib.Analysis.Convex.Normed import Mathlib.Analysis.Convex.Strict import Mathlib.Analysis.Normed.Order.Basic -import Mathlib.Analysis.NormedSpace.AddTorsor +import Mathlib.Analysis.Normed.Affine.AddTorsor import Mathlib.Analysis.NormedSpace.Pointwise import Mathlib.Analysis.Normed.Module.Ray diff --git a/Mathlib/Analysis/NormedSpace/AddTorsor.lean b/Mathlib/Analysis/Normed/Affine/AddTorsor.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/AddTorsor.lean rename to Mathlib/Analysis/Normed/Affine/AddTorsor.lean diff --git a/Mathlib/Analysis/NormedSpace/AddTorsorBases.lean b/Mathlib/Analysis/Normed/Affine/AddTorsorBases.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/AddTorsorBases.lean rename to Mathlib/Analysis/Normed/Affine/AddTorsorBases.lean diff --git a/Mathlib/Analysis/NormedSpace/ContinuousAffineMap.lean b/Mathlib/Analysis/Normed/Affine/ContinuousAffineMap.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/ContinuousAffineMap.lean rename to Mathlib/Analysis/Normed/Affine/ContinuousAffineMap.lean index 153bfd6cec084..5e63c0037ade0 100644 --- a/Mathlib/Analysis/NormedSpace/ContinuousAffineMap.lean +++ b/Mathlib/Analysis/Normed/Affine/ContinuousAffineMap.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ -import Mathlib.Analysis.NormedSpace.AffineIsometry +import Mathlib.Analysis.Normed.Affine.Isometry import Mathlib.Topology.Algebra.ContinuousAffineMap import Mathlib.Analysis.NormedSpace.OperatorNorm.NormedSpace diff --git a/Mathlib/Analysis/NormedSpace/AffineIsometry.lean b/Mathlib/Analysis/Normed/Affine/Isometry.lean similarity index 100% rename from Mathlib/Analysis/NormedSpace/AffineIsometry.lean rename to Mathlib/Analysis/Normed/Affine/Isometry.lean diff --git a/Mathlib/Analysis/NormedSpace/MazurUlam.lean b/Mathlib/Analysis/Normed/Affine/MazurUlam.lean similarity index 99% rename from Mathlib/Analysis/NormedSpace/MazurUlam.lean rename to Mathlib/Analysis/Normed/Affine/MazurUlam.lean index 31d9886ff576c..0ccabea54df9c 100644 --- a/Mathlib/Analysis/NormedSpace/MazurUlam.lean +++ b/Mathlib/Analysis/Normed/Affine/MazurUlam.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Topology.Instances.RealVectorSpace -import Mathlib.Analysis.NormedSpace.AffineIsometry +import Mathlib.Analysis.Normed.Affine.Isometry /-! # Mazur-Ulam Theorem diff --git a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean index 69ee742513d76..6fbb936e9f60b 100644 --- a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean +++ b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean @@ -5,8 +5,8 @@ Authors: Sébastien Gouëzel -/ import Mathlib.Analysis.Asymptotics.AsymptoticEquivalent import Mathlib.Analysis.Normed.Group.Lemmas -import Mathlib.Analysis.NormedSpace.AddTorsor -import Mathlib.Analysis.NormedSpace.AffineIsometry +import Mathlib.Analysis.Normed.Affine.AddTorsor +import Mathlib.Analysis.Normed.Affine.Isometry import Mathlib.Analysis.NormedSpace.OperatorNorm.NormedSpace import Mathlib.Analysis.NormedSpace.RieszLemma import Mathlib.Analysis.NormedSpace.Pointwise diff --git a/Mathlib/Analysis/Normed/Operator/Banach.lean b/Mathlib/Analysis/Normed/Operator/Banach.lean index b7e93269d948e..cacdacc7a65c6 100644 --- a/Mathlib/Analysis/Normed/Operator/Banach.lean +++ b/Mathlib/Analysis/Normed/Operator/Banach.lean @@ -6,7 +6,7 @@ Authors: Sébastien Gouëzel import Mathlib.Topology.Baire.Lemmas import Mathlib.Topology.Baire.CompleteMetrizable import Mathlib.Analysis.NormedSpace.OperatorNorm.NormedSpace -import Mathlib.Analysis.NormedSpace.AffineIsometry +import Mathlib.Analysis.Normed.Affine.Isometry import Mathlib.Analysis.Normed.Group.InfiniteSum /-! diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean index 3d766303aed90..9b74a140ec8b5 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean @@ -6,7 +6,7 @@ Authors: Joseph Myers, Manuel Candales import Mathlib.Analysis.Convex.Between import Mathlib.Analysis.Normed.Group.AddTorsor import Mathlib.Geometry.Euclidean.Angle.Unoriented.Basic -import Mathlib.Analysis.NormedSpace.AffineIsometry +import Mathlib.Analysis.Normed.Affine.Isometry /-! # Angles between points diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean index dfd2890057957..75c1550c86990 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean @@ -33,7 +33,7 @@ This file defines affine maps. `P` or `V`. This file only provides purely algebraic definitions and results. Those depending on analysis or -topology are defined elsewhere; see `Analysis.NormedSpace.AddTorsor` and +topology are defined elsewhere; see `Analysis.Normed.Affine.AddTorsor` and `Topology.Algebra.Affine`. ## References diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean index fc963c50a4cdc..e8ffe0b27a7f5 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean @@ -35,7 +35,8 @@ from `P`) in most cases. As for modules, `k` is an explicit argument rather than `V`. This file only provides purely algebraic definitions and results. Those depending on analysis or -topology are defined elsewhere; see `Analysis.NormedSpace.AddTorsor` and `Topology.Algebra.Affine`. +topology are defined elsewhere; see `Analysis.Normed.Affine.AddTorsor` and +`Topology.Algebra.Affine`. ## References diff --git a/Mathlib/Topology/UrysohnsLemma.lean b/Mathlib/Topology/UrysohnsLemma.lean index 5226cd1fb3438..d6e4967be156a 100644 --- a/Mathlib/Topology/UrysohnsLemma.lean +++ b/Mathlib/Topology/UrysohnsLemma.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Analysis.NormedSpace.AddTorsor +import Mathlib.Analysis.Normed.Affine.AddTorsor import Mathlib.LinearAlgebra.AffineSpace.Ordered import Mathlib.Topology.ContinuousFunction.Basic import Mathlib.Topology.GDelta From 815eabcd67f20019a87c6ee332abca83ebeaa4fb Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Mon, 12 Aug 2024 10:09:28 +0000 Subject: [PATCH 200/359] feat(CategoryTheory): Finality of functors from filtered categories into PUnit (#15641) --- Mathlib/CategoryTheory/Filtered/Final.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/Mathlib/CategoryTheory/Filtered/Final.lean b/Mathlib/CategoryTheory/Filtered/Final.lean index 5cfc7edba0acc..eb3ef28ce871c 100644 --- a/Mathlib/CategoryTheory/Filtered/Final.lean +++ b/Mathlib/CategoryTheory/Filtered/Final.lean @@ -230,4 +230,24 @@ theorem Functor.initial_iff_isCofiltered_costructuredArrow [IsCofilteredOrEmpty end LocallySmall +/-- If `C` is filtered, then every functor `F : C ⥤ Discrete PUnit` is final. -/ +theorem Functor.final_of_isFiltered_of_pUnit {C : Type u₁} [Category.{v₁} C] + [IsFiltered C] (F : C ⥤ Discrete PUnit) : + Final F := by + refine final_of_exists_of_isFiltered F (fun _ => ?_) (fun {_} {c} _ _ => ?_) + · use Classical.choice IsFiltered.nonempty + exact ⟨Discrete.eqToHom (by simp)⟩ + · use c; use 𝟙 c + apply Subsingleton.elim + +/-- If `C` is cofiltered, then every functor `F : C ⥤ Discrete PUnit` is initial. -/ +theorem Functor.initial_of_isCofiltered_pUnit {C : Type u₁} [Category.{v₁} C] + [IsCofiltered C] (F : C ⥤ Discrete PUnit) : + Initial F := by + refine initial_of_exists_of_isCofiltered F (fun _ => ?_) (fun {_} {c} _ _ => ?_) + · use Classical.choice IsCofiltered.nonempty + exact ⟨Discrete.eqToHom (by simp)⟩ + · use c; use 𝟙 c + apply Subsingleton.elim + end CategoryTheory From 81b928ca2905b3abbed5ef4c41dd561ac6e98055 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Mon, 12 Aug 2024 10:18:59 +0000 Subject: [PATCH 201/359] feat(CategoryTheory): Equivalence between certain comma categories and products (#15648) --- Mathlib/CategoryTheory/Closed/Cartesian.lean | 4 +- Mathlib/CategoryTheory/Closed/Zero.lean | 2 +- Mathlib/CategoryTheory/Comma/Basic.lean | 49 +++++++++++++++++++ .../CategoryTheory/Dialectica/Monoidal.lean | 4 +- .../Monoidal/OfHasFiniteProducts.lean | 4 +- 5 files changed, 56 insertions(+), 7 deletions(-) diff --git a/Mathlib/CategoryTheory/Closed/Cartesian.lean b/Mathlib/CategoryTheory/Closed/Cartesian.lean index 882c57160c1ef..4253d1f338112 100644 --- a/Mathlib/CategoryTheory/Closed/Cartesian.lean +++ b/Mathlib/CategoryTheory/Closed/Cartesian.lean @@ -232,8 +232,8 @@ The typeclass argument is explicit: any instance can be used. -/ def expTerminalIsoSelf [Exponentiable (⊤_ C)] : (⊤_ C) ⟹ X ≅ X := Yoneda.ext ((⊤_ C) ⟹ X) X - (fun {Y} f => (prod.leftUnitor Y).inv ≫ CartesianClosed.uncurry f) - (fun {Y} f => CartesianClosed.curry ((prod.leftUnitor Y).hom ≫ f)) + (fun {Y} f => (Limits.prod.leftUnitor Y).inv ≫ CartesianClosed.uncurry f) + (fun {Y} f => CartesianClosed.curry ((Limits.prod.leftUnitor Y).hom ≫ f)) (fun g => by rw [curry_eq_iff, Iso.hom_inv_id_assoc]) (fun g => by simp) diff --git a/Mathlib/CategoryTheory/Closed/Zero.lean b/Mathlib/CategoryTheory/Closed/Zero.lean index b9d57caa6bef0..4c6d8edac4a37 100644 --- a/Mathlib/CategoryTheory/Closed/Zero.lean +++ b/Mathlib/CategoryTheory/Closed/Zero.lean @@ -37,7 +37,7 @@ then each homset has exactly one element. def uniqueHomsetOfInitialIsoTerminal [HasInitial C] (i : ⊥_ C ≅ ⊤_ C) (X Y : C) : Unique (X ⟶ Y) := Equiv.unique <| calc - (X ⟶ Y) ≃ (X ⨯ ⊤_ C ⟶ Y) := Iso.homCongr (prod.rightUnitor _).symm (Iso.refl _) + (X ⟶ Y) ≃ (X ⨯ ⊤_ C ⟶ Y) := Iso.homCongr (Limits.prod.rightUnitor _).symm (Iso.refl _) _ ≃ (X ⨯ ⊥_ C ⟶ Y) := (Iso.homCongr (prod.mapIso (Iso.refl _) i.symm) (Iso.refl _)) _ ≃ (⊥_ C ⟶ Y ^^ X) := (exp.adjunction _).homEquiv _ _ diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index 8754946c6e5c6..f8f1cea64b759 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -6,6 +6,7 @@ Authors: Scott Morrison, Johan Commelin, Bhavik Mehta import Mathlib.CategoryTheory.Iso import Mathlib.CategoryTheory.Functor.Category import Mathlib.CategoryTheory.EqToHom +import Mathlib.CategoryTheory.Products.Unitor /-! # Comma categories @@ -409,6 +410,54 @@ def post (L : A ⥤ T) (R : B ⥤ T) (F : T ⥤ C) : Comma L R ⥤ Comma (L ⋙ right := f.right w := by simp only [Functor.comp_map, ← F.map_comp, f.w] } +/-- The canonical functor from the product of two categories to the comma category of their +respective functors into `Discrete PUnit`. -/ +@[simps] +def fromProd (L : A ⥤ Discrete PUnit) (R : B ⥤ Discrete PUnit) : + A × B ⥤ Comma L R where + obj X := + { left := X.1 + right := X.2 + hom := Discrete.eqToHom rfl } + map {X} {Y} f := + { left := f.1 + right := f.2 } + +/-- Taking the comma category of two functors into `Discrete PUnit` results in something +is equivalent to their product. -/ +@[simps!] +def equivProd (L : A ⥤ Discrete PUnit) (R : B ⥤ Discrete PUnit) : + Comma L R ≌ A × B := + Equivalence.mk ((fst L R).prod' (snd L R)) (fromProd L R) + { hom := 𝟙 _, inv := 𝟙 _ } + { hom := 𝟙 _, inv := 𝟙 _ } + +/-- Taking the comma category of a functor into `A ⥤ Discrete PUnit` and the identity +`Discrete PUnit ⥤ Discrete PUnit` results in a category equivalent to `A`. -/ +@[simps!] +def toPUnitIdEquiv (L : A ⥤ Discrete PUnit) (R : Discrete PUnit ⥤ Discrete PUnit) : + Comma L R ≌ A := + (equivProd L _).trans (prod.rightUnitorEquivalence A) + +@[simp] +theorem toPUnitIdEquiv_functor_iso {L : A ⥤ Discrete PUnit} + {R : Discrete PUnit ⥤ Discrete PUnit} : + (toPUnitIdEquiv L R).functor = fst L R := + rfl + +/-- Taking the comma category of the identity `Discrete PUnit ⥤ Discrete PUnit` +and a functor `B ⥤ Discrete PUnit` results in a category equivalent to `B`. -/ +@[simps!] +def toIdPUnitEquiv (L : Discrete PUnit ⥤ Discrete PUnit) (R : B ⥤ Discrete PUnit) : + Comma L R ≌ B := + (equivProd _ R).trans (prod.leftUnitorEquivalence B) + +@[simp] +theorem toIdPUnitEquiv_functor_iso {L : Discrete PUnit ⥤ Discrete PUnit} + {R : B ⥤ Discrete PUnit} : + (toIdPUnitEquiv L R).functor = snd L R := + rfl + end end Comma diff --git a/Mathlib/CategoryTheory/Dialectica/Monoidal.lean b/Mathlib/CategoryTheory/Dialectica/Monoidal.lean index 89918b1718d9d..201e37a4b3f3c 100644 --- a/Mathlib/CategoryTheory/Dialectica/Monoidal.lean +++ b/Mathlib/CategoryTheory/Dialectica/Monoidal.lean @@ -58,11 +58,11 @@ local notation "π(" a ", " b ")" => prod.lift a b /-- Left unit cancellation `1 ⊗ X ≅ X` in `Dial C`. -/ @[simps!] def leftUnitor (X : Dial C) : tensorObj tensorUnit X ≅ X := - isoMk (prod.leftUnitor _) (prod.leftUnitor _) <| by simp [Subobject.pullback_top] + isoMk (Limits.prod.leftUnitor _) (Limits.prod.leftUnitor _) <| by simp [Subobject.pullback_top] /-- Right unit cancellation `X ⊗ 1 ≅ X` in `Dial C`. -/ @[simps!] def rightUnitor (X : Dial C) : tensorObj X tensorUnit ≅ X := - isoMk (prod.rightUnitor _) (prod.rightUnitor _) <| by simp [Subobject.pullback_top] + isoMk (Limits.prod.rightUnitor _) (Limits.prod.rightUnitor _) <| by simp [Subobject.pullback_top] /-- The associator for tensor, `(X ⊗ Y) ⊗ Z ≅ X ⊗ (Y ⊗ Z)` in `Dial C`. -/ @[simps!] diff --git a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean index b8482d2c16501..043177863cb9e 100644 --- a/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean +++ b/Mathlib/CategoryTheory/Monoidal/OfHasFiniteProducts.lean @@ -51,8 +51,8 @@ def monoidalOfHasFiniteProducts [HasTerminal C] [HasBinaryProducts C] : Monoidal tensorHom := fun f g ↦ Limits.prod.map f g tensorUnit := ⊤_ C associator := prod.associator - leftUnitor := fun P ↦ prod.leftUnitor P - rightUnitor := fun P ↦ prod.rightUnitor P + leftUnitor := fun P ↦ Limits.prod.leftUnitor P + rightUnitor := fun P ↦ Limits.prod.rightUnitor P } .ofTensorHom (pentagon := prod.pentagon) From a9c4c79d8f2b236dbfd97e1d59f67932e55c3972 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 12 Aug 2024 11:58:20 +0000 Subject: [PATCH 202/359] chore: update Mathlib dependencies 2024-08-12 (#15728) Co-authored-by: mathlib-bors <150093616+mathlib-bors@users.noreply.github.com> --- Mathlib/Data/Bool/Basic.lean | 1 + Mathlib/Data/Option/NAry.lean | 1 + Mathlib/Init/Set.lean | 1 - Mathlib/Logic/Function/Basic.lean | 1 + Mathlib/Tactic/ByContra.lean | 1 + Mathlib/Tactic/DeprecateMe.lean | 2 +- Mathlib/Tactic/ITauto.lean | 1 + .../Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean | 1 + .../Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean | 1 + Mathlib/Tactic/Widget/InteractiveUnfold.lean | 1 + lake-manifest.json | 4 ++-- scripts/noshake.json | 7 ++++++- 12 files changed, 17 insertions(+), 5 deletions(-) diff --git a/Mathlib/Data/Bool/Basic.lean b/Mathlib/Data/Bool/Basic.lean index 231644e90840f..da700043cf44d 100644 --- a/Mathlib/Data/Bool/Basic.lean +++ b/Mathlib/Data/Bool/Basic.lean @@ -3,6 +3,7 @@ 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 -/ +import Batteries.Tactic.Init import Mathlib.Logic.Function.Defs import Mathlib.Order.Defs diff --git a/Mathlib/Data/Option/NAry.lean b/Mathlib/Data/Option/NAry.lean index 870d0d750bd99..855ae63c32d0a 100644 --- a/Mathlib/Data/Option/NAry.lean +++ b/Mathlib/Data/Option/NAry.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Logic.Function.Defs +import Batteries.Tactic.Init /-! # Binary map of options diff --git a/Mathlib/Init/Set.lean b/Mathlib/Init/Set.lean index d37d9d6a5706f..4999dfe196e03 100644 --- a/Mathlib/Init/Set.lean +++ b/Mathlib/Init/Set.lean @@ -3,7 +3,6 @@ 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 Lean.Parser.Term import Batteries.Util.ExtendedBinder /-! diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index 2eae7281fd753..abebf7a7f96df 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro import Mathlib.Logic.Nonempty import Mathlib.Init.Set import Mathlib.Logic.Basic +import Batteries.Tactic.Init /-! # Miscellaneous function constructions and lemmas diff --git a/Mathlib/Tactic/ByContra.lean b/Mathlib/Tactic/ByContra.lean index a6f16134c977c..fa2859c54f3c9 100644 --- a/Mathlib/Tactic/ByContra.lean +++ b/Mathlib/Tactic/ByContra.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Kevin Buzzard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Buzzard -/ +import Batteries.Tactic.Init import Mathlib.Tactic.PushNeg /-! diff --git a/Mathlib/Tactic/DeprecateMe.lean b/Mathlib/Tactic/DeprecateMe.lean index 0320b7cb89a34..3fb56313bec00 100644 --- a/Mathlib/Tactic/DeprecateMe.lean +++ b/Mathlib/Tactic/DeprecateMe.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ - +import Lean.Meta.Tactic.TryThis import Mathlib.Lean.Expr.Basic import Mathlib.Tactic.Lemma diff --git a/Mathlib/Tactic/ITauto.lean b/Mathlib/Tactic/ITauto.lean index 993207dd987f7..fdb027224a4f9 100644 --- a/Mathlib/Tactic/ITauto.lean +++ b/Mathlib/Tactic/ITauto.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import Batteries.Logic import Batteries.Tactic.Exact +import Batteries.Tactic.Init import Mathlib.Tactic.Hint import Mathlib.Tactic.DeriveToExpr import Mathlib.Util.AtomM diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean index a6e84e630d1a3..fc26e112464dd 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Vasily Nesterov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Vasily Nesterov -/ +import Lean.Data.HashMap import Batteries.Data.Rat.Basic /-! diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean index 1c67ffae264eb..8198ec0a53234 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Vasily Nesterov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Vasily Nesterov -/ +import Lean.Meta.Basic import Mathlib.Tactic.Linarith.Oracle.SimplexAlgorithm.SimplexAlgorithm import Mathlib.Tactic.Linarith.Oracle.SimplexAlgorithm.Gauss diff --git a/Mathlib/Tactic/Widget/InteractiveUnfold.lean b/Mathlib/Tactic/Widget/InteractiveUnfold.lean index 5b4ea1885a770..d6b02f009dd2a 100644 --- a/Mathlib/Tactic/Widget/InteractiveUnfold.lean +++ b/Mathlib/Tactic/Widget/InteractiveUnfold.lean @@ -7,6 +7,7 @@ import Batteries.Lean.Position import Mathlib.Tactic.Widget.SelectPanelUtils import Mathlib.Lean.GoalsLocation import Mathlib.Lean.Meta.KAbstractPositions +import Lean.Util.FoldConsts /-! diff --git a/lake-manifest.json b/lake-manifest.json index 23a359909f284..d771a7dee19b1 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "dc167d260ff7ee9849b436037add06bed15104be", + "rev": "ad26fe1ebccc9d5b7ca9111d5daf9b4488374415", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "0444234b4216e944d5be2ce42a25d7410c67876f", + "rev": "776a5a8f9c789395796e442d78a9d4cb9c4c9d03", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", diff --git a/scripts/noshake.json b/scripts/noshake.json index e2defe13f3e81..021b8a75f57ef 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -235,7 +235,8 @@ "Mathlib.Tactic.Linter.UnusedTactic": ["Batteries.Tactic.Unreachable"], "Mathlib.Tactic.Lemma": ["Lean.Parser.Command"], "Mathlib.Tactic.IrreducibleDef": ["Mathlib.Data.Subtype"], - "Mathlib.Tactic.ITauto": ["Batteries.Logic", "Mathlib.Init.Logic"], + "Mathlib.Tactic.ITauto": + ["Batteries.Logic", "Batteries.Tactic.Init", "Mathlib.Init.Logic"], "Mathlib.Tactic.Group": ["Mathlib.Algebra.Group.Commutator"], "Mathlib.Tactic.GCongr.Core": ["Mathlib.Order.Defs"], "Mathlib.Tactic.GCongr": ["Mathlib.Tactic.Positivity.Core"], @@ -279,6 +280,7 @@ "Mathlib.Tactic.CategoryTheory.BicategoricalComp"], "Mathlib.Tactic.CC.Addition": ["Mathlib.Logic.Basic", "Mathlib.Tactic.CC.Lemmas"], + "Mathlib.Tactic.ByContra": ["Batteries.Tactic.Init"], "Mathlib.Tactic.Bound.Init": ["Aesop.Frontend.Command"], "Mathlib.Tactic.Bound.Attribute": ["Mathlib.Algebra.Group.ZeroOne", "Mathlib.Tactic.Bound.Init"], @@ -312,6 +314,7 @@ "Mathlib.MeasureTheory.MeasurableSpace.Defs": ["Mathlib.Tactic.FunProp.Attr"], "Mathlib.Logic.Nontrivial.Defs": ["Mathlib.Init.Logic"], "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": ["Mathlib.Data.Int.Notation"], "Mathlib.LinearAlgebra.Matrix.Transvection": ["Mathlib.Data.Matrix.DMatrix"], @@ -338,6 +341,7 @@ "Mathlib.Data.Seq.Parallel": ["Mathlib.Init.Data.Prod"], "Mathlib.Data.PEquiv": ["Batteries.Tactic.Congr"], "Mathlib.Data.Ordmap.Ordnode": ["Mathlib.Data.Option.Basic"], + "Mathlib.Data.Option.NAry": ["Batteries.Tactic.Init"], "Mathlib.Data.Nat.Find": ["Batteries.WF"], "Mathlib.Data.Multiset.Bind": ["Mathlib.Algebra.GroupWithZero.Action.Defs"], "Mathlib.Data.List.TFAE": ["Batteries.Tactic.Classical"], @@ -352,6 +356,7 @@ "Mathlib.Data.FunLike.Basic": ["Mathlib.Logic.Function.Basic"], "Mathlib.Data.Finset.Basic": ["Mathlib.Data.Finset.Attr"], "Mathlib.Data.DFinsupp.Notation": ["Mathlib.Data.Finsupp.Notation"], + "Mathlib.Data.Bool.Basic": ["Batteries.Tactic.Init"], "Mathlib.Control.Traversable.Instances": ["Mathlib.Control.Applicative"], "Mathlib.Control.Monad.Cont": ["Batteries.Tactic.Congr"], "Mathlib.Control.Basic": ["Mathlib.Logic.Function.Defs"], From b1f62bdd085f55bdcbf1294d5ffcd5a41ea95501 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 12 Aug 2024 12:57:15 +0000 Subject: [PATCH 203/359] chore: remove unused argument from Behrend.map_injOn (#15735) --- .../Additive/AP/Three/Behrend.lean | 17 +++++++---------- 1 file changed, 7 insertions(+), 10 deletions(-) diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean index 6fb6211e86b9e..22b62ad1ed1aa 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean @@ -153,10 +153,10 @@ theorem map_eq_iff {x₁ x₂ : Fin n.succ → ℕ} (hx₁ : ∀ i, x₁ i < d) exact ⟨this, h.resolve_right (pos_of_gt (hx₁ 0)).ne'⟩ theorem map_injOn : {x : Fin n → ℕ | ∀ i, x i < d}.InjOn (map d) := by + clear x intro x₁ hx₁ x₂ hx₂ h induction' n with n ih · simp [eq_iff_true_of_subsingleton] - rw [forall_const] at ih ext i have x := (map_eq_iff hx₁ hx₂).1 h refine Fin.cases x.1 (congr_fun <| ih (fun _ => ?_) (fun _ => ?_) x.2) i @@ -183,13 +183,12 @@ theorem threeAPFree_image_sphere : rw [coe_image] apply ThreeAPFree.image' (α := Fin n → ℕ) (β := ℕ) (s := sphere n d k) (map (2 * d - 1)) (map_injOn.mono _) threeAPFree_sphere - · rw [Set.add_subset_iff] - rintro a ha b hb i - have hai := mem_box.1 (sphere_subset_box ha) i - have hbi := mem_box.1 (sphere_subset_box hb) i - rw [lt_tsub_iff_right, ← succ_le_iff, two_mul] - exact (add_add_add_comm _ _ 1 1).trans_le (_root_.add_le_add hai hbi) - · exact x + rw [Set.add_subset_iff] + rintro a ha b hb i + have hai := mem_box.1 (sphere_subset_box ha) i + have hbi := mem_box.1 (sphere_subset_box hb) i + rw [lt_tsub_iff_right, ← succ_le_iff, two_mul] + exact (add_add_add_comm _ _ 1 1).trans_le (_root_.add_le_add hai hbi) theorem sum_sq_le_of_mem_box (hx : x ∈ box n d) : ∑ i : Fin n, x i ^ 2 ≤ n * (d - 1) ^ 2 := by rw [mem_box] at hx @@ -212,13 +211,11 @@ theorem card_sphere_le_rothNumberNat (n d k : ℕ) : cases d · simp apply threeAPFree_image_sphere.le_rothNumberNat _ _ (card_image_of_injOn _) - · intro; assumption · simp only [subset_iff, mem_image, and_imp, forall_exists_index, mem_range, forall_apply_eq_imp_iff₂, sphere, mem_filter] rintro _ x hx _ rfl exact (map_le_of_mem_box hx).trans_lt sum_lt apply map_injOn.mono fun x => ?_ - · intro; assumption simp only [mem_coe, sphere, mem_filter, mem_box, and_imp, two_mul] exact fun h _ i => (h i).trans_le le_self_add From ac8cc94fcdf51cb273292fe3273534e9acb62275 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Mon, 12 Aug 2024 13:06:49 +0000 Subject: [PATCH 204/359] feat(Algebra/Homology): action of the flip of a bifunctor on homological complexes (#11713) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We construct an isomorphism `mapBifunctor K₂ K₁ F.flip c ≅ mapBifunctor K₁ K₂ F c` when `F` is a bifunctor, and `K₁` and `K₂` are two homological complexes. This is under the assumption of a certain instance of `TotalComplexShapeSymmetry` which provides compatible signs for this isomorphism. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Algebra/Homology/BifunctorFlip.lean | 53 +++++++++++++++++++ .../Algebra/Homology/ComplexShapeSigns.lean | 31 +++++++++++ .../Homology/TotalComplexSymmetry.lean | 16 +++++- .../Limits/Preserves/Shapes/Zero.lean | 6 +++ 5 files changed, 106 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Algebra/Homology/BifunctorFlip.lean diff --git a/Mathlib.lean b/Mathlib.lean index eb1c20596def4..f7fdf1692a312 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -320,6 +320,7 @@ import Mathlib.Algebra.HierarchyDesign import Mathlib.Algebra.Homology.Additive import Mathlib.Algebra.Homology.Augment import Mathlib.Algebra.Homology.Bifunctor +import Mathlib.Algebra.Homology.BifunctorFlip import Mathlib.Algebra.Homology.BifunctorHomotopy import Mathlib.Algebra.Homology.BifunctorShift import Mathlib.Algebra.Homology.CommSq diff --git a/Mathlib/Algebra/Homology/BifunctorFlip.lean b/Mathlib/Algebra/Homology/BifunctorFlip.lean new file mode 100644 index 0000000000000..c6149ca1d5248 --- /dev/null +++ b/Mathlib/Algebra/Homology/BifunctorFlip.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.Algebra.Homology.Bifunctor +import Mathlib.Algebra.Homology.TotalComplexSymmetry + +/-! +# Action of the flip of a bifunctor on homological complexes + +Given `K₁ : HomologicalComplex C₁ c₁`, `K₂ : HomologicalComplex C₂ c₂`, +a bifunctor `F : C₁ ⥤ C₂ ⥤ D`, and a complex shape `c` with +`[TotalComplexShape c₁ c₂ c]` and `[TotalComplexShape c₂ c₁ c]`, we define +an isomorphism `mapBifunctor K₂ K₁ F.flip c ≅ mapBifunctor K₁ K₂ F c` +under the additional assumption `[TotalComplexShapeSymmetry c₁ c₂ c]`. + +-/ + +open CategoryTheory Limits + +variable {C₁ C₂ D : Type*} [Category C₁] [Category C₂] [Category D] + +namespace HomologicalComplex + +variable {I₁ I₂ J : Type*} {c₁ : ComplexShape I₁} {c₂ : ComplexShape I₂} + [HasZeroMorphisms C₁] [HasZeroMorphisms C₂] [Preadditive D] + (K₁ : HomologicalComplex C₁ c₁) (K₂ : HomologicalComplex C₂ c₂) + (F : C₁ ⥤ C₂ ⥤ D) [F.PreservesZeroMorphisms] [∀ X₁, (F.obj X₁).PreservesZeroMorphisms] + (c : ComplexShape J) [TotalComplexShape c₁ c₂ c] [TotalComplexShape c₂ c₁ c] + [TotalComplexShapeSymmetry c₁ c₂ c] [DecidableEq J] + +lemma hasMapBifunctor_flip_iff : + HasMapBifunctor K₂ K₁ F.flip c ↔ HasMapBifunctor K₁ K₂ F c := + (((F.mapBifunctorHomologicalComplex c₁ c₂).obj K₁).obj K₂).flip_hasTotal_iff c + +variable [HasMapBifunctor K₁ K₂ F c] + +instance : HasMapBifunctor K₂ K₁ F.flip c := by + rw [hasMapBifunctor_flip_iff] + infer_instance + +/-- The canonical isomorphism `mapBifunctor K₂ K₁ F.flip c ≅ mapBifunctor K₁ K₂ F c`. -/ +noncomputable def mapBifunctorFlipIso : + mapBifunctor K₂ K₁ F.flip c ≅ mapBifunctor K₁ K₂ F c := + (((F.mapBifunctorHomologicalComplex c₁ c₂).obj K₁).obj K₂).totalFlipIso c + +lemma mapBifunctorFlipIso_flip + [TotalComplexShapeSymmetry c₂ c₁ c] [TotalComplexShapeSymmetrySymmetry c₁ c₂ c] : + mapBifunctorFlipIso K₂ K₁ F.flip c = (mapBifunctorFlipIso K₁ K₂ F c).symm := + (((F.mapBifunctorHomologicalComplex c₁ c₂).obj K₁).obj K₂).flip_totalFlipIso c + +end HomologicalComplex diff --git a/Mathlib/Algebra/Homology/ComplexShapeSigns.lean b/Mathlib/Algebra/Homology/ComplexShapeSigns.lean index 1b6767427cb44..2d544038dc5c3 100644 --- a/Mathlib/Algebra/Homology/ComplexShapeSigns.lean +++ b/Mathlib/Algebra/Homology/ComplexShapeSigns.lean @@ -210,6 +210,15 @@ lemma π_symm (i₁ : I₁) (i₂ : I₂) : π c₂ c₁ c₁₂ ⟨i₂, i₁⟩ = π c₁ c₂ c₁₂ ⟨i₁, i₂⟩ := by apply TotalComplexShapeSymmetry.symm +/-- The symmetry bijection `(π c₂ c₁ c₁₂ ⁻¹' {j}) ≃ (π c₁ c₂ c₁₂ ⁻¹' {j})`. -/ +@[simps] +def symmetryEquiv (j : I₁₂) : + (π c₂ c₁ c₁₂ ⁻¹' {j}) ≃ (π c₁ c₂ c₁₂ ⁻¹' {j}) where + toFun := fun ⟨⟨i₂, i₁⟩, h⟩ => ⟨⟨i₁, i₂⟩, by simpa [π_symm] using h⟩ + invFun := fun ⟨⟨i₁, i₂⟩, h⟩ => ⟨⟨i₂, i₁⟩, by simpa [π_symm] using h⟩ + left_inv _ := rfl + right_inv _ := rfl + variable {c₁} lemma σ_ε₁ {i₁ i₁' : I₁} (h₁ : c₁.Rel i₁ i₁') (i₂ : I₂) : @@ -238,6 +247,28 @@ instance : TotalComplexShapeSymmetry (up ℤ) (up ℤ) (up ℤ) where end ComplexShape +/-- The obvious `TotalComplexShapeSymmetry c₂ c₁ c₁₂` deduced from a +`TotalComplexShapeSymmetry c₁ c₂ c₁₂`. -/ +def TotalComplexShapeSymmetry.symmetry [TotalComplexShape c₁ c₂ c₁₂] + [TotalComplexShape c₂ c₁ c₁₂] [TotalComplexShapeSymmetry c₁ c₂ c₁₂] : + TotalComplexShapeSymmetry c₂ c₁ c₁₂ where + symm i₂ i₁ := (ComplexShape.π_symm c₁ c₂ c₁₂ i₁ i₂).symm + σ i₂ i₁ := ComplexShape.σ c₁ c₂ c₁₂ i₁ i₂ + σ_ε₁ {i₂ i₂'} h₂ i₁ := by + dsimp + apply mul_right_cancel (b := ComplexShape.ε₂ c₁ c₂ c₁₂ (i₁, i₂)) + rw [mul_assoc] + nth_rw 2 [mul_comm] + rw [← mul_assoc, ComplexShape.σ_ε₂ c₁ c₁₂ i₁ h₂, mul_comm, ← mul_assoc, + Int.units_mul_self, one_mul, mul_comm, ← mul_assoc, Int.units_mul_self, one_mul] + σ_ε₂ i₂ i₁ i₁' h₁ := by + dsimp + apply mul_right_cancel (b := ComplexShape.ε₁ c₁ c₂ c₁₂ (i₁, i₂)) + rw [mul_assoc] + nth_rw 2 [mul_comm] + rw [← mul_assoc, ComplexShape.σ_ε₁ c₂ c₁₂ h₁ i₂, mul_comm, ← mul_assoc, + Int.units_mul_self, one_mul, mul_comm, ← mul_assoc, Int.units_mul_self, one_mul] + /-- This typeclass expresses that the signs given by `[TotalComplexShapeSymmetry c₁ c₂ c₁₂]` and by `[TotalComplexShapeSymmetry c₂ c₁ c₁₂]` are compatible. -/ class TotalComplexShapeSymmetrySymmetry [TotalComplexShape c₁ c₂ c₁₂] diff --git a/Mathlib/Algebra/Homology/TotalComplexSymmetry.lean b/Mathlib/Algebra/Homology/TotalComplexSymmetry.lean index 5ef667317230e..8eada119b47c9 100644 --- a/Mathlib/Algebra/Homology/TotalComplexSymmetry.lean +++ b/Mathlib/Algebra/Homology/TotalComplexSymmetry.lean @@ -28,7 +28,21 @@ variable {C I₁ I₂ J : Type*} [Category C] [Preadditive C] {c₁ : ComplexShape I₁} {c₂ : ComplexShape I₂} (K : HomologicalComplex₂ C c₁ c₂) (c : ComplexShape J) [TotalComplexShape c₁ c₂ c] [TotalComplexShape c₂ c₁ c] [TotalComplexShapeSymmetry c₁ c₂ c] - [K.HasTotal c] [K.flip.HasTotal c] [DecidableEq J] + +instance [K.HasTotal c] : K.flip.HasTotal c := fun j => + hasCoproduct_of_equiv_of_iso (K.toGradedObject.mapObjFun (ComplexShape.π c₁ c₂ c) j) _ + (ComplexShape.symmetryEquiv c₁ c₂ c j) (fun _ => Iso.refl _) + +lemma flip_hasTotal_iff : K.flip.HasTotal c ↔ K.HasTotal c := by + constructor + · intro + change K.flip.flip.HasTotal c + have := TotalComplexShapeSymmetry.symmetry c₁ c₂ c + infer_instance + · intro + infer_instance + +variable [K.HasTotal c] [DecidableEq J] attribute [local simp] smul_smul diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean index e506226cfcf15..f13d94165852e 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Zero.lean @@ -109,6 +109,12 @@ lemma preservesZeroMorphisms_of_iso {F₁ F₂ : C ⥤ D} [F₁.PreservesZeroMor instance preservesZeroMorphisms_evaluation_obj (j : D) : PreservesZeroMorphisms ((evaluation D C).obj j) where +instance (F : C ⥤ D ⥤ E) [∀ X, (F.obj X).PreservesZeroMorphisms] : + F.flip.PreservesZeroMorphisms where + +instance (F : C ⥤ D ⥤ E) [F.PreservesZeroMorphisms] (Y : D) : + (F.flip.obj Y).PreservesZeroMorphisms where + end ZeroMorphisms section ZeroObject From 58d29b1d108bbac0573678c9d5269bd635456d30 Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Mon, 12 Aug 2024 13:06:50 +0000 Subject: [PATCH 205/359] feat(Combinatorics/SimpleGraph/Matching): Add matching free graphs (#15357) Add supporting lemmas for graphs free of matchings in preperation for Tuttes theorem --- .../Combinatorics/SimpleGraph/Matching.lean | 25 +++++++++++++++++++ Mathlib/Data/Fintype/Order.lean | 12 +++++++++ 2 files changed, 37 insertions(+) diff --git a/Mathlib/Combinatorics/SimpleGraph/Matching.lean b/Mathlib/Combinatorics/SimpleGraph/Matching.lean index 2b7ea6ab503ed..70951275d32d1 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Matching.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Matching.lean @@ -5,6 +5,7 @@ Authors: Alena Gusakov, Arthur Paulino, Kyle Miller, Pim Otte -/ import Mathlib.Combinatorics.SimpleGraph.DegreeSum import Mathlib.Combinatorics.SimpleGraph.Connectivity.WalkCounting +import Mathlib.Data.Fintype.Order /-! # Matchings @@ -228,4 +229,28 @@ lemma odd_matches_node_outside {u : Set V} {c : ConnectedComponent (Subgraph.del end Finite end ConnectedComponent + +/-- +A graph is matching free if it has no perfect matching. It does not make much sense to +consider a graph being free of just matchings, because any non-trivial graph has those. +-/ +def IsMatchingFree (G : SimpleGraph V) := ∀ M : Subgraph G, ¬ M.IsPerfectMatching + +lemma IsMatchingFree.mono {G G' : SimpleGraph V} (h : G ≤ G') (hmf : G'.IsMatchingFree) : + G.IsMatchingFree := by + intro x + by_contra! hc + apply hmf (x.map (SimpleGraph.Hom.ofLE h)) + refine ⟨hc.1.map_ofLE h, ?_⟩ + intro v + simp only [Subgraph.map_verts, Hom.coe_ofLE, id_eq, Set.image_id'] + exact hc.2 v + +lemma exists_maximal_isMatchingFree [Fintype V] [DecidableEq V] + (h : G.IsMatchingFree) : ∃ Gmax : SimpleGraph V, + G ≤ Gmax ∧ Gmax.IsMatchingFree ∧ ∀ G', G' > Gmax → ∃ M : Subgraph G', M.IsPerfectMatching := by + simp_rw [← @not_forall_not _ Subgraph.IsPerfectMatching] + obtain ⟨Gmax, hGmax⟩ := Finite.exists_le_maximal h + exact ⟨Gmax, ⟨hGmax.1, ⟨hGmax.2.prop, fun _ h' ↦ hGmax.2.not_prop_of_gt h'⟩⟩⟩ + end SimpleGraph diff --git a/Mathlib/Data/Fintype/Order.lean b/Mathlib/Data/Fintype/Order.lean index de38202522c22..daef09e27fc55 100644 --- a/Mathlib/Data/Fintype/Order.lean +++ b/Mathlib/Data/Fintype/Order.lean @@ -154,6 +154,18 @@ end Nonempty end Fintype +/-! ### Properties for PartialOrders -/ + +lemma Finite.exists_ge_minimal {α} [Finite α] [PartialOrder α] {a : α} {p : α → Prop} (h : p a) : + ∃ b, b ≤ a ∧ Minimal p b := by + obtain ⟨b, ⟨hba, hb⟩, hbmin⟩ := + Set.Finite.exists_minimal_wrt id {x | x ≤ a ∧ p x} (Set.toFinite _) ⟨a, rfl.le, h⟩ + exact ⟨b, hba, hb, fun x hx hxb ↦ (hbmin x ⟨hxb.trans hba, hx⟩ hxb).le⟩ + +lemma Finite.exists_le_maximal {α} [Finite α] [PartialOrder α] {a : α} {p : α → Prop} (h : p a) : + ∃ b, a ≤ b ∧ Maximal p b := + Finite.exists_ge_minimal (α := αᵒᵈ) h + /-! ### Concrete instances -/ noncomputable instance Fin.completeLinearOrder {n : ℕ} [NeZero n] : CompleteLinearOrder (Fin n) := From a83f96ceba2f38233abd913ccb8cc8f7f3e09dd1 Mon Sep 17 00:00:00 2001 From: grhkm21 Date: Mon, 12 Aug 2024 13:06:52 +0000 Subject: [PATCH 206/359] chore: Fix enccard typo in Cardinal file (#15709) Fix typo in theorem name Moves: - Function.Embedding.enccard_le -> Function.Embedding.encard_le --- Mathlib/Data/Set/Card.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index 9984499e1f215..d8d7d4405bf96 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -389,7 +389,7 @@ theorem _root_.Function.Injective.encard_image (hf : f.Injective) (s : Set α) : (f '' s).encard = s.encard := hf.injOn.encard_image -theorem _root_.Function.Embedding.enccard_le (e : s ↪ t) : s.encard ≤ t.encard := by +theorem _root_.Function.Embedding.encard_le (e : s ↪ t) : s.encard ≤ t.encard := by rw [← encard_univ_coe, ← e.injective.encard_image, ← Subtype.coe_injective.encard_image] exact encard_mono (by simp) From 40b6b147bd382cc0a0622f0c0934a30197032944 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Mon, 12 Aug 2024 13:06:54 +0000 Subject: [PATCH 207/359] feat(Algebra/Homology): develop the API for the action of bifunctors on homological complexes (#15712) --- Mathlib/Algebra/Homology/Bifunctor.lean | 121 ++++++++++++++++++ .../Algebra/Homology/BifunctorHomotopy.lean | 13 +- Mathlib/Algebra/Homology/TotalComplex.lean | 2 +- .../Algebra/Homology/TotalComplexShift.lean | 4 +- .../Homology/TotalComplexSymmetry.lean | 3 +- .../GradedObject/Trifunctor.lean | 67 +++++++++- 6 files changed, 192 insertions(+), 18 deletions(-) diff --git a/Mathlib/Algebra/Homology/Bifunctor.lean b/Mathlib/Algebra/Homology/Bifunctor.lean index bb5c93ccda6be..ab3eea25ef2c6 100644 --- a/Mathlib/Algebra/Homology/Bifunctor.lean +++ b/Mathlib/Algebra/Homology/Bifunctor.lean @@ -137,6 +137,127 @@ lemma ιMapBifunctorOrZero_eq_zero (i₁ : I₁) (i₂ : I₂) (j : J) section +variable {K₁ K₂ F c} +variable {A : D} {j : J} + (f : ∀ (i₁ : I₁) (i₂ : I₂) (_ : ComplexShape.π c₁ c₂ c ⟨i₁, i₂⟩ = j), + (F.obj (K₁.X i₁)).obj (K₂.X i₂) ⟶ A) + +/-- Constructor for morphisms from `(mapBifunctor K₁ K₂ F c).X j`. -/ +noncomputable def mapBifunctorDesc : (mapBifunctor K₁ K₂ F c).X j ⟶ A := + HomologicalComplex₂.totalDesc _ f + +@[reassoc (attr := simp)] +lemma ι_mapBifunctorDesc (i₁ : I₁) (i₂ : I₂) (h : ComplexShape.π c₁ c₂ c ⟨i₁, i₂⟩ = j) : + ιMapBifunctor K₁ K₂ F c i₁ i₂ j h ≫ mapBifunctorDesc f = f i₁ i₂ h := by + apply HomologicalComplex₂.ι_totalDesc + +end + +namespace mapBifunctor + +variable {K₁ K₂ F c} in +@[ext] +lemma hom_ext {Y : D} {j : J} {f g : (mapBifunctor K₁ K₂ F c).X j ⟶ Y} + (h : ∀ (i₁ : I₁) (i₂ : I₂) (h : ComplexShape.π c₁ c₂ c ⟨i₁, i₂⟩ = j), + ιMapBifunctor K₁ K₂ F c i₁ i₂ j h ≫ f = ιMapBifunctor K₁ K₂ F c i₁ i₂ j h ≫ g) : + f = g := + HomologicalComplex₂.total.hom_ext _ h + +section + +variable (j j' : J) + +/-- The first differential on `mapBifunctor K₁ K₂ F c` -/ +noncomputable def D₁ : + (mapBifunctor K₁ K₂ F c).X j ⟶ (mapBifunctor K₁ K₂ F c).X j' := + (((F.mapBifunctorHomologicalComplex c₁ c₂).obj K₁).obj K₂).D₁ c j j' + +/-- The second differential on `mapBifunctor K₁ K₂ F c` -/ +noncomputable def D₂ : + (mapBifunctor K₁ K₂ F c).X j ⟶ (mapBifunctor K₁ K₂ F c).X j' := + (((F.mapBifunctorHomologicalComplex c₁ c₂).obj K₁).obj K₂).D₂ c j j' + +lemma d_eq : + (mapBifunctor K₁ K₂ F c).d j j' = D₁ K₁ K₂ F c j j' + D₂ K₁ K₂ F c j j' := rfl + +end + +section + +variable (i₁ : I₁) (i₂ : I₂) (j : J) + +/-- The first differential on a summand of `mapBifunctor K₁ K₂ F c` -/ +noncomputable def d₁ : + (F.obj (K₁.X i₁)).obj (K₂.X i₂) ⟶ (mapBifunctor K₁ K₂ F c).X j := + (((F.mapBifunctorHomologicalComplex c₁ c₂).obj K₁).obj K₂).d₁ c i₁ i₂ j + +/-- The second differential on a summand of `mapBifunctor K₁ K₂ F c` -/ +noncomputable def d₂ : + (F.obj (K₁.X i₁)).obj (K₂.X i₂) ⟶ (mapBifunctor K₁ K₂ F c).X j := + (((F.mapBifunctorHomologicalComplex c₁ c₂).obj K₁).obj K₂).d₂ c i₁ i₂ j + +lemma d₁_eq_zero (h : ¬ c₁.Rel i₁ (c₁.next i₁)): + d₁ K₁ K₂ F c i₁ i₂ j = 0 := + HomologicalComplex₂.d₁_eq_zero _ _ _ _ _ h + +lemma d₂_eq_zero (h : ¬ c₂.Rel i₂ (c₂.next i₂)): + d₂ K₁ K₂ F c i₁ i₂ j = 0 := + HomologicalComplex₂.d₂_eq_zero _ _ _ _ _ h + +lemma d₁_eq_zero' {i₁ i₁' : I₁} (h : c₁.Rel i₁ i₁') (i₂ : I₂) (j : J) + (h' : ComplexShape.π c₁ c₂ c ⟨i₁', i₂⟩ ≠ j) : + d₁ K₁ K₂ F c i₁ i₂ j = 0 := + HomologicalComplex₂.d₁_eq_zero' _ _ h _ _ h' + +lemma d₂_eq_zero' (i₁ : I₁) {i₂ i₂' : I₂} (h : c₂.Rel i₂ i₂') (j : J) + (h' : ComplexShape.π c₁ c₂ c ⟨i₁, i₂'⟩ ≠ j) : + d₂ K₁ K₂ F c i₁ i₂ j = 0 := + HomologicalComplex₂.d₂_eq_zero' _ _ _ h _ h' + +lemma d₁_eq' {i₁ i₁' : I₁} (h : c₁.Rel i₁ i₁') (i₂ : I₂) (j : J) : + d₁ K₁ K₂ F c i₁ i₂ j = ComplexShape.ε₁ c₁ c₂ c ⟨i₁, i₂⟩ • + ((F.map (K₁.d i₁ i₁')).app (K₂.X i₂) ≫ ιMapBifunctorOrZero K₁ K₂ F c i₁' i₂ j) := + HomologicalComplex₂.d₁_eq' _ _ h _ _ + +lemma d₂_eq' (i₁ : I₁) {i₂ i₂' : I₂} (h : c₂.Rel i₂ i₂') (j : J) : + d₂ K₁ K₂ F c i₁ i₂ j = ComplexShape.ε₂ c₁ c₂ c ⟨i₁, i₂⟩ • + ((F.obj (K₁.X i₁)).map (K₂.d i₂ i₂') ≫ ιMapBifunctorOrZero K₁ K₂ F c i₁ i₂' j) := + HomologicalComplex₂.d₂_eq' _ _ _ h _ + +lemma d₁_eq {i₁ i₁' : I₁} (h : c₁.Rel i₁ i₁') (i₂ : I₂) (j : J) + (h' : ComplexShape.π c₁ c₂ c ⟨i₁', i₂⟩ = j) : + d₁ K₁ K₂ F c i₁ i₂ j = ComplexShape.ε₁ c₁ c₂ c ⟨i₁, i₂⟩ • + ((F.map (K₁.d i₁ i₁')).app (K₂.X i₂) ≫ ιMapBifunctor K₁ K₂ F c i₁' i₂ j h') := + HomologicalComplex₂.d₁_eq _ _ h _ _ h' + +lemma d₂_eq (i₁ : I₁) {i₂ i₂' : I₂} (h : c₂.Rel i₂ i₂') (j : J) + (h' : ComplexShape.π c₁ c₂ c ⟨i₁, i₂'⟩ = j) : + d₂ K₁ K₂ F c i₁ i₂ j = ComplexShape.ε₂ c₁ c₂ c ⟨i₁, i₂⟩ • + ((F.obj (K₁.X i₁)).map (K₂.d i₂ i₂') ≫ ιMapBifunctor K₁ K₂ F c i₁ i₂' j h') := + HomologicalComplex₂.d₂_eq _ _ _ h _ h' + +end + +section + +variable (j j' : J) (i₁ : I₁) (i₂ : I₂) (h : ComplexShape.π c₁ c₂ c (i₁, i₂) = j) + +@[reassoc (attr := simp)] +lemma ι_D₁ : + ιMapBifunctor K₁ K₂ F c i₁ i₂ j h ≫ D₁ K₁ K₂ F c j j' = d₁ K₁ K₂ F c i₁ i₂ j' := by + apply HomologicalComplex₂.ι_D₁ + +@[reassoc (attr := simp)] +lemma ι_D₂ : + ιMapBifunctor K₁ K₂ F c i₁ i₂ j h ≫ D₂ K₁ K₂ F c j j' = d₂ K₁ K₂ F c i₁ i₂ j' := by + apply HomologicalComplex₂.ι_D₂ + +end + +end mapBifunctor + +section + variable {K₁ K₂ L₁ L₂} /-- The morphism `mapBifunctor K₁ K₂ F c ⟶ mapBifunctor L₁ L₂ F c` induced by diff --git a/Mathlib/Algebra/Homology/BifunctorHomotopy.lean b/Mathlib/Algebra/Homology/BifunctorHomotopy.lean index 286bd40522c88..033002ed35b6b 100644 --- a/Mathlib/Algebra/Homology/BifunctorHomotopy.lean +++ b/Mathlib/Algebra/Homology/BifunctorHomotopy.lean @@ -93,13 +93,14 @@ lemma comm₁ (j : J) : (mapBifunctor L₁ L₂ F c).d (c.prev j) j + (mapBifunctorMap f₁' f₂ F c).f j := by ext i₁ i₂ h - simp? [h₁.comm i₁, dFrom, fromNext, toPrev, dTo] says - simp only [Functor.mapBifunctorHomologicalComplex_obj_obj_X_X, ι_mapBifunctorMap, - h₁.comm i₁, dNext_eq_dFrom_fromNext, dFrom, fromNext, AddMonoidHom.mk'_apply, - prevD_eq_toPrev_dTo, toPrev, dTo, Functor.map_add, Functor.map_comp, NatTrans.app_add, - NatTrans.comp_app, Preadditive.add_comp, assoc, HomologicalComplex₂.total_d, + simp? [HomologicalComplex₂.total_d, h₁.comm i₁, dFrom, fromNext, toPrev, dTo] says + simp only [ι_mapBifunctorMap, h₁.comm i₁, dNext_eq_dFrom_fromNext, dFrom, fromNext, + AddMonoidHom.mk'_apply, prevD_eq_toPrev_dTo, toPrev, dTo, Functor.map_add, + Functor.map_comp, NatTrans.app_add, NatTrans.comp_app, + Preadditive.add_comp, assoc, HomologicalComplex₂.total_d, Functor.mapBifunctorHomologicalComplex_obj_obj_toGradedObject, Preadditive.comp_add, - HomologicalComplex₂.ι_D₁_assoc, HomologicalComplex₂.ι_D₂_assoc, add_left_inj] + HomologicalComplex₂.ι_D₁_assoc, Functor.mapBifunctorHomologicalComplex_obj_obj_X_X, + HomologicalComplex₂.ι_D₂_assoc, add_left_inj] have : ∀ {X Y : D} (a b c d e f : X ⟶ Y), a = c → b = e → f = -d → a + b = c + d + (e + f) := by rintro X Y a b _ d _ _ rfl rfl rfl; abel apply this diff --git a/Mathlib/Algebra/Homology/TotalComplex.lean b/Mathlib/Algebra/Homology/TotalComplex.lean index 7bb57bddef5e1..3a338452f6582 100644 --- a/Mathlib/Algebra/Homology/TotalComplex.lean +++ b/Mathlib/Algebra/Homology/TotalComplex.lean @@ -243,7 +243,7 @@ lemma D₁_D₂ (i₁₂ i₁₂' i₁₂'' : I₁₂) : K.D₁ c₁₂ i₁₂ i₁₂' ≫ K.D₂ c₁₂ i₁₂' i₁₂'' = - K.D₂ c₁₂ i₁₂ i₁₂' ≫ K.D₁ c₁₂ i₁₂' i₁₂'' := by simp /-- The total complex of a bicomplex. -/ -@[simps d] +@[simps (config := .lemmasOnly) d] noncomputable def total : HomologicalComplex C c₁₂ where X := K.toGradedObject.mapObj (ComplexShape.π c₁ c₂ c₁₂) d i₁₂ i₁₂' := K.D₁ c₁₂ i₁₂ i₁₂' + K.D₂ c₁₂ i₁₂ i₁₂' diff --git a/Mathlib/Algebra/Homology/TotalComplexShift.lean b/Mathlib/Algebra/Homology/TotalComplexShift.lean index 320dda9418c3e..756aa47f0785d 100644 --- a/Mathlib/Algebra/Homology/TotalComplexShift.lean +++ b/Mathlib/Algebra/Homology/TotalComplexShift.lean @@ -188,7 +188,7 @@ noncomputable def totalShift₁Iso : HomologicalComplex.Hom.isoOfComponents (fun n => K.totalShift₁XIso x n (n + x) rfl) (fun n n' _ => by dsimp - simp only [Preadditive.add_comp, Preadditive.comp_add, smul_add, + simp only [total_d, Preadditive.add_comp, Preadditive.comp_add, smul_add, Linear.comp_units_smul, K.D₁_totalShift₁XIso_hom x n n' _ _ rfl rfl, K.D₂_totalShift₁XIso_hom x n n' _ _ rfl rfl]) @@ -300,7 +300,7 @@ noncomputable def totalShift₂Iso : HomologicalComplex.Hom.isoOfComponents (fun n => K.totalShift₂XIso y n (n + y) rfl) (fun n n' _ => by dsimp - simp only [Preadditive.add_comp, Preadditive.comp_add, smul_add, + simp only [total_d, Preadditive.add_comp, Preadditive.comp_add, smul_add, Linear.comp_units_smul, K.D₁_totalShift₂XIso_hom y n n' _ _ rfl rfl, K.D₂_totalShift₂XIso_hom y n n' _ _ rfl rfl]) diff --git a/Mathlib/Algebra/Homology/TotalComplexSymmetry.lean b/Mathlib/Algebra/Homology/TotalComplexSymmetry.lean index 8eada119b47c9..be865c8ddd218 100644 --- a/Mathlib/Algebra/Homology/TotalComplexSymmetry.lean +++ b/Mathlib/Algebra/Homology/TotalComplexSymmetry.lean @@ -101,8 +101,7 @@ lemma totalFlipIsoX_hom_D₂ (j j' : J) : bicomplex when we have `[TotalComplexShapeSymmetry c₁ c₂ c]`. -/ noncomputable def totalFlipIso : K.flip.total c ≅ K.total c := HomologicalComplex.Hom.isoOfComponents (K.totalFlipIsoX c) (fun j j' _ => by - dsimp - simp only [Preadditive.comp_add, totalFlipIsoX_hom_D₁, + simp only [total_d, Preadditive.comp_add, totalFlipIsoX_hom_D₁, totalFlipIsoX_hom_D₂, Preadditive.add_comp] rw [add_comm]) diff --git a/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean b/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean index 1987e4374302f..5568b3f19475a 100644 --- a/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean +++ b/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean @@ -29,10 +29,13 @@ open Category Limits variable {C₁ C₂ C₃ C₄ C₁₂ C₂₃ : Type*} [Category C₁] [Category C₂] [Category C₃] [Category C₄] [Category C₁₂] [Category C₂₃] - (F F' : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄) namespace GradedObject +section + +variable (F F' : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄) + /-- Auxiliary definition for `mapTrifunctor`. -/ @[simps] def mapTrifunctorObj {I₁ : Type*} (X₁ : GradedObject I₁ C₁) (I₂ I₃ : Type*) : @@ -61,10 +64,11 @@ def mapTrifunctor (I₁ I₂ I₃ : Type*) : simp only [← NatTrans.comp_app] congr 1 rw [NatTrans.naturality] } +end section -variable {F F'} +variable {F F' : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄} /-- The natural transformation `mapTrifunctor F I₁ I₂ I₃ ⟶ mapTrifunctor F' I₁ I₂ I₃` induced by a natural transformation `F ⟶ F` of trifunctors. -/ @@ -103,6 +107,7 @@ end section +variable (F : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₄) variable {I₁ I₂ I₃ J : Type*} (p : I₁ × I₂ × I₃ → J) /-- Given a trifunctor `F : C₁ ⥤ C₂ ⥤ C₃ ⥤ C₃`, graded objects `X₁ : GradedObject I₁ C₁`, @@ -379,10 +384,11 @@ lemma ι_mapBifunctorComp₁₂MapObjIso_inv (i₁ : I₁) (i₂ : I₂) (i₃ : (isColimitCofan₃MapBifunctor₁₂BifunctorMapObj F₁₂ G ρ₁₂ X₁ X₂ X₃ j) _ h variable {X₁ X₂ X₃ F₁₂ G ρ₁₂} +variable {j : J} {A : C₄} @[ext] -lemma mapBifunctor₁₂BifunctorMapObj_ext {j : J} {A : C₄} - (f g : mapBifunctorMapObj G ρ₁₂.q (mapBifunctorMapObj F₁₂ ρ₁₂.p X₁ X₂) X₃ j ⟶ A) +lemma mapBifunctor₁₂BifunctorMapObj_ext {A : C₄} + {f g : mapBifunctorMapObj G ρ₁₂.q (mapBifunctorMapObj F₁₂ ρ₁₂.p X₁ X₂) X₃ j ⟶ A} (h : ∀ (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) (h : r ⟨i₁, i₂, i₃⟩ = j), ιMapBifunctor₁₂BifunctorMapObj F₁₂ G ρ₁₂ X₁ X₂ X₃ i₁ i₂ i₃ j h ≫ f = ιMapBifunctor₁₂BifunctorMapObj F₁₂ G ρ₁₂ X₁ X₂ X₃ i₁ i₂ i₃ j h ≫ g) : f = g := by @@ -390,6 +396,28 @@ lemma mapBifunctor₁₂BifunctorMapObj_ext {j : J} {A : C₄} rintro ⟨i, hi⟩ exact h _ _ _ hi +section + +variable (f : ∀ (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) (_ : r ⟨i₁, i₂, i₃⟩ = j), + (G.obj ((F₁₂.obj (X₁ i₁)).obj (X₂ i₂))).obj (X₃ i₃) ⟶ A) + +/-- Constructor for morphisms from +`mapBifunctorMapObj G ρ₁₂.q (mapBifunctorMapObj F₁₂ ρ₁₂.p X₁ X₂) X₃ j`. -/ +noncomputable def mapBifunctor₁₂BifunctorDesc : + mapBifunctorMapObj G ρ₁₂.q (mapBifunctorMapObj F₁₂ ρ₁₂.p X₁ X₂) X₃ j ⟶ A := + Cofan.IsColimit.desc (isColimitCofan₃MapBifunctor₁₂BifunctorMapObj F₁₂ G ρ₁₂ X₁ X₂ X₃ j) + (fun i ↦ f i.1.1 i.1.2.1 i.1.2.2 i.2) + +@[reassoc (attr := simp)] +lemma ι_mapBifunctor₁₂BifunctorDesc + (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) (h : r ⟨i₁, i₂, i₃⟩ = j) : + ιMapBifunctor₁₂BifunctorMapObj F₁₂ G ρ₁₂ X₁ X₂ X₃ i₁ i₂ i₃ j h ≫ + mapBifunctor₁₂BifunctorDesc f = f i₁ i₂ i₃ h := + Cofan.IsColimit.fac + (isColimitCofan₃MapBifunctor₁₂BifunctorMapObj F₁₂ G ρ₁₂ X₁ X₂ X₃ j) _ ⟨_, h⟩ + +end + end section @@ -528,11 +556,12 @@ lemma ι_mapBifunctorComp₂₃MapObjIso_inv (i₁ : I₁) (i₂ : I₂) (i₃ : CofanMapObjFun.inj_iso_hom (isColimitCofan₃MapBifunctorBifunctor₂₃MapObj F G₂₃ ρ₂₃ X₁ X₂ X₃ j) _ h -variable {X₁ X₂ X₃ F₁₂ G ρ₁₂} +variable {X₁ X₂ X₃ F G₂₃ ρ₂₃} +variable {j : J} {A : C₄} @[ext] -lemma mapBifunctorBifunctor₂₃MapObj_ext {j : J} {A : C₄} - (f g : mapBifunctorMapObj F ρ₂₃.q X₁ (mapBifunctorMapObj G₂₃ ρ₂₃.p X₂ X₃) j ⟶ A) +lemma mapBifunctorBifunctor₂₃MapObj_ext + {f g : mapBifunctorMapObj F ρ₂₃.q X₁ (mapBifunctorMapObj G₂₃ ρ₂₃.p X₂ X₃) j ⟶ A} (h : ∀ (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) (h : r ⟨i₁, i₂, i₃⟩ = j), ιMapBifunctorBifunctor₂₃MapObj F G₂₃ ρ₂₃ X₁ X₂ X₃ i₁ i₂ i₃ j h ≫ f = ιMapBifunctorBifunctor₂₃MapObj F G₂₃ ρ₂₃ X₁ X₂ X₃ i₁ i₂ i₃ j h ≫ g) : f = g := by @@ -540,6 +569,30 @@ lemma mapBifunctorBifunctor₂₃MapObj_ext {j : J} {A : C₄} rintro ⟨i, hi⟩ exact h _ _ _ hi +section + +variable + (f : ∀ (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) (_ : r ⟨i₁, i₂, i₃⟩ = j), + (F.obj (X₁ i₁)).obj ((G₂₃.obj (X₂ i₂)).obj (X₃ i₃)) ⟶ A) + +/-- Constructor for morphisms from +`mapBifunctorMapObj F ρ₂₃.q X₁ (mapBifunctorMapObj G₂₃ ρ₂₃.p X₂ X₃) j`. -/ +noncomputable def mapBifunctorBifunctor₂₃Desc : + mapBifunctorMapObj F ρ₂₃.q X₁ (mapBifunctorMapObj G₂₃ ρ₂₃.p X₂ X₃) j ⟶ A := + Cofan.IsColimit.desc (isColimitCofan₃MapBifunctorBifunctor₂₃MapObj F G₂₃ ρ₂₃ X₁ X₂ X₃ j) + (fun i ↦ f i.1.1 i.1.2.1 i.1.2.2 i.2) + +@[reassoc (attr := simp)] +lemma ι_mapBifunctorBifunctor₂₃Desc + (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) (h : r ⟨i₁, i₂, i₃⟩ = j) : + ιMapBifunctorBifunctor₂₃MapObj F G₂₃ ρ₂₃ X₁ X₂ X₃ i₁ i₂ i₃ j h ≫ + mapBifunctorBifunctor₂₃Desc f = f i₁ i₂ i₃ h := + Cofan.IsColimit.fac + (isColimitCofan₃MapBifunctorBifunctor₂₃MapObj F G₂₃ ρ₂₃ X₁ X₂ X₃ j) _ ⟨_, h⟩ + +end + + end end GradedObject From f9b79d693a11c915bba0e93c8f55925aa3b1b519 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 12 Aug 2024 13:06:56 +0000 Subject: [PATCH 208/359] chore(SetTheory/Game/PGame): Add note on `NatCast` (#15713) I explain that this is not in fact the usual definition for natural numbers as games (a fact that just cost me a few hours of dev time...) --- Mathlib/SetTheory/Game/PGame.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index 3f2affe6986e1..6238d450593da 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -1267,7 +1267,11 @@ instance : Add PGame.{u} := · exact fun i => IHxr i y · exact IHyr⟩ -/-- The pre-game `((0+1)+⋯)+1`. -/ +/-- The pre-game `((0 + 1) + ⋯) + 1`. + +Note that this is **not** the usual recursive definition `n = {0, 1, … | }`. For instance, +`2 = 0 + 1 + 1 = {0 + 0 + 1, 0 + 1 + 0 | }` does not contain any left option equivalent to `0`. For +an implementation of said definition, see `Ordinal.toPGame`. -/ instance : NatCast PGame := ⟨Nat.unaryCast⟩ From 46a72deeb18d07ecf8728f8c91cc6bba41ee64c7 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 12 Aug 2024 13:31:34 +0000 Subject: [PATCH 209/359] feat(Ergodic): the left action of a group on itself is ergodic (#14596) --- Mathlib.lean | 1 + Mathlib/Dynamics/Ergodic/Action/Regular.lean | 49 +++++++++++++++++++ .../Constructions/Prod/Basic.lean | 10 ++++ 3 files changed, 60 insertions(+) create mode 100644 Mathlib/Dynamics/Ergodic/Action/Regular.lean diff --git a/Mathlib.lean b/Mathlib.lean index f7fdf1692a312..152422eef1ca1 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2546,6 +2546,7 @@ import Mathlib.Dynamics.BirkhoffSum.Basic import Mathlib.Dynamics.BirkhoffSum.NormedSpace import Mathlib.Dynamics.Circle.RotationNumber.TranslationNumber import Mathlib.Dynamics.Ergodic.Action.Basic +import Mathlib.Dynamics.Ergodic.Action.Regular import Mathlib.Dynamics.Ergodic.AddCircle import Mathlib.Dynamics.Ergodic.Conservative import Mathlib.Dynamics.Ergodic.Ergodic diff --git a/Mathlib/Dynamics/Ergodic/Action/Regular.lean b/Mathlib/Dynamics/Ergodic/Action/Regular.lean new file mode 100644 index 0000000000000..7abce701f315e --- /dev/null +++ b/Mathlib/Dynamics/Ergodic/Action/Regular.lean @@ -0,0 +1,49 @@ +/- +Copyright (c) 2024 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Dynamics.Ergodic.Action.Basic +import Mathlib.MeasureTheory.Group.Prod + +/-! +# Regular action of a group on itself is ergodic + +In this file we prove that the left and right actions of a group on itself are ergodic. +-/ + +open MeasureTheory Measure Filter Set +open scoped Pointwise + +variable {G : Type*} [Group G] [MeasurableSpace G] [MeasurableMul₂ G] [MeasurableInv G] + {μ : Measure G} [SFinite μ] + +@[to_additive] +instance [μ.IsMulLeftInvariant] : ErgodicSMul G G μ := by + refine ⟨fun {s} hsm hs ↦ ?_⟩ + suffices (∃ᵐ x ∂μ, x ∈ s) → ∀ᵐ x ∂μ, x ∈ s by + simp only [eventuallyConst_set, ← not_frequently] + exact or_not_of_imp this + intro hμs + obtain ⟨a, has, ha⟩ : ∃ a ∈ s, ∀ᵐ b ∂μ, (b * a ∈ s ↔ a ∈ s) := by + refine (hμs.and_eventually ?_).exists + rw [ae_ae_comm] + · exact ae_of_all _ fun b ↦ (hs b).mem_iff + · exact ((hsm.preimage <| measurable_snd.mul measurable_fst).mem.iff + (hsm.preimage measurable_fst).mem).setOf + simpa [has] using (MeasureTheory.quasiMeasurePreserving_mul_right μ a⁻¹).ae ha + +@[to_additive] +instance [μ.IsMulRightInvariant] : ErgodicSMul Gᵐᵒᵖ G μ := by + refine ⟨fun {s} hsm hs ↦ ?_⟩ + suffices (∃ᵐ x ∂μ, x ∈ s) → ∀ᵐ x ∂μ, x ∈ s by + simp only [eventuallyConst_set, ← not_frequently] + exact or_not_of_imp this + intro hμs + obtain ⟨a, has, ha⟩ : ∃ a ∈ s, ∀ᵐ b ∂μ, (a * b ∈ s ↔ a ∈ s) := by + refine (hμs.and_eventually ?_).exists + rw [ae_ae_comm] + · exact ae_of_all _ fun b ↦ (hs ⟨b⟩).mem_iff + · exact ((hsm.preimage <| measurable_fst.mul measurable_snd).mem.iff + (hsm.preimage measurable_fst).mem).setOf + simpa [has] using (quasiMeasurePreserving_mul_left μ a⁻¹).ae ha diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean index ebf5efd059fd2..e219a86ee54da 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean @@ -474,6 +474,10 @@ theorem ae_ae_eq_of_ae_eq_uncurry {γ : Type*} {f g : α → β → γ} (h : uncurry f =ᵐ[μ.prod ν] uncurry g) : ∀ᵐ x ∂μ, f x =ᵐ[ν] g x := ae_ae_eq_curry_of_prod h +theorem ae_prod_iff_ae_ae {p : α × β → Prop} (hp : MeasurableSet {x | p x}) : + (∀ᵐ z ∂μ.prod ν, p z) ↔ ∀ᵐ x ∂μ, ∀ᵐ y ∂ν, p (x, y) := + measure_prod_null hp.compl + theorem ae_prod_mem_iff_ae_ae_mem {s : Set (α × β)} (hs : MeasurableSet s) : (∀ᵐ z ∂μ.prod ν, z ∈ s) ↔ ∀ᵐ x ∂μ, ∀ᵐ y ∂ν, (x, y) ∈ s := measure_prod_null hs.compl @@ -637,6 +641,12 @@ theorem prod_apply_symm {s : Set (α × β)} (hs : MeasurableSet s) : rw [← prod_swap, map_apply measurable_swap hs, prod_apply (measurable_swap hs)] rfl +theorem ae_ae_comm {p : α → β → Prop} (h : MeasurableSet {x : α × β | p x.1 x.2}) : + (∀ᵐ x ∂μ, ∀ᵐ y ∂ν, p x y) ↔ ∀ᵐ y ∂ν, ∀ᵐ x ∂μ, p x y := calc + _ ↔ ∀ᵐ x ∂μ.prod ν, p x.1 x.2 := .symm <| ae_prod_iff_ae_ae h + _ ↔ ∀ᵐ x ∂ν.prod μ, p x.2 x.1 := by rw [← prod_swap, ae_map_iff (by fun_prop) h]; rfl + _ ↔ ∀ᵐ y ∂ν, ∀ᵐ x ∂μ, p x y := ae_prod_iff_ae_ae <| measurable_swap h + /-- If `s ×ˢ t` is a null measurable set and `ν t ≠ 0`, then `s` is a null measurable set. -/ lemma _root_.MeasureTheory.NullMeasurableSet.left_of_prod {s : Set α} {t : Set β} (h : NullMeasurableSet (s ×ˢ t) (μ.prod ν)) (ht : ν t ≠ 0) : NullMeasurableSet s μ := by From 639e997696fc203081e8b90a26d271a1f1fc75ce Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Mon, 12 Aug 2024 13:31:36 +0000 Subject: [PATCH 210/359] feat(CategoryTheory/Sites): gluing of sections for Mayer-Vietoris squares (#14957) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Markus Himmel Co-authored-by: Markus Himmel Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- .../Limits/Preserves/Shapes/Square.lean | 53 ++++++++++- .../Sites/MayerVietorisSquare.lean | 94 ++++++++++++++++++- 2 files changed, 142 insertions(+), 5 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Square.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Square.lean index f405241234d0a..fb95223d7a2e0 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Square.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Square.lean @@ -5,7 +5,7 @@ Authors: Joël Riou -/ import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Square import Mathlib.CategoryTheory.Limits.Yoneda -import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Pullbacks +import Mathlib.CategoryTheory.Limits.Preserves.Ulift /-! # Preservations of pullback/pushout squares @@ -39,10 +39,30 @@ lemma IsPullback.map (h : sq.IsPullback) (F : C ⥤ D) [PreservesLimit (cospan s (sq.map F).IsPullback := Square.IsPullback.mk _ (isLimitPullbackConeMapOfIsLimit F sq.fac h.isLimit) +lemma IsPullback.of_map (F : C ⥤ D) [ReflectsLimit (cospan sq.f₂₄ sq.f₃₄) F] + (h : (sq.map F).IsPullback) : sq.IsPullback := + CategoryTheory.IsPullback.of_map F sq.fac h + +variable (sq) in +lemma IsPullback.map_iff (F : C ⥤ D) [PreservesLimit (cospan sq.f₂₄ sq.f₃₄) F] + [ReflectsLimit (cospan sq.f₂₄ sq.f₃₄) F] : + (sq.map F).IsPullback ↔ sq.IsPullback := + ⟨fun h ↦ of_map F h, fun h ↦ h.map F⟩ + lemma IsPushout.map (h : sq.IsPushout) (F : C ⥤ D) [PreservesColimit (span sq.f₁₂ sq.f₁₃) F] : (sq.map F).IsPushout := Square.IsPushout.mk _ (isColimitPushoutCoconeMapOfIsColimit F sq.fac h.isColimit) +lemma IsPushout.of_map (F : C ⥤ D) [ReflectsColimit (span sq.f₁₂ sq.f₁₃) F] + (h : (sq.map F).IsPushout) : sq.IsPushout := + CategoryTheory.IsPushout.of_map F sq.fac h + +variable (sq) in +lemma IsPushout.map_iff (F : C ⥤ D) [PreservesColimit (span sq.f₁₂ sq.f₁₃) F] + [ReflectsColimit (span sq.f₁₂ sq.f₁₃) F] : + (sq.map F).IsPushout ↔ sq.IsPushout := + ⟨fun h ↦ of_map F h, fun h ↦ h.map F⟩ + variable (sq) lemma isPullback_iff_map_coyoneda_isPullback : @@ -56,6 +76,37 @@ lemma isPushout_iff_op_map_yoneda_isPullback : ((sq.pushoutCocone.isColimitYonedaEquiv).symm (fun X ↦ IsLimit.ofIsoLimit (h X).isLimit (PullbackCone.ext (Iso.refl _))))⟩ +section + +variable {sq₁ : Square (Type v)} {sq₂ : Square (Type u)} + (e₁ : sq₁.X₁ ≃ sq₂.X₁) (e₂ : sq₁.X₂ ≃ sq₂.X₂) + (e₃ : sq₁.X₃ ≃ sq₂.X₃) (e₄ : sq₁.X₄ ≃ sq₂.X₄) + (comm₁₂ : e₂ ∘ sq₁.f₁₂ = sq₂.f₁₂ ∘ e₁) + (comm₁₃ : e₃ ∘ sq₁.f₁₃ = sq₂.f₁₃ ∘ e₁) + (comm₂₄ : e₄ ∘ sq₁.f₂₄ = sq₂.f₂₄ ∘ e₂) + (comm₃₄ : e₄ ∘ sq₁.f₃₄ = sq₂.f₃₄ ∘ e₃) + +variable (sq₁ sq₂) in +lemma IsPullback.iff_of_equiv : sq₁.IsPullback ↔ sq₂.IsPullback := by + rw [← IsPullback.map_iff sq₁ uliftFunctor.{max u v}, + ← IsPullback.map_iff sq₂ uliftFunctor.{max u v}] + refine iff_of_iso (Square.isoMk + (((Equiv.trans Equiv.ulift e₁).trans Equiv.ulift.symm).toIso) + (((Equiv.trans Equiv.ulift e₂).trans Equiv.ulift.symm).toIso) + (((Equiv.trans Equiv.ulift e₃).trans Equiv.ulift.symm).toIso) + (((Equiv.trans Equiv.ulift e₄).trans Equiv.ulift.symm).toIso) + ?_ ?_ ?_ ?_) + all_goals ext; apply ULift.down_injective + · simpa [types_comp, uliftFunctor_map] using congrFun comm₁₂ _ + · simpa [types_comp, uliftFunctor_map] using congrFun comm₁₃ _ + · simpa [types_comp, uliftFunctor_map] using congrFun comm₂₄ _ + · simpa [types_comp, uliftFunctor_map] using congrFun comm₃₄ _ + +lemma IsPullback.of_equiv (h₁ : sq₁.IsPullback) : sq₂.IsPullback := + (iff_of_equiv sq₁ sq₂ e₁ e₂ e₃ e₄ comm₁₂ comm₁₃ comm₂₄ comm₃₄).1 h₁ + +end + end Square end CategoryTheory diff --git a/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean b/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean index f6b461edc0b5d..d4273af81064a 100644 --- a/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean +++ b/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Square +import Mathlib.CategoryTheory.Limits.Preserves.Shapes.Square import Mathlib.CategoryTheory.Limits.Shapes.Types import Mathlib.CategoryTheory.Sites.Sheafification @@ -35,19 +35,23 @@ squares in the case of the Nisnevich topology on schemes (in which case the closed (reduced) subscheme `X₄ - X₂`, and `X₁` shall be the pullback of `f₂₄` and `f₃₄`.). +Given a Mayer-Vietoris square `S` and a presheaf `P` on `C`, +we introduce a sheaf condition `S.SheafCondition P` and show +that it is indeed satisfied by sheaves. + ## TODO -* show that when "evaluating" a sheaf on a Mayer-Vietoris -square, one obtains a pullback square * provide constructors for `MayerVietorisSquare` ## References * https://stacks.math.columbia.edu/tag/08GL -/ -universe v u +universe v v' u u' namespace CategoryTheory +open Limits Opposite + namespace GrothendieckTopology variable {C : Type u} [Category.{v} C] @@ -62,6 +66,88 @@ structure MayerVietorisSquare extends Square C where /-- the square becomes a pushout square in the category of sheaves of types -/ isPushout : (toSquare.map (yoneda ⋙ presheafToSheaf J _)).IsPushout +namespace MayerVietorisSquare + +variable {J} +variable (S : J.MayerVietorisSquare) + +/-- The condition that a Mayer-Vietoris square becomes a pullback square +when we evaluate a presheaf on it. --/ +def SheafCondition {A : Type u'} [Category.{v'} A] (P : Cᵒᵖ ⥤ A) : Prop := + (S.toSquare.op.map P).IsPullback + +lemma sheafCondition_iff_comp_coyoneda {A : Type u'} [Category.{v'} A] (P : Cᵒᵖ ⥤ A) : + S.SheafCondition P ↔ ∀ (X : Aᵒᵖ), S.SheafCondition (P ⋙ coyoneda.obj X) := + Square.isPullback_iff_map_coyoneda_isPullback (S.op.map P) + +/-- Given a Mayer-Vietoris square `S` and a presheaf of types, this is the +map from `P.obj (op S.X₄)` to the explicit fibre product of +`P.map S.f₁₂.op` and `P.map S.f₁₃.op`. -/ +abbrev toPullbackObj (P : Cᵒᵖ ⥤ Type v') : + P.obj (op S.X₄) → Types.PullbackObj (P.map S.f₁₂.op) (P.map S.f₁₃.op) := + (S.toSquare.op.map P).pullbackCone.toPullbackObj + +lemma sheafCondition_iff_bijective_toPullbackObj (P : Cᵒᵖ ⥤ Type v') : + S.SheafCondition P ↔ Function.Bijective (S.toPullbackObj P) := by + have := (S.toSquare.op.map P).pullbackCone.isLimitEquivBijective + exact ⟨fun h ↦ this h.isLimit, fun h ↦ Square.IsPullback.mk _ (this.symm h)⟩ + +namespace SheafCondition + +variable {S} +variable {P : Cᵒᵖ ⥤ Type v'} (h : S.SheafCondition P) + +lemma bijective_toPullbackObj : Function.Bijective (S.toPullbackObj P) := by + rwa [← sheafCondition_iff_bijective_toPullbackObj] + +lemma ext {x y : P.obj (op S.X₄)} + (h₁ : P.map S.f₂₄.op x = P.map S.f₂₄.op y) + (h₂ : P.map S.f₃₄.op x = P.map S.f₃₄.op y) : x = y := + h.bijective_toPullbackObj.injective (by ext <;> assumption) + +variable (u : P.obj (op S.X₂)) (v : P.obj (op S.X₃)) + (huv : P.map S.f₁₂.op u = P.map S.f₁₃.op v) + +/-- If `S` is a Mayer-Vietoris square, and `P` is a presheaf +which satisfies the sheaf condition with respect to `S`, then +elements of `P` over `S.X₂` and `S.X₃` can be glued if the +coincide over `S.X₁`. -/ +noncomputable def glue : P.obj (op S.X₄) := + (PullbackCone.IsLimit.equivPullbackObj h.isLimit).symm ⟨⟨u, v⟩, huv⟩ + +@[simp] +lemma map_f₂₄_op_glue : P.map S.f₂₄.op (h.glue u v huv) = u := + PullbackCone.IsLimit.equivPullbackObj_symm_apply_fst h.isLimit _ + +@[simp] +lemma map_f₃₄_op_glue : P.map S.f₃₄.op (h.glue u v huv) = v := + PullbackCone.IsLimit.equivPullbackObj_symm_apply_snd h.isLimit _ + +end SheafCondition + +private lemma sheafCondition_of_sheaf' (F : Sheaf J (Type v)) : + S.SheafCondition F.val := by + refine (S.isPushout.op.map (yoneda.obj F)).of_equiv + (((sheafificationAdjunction J (Type v)).homEquiv _ _).trans yonedaEquiv) + (((sheafificationAdjunction J (Type v)).homEquiv _ _).trans yonedaEquiv) + (((sheafificationAdjunction J (Type v)).homEquiv _ _).trans yonedaEquiv) + (((sheafificationAdjunction J (Type v)).homEquiv _ _).trans yonedaEquiv) ?_ ?_ ?_ ?_ + all_goals + ext x + dsimp + rw [yonedaEquiv_naturality] + erw [Adjunction.homEquiv_naturality_left] + rfl + +lemma sheafCondition_of_sheaf {A : Type u'} [Category.{v} A] + (F : Sheaf J A) : S.SheafCondition F.val := by + rw [sheafCondition_iff_comp_coyoneda] + intro X + exact S.sheafCondition_of_sheaf' + ⟨_, (isSheaf_iff_isSheaf_of_type _ _).2 (F.cond X.unop)⟩ + +end MayerVietorisSquare + end GrothendieckTopology end CategoryTheory From e04c5f0c9b077f0ec777e3fb2c11395d9ff6839e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Mon, 12 Aug 2024 14:21:50 +0000 Subject: [PATCH 211/359] feat(Algebra/Homology): signs for the associativity compatibility of the total complex (#15719) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR introduces a typeclass `ComplexShape.Associative` which takes as inputs six complex shapes `c₁`, `c₂`, `c₃`, `c₁₂`, `c₂₃` and `c`, and asserts compatibilities of signs for the associativity of the total complex of a triple complex. It shall also be used in the construction of the associator isomorphism for the tensor product of homological complexes. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- .../Algebra/Homology/ComplexShapeSigns.lean | 87 ++++++++++++++++++- .../Algebra/Homology/TotalComplexShift.lean | 9 +- 2 files changed, 88 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Homology/ComplexShapeSigns.lean b/Mathlib/Algebra/Homology/ComplexShapeSigns.lean index 2d544038dc5c3..aa22deaab5726 100644 --- a/Mathlib/Algebra/Homology/ComplexShapeSigns.lean +++ b/Mathlib/Algebra/Homology/ComplexShapeSigns.lean @@ -5,6 +5,7 @@ Authors: Joël Riou -/ import Mathlib.Algebra.Homology.ComplexShape import Mathlib.Algebra.Ring.NegOnePow +import Mathlib.CategoryTheory.GradedObject.Trifunctor /-! Signs in constructions on homological complexes @@ -21,12 +22,11 @@ In particular, we construct an instance of `TotalComplexShape c c c` when `c : C and `I` is an additive monoid equipped with a group homomorphism `ε' : Multiplicative I → ℤˣ` satisfying certain properties (see `ComplexShape.TensorSigns`). -TODO @joelriou: add more classes for the associativity of the total complex, etc. - -/ -variable {I₁ I₂ I₁₂ : Type*} - (c₁ : ComplexShape I₁) (c₂ : ComplexShape I₂) (c₁₂ : ComplexShape I₁₂) +variable {I₁ I₂ I₃ I₁₂ I₂₃ J : Type*} + (c₁ : ComplexShape I₁) (c₂ : ComplexShape I₂) (c₃ : ComplexShape I₃) + (c₁₂ : ComplexShape I₁₂) (c₂₃ : ComplexShape I₂₃) (c : ComplexShape J) /-- A total complex shape for three complexes shapes `c₁`, `c₂`, `c₁₂` on three types `I₁`, `I₂` and `I₁₂` consists of the data and properties that will allow the construction @@ -183,6 +183,85 @@ lemma ε_up_ℤ (n : ℤ) : (ComplexShape.up ℤ).ε n = n.negOnePow := rfl end +section + +variable (c₁ c₂) +variable [TotalComplexShape c₁₂ c₃ c] [TotalComplexShape c₂ c₃ c₂₃] [TotalComplexShape c₁ c₂₃ c] + +/-- When we have six complex shapes `c₁`, `c₂`, `c₃`, `c₁₂`, `c₂₃`, `c`, and total functors +`HomologicalComplex₂ C c₁ c₂ ⥤ HomologicalComplex C c₁₂`, +`HomologicalComplex₂ C c₁₂ c₃ ⥤ HomologicalComplex C c`, +`HomologicalComplex₂ C c₂ c₃ ⥤ HomologicalComplex C c₂₃`, +`HomologicalComplex₂ C c₁ c₂₂₃ ⥤ HomologicalComplex C c`, we get two ways to +compute the total complex of a triple complex in `HomologicalComplex₃ C c₁ c₂ c₃`, then +under this assumption `[Associative c₁ c₂ c₃ c₁₂ c₂₃ c]`, these two complexes +canonically identify (without introducing signs). -/ +class Associative : Prop where + assoc (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) : + π c₁₂ c₃ c ⟨π c₁ c₂ c₁₂ ⟨i₁, i₂⟩, i₃⟩ = π c₁ c₂₃ c ⟨i₁, π c₂ c₃ c₂₃ ⟨i₂, i₃⟩⟩ + ε₁_eq_mul (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) : + ε₁ c₁ c₂₃ c (i₁, π c₂ c₃ c₂₃ (i₂, i₃)) = + ε₁ c₁₂ c₃ c (π c₁ c₂ c₁₂ (i₁, i₂), i₃) * ε₁ c₁ c₂ c₁₂ (i₁, i₂) + ε₂_ε₁ (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) : + ε₂ c₁ c₂₃ c (i₁, π c₂ c₃ c₂₃ (i₂, i₃)) * ε₁ c₂ c₃ c₂₃ (i₂, i₃) = + ε₁ c₁₂ c₃ c (π c₁ c₂ c₁₂ (i₁, i₂), i₃) * ε₂ c₁ c₂ c₁₂ (i₁, i₂) + ε₂_eq_mul (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) : + ε₂ c₁₂ c₃ c (π c₁ c₂ c₁₂ (i₁, i₂), i₃) = + (ε₂ c₁ c₂₃ c (i₁, π c₂ c₃ c₂₃ (i₂, i₃)) * ε₂ c₂ c₃ c₂₃ (i₂, i₃)) + +variable [Associative c₁ c₂ c₃ c₁₂ c₂₃ c] + +lemma assoc (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) : + π c₁₂ c₃ c ⟨π c₁ c₂ c₁₂ ⟨i₁, i₂⟩, i₃⟩ = π c₁ c₂₃ c ⟨i₁, π c₂ c₃ c₂₃ ⟨i₂, i₃⟩⟩ := by + apply Associative.assoc + +lemma associative_ε₁_eq_mul (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) : + ε₁ c₁ c₂₃ c (i₁, π c₂ c₃ c₂₃ (i₂, i₃)) = + ε₁ c₁₂ c₃ c (π c₁ c₂ c₁₂ (i₁, i₂), i₃) * ε₁ c₁ c₂ c₁₂ (i₁, i₂) := by + apply Associative.ε₁_eq_mul + +lemma associative_ε₂_ε₁ (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) : + ε₂ c₁ c₂₃ c (i₁, π c₂ c₃ c₂₃ (i₂, i₃)) * ε₁ c₂ c₃ c₂₃ (i₂, i₃) = + ε₁ c₁₂ c₃ c (π c₁ c₂ c₁₂ (i₁, i₂), i₃) * ε₂ c₁ c₂ c₁₂ (i₁, i₂) := by + apply Associative.ε₂_ε₁ + +lemma associative_ε₂_eq_mul (i₁ : I₁) (i₂ : I₂) (i₃ : I₃) : + ε₂ c₁₂ c₃ c (π c₁ c₂ c₁₂ (i₁, i₂), i₃) = + (ε₂ c₁ c₂₃ c (i₁, π c₂ c₃ c₂₃ (i₂, i₃)) * ε₂ c₂ c₃ c₂₃ (i₂, i₃)) := by + apply Associative.ε₂_eq_mul + +/-- The map `I₁ × I₂ × I₃ → j` that is obtained using `TotalComplexShape c₁ c₂ c₁₂` +and `TotalComplexShape c₁₂ c₃ c` when `c₁ : ComplexShape I₁`, `c₂ : ComplexShape I₂`, +`c₃ : ComplexShape I₃`, `c₁₂ : ComplexShape I₁₂` and `c : ComplexShape J`. -/ +def r : I₁ × I₂ × I₃ → J := fun ⟨i₁, i₂, i₃⟩ ↦ π c₁₂ c₃ c ⟨π c₁ c₂ c₁₂ ⟨i₁, i₂⟩, i₃⟩ + +open CategoryTheory + +/-- The `GradedObject.BifunctorComp₁₂IndexData` which arises from complex shapes. -/ +@[reducible] +def ρ₁₂ : GradedObject.BifunctorComp₁₂IndexData (r c₁ c₂ c₃ c₁₂ c) where + I₁₂ := I₁₂ + p := π c₁ c₂ c₁₂ + q := π c₁₂ c₃ c + hpq _ := rfl + +/-- The `GradedObject.BifunctorComp₂₃IndexData` which arises from complex shapes. -/ +@[reducible] +def ρ₂₃ : GradedObject.BifunctorComp₂₃IndexData (r c₁ c₂ c₃ c₁₂ c) where + I₂₃ := I₂₃ + p := π c₂ c₃ c₂₃ + q := π c₁ c₂₃ c + hpq := fun ⟨i₁, i₂, i₃⟩ ↦ (assoc c₁ c₂ c₃ c₁₂ c₂₃ c i₁ i₂ i₃).symm + +end + +instance {I : Type*} [AddMonoid I] (c : ComplexShape I) [c.TensorSigns] : + Associative c c c c c c where + assoc := add_assoc + ε₁_eq_mul _ _ _ := by dsimp; rw [one_mul] + ε₂_ε₁ _ _ _ := by dsimp; rw [one_mul, mul_one] + ε₂_eq_mul _ _ _ := by dsimp; rw [ε_add] + end ComplexShape /-- A total complex shape symmetry contains the data and properties which allow the diff --git a/Mathlib/Algebra/Homology/TotalComplexShift.lean b/Mathlib/Algebra/Homology/TotalComplexShift.lean index 756aa47f0785d..8d32945310df2 100644 --- a/Mathlib/Algebra/Homology/TotalComplexShift.lean +++ b/Mathlib/Algebra/Homology/TotalComplexShift.lean @@ -151,7 +151,7 @@ lemma D₁_totalShift₁XIso_hom (n₀ n₁ n₀' n₁' : ℤ) (h₀ : n₀ + x ((shiftFunctor₁ C x).obj K).d₁_eq _ rfl _ _ (by dsimp; omega), K.d₁_eq _ (show p + x + 1 = p + 1 + x by omega) _ _ (by dsimp; omega)] dsimp - rw [one_smul, assoc, ι_totalDesc, one_smul, Linear.units_smul_comp] + rw [one_smul, Category.assoc, ι_totalDesc, one_smul, Linear.units_smul_comp] · rw [D₁_shape _ _ _ _ h, zero_comp, D₁_shape, comp_zero, smul_zero] intro h' apply h @@ -170,7 +170,7 @@ lemma D₂_totalShift₁XIso_hom (n₀ n₁ n₀' n₁' : ℤ) (h₀ : n₀ + x rw [ι_D₂_assoc, Linear.comp_units_smul, ι_totalDesc_assoc, ι_D₂, ((shiftFunctor₁ C x).obj K).d₂_eq _ _ rfl _ (by dsimp; omega), K.d₂_eq _ _ rfl _ (by dsimp; omega), smul_smul, - Linear.units_smul_comp, assoc, ι_totalDesc] + Linear.units_smul_comp, Category.assoc, ι_totalDesc] dsimp congr 1 rw [add_comm p, Int.negOnePow_add, ← mul_assoc, Int.units_mul_self, one_mul] @@ -257,7 +257,8 @@ lemma D₁_totalShift₂XIso_hom (n₀ n₁ n₀' n₁' : ℤ) (h₀ : n₀ + y ι_D₁, smul_smul, ((shiftFunctor₂ C y).obj K).d₁_eq _ rfl _ _ (by dsimp; omega), K.d₁_eq _ rfl _ _ (by dsimp; omega)] dsimp - rw [one_smul, one_smul, assoc, ι_totalDesc, Linear.comp_units_smul, ← Int.negOnePow_add] + rw [one_smul, one_smul, Category.assoc, ι_totalDesc, Linear.comp_units_smul, + ← Int.negOnePow_add] congr 2 linarith · rw [D₁_shape _ _ _ _ h, zero_comp, D₁_shape, comp_zero, smul_zero] @@ -278,7 +279,7 @@ lemma D₂_totalShift₂XIso_hom (n₀ n₁ n₀' n₁' : ℤ) (h₀ : n₀ + y rw [ι_D₂_assoc, Linear.comp_units_smul, ι_totalDesc_assoc, Linear.units_smul_comp, smul_smul, ι_D₂, ((shiftFunctor₂ C y).obj K).d₂_eq _ _ rfl _ (by dsimp; omega), K.d₂_eq _ _ (show q + y + 1 = q + 1 + y by omega) _ (by dsimp; omega), - Linear.units_smul_comp, assoc, smul_smul, ι_totalDesc] + Linear.units_smul_comp, Category.assoc, smul_smul, ι_totalDesc] dsimp rw [Linear.units_smul_comp, Linear.comp_units_smul, smul_smul, smul_smul, ← Int.negOnePow_add, ← Int.negOnePow_add, ← Int.negOnePow_add, From 9bb9d9ecd2eaba634fbb68b418d0d74da0ce8d1a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Tom=C3=A1=C5=A1=20Sk=C5=99ivan?= Date: Mon, 12 Aug 2024 15:04:33 +0000 Subject: [PATCH 212/359] fix(fun_prop): few bug fixes for fun_prop (#11092) Bug fixes 1. Fix infinite loop triggered by `(by fun_prop : ?m)` 2. Support for multiple lambda theorems. For example `ContinuousOn` has multiple versions of composition theorem. 3. Support for special constant lambda theorem. For example `IsLinearMap R fun _ => 0` should be registered as constant lambda theorem 4. Improved error messages when trying `fun_prop` on non-fun_prop goal 5. Improved error messages when missing lambda theorems --- Mathlib.lean | 1 - .../Constructions/BorelSpace/Basic.lean | 1 + Mathlib/Tactic.lean | 1 - Mathlib/Tactic/FunProp.lean | 335 ++++++++++++++++++ Mathlib/Tactic/FunProp/AEMeasurable.lean | 15 - Mathlib/Tactic/FunProp/Attr.lean | 7 +- Mathlib/Tactic/FunProp/Core.lean | 265 ++++++++------ Mathlib/Tactic/FunProp/Decl.lean | 6 + Mathlib/Tactic/FunProp/Elab.lean | 74 ++-- Mathlib/Tactic/FunProp/FunctionData.lean | 15 +- Mathlib/Tactic/FunProp/Theorems.lean | 142 +++----- Mathlib/Tactic/FunProp/Types.lean | 105 +++--- test/fun_prop2.lean | 5 +- test/fun_prop_dev.lean | 98 ++++- 14 files changed, 768 insertions(+), 302 deletions(-) delete mode 100644 Mathlib/Tactic/FunProp/AEMeasurable.lean diff --git a/Mathlib.lean b/Mathlib.lean index 152422eef1ca1..54c7940c8bc63 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4078,7 +4078,6 @@ import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.FinCases import Mathlib.Tactic.Find import Mathlib.Tactic.FunProp -import Mathlib.Tactic.FunProp.AEMeasurable import Mathlib.Tactic.FunProp.Attr import Mathlib.Tactic.FunProp.ContDiff import Mathlib.Tactic.FunProp.Core diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 6aecc1704584b..74681e526e72b 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -467,6 +467,7 @@ theorem Continuous.measurable {f : α → γ} (hf : Continuous f) : Measurable f /-- A continuous function from an `OpensMeasurableSpace` to a `BorelSpace` is ae-measurable. -/ +@[fun_prop] theorem Continuous.aemeasurable {f : α → γ} (h : Continuous f) {μ : Measure α} : AEMeasurable f μ := h.measurable.aemeasurable diff --git a/Mathlib/Tactic.lean b/Mathlib/Tactic.lean index 8aa3d3c995b74..dd7ed5f122df3 100644 --- a/Mathlib/Tactic.lean +++ b/Mathlib/Tactic.lean @@ -69,7 +69,6 @@ import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.FinCases import Mathlib.Tactic.Find import Mathlib.Tactic.FunProp -import Mathlib.Tactic.FunProp.AEMeasurable import Mathlib.Tactic.FunProp.Attr import Mathlib.Tactic.FunProp.ContDiff import Mathlib.Tactic.FunProp.Core diff --git a/Mathlib/Tactic/FunProp.lean b/Mathlib/Tactic/FunProp.lean index 24b8c3de402c1..bc97ea7cc4def 100644 --- a/Mathlib/Tactic/FunProp.lean +++ b/Mathlib/Tactic/FunProp.lean @@ -1,3 +1,8 @@ +/- +Copyright (c) 2024 Tomas Skrivan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Tomas Skrivan +-/ import Mathlib.Tactic.FunProp.Attr import Mathlib.Tactic.FunProp.Core import Mathlib.Tactic.FunProp.Decl @@ -9,3 +14,333 @@ import Mathlib.Tactic.FunProp.StateList import Mathlib.Tactic.FunProp.Theorems import Mathlib.Tactic.FunProp.ToBatteries import Mathlib.Tactic.FunProp.Types + +/-! +# Tactic `fun_prop` for proving function properties like `Continuous f`, `Differentiable ℝ f`, ... + +**Basic use:** +Using the `fun_prop` tactic should be as simple as: +```lean +example : Continuous (fun x : ℝ => x * sin x) := by fun_prop +``` +Mathlib sets up `fun_prop` for many different properties like `Continuous`, `Measurable`, +`Differentiable`, `ContDiff`, etc., so `fun_prop` should work for such goals. The basic idea behind +`fun_prop` is that it decomposes the function into a composition of elementary functions and then +checks if every single elementary function is, e.g., `Continuous`. + +For `ContinuousAt/On/Within` variants, one has to specify a tactic to solve potential side goals +with `disch:=`. For example: +```lean +example (y : ℝ) (hy : y ≠ 0) : ContinuousAt (fun x : ℝ => 1/x) y := by fun_prop (disch:=assumption) +``` + +**Basic debugging:** +The most common issue is that a function is missing the appropriate theorem. For example: +```lean +import Mathlib.Data.Complex.Exponential +example : Continuous (fun x : ℝ => x * Real.sin x) := by fun_prop +``` +Fails with the error: +```lean +`fun_prop` was unable to prove `Continuous fun x => x * x.sin` + +Issues: + No theorems found for `Real.sin` in order to prove `Continuous fun x => x.sin` +``` +This can be easily fixed by importing `Mathlib.Analysis.SpecialFunctions.Trigonometric.Basic` + where the theorem `Real.continuous_sin` is marked with the `fun_prop` attribute. + +When the issue is not simply a few missing theorems, you can turn on the option: +```lean +set_option trace.Meta.Tactic.fun_prop true +``` +This will display the trace of how `fun_prop` steps through the whole expression. + +**Basic setup for a new function property:** + +To set up `fun_prop` for a new function property, you have to: + +1. Mark the function property with the `fun_prop` attribute when defining it: +```lean +@[fun_prop] +def Continuous (f : X → Y) := ... +``` +or later on with: +```lean +attribute [fun_prop] Continuous +``` + +2. Mark basic lambda calculus theorems. The bare minimum is the identity, constant, and composition +theorems: +```lean +@[fun_prop] +theorem continuous_id : Continuous (fun x => x) := ... + +@[fun_prop] +theorem continuous_const (y : Y) : Continuous (fun x => y) := ... + +@[fun_prop] +theorem continuous_comp (f : Y → Z) (g : X → Y) (hf : Continuous f) (hg : Continuous g) : + Continuous (fun x => f (g x)) := ... +``` +The constant theorem is not absolutely necessary as, for example, `IsLinearMap ℝ (fun x => y)` does + not hold, but we almost certainly want to mark it if it is available. + +You should also provide theorems for `Prod.mk`, `Prod.fst`, and `Prod.snd`: +```lean +@[fun_prop] +theorem continuous_fst (f : X → Y × Z) (hf : Continuous f) : Continuous (fun x => (f x).fst) := ... +@[fun_prop] +theorem continuous_snd (f : X → Y × Z) (hf : Continuous f) : Continuous (fun x => (f x).snd) := ... +@[fun_prop] +theorem continuous_prod_mk (f : X → Y) (g : X → Z) (hf : Continuous f) (hg : Continuous g) : + Continuous (fun x => Prod.mk (f x) (g x)) := ... +``` + +3. Mark function theorems. They can be stated simply as: +```lean +@[fun_prop] +theorem continuous_neg : Continuous (fun x => - x) := ... + +@[fun_prop] +theorem continuous_add : Continuous (fun x : X × X => x.1 + x.2) := ... +``` +where functions of multiple arguments have to be appropriately uncurried. Alternatively, they can +be stated in compositional form as: +```lean +@[fun_prop] +theorem continuous_neg (f : X → Y) (hf : Continuous f) : Continuous (fun x => - f x) := ... + +@[fun_prop] +theorem continuous_add (f g : X → Y) (hf : Continuous f) (hg : Continuous g) : + Continuous (fun x => f x + g x) := ... +``` +It is enough to provide function theorems in either form. It is mainly a matter of convenience. + +With such a basic setup, you should be able to prove the continuity of basic algebraic expressions. + +When marking theorems, it is a good idea to check that a theorem has been registered correctly. +You can do this by turning on the `Meta.Tactic.fun_prop.attr` option. For example: +(note that the notation `f x + g x` is just syntactic sugar for `@HAdd.hAdd X Y Y _ (f x) (g x)`) +```lean +set_option trace.Meta.Tactic.fun_prop.attr true +@[fun_prop] +theorem continuous_add (f g : X → Y) (hf : Continuous f) (hg : Continuous g) : + Continuous (fun x => @HAdd.hAdd X Y Y _ (f x) (g x)) := ... +``` +displays: +```lean +[Meta.Tactic.fun_prop.attr] function theorem: continuous_add + function property: Continuous + function name: HAdd.hAdd + main arguments: [4, 5] + applied arguments: 6 + form: compositional form +``` +This indicates that the theorem `continuous_add` states the continuity of `HAdd.hAdd` in the 4th and + 5th arguments and the theorem is in compositional form. + +### Advanced + +### Type of Theorems + +There are four types of theorems that are used a bit differently. + +- Lambda Theorems: + These theorems are about basic lambda calculus terms like identity, constant, composition, etc. + They are used when a bigger expression is decomposed into a sequence of function compositions. + They are also used when, for example, we know that a function is continuous in the `x` and `y` + variables, but we want continuity only in `x`. + + There are five of them, and they should be formulated in the following way: + + - Identity Theorem + ```lean + @[fun_prop] + theorem continuous_id : Continuous (fun (x : X) => x) := .. + ``` + + - Constant Theorem + ```lean + @[fun_prop] + theorem continuous_const (y : Y) : Continuous (fun (x : X) => y) := .. + ``` + + - Composition Theorem + ```lean + @[fun_prop] + theorem continuous_comp (f : Y → Z) (g : X → Y) (hf : Continuous f) (hg : Continuous g) : + Continuous (fun (x : X) => f (g x)) := .. + ``` + + - Apply Theorem + It can be either non-dependent version + ```lean + @[fun_prop] + theorem continuous_apply (a : α) : Continuous (fun f : (α → X) => f a) := .. + ``` + or dependent version + ```lean + @[fun_prop] + theorem continuous_apply (a : α) : Continuous (fun f : ((a' : α) → E a') => f a) := .. + ``` + + - Pi Theorem + ```lean + @[fun_prop] + theorem continuous_pi (f : X → α → Y) (hf : ∀ a, Continuous (f x a)) : + Continuous (fun x a => f x a) := .. + ``` + + Not all of these theorems have to be provided, but at least the identity and composition + theorems should be. + + You should also add theorems for `Prod.mk`, `Prod.fst`, and `Prod.snd`. Technically speaking, + they fall under the *function theorems* category, but without them, `fun_prop` can't function + properly. We are mentioning them as they are used together with *lambda theorems* to break + complicated expressions into a composition of simpler functions. + ```lean + @[fun_prop] + theorem continuous_fst (f : X → Y × Z) (hf : Continuous f) : + Continuous (fun x => (f x).fst) := ... + @[fun_prop] + theorem continuous_snd (f : X → Y × Z) (hf : Continuous f) : + Continuous (fun x => (f x).snd) := ... + @[fun_prop] + theorem continuous_prod_mk (f : X → Y) (g : X → Z) (hf : Continuous f) (hg : Continuous g) : + Continuous (fun x => (f x, g x)) := ... + ``` + +- Function Theorems: + When `fun_prop` breaks complicated expression apart using *lambda theorems* it then uses + *function theorems* to prove that each piece is, for example, continuous. + + The function theorem for `Neg.neg` and `Continuous` can be stated as: + ```lean + @[fun_prop] + theorem continuous_neg : Continuous (fun x => - x) := ... + ``` + or as: + ```lean + @[fun_prop] + theorem continuous_neg (f : X → Y) (hf : Continuous f) : Continuous (fun x => - f x) := ... + ``` + The first form is called *uncurried form* and the second form is called *compositional form*. + You can provide either form; it is mainly a matter of convenience. You can check if the form of + a theorem has been correctly detected by turning on the option: + ```lean + set_option Meta.Tactic.fun_prop.attr true + ``` + If you really care that the resulting proof term is as short as possible, it is a good idea to + provide both versions. + + One exception to this rule is the theorem for `Prod.mk`, which has to be stated in compositional + form. This because this theorem together with *lambda theorems* is used to break expression to + smaller pieces and `fun_prop` assumes it is written in compositional form. + + The reason the first form is called *uncurried* is because if we have a function of multiple + arguments, we have to uncurry the function: + ```lean + @[fun_prop] + theorem continuous_add : Continuous (fun (x : X × X) => x.1 + x.2) := ... + ``` + and the *compositional form* of this theorem is: + ```lean + @[fun_prop] + theorem continuous_add (f g : X → Y) (hf : Continuous f) (hg : Continuous g) : + Continuous (fun x => f x + g x) := ... + ``` + + When dealing with functions with multiple arguments, you need to state, e.g., continuity only + in the maximal set of arguments. Once we state that addition is jointly continuous in both + arguments, we do not need to add new theorems that addition is also continuous only in the first + or only in the second argument. This is automatically inferred using lambda theorems. + + +- Morphism Theorems: + The `fun_prop` tactic can also deal with bundled morphisms. For example, we can state that every + continuous linear function is indeed continuous: + ```lean + @[fun_prop] + theorem continuous_clm_eval (f : X →L[𝕜] Y) : Continuous 𝕜 (fun x => f x) := ... + ``` + In this case, the head of the function body `f x` is `DFunLike.coe`. This function is + treated differently and its theorems are tracked separately. + + `DFunLike.coe` has two arguments: the morphism `f` and the argument `x`. One difference is that + theorems talking about the argument `f` have to be stated in the compositional form: + ```lean + @[fun_prop] + theorem continuous_clm_apply (f : X → Y →L[𝕜] Z) (hf : Continuous f) (y : Y) : + Continuous 𝕜 (fun x => f x y) := ... + ``` + Note that without notation and coercion, the function looks like + `fun x => DFunLike.coe (f x) y`. + + In fact, not only `DFunLike.coe` but any function coercion is treated this way. Such function + coercion has to be registered with `Lean.Meta.registerCoercion` with coercion type `.coeFun`. + Here is an example of custom structure `MyFunLike` that that should be considered as bundled + morphism by `fun_prop`: + ```lean + structure MyFunLike (α β : Type) where + toFun : α → β + + instance {α β} : CoeFun (MyFunLike α β) (fun _ => α → β) := ⟨MyFunLike.toFun⟩ + + #eval Lean.Elab.Command.liftTermElabM do + Lean.Meta.registerCoercion ``MyFunLike.toFun + (.some { numArgs := 3, coercee := 2, type := .coeFun }) + ``` + +- Transition Theorems: + Transition theorems allow `fun_prop` to infer one function property from another. + For example, a theorem like: + ```lean + @[fun_prop] + theorem differentiable_continuous (f : X → Y) (hf : Differentiable 𝕜 f) : + Continuous f := ... + ``` + There are two features of these theorems: they mix different function properties and the + conclusion is about a free variable `f`. + + Transition theorems are the most dangerous theorems as they considerably increase the search + space since they do not simplify the function in question. For this reason, `fun_prop` only + applies transition theorems to functions that can't be written as a non-trivial composition of + two functions (`f = f ∘ id`, `f = id ∘ f` is considered to be a trivial composition). + + For this reason, it is recommended to state *function theorems* for every property. For example, + if you have a theorem: + ```lean + @[fun_prop] + theorem differentiable_neg : Differentiable ℝ (fun x => -x) := ... + ``` + you should also state the continuous theorem: + ```lean + @[fun_prop] + theorem continuous_neg : Continuous ℝ (fun x => -x) := ... + ``` + even though `fun_prop` can already prove `continuous_neg` from `differentiable_continuous` and + `differentiable_neg`. Doing this will have a considerable impact on `fun_prop` speed. + + By default, `fun_prop` will not apply more then one transitions theorems consecutivelly. For + example, it won't prove `AEMeasurable f` from `Continuous f` by using transition theorems + `Measurable.aemeasurable` and `Continuous.measurable`. You can enable this by running + `fun_prop (config:={maxTransitionDepth:=2})`. + Ideally `fun_prop` theorems should be transitivelly closed i.e. if `Measurable.aemeasurable` and + `Continuous.measurable` are `fun_prop` theorems then `Continuous.aemeasurable` should be too. + + Transition theorems do not have to be between two completely different properties. They can be + between the same property differing by a parameter. Consider this example: + ```lean + example (f : X → Y) (hf : ContDiff ℝ ∞ f) : ContDiff ℝ 2 (fun x => f x + f x) := by + fun_prop (disch := aesop) + ``` + which is first reduced to `ContDiff ℝ 2 f` using lambda theorems and then the transition + theorem: + ```lean + @[fun_prop] + theorem contdiff_le (f : ContDiff 𝕜 n f) (h : m ≤ n) : ContDiff 𝕜 m f := ... + ``` + is used together with `aesop` to discharge the `2 ≤ ∞` subgoal. + +-/ diff --git a/Mathlib/Tactic/FunProp/AEMeasurable.lean b/Mathlib/Tactic/FunProp/AEMeasurable.lean deleted file mode 100644 index fb72bae847eb8..0000000000000 --- a/Mathlib/Tactic/FunProp/AEMeasurable.lean +++ /dev/null @@ -1,15 +0,0 @@ -/- -Copyright (c) 2024 Tomáš Skřivan. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Tomáš Skřivan --/ -import Mathlib.MeasureTheory.Measure.AEMeasurable - -/-! -## `fun_prop` minimal setup for AEMeasurable --/ - -open Mathlib - --- lambda rules: these two are currently missing --- attribute [fun_prop] AEMeasurable_apply AEMeasurable_pi diff --git a/Mathlib/Tactic/FunProp/Attr.lean b/Mathlib/Tactic/FunProp/Attr.lean index e3fa55a39e7af..314b0c2f80965 100644 --- a/Mathlib/Tactic/FunProp/Attr.lean +++ b/Mathlib/Tactic/FunProp/Attr.lean @@ -12,18 +12,13 @@ import Mathlib.Tactic.FunProp.Theorems ## `funProp` attribute -/ - namespace Mathlib open Lean Meta namespace Meta.FunProp --- TODO: add support for specifying priority and discharger --- open Lean.Parser.Tactic --- syntax (name:=Attr.funProp) "funProp" (prio)? (discharger)? : attr - private def funPropHelpString : String := -"`funProp` tactic to prove function properties like `Continuous`, `Differentiable`, `IsLinearMap`" +"`fun_prop` tactic to prove function properties like `Continuous`, `Differentiable`, `IsLinearMap`" /-- Initialization of `funProp` attribute -/ initialize funPropAttr : Unit ← diff --git a/Mathlib/Tactic/FunProp/Core.lean b/Mathlib/Tactic/FunProp/Core.lean index 74e9c8a628be3..1459b86032271 100644 --- a/Mathlib/Tactic/FunProp/Core.lean +++ b/Mathlib/Tactic/FunProp/Core.lean @@ -11,7 +11,7 @@ import Mathlib.Lean.Expr.Basic import Batteries.Tactic.Exact /-! -## `funProp` core tactic algorithm +# Tactic `fun_prop` for proving function properties like `Continuous f`, `Differentiable ℝ f`, ... -/ namespace Mathlib @@ -39,7 +39,7 @@ sythesized value{indentExpr val}\nis not definitionally equal to{indentExpr x}" return false -/-- Synthesize arguments `xs` either with typeclass synthesis, with funProp or with discharger. -/ +/-- Synthesize arguments `xs` either with typeclass synthesis, with `fun_prop` or with discharger. -/ def synthesizeArgs (thmId : Origin) (xs : Array Expr) (bis : Array BinderInfo) (funProp : Expr → FunPropM (Option Result)) : FunPropM Bool := do @@ -48,8 +48,8 @@ def synthesizeArgs (thmId : Origin) (xs : Array Expr) (bis : Array BinderInfo) let type ← inferType x if bi.isInstImplicit then unless (← synthesizeInstance thmId x type) do - logError s!"Failed to synthesize instance {← ppExpr type} \ - when applying theorem {← ppOrigin' thmId}." + logError s!"Failed to synthesize instance `{← ppExpr type}` \ + when applying theorem `{← ppOrigin' thmId}`." return false else if (← instantiateMVars x).isMVar then @@ -69,9 +69,9 @@ def synthesizeArgs (thmId : Origin) (xs : Array Expr) (bis : Array BinderInfo) return false else -- try user provided discharger - let cfg : Config ← read + let ctx : Context ← read if (← isProp type) then - if let .some proof ← cfg.disch type then + if let .some proof ← ctx.disch type then if (← isDefEq x proof) then continue else do @@ -79,8 +79,8 @@ def synthesizeArgs (thmId : Origin) (xs : Array Expr) (bis : Array BinderInfo) "{← ppOrigin thmId}, failed to assign proof{indentExpr type}" return false else - logError s!"Failed to prove necessary assumption {← ppExpr type} \ - when applying theorem {← ppOrigin' thmId}." + logError s!"Failed to prove necessary assumption `{← ppExpr type}` \ + when applying theorem `{← ppOrigin' thmId}`." if ¬(← isProp type) then postponed := postponed.push x @@ -93,7 +93,7 @@ def synthesizeArgs (thmId : Origin) (xs : Array Expr) (bis : Array BinderInfo) for x in postponed do if (← instantiateMVars x).isMVar then logError s!"Failed to infer `({← ppExpr x} : {← ppExpr (← inferType x)})` \ - when applying theorem {← ppOrigin' thmId}." + when applying theorem `{← ppOrigin' thmId}`." trace[Meta.Tactic.fun_prop] "{← ppOrigin thmId}, failed to infer `({← ppExpr x} : {← ppExpr (← inferType x)})`" @@ -108,16 +108,12 @@ def tryTheoremCore (xs : Array Expr) (bis : Array BinderInfo) (val : Expr) (type withTraceNode `Meta.Tactic.fun_prop (fun r => return s!"[{ExceptToEmoji.toEmoji r}] applying: {← ppOrigin' thmId}") do - -- add theorem to the stack - withTheorem thmId do - if (← isDefEq type e) then if ¬(← synthesizeArgs thmId xs bis funProp) then return none let proof ← instantiateMVars (mkAppN val xs) - trace[Meta.Tactic.fun_prop.apply] "{← ppOrigin thmId}, \n{e}" return .some { proof := proof } else trace[Meta.Tactic.fun_prop] "failed to unify {← ppOrigin thmId}\n{type}\nwith\n{e}" @@ -144,95 +140,142 @@ def tryTheoremWithHint? (e : Expr) (thmOrigin : Origin) tryTheoremCore xs bis thmProof type e thmOrigin funProp + -- `simp` introduces new meta variable context depth for some reason + -- This is probably to avoid mvar assignment when trying a theorem fails + -- + -- However, in `fun_prop` case this is not completely desirable + -- For example, I want to be able to solve a goal with mvars like `ContDiff ℝ ?n f` using local + -- hypothesis `(h : ContDiff ℝ ∞ f)` and assign `∞` to the mvar `?n`. + -- + -- This could be problematic if there are two local hypothesis `(hinf : ContDiff ℝ ∞ f)` and + -- `(h1 : ContDiff ℝ 1 f)` and appart from solving `ContDiff ℝ ?n f` there is also a subgoal + -- `2 ≤ ?n`. If `fun_prop` decides to try `h1` first it would assign `1` to `?n` and then there + -- is no hope solving `2 ≤ 1` and it won't be able to apply `hinf` after trying `h1` as `n?` is + -- assigned already. Ideally `fun_prop` would roll back the `MetaM.State`. This issue did not + -- come up yet so I didn't bother and I'm worried about the performance impact. if newMCtxDepth then withNewMCtxDepth go else go -/-- Try to apply a theorem -/ +/-- Try to apply a theorem `thmOrigin` to the goal `e`. -/ def tryTheorem? (e : Expr) (thmOrigin : Origin) (funProp : Expr → FunPropM (Option Result)) (newMCtxDepth : Bool := false) : FunPropM (Option Result) := tryTheoremWithHint? e thmOrigin #[] funProp newMCtxDepth -/-- Apply lambda calculus rule P fun x => x` -/ -def applyIdRule (funPropDecl : FunPropDecl) (e X : Expr) +/-- +Try to prove `e` using using *identity lambda theorem*. + +For example, `e = q(Continuous fun x => x)` and `funPropDecl` is `FunPropDecl` for `Continuous`. +-/ +def applyIdRule (funPropDecl : FunPropDecl) (e : Expr) (funProp : Expr → FunPropM (Option Result)) : FunPropM (Option Result) := do - let ext := lambdaTheoremsExt.getState (← getEnv) - let .some thm := ext.theorems.find? (funPropDecl.funPropName, .id) - | trace[Meta.Tactic.fun_prop] - "missing identity rule to prove `{← ppExpr e}`" - return none - let .id id_X := thm.thmArgs | return none + let thms ← getLambdaTheorems funPropDecl.funPropName .id + if thms.size = 0 then + let msg := s!"missing identity rule to prove `{← ppExpr e}`" + logError msg + trace[Meta.Tactic.fun_prop] msg + return none + + for thm in thms do + if let .some r ← tryTheoremWithHint? e (.decl thm.thmName) #[] funProp then + return r + + return none - tryTheoremWithHint? e (.decl thm.thmName) #[(id_X,X)] funProp +/-- +Try to prove `e` using using *constant lambda theorem*. -/-- Apply lambda calculus rule P fun x => y` -/ -def applyConstRule (funPropDecl : FunPropDecl) (e X y : Expr) +For example, `e = q(Continuous fun x => y)` and `funPropDecl` is `FunPropDecl` for `Continuous`. +-/ +def applyConstRule (funPropDecl : FunPropDecl) (e : Expr) (funProp : Expr → FunPropM (Option Result)) : FunPropM (Option Result) := do + let thms ← getLambdaTheorems funPropDecl.funPropName .const + if thms.size = 0 then + let msg := s!"missing constant rule to prove `{← ppExpr e}`" + logError msg + trace[Meta.Tactic.fun_prop] msg + return none - let ext := lambdaTheoremsExt.getState (← getEnv) - let .some thm := ext.theorems.find? (funPropDecl.funPropName, .const) - | trace[Meta.Tactic.fun_prop] - "missing constant rule to prove `{← ppExpr e}`" - return none - let .const id_X id_y := thm.thmArgs | return none + for thm in thms do + let .const := thm.thmArgs | return none + if let .some r ← tryTheorem? e (.decl thm.thmName) funProp then + return r - tryTheoremWithHint? e (.decl thm.thmName) #[(id_X,X),(id_y,y)] funProp + return none + +/-- +Try to prove `e` using using *apply lambda theorem*. -/-- Apply lambda calculus rule P fun f => f i` -/ -def applyProjRule (funPropDecl : FunPropDecl) (e x XY : Expr) +For example, `e = q(Continuous fun f => f x)` and `funPropDecl` is `FunPropDecl` for `Continuous`. +-/ +def applyApplyRule (funPropDecl : FunPropDecl) (e : Expr) (funProp : Expr → FunPropM (Option Result)) : FunPropM (Option Result) := do - let ext := lambdaTheoremsExt.getState (← getEnv) - let .forallE n X Y _ := XY | return none - - -- non dependent case - if ¬(Y.hasLooseBVars) then - if let .some thm := ext.theorems.find? (funPropDecl.funPropName, .proj) then - let .proj id_x id_Y := thm.thmArgs | return none - return ← tryTheoremWithHint? e (.decl thm.thmName) #[(id_x,x),(id_Y,Y)] funProp - - -- dependent case - -- can also handle non-dependent cases if non-dependent theorem is not available - let Y := Expr.lam n X Y default - - let .some thm := ext.theorems.find? (funPropDecl.funPropName, .projDep) - | trace[Meta.Tactic.fun_prop] - "missing projection rule to prove `{← ppExpr e}`" - return none - let .projDep id_x id_Y := thm.thmArgs | return none + let thms := (← getLambdaTheorems funPropDecl.funPropName .apply) + + for thm in thms do + if let .some r ← tryTheoremWithHint? e (.decl thm.thmName) #[] funProp then + return r + + return none - tryTheoremWithHint? e (.decl thm.thmName) #[(id_x,x),(id_Y,Y)] funProp +/-- +Try to prove `e` using *composition lambda theorem*. -/-- Apply lambda calculus rule `P f → P g → P fun x => f (g x)` -/ +For example, `e = q(Continuous fun x => f (g x))` and `funPropDecl` is `FunPropDecl` for +`Continuous` + +You also have to provide the functions `f` and `g`. -/ def applyCompRule (funPropDecl : FunPropDecl) (e f g : Expr) (funProp : Expr → FunPropM (Option Result)) : FunPropM (Option Result) := do - let ext := lambdaTheoremsExt.getState (← getEnv) - let .some thm := ext.theorems.find? (funPropDecl.funPropName, .comp) - | trace[Meta.Tactic.fun_prop] - "missing composition rule to prove `{← ppExpr e}`" - return none - let .comp id_f id_g := thm.thmArgs | return none + let thms ← getLambdaTheorems funPropDecl.funPropName .comp + if thms.size = 0 then + let msg := s!"missing composition rule to prove `{← ppExpr e}`" + logError msg + trace[Meta.Tactic.fun_prop] msg + return none + + for thm in thms do + let .comp id_f id_g := thm.thmArgs | return none + if let .some r ← tryTheoremWithHint? e (.decl thm.thmName) #[(id_f,f),(id_g,g)] funProp then + return r - tryTheoremWithHint? e (.decl thm.thmName) #[(id_f,f),(id_g,g)] funProp + return none + +/-- +Try to prove `e` using *pi lambda theorem*. -/-- Apply lambda calculus rule `∀ y, P (f · y) → P fun x y => f x y` -/ -def applyPiRule (funPropDecl : FunPropDecl) (e f : Expr) +For example, `e = q(Continuous fun x y => f x y)` and `funPropDecl` is `FunPropDecl` for +`Continuous` +-/ +def applyPiRule (funPropDecl : FunPropDecl) (e : Expr) (funProp : Expr → FunPropM (Option Result)) : FunPropM (Option Result) := do - let ext := lambdaTheoremsExt.getState (← getEnv) - let .some thm := ext.theorems.find? (funPropDecl.funPropName, .pi) - | trace[Meta.Tactic.fun_prop] - "missing pi rule to prove `{← ppExpr e}`" - return none - let .pi id_f := thm.thmArgs | return none + let thms ← getLambdaTheorems funPropDecl.funPropName .pi + if thms.size = 0 then + let msg := s!"missing pi rule to prove `{← ppExpr e}`" + logError msg + trace[Meta.Tactic.fun_prop] msg + return none - tryTheoremWithHint? e (.decl thm.thmName) #[(id_f,f)] funProp + for thm in thms do + if let .some r ← tryTheoremWithHint? e (.decl thm.thmName) #[] funProp then + return r + + return none -/-- Prove function property of `fun x => let y := g x; f x y`. -/ +/-- +Try to prove `e = q(P (fun x => let y := φ x; ψ x y)`. + +For example, + - `funPropDecl` is `FunPropDecl` for `Continuous` + - `e = q(Continuous fun x => let y := φ x; ψ x y)` + - `f = q(fun x => let y := φ x; ψ x y)` +-/ def letCase (funPropDecl : FunPropDecl) (e : Expr) (f : Expr) (funProp : Expr → FunPropM (Option Result)) : FunPropM (Option Result) := do @@ -274,15 +317,15 @@ def letCase (funPropDecl : FunPropDecl) (e : Expr) (f : Expr) | _ => throwError "expected expression of the form `fun x => lam y := ..; ..`" -/-- Prove function property of using "morphism theorems" e.g. bundled linear map is linear map. -/ +/-- Prove function property of using *morphism theorems*. -/ def applyMorRules (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) (funProp : Expr → FunPropM (Option Result)) : FunPropM (Option Result) := do - trace[Meta.Tactic.fun_prop.step] "applying morphism theoresm to {← ppExpr e}" + trace[Debug.Meta.Tactic.fun_prop] "applying morphism theorems to {← ppExpr e}" match ← fData.isMorApplication with | .none => throwError "fun_prop bug: ivalid use of mor rules on {← ppExpr e}" | .underApplied => - applyPiRule funPropDecl e (← fData.toExpr) funProp + applyPiRule funPropDecl e funProp | .overApplied => let .some (f,g) ← fData.peeloffArgDecomposition | return none applyCompRule funPropDecl e f g funProp @@ -299,12 +342,13 @@ def applyMorRules (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) if let .some r ← tryTheorem? e (.decl c.thmName) funProp then return r - trace[Meta.Tactic.fun_prop.step] "no theorem matched" + trace[Debug.Meta.Tactic.fun_prop] "no theorem matched" return none -/-- Prove function property of using "transition theorems" e.g. continuity from linearity. -/ +/-- Prove function property of using *transition theorems*. -/ def applyTransitionRules (e : Expr) (funProp : Expr → FunPropM (Option Result)) : FunPropM (Option Result) := do + withIncreasedTransitionDepth do let ext := transitionTheoremsExt.getState (← getEnv) let candidates ← ext.theorems.getMatchWithScore e false { iota := false, zeta := false } @@ -314,16 +358,20 @@ def applyTransitionRules (e : Expr) (funProp : Expr → FunPropM (Option Result) "candidate transition theorems: {← candidates.mapM fun c => ppOrigin (.decl c.thmName)}" for c in candidates do - if ← previouslyUsedThm (.decl c.thmName) then - trace[Meta.Tactic.fun_prop] "skipping {c.thmName} to prevent potential infinite loop" - else - if let .some r ← tryTheorem? e (.decl c.thmName) funProp then - return r + if let .some r ← tryTheorem? e (.decl c.thmName) funProp then + return r - trace[Meta.Tactic.fun_prop.step] "no theorem matched" + trace[Debug.Meta.Tactic.fun_prop] "no theorem matched" return none -/-- Try to remove applied argument. -/ +/-- Try to remove applied argument i.e. prove `P (fun x => f x y)` from `P (fun x => f x)`. + +For example + - `funPropDecl` is `FunPropDecl` for `Continuous` + - `e = q(Continuous fun x => foo (bar x) y)` + - `fData` contains info on `fun x => foo (bar x) y` + This tries to prove `Continuous fun x => foo (bar x) y` from `Continuous fun x => foo (bar x)` + -/ def removeArgRule (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) (funProp : Expr → FunPropM (Option Result)) : FunPropM (Option Result) := do @@ -352,9 +400,10 @@ def bvarAppCase (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) if let .some (f, g) ← fData.nontrivialDecomposition then applyCompRule funPropDecl e f g funProp else - applyProjRule funPropDecl e fData.args[0]!.expr (← fData.domainType) funProp + applyApplyRule funPropDecl e funProp -/-- Get candidate theorems from the environment for function property `funPropDecl` and +/-- +Get candidate theorems from the environment for function property `funPropDecl` and function `funName`. -/ def getDeclTheorems (funPropDecl : FunPropDecl) (funName : Name) (mainArgs : Array Nat) (appliedArgs : Nat) : MetaM (Array FunctionTheorem) := do @@ -373,7 +422,8 @@ def getDeclTheorems (funPropDecl : FunPropDecl) (funName : Name) -- todo: sorting and filtering return thms -/-- Get candidate theorems from the local context for function property `funPropDecl` and +/-- +Get candidate theorems from the local context for function property `funPropDecl` and function `funName`. -/ def getLocalTheorems (funPropDecl : FunPropDecl) (funOrigin : Origin) (mainArgs : Array Nat) (appliedArgs : Nat) : FunPropM (Array FunctionTheorem) := do @@ -424,7 +474,7 @@ def getLocalTheorems (funPropDecl : FunPropDecl) (funOrigin : Origin) return thms -/-- Try to apply theorems `thms` to `e` -/ +/-- Try to apply *function theorems* `thms` to `e`. -/ def tryTheorems (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) (thms : Array FunctionTheorem) (funProp : Expr → FunPropM (Option Result)) : FunPropM (Option Result) := do @@ -436,7 +486,7 @@ def tryTheorems (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) for thm in thms do - trace[Meta.Tactic.fun_prop.step] s!"trying theorem {← ppOrigin' thm.thmOrigin}" + trace[Debug.Meta.Tactic.fun_prop] s!"trying theorem {← ppOrigin' thm.thmOrigin}" match compare thm.appliedArgs fData.args.size with | .lt => @@ -446,7 +496,7 @@ def tryTheorems (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) continue | .gt => trace[Meta.Tactic.fun_prop] s!"adding argument to later use {← ppOrigin' thm.thmOrigin}" - if let .some r ← applyPiRule funPropDecl e (← fData.toExpr) funProp then + if let .some r ← applyPiRule funPropDecl e funProp then return r continue | .eq => @@ -463,19 +513,19 @@ def tryTheorems (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) if let .some r ← tryTheorem? e thm.thmOrigin funProp then return r | .some (.some (f,g)) => - trace[Meta.Tactic.fun_prop.step] + trace[Debug.Meta.Tactic.fun_prop] s!"decomposing to later use {←ppOrigin' thm.thmOrigin}" - trace[Meta.Tactic.fun_prop.step] + trace[Debug.Meta.Tactic.fun_prop] s!"decomposition: {← ppExpr f} ∘ {← ppExpr g}" if let .some r ← applyCompRule funPropDecl e f g funProp then return r | _ => continue else - trace[Meta.Tactic.fun_prop.step] + trace[Debug.Meta.Tactic.fun_prop] s!"decomposing in args {thm.mainArgs} to later use {←ppOrigin' thm.thmOrigin}" let .some (f,g) ← fData.decompositionOverArgs thm.mainArgs | continue - trace[Meta.Tactic.fun_prop.step] + trace[Debug.Meta.Tactic.fun_prop] s!"decomposition: {← ppExpr f} ∘ {← ppExpr g}" if let .some r ← applyCompRule funPropDecl e f g funProp then return r @@ -513,7 +563,7 @@ def fvarAppCase (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) return r if thms.size = 0 then - logError s!"No theorems found for `{← ppExpr (.fvar id)}` in order to prove {← ppExpr e}" + logError s!"No theorems found for `{← ppExpr (.fvar id)}` in order to prove `{← ppExpr e}`" return none @@ -541,6 +591,10 @@ def constAppCase (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) if let .some r ← tryTheorems funPropDecl e fData localThms funProp then return r + -- log error if no global or local theorems were found + if globalThms.size = 0 && localThms.size = 0 then + logError s!"No theorems found for `{funName}` in order to prove `{← ppExpr e}`" + if (← fData.isMorApplication) != .none then if let .some r ← applyMorRules funPropDecl e fData funProp then return r @@ -552,9 +606,6 @@ def constAppCase (funPropDecl : FunPropDecl) (e : Expr) (fData : FunctionData) if let .some r ← applyTransitionRules e funProp then return r - if globalThms.size = 0 && - localThms.size = 0 then - logError s!"No theorems found for `{funName}` in order to prove {← ppExpr e}" return none @@ -568,9 +619,10 @@ mutual /-- Main `funProp` function. Returns proof of `e`. -/ partial def funProp (e : Expr) : FunPropM (Option Result) := do + let e ← instantiateMVars e -- check cache if let .some { expr := _, proof? := .some proof } := (← get).cache.find? e then - trace[Meta.Tactic.fun_prop.cache] "cached result for {e}" + trace[Debug.Meta.Tactic.fun_prop] "cached result for {e}" return .some { proof := proof } else -- take care of forall and let binders and run main @@ -586,7 +638,6 @@ mutual | return none cacheResult e {proof := ← mkLambdaFVars xs r.proof } | .mdata _ e' => funProp e' - | .mvar _ => instantiateMVars e >>= funProp | _ => let .some r ← main e | return none @@ -603,8 +654,7 @@ mutual withTraceNode `Meta.Tactic.fun_prop (fun r => do pure s!"[{ExceptToEmoji.toEmoji r}] {← ppExpr e}") do - -- if function starts with let bindings move them the top of `e` and try - -- again + -- if function starts with let bindings move them the top of `e` and try again if f.isLet then return ← letTelescope f fun xs b => do let e' := e.setArg funPropDecl.funArgId b @@ -612,20 +662,20 @@ mutual match ← getFunctionData? f (← unfoldNamePred) {zeta := false} with | .letE f => - trace[Meta.Tactic.fun_prop.step] "let case on {← ppExpr f}" + trace[Debug.Meta.Tactic.fun_prop] "let case on {← ppExpr f}" let e := e.setArg funPropDecl.funArgId f -- update e with reduced f letCase funPropDecl e f funProp | .lam f => - trace[Meta.Tactic.fun_prop.step] "pi case on {← ppExpr f}" + trace[Debug.Meta.Tactic.fun_prop] "pi case on {← ppExpr f}" let e := e.setArg funPropDecl.funArgId f -- update e with reduced f - applyPiRule funPropDecl e f funProp + applyPiRule funPropDecl e funProp | .data fData => let e := e.setArg funPropDecl.funArgId (← fData.toExpr) -- update e with reduced f if fData.isIdentityFun then - applyIdRule funPropDecl e (← fData.domainType) funProp + applyIdRule funPropDecl e funProp else if fData.isConstantFun then - applyConstRule funPropDecl e (← fData.domainType) (Mor.mkAppN fData.fn fData.args) funProp + applyConstRule funPropDecl e funProp else match fData.fn with | .fvar id => @@ -633,11 +683,10 @@ mutual bvarAppCase funPropDecl e fData funProp else fvarAppCase funPropDecl e fData funProp - | .mvar .. => funProp (← instantiateMVars e) | .const .. | .proj .. => do constAppCase funPropDecl e fData funProp | _ => - trace[Meta.Tactic.fun_prop.step] "unknown case, ctor: {f.ctorName}\n{e}" + trace[Debug.Meta.Tactic.fun_prop] "unknown case, ctor: {f.ctorName}\n{e}" return none end diff --git a/Mathlib/Tactic/FunProp/Decl.lean b/Mathlib/Tactic/FunProp/Decl.lean index 690077cb3a1b2..8ef6ae4bb0367 100644 --- a/Mathlib/Tactic/FunProp/Decl.lean +++ b/Mathlib/Tactic/FunProp/Decl.lean @@ -99,6 +99,7 @@ fun_prop bug: expression {← ppExpr e} matches multiple function properties {decls.map (fun d => d.funPropName)}" let decl := decls[0]! + unless decl.funArgId < e.getAppNumArgs do return none let f := e.getArg! decl.funArgId return (decl,f) @@ -106,6 +107,11 @@ fun_prop bug: expression {← ppExpr e} matches multiple function properties /-- Is `e` a function property statement? -/ def isFunProp (e : Expr) : MetaM Bool := do return (← getFunProp? e).isSome +/-- Is `e` a `fun_prop` goal? For example `∀ y z, Continuous fun x => f x y z` -/ +def isFunPropGoal (e : Expr) : MetaM Bool := do + forallTelescope e fun _ b => + return (← getFunProp? b).isSome + /-- Returns function property declaration from `e = P f`. -/ def getFunPropDecl? (e : Expr) : MetaM (Option FunPropDecl) := do match ← getFunProp? e with diff --git a/Mathlib/Tactic/FunProp/Elab.lean b/Mathlib/Tactic/FunProp/Elab.lean index 1500bc79931fa..37a1f37478417 100644 --- a/Mathlib/Tactic/FunProp/Elab.lean +++ b/Mathlib/Tactic/FunProp/Elab.lean @@ -17,9 +17,12 @@ namespace Meta.FunProp open Lean.Parser.Tactic +/-- `fun_prop` config elaborator -/ +declare_config_elab elabFunPropConfig FunProp.Config + /-- Tactic to prove function properties -/ syntax (name := funPropTacStx) - "fun_prop" (discharger)? (" [" withoutPosition(ident,*,?) "]")? : tactic + "fun_prop" (config)? (discharger)? (" [" withoutPosition(ident,*,?) "]")? : tactic private def emptyDischarge : Expr → MetaM (Option Expr) := fun e => @@ -30,47 +33,54 @@ private def emptyDischarge : Expr → MetaM (Option Expr) := /-- Tactic to prove function properties -/ @[tactic funPropTacStx] def funPropTac : Tactic - | `(tactic| fun_prop $[$d]? $[[$names,*]]?) => do - - let disch ← show MetaM (Expr → MetaM (Option Expr)) from do - match d with - | none => pure emptyDischarge - | some d => - match d with - | `(discharger| (discharger := $tac)) => pure <| tacticToDischarge (← `(tactic| ($tac))) - | _ => pure emptyDischarge - - let namesToUnfold : Array Name := - match names with - | none => #[] - | .some ns => ns.getElems.map (fun n => n.getId) - - let namesToUnfold := namesToUnfold.append defaultNamesToUnfold + | `(tactic| fun_prop $[$cfg]? $[$d]? $[[$names,*]]?) => do let goal ← getMainGoal goal.withContext do let goalType ← goal.getType - let cfg : Config := {disch := disch, constToUnfold := .ofArray namesToUnfold _} - let (r?, s) ← funProp goalType cfg |>.run {} + -- the whnf and telescope is here because the goal can be + -- `∀ y, let f := fun x => x + y; Continuous fun x => x + f x` + -- However it is still not complete solution. How should we deal with mix of let and forall? + withReducible <| forallTelescopeReducing (← whnfR goalType) fun _ type => do + unless (← getFunProp? type).isSome do + let hint := + if let .some n := type.getAppFn.constName? + then s!" Maybe you forgot marking `{n}` with `@[fun_prop]`." + else "" + throwError "`{← ppExpr type}` is not a `fun_prop` goal!{hint}" + + let cfg := cfg.map (fun c => mkNullNode #[c.raw]) |>.getD (mkNullNode #[]) + let cfg ← elabFunPropConfig cfg + + let disch ← show MetaM (Expr → MetaM (Option Expr)) from do + match d with + | none => pure emptyDischarge + | some d => + match d with + | `(discharger| (discharger:=$tac)) => pure <| tacticToDischarge (← `(tactic| ($tac))) + | _ => pure emptyDischarge + + let namesToUnfold : Array Name := + match names with + | none => #[] + | .some ns => ns.getElems.map (fun n => n.getId) + + let namesToUnfold := namesToUnfold.append defaultNamesToUnfold + + let ctx : Context := + { config := cfg, + disch := disch + constToUnfold := .ofArray namesToUnfold _} + let (r?, s) ← funProp goalType ctx |>.run {} if let .some r := r? then goal.assign r.proof else let mut msg := s!"`fun_prop` was unable to prove `{← Meta.ppExpr goalType}`\n\n" - if d.isSome then - msg := msg ++ "Try running with a different discharger tactic like \ - `aesop`, `assumption`, `linarith`, `omega` etc.\n" - else - msg := msg ++ "Try running with discharger `fun_prop (disch:=aesop)` or with a different \ - discharger tactic like `assumption`, `linarith`, `omega`.\n" - - msg := msg ++ "Sometimes it is useful to run `fun_prop (disch:=trace_state; sorry)` \ - which will print all the necessary subgoals for `fun_prop` to succeed.\n" - msg := msg ++ "\nPotential issues to fix:" + + msg := msg ++ "Issues:" msg := s.msgLog.foldl (init := msg) (fun msg m => msg ++ "\n " ++ m) - msg := msg ++ "\n\nFor more detailed information use \ - `set_option trace.Meta.Tactic.fun_prop true`" - throwError msg + throwError msg | _ => throwUnsupportedSyntax diff --git a/Mathlib/Tactic/FunProp/FunctionData.lean b/Mathlib/Tactic/FunProp/FunctionData.lean index 26a9f42315b07..843b3963b807e 100644 --- a/Mathlib/Tactic/FunProp/FunctionData.lean +++ b/Mathlib/Tactic/FunProp/FunctionData.lean @@ -76,6 +76,17 @@ def getFunctionData (f : Expr) : MetaM FunctionData := do Mor.withApp b fun fn args => do + let mut fn := fn + let mut args := args + + -- revert projection in fn + if let .proj n i x := fn then + let .some info := getStructureInfo? (← getEnv) n | unreachable! + let .some projName := info.getProjFn? i | unreachable! + let p ← mkAppM projName #[x] + fn := p.getAppFn + args := p.getAppArgs.map (fun a => {expr:=a}) ++ args + let mainArgs := args |>.mapIdx (fun i ⟨arg,_⟩ => if arg.containsFVar xId then some i.1 else none) |>.filterMap id @@ -116,7 +127,9 @@ def getFunctionData? (f : Expr) else pure false - let .forallE xName xType _ _ ← inferType f | throwError "fun_prop bug: function expected" + let .forallE xName xType _ _ ← instantiateMVars (← inferType f) + | throwError m!"fun_prop bug: function expected, got `{f} : {← inferType f}, \ + type ctor {(← inferType f).ctorName}" withLocalDeclD xName xType fun x => do let fx' ← Mor.whnfPred (f.beta #[x]).eta unfold cfg let f' ← mkLambdaFVars #[x] fx' diff --git a/Mathlib/Tactic/FunProp/Theorems.lean b/Mathlib/Tactic/FunProp/Theorems.lean index 9993eef51a215..4faaad35f30b2 100644 --- a/Mathlib/Tactic/FunProp/Theorems.lean +++ b/Mathlib/Tactic/FunProp/Theorems.lean @@ -18,104 +18,73 @@ open Lean Meta namespace Meta.FunProp -/-- Stores important argument indices of lambda theorems - -For example -``` -theorem Continuous_const {α β} [TopologicalSpace α] [TopologicalSpace β] (y : β) : - Continuous fun _ : α => y -``` -is represented by -``` - .const 0 4 -``` +/-- Tag for one of the 5 basic lambda theorems, that also hold extra data for composition theorem -/ inductive LambdaTheoremArgs - | id (X : Nat) - | const (X y : Nat) - | proj (x Y : Nat) - | projDep (x Y : Nat) - | comp (f g : Nat) - | letE (f g : Nat) - | pi (f : Nat) + /-- Identity theorem e.g. `Continuous fun x => x` -/ + | id + /-- Constant theorem e.g. `Continuous fun x => y` -/ + | const + /-- Apply theorem e.g. `Continuous fun (f : (x : X) → Y x => f x)` -/ + | apply + /-- Composition theorem e.g. `Continuous f → Continuous g → Continuous fun x => f (g x)` + + The numbers `fArgId` and `gArgId` store the argument index for `f` and `g` in the composition + theorem. -/ + | comp (fArgId gArgId : Nat) + /-- Pi theorem e.g. `∀ y, Continuous (f · y) → Continuous fun x y => f x y` -/ + | pi deriving Inhabited, BEq, Repr, Hashable -/-- There are 5(+1) basic lambda theorems - -- id `Continuous fun x => x` -- const `Continuous fun x => y` -- proj `Continuous fun (f : X → Y) => f x` -- projDep `Continuous fun (f : (x : X) → Y x => f x)` -- comp `Continuous f → Continuous g → Continuous fun x => f (g x)` -- letE `Continuous f → Continuous g → Continuous fun x => let y := g x; f x y` -- pi `∀ y, Continuous (f · y) → Continuous fun x y => f x y` -/ +/-- Tag for one of the 5 basic lambda theorems -/ inductive LambdaTheoremType - | id | const | proj| projDep | comp | letE | pi + /-- Identity theorem e.g. `Continuous fun x => x` -/ + | id + /-- Constant theorem e.g. `Continuous fun x => y` -/ + | const + /-- Apply theorem e.g. `Continuous fun (f : (x : X) → Y x => f x)` -/ + | apply + /-- Composition theorem e.g. `Continuous f → Continuous g → Continuous fun x => f (g x)` -/ + | comp + /-- Pi theorem e.g. `∀ y, Continuous (f · y) → Continuous fun x y => f x y` -/ + | pi deriving Inhabited, BEq, Repr, Hashable -/-- -/ +/-- Convert `LambdaTheoremArgs` to `LambdaTheoremType`. -/ def LambdaTheoremArgs.type (t : LambdaTheoremArgs) : LambdaTheoremType := match t with - | .id .. => .id - | .const .. => .const - | .proj .. => .proj - | .projDep .. => .projDep + | .id => .id + | .const => .const | .comp .. => .comp - | .letE .. => .letE - | .pi .. => .pi + | .apply => .apply + | .pi => .pi -set_option linter.unusedVariables false in -/-- -/ +/-- Decides whether `f` is a function corresponding to one of the lambda theorems. -/ def detectLambdaTheoremArgs (f : Expr) (ctxVars : Array Expr) : MetaM (Option LambdaTheoremArgs) := do -- eta expand but beta reduce body - let f ← forallTelescope (← inferType f) fun xs b => + let f ← forallTelescope (← inferType f) fun xs _ => mkLambdaFVars xs (mkAppN f xs).headBeta match f with - | .lam xName xType xBody xBi => + | .lam _ _ xBody _ => + unless xBody.hasLooseBVars do return .some .const match xBody with - | .bvar 0 => - -- fun x => x - let .some argId_X := ctxVars.findIdx? (fun x => x == xType) | return none - return .some (.id argId_X) - | .fvar yId => - -- fun x => y - let .some argId_X := ctxVars.findIdx? (fun x => x == xType) | return none - let .some argId_y := ctxVars.findIdx? (fun x => x == (.fvar yId)) | return none - return .some (.const argId_X argId_y) - | .app (.bvar 0) (.fvar xId) => - -- fun f => f x - let fType := xType - let .some argId_x := ctxVars.findIdx? (fun x => x == (.fvar xId)) | return none - match fType with - | .forallE xName' xType' (.fvar yId) xBi' => - let .some argId_Y := ctxVars.findIdx? (fun x => x == (.fvar yId)) | return none - return .some <| .proj argId_x argId_Y - | .forallE xName' xType' (.app (.fvar yId) (.bvar 0)) xBi' => - let .some argId_Y := ctxVars.findIdx? (fun x => x == (.fvar yId)) | return none - return .some <| .projDep argId_x argId_Y - | _ => return none + | .bvar 0 => return .some .id + | .app (.bvar 0) (.fvar _) => return .some .apply | .app (.fvar fId) (.app (.fvar gId) (.bvar 0)) => -- fun x => f (g x) let .some argId_f := ctxVars.findIdx? (fun x => x == (.fvar fId)) | return none let .some argId_g := ctxVars.findIdx? (fun x => x == (.fvar gId)) | return none return .some <| .comp argId_f argId_g - | .letE yName yType (.app (.fvar gId) (.bvar 0)) - (.app (.app (.fvar fId) (.bvar 1)) (.bvar 0)) dep => - let .some argId_f := ctxVars.findIdx? (fun x => x == (.fvar fId)) | return none - let .some argId_g := ctxVars.findIdx? (fun x => x == (.fvar gId)) | return none - return .some <| .letE argId_f argId_g - | .lam Name yType (.app (.app (.fvar fId) (.bvar 1)) (.bvar 0)) yBi => - -- fun x y => f x y - let .some argId_f := ctxVars.findIdx? (fun x => x == (.fvar fId)) | return none - return .some <| .pi argId_f + | .lam _ _ (.app (.app (.fvar _) (.bvar 1)) (.bvar 0)) _ => + return .some .pi | _ => return none | _ => return none -/-- -/ +/-- Structure holding information about lambda theorem. -/ structure LambdaTheorem where /-- Name of function property -/ funPropName : Name @@ -125,33 +94,35 @@ structure LambdaTheorem where thmArgs : LambdaTheoremArgs deriving Inhabited, BEq -/-- -/ +/-- Collection of lambda theorems -/ structure LambdaTheorems where /-- map: function property name × theorem type → lambda theorem -/ - theorems : HashMap (Name × LambdaTheoremType) LambdaTheorem := {} + theorems : HashMap (Name × LambdaTheoremType) (Array LambdaTheorem) := {} deriving Inhabited -/-- return proof of lambda theorem -/ +/-- Return proof of lambda theorem -/ def LambdaTheorem.getProof (thm : LambdaTheorem) : MetaM Expr := do mkConstWithFreshMVarLevels thm.thmName -/-- -/ +/-- Environment extension storing lambda theorems. -/ abbrev LambdaTheoremsExt := SimpleScopedEnvExtension LambdaTheorem LambdaTheorems -/-- Extension storing all lambda theorems. -/ +/-- Environment extension storing all lambda theorems. -/ initialize lambdaTheoremsExt : LambdaTheoremsExt ← registerSimpleScopedEnvExtension { name := by exact decl_name% initial := {} addEntry := fun d e => - {d with theorems := d.theorems.insert (e.funPropName, e.thmArgs.type) e} + {d with theorems := + let es := d.theorems.findD (e.funPropName, e.thmArgs.type) #[] + d.theorems.insert (e.funPropName, e.thmArgs.type) (es.push e)} } -/-- -/ -def getLambdaTheorem (funPropName : Name) (type : LambdaTheoremType) : - CoreM (Option LambdaTheorem) := do - return (lambdaTheoremsExt.getState (← getEnv)).theorems.find? (funPropName,type) +/-- Get lambda theorems for particular function property `funPropName`. -/ +def getLambdaTheorems (funPropName : Name) (type : LambdaTheoremType) : + CoreM (Array LambdaTheorem) := do + return (lambdaTheoremsExt.getState (← getEnv)).theorems.findD (funPropName,type) #[] -------------------------------------------------------------------------------- @@ -172,6 +143,10 @@ inductive TheoremForm where | uncurried | comp deriving Inhabited, BEq, Repr +/-- TheoremForm to string -/ +instance : ToString TheoremForm := + ⟨fun x => match x with | .uncurried => "uncurried" | .comp => "compositional"⟩ + /-- theorem about specific function (either declared constant or free variable) -/ structure FunctionTheorem where /-- function property name -/ @@ -321,7 +296,8 @@ inductive Theorem where | transition (thm : GeneralTheorem) -/-- -/ +/-- For a theorem declaration `declName` return `fun_prop` theorem. It correctly detects which +type of theorem it is. -/ def getTheoremFromConst (declName : Name) (prio : Nat := eval_prio default) : MetaM Theorem := do let info ← getConstInfo declName forallTelescope info.type fun xs b => do @@ -384,7 +360,7 @@ def getTheoremFromConst (declName : Name) (prio : Nat := eval_prio default) : Me throwError "unrecognized theoremType `{← ppExpr b}`" -/-- -/ +/-- Register theorem `declName` with `fun_prop`. -/ def addTheorem (declName : Name) (attrKind : AttributeKind := .global) (prio : Nat := eval_prio default) : MetaM Unit := do match (← getTheoremFromConst declName prio) with @@ -401,7 +377,7 @@ function property: {thm.funPropName} function name: {thm.funOrigin.name} main arguments: {thm.mainArgs} applied arguments: {thm.appliedArgs} -form: {repr thm.form}" +form: {toString thm.form} form" functionTheoremsExt.add thm attrKind | .mor thm => trace[Meta.Tactic.fun_prop.attr] "\ diff --git a/Mathlib/Tactic/FunProp/Types.lean b/Mathlib/Tactic/FunProp/Types.lean index 00daaab6c8b4d..398048a9db06a 100644 --- a/Mathlib/Tactic/FunProp/Types.lean +++ b/Mathlib/Tactic/FunProp/Types.lean @@ -18,15 +18,10 @@ open Lean Meta namespace Meta.FunProp - -initialize registerTraceClass `Meta.Tactic.fun_prop.attr initialize registerTraceClass `Meta.Tactic.fun_prop -initialize registerTraceClass `Meta.Tactic.fun_prop.step -initialize registerTraceClass `Meta.Tactic.fun_prop.unify -initialize registerTraceClass `Meta.Tactic.fun_prop.discharge -initialize registerTraceClass `Meta.Tactic.fun_prop.apply -initialize registerTraceClass `Meta.Tactic.fun_prop.unfold -initialize registerTraceClass `Meta.Tactic.fun_prop.cache +initialize registerTraceClass `Meta.Tactic.fun_prop.attr +initialize registerTraceClass `Debug.Meta.Tactic.fun_prop + /-- Indicated origin of a function or a statement. -/ inductive Origin where @@ -72,62 +67,55 @@ def defaultNamesToUnfold : Array Name := /-- `fun_prop` configuration -/ structure Config where + /-- Maximum number of transitions between function properties. For example infering continuity + from differentiability and then differentiability from smoothness (`ContDiff ℝ ∞`) requires + `maxTransitionDepth = 2`. The default value of one expects that transition theorems are + transitively closed e.g. there is a transition theorem that infers continuity directly from + smoothenss. + + Setting `maxTransitionDepth` to zero will disable all transition theorems. This can be very + useful when `fun_prop` should fail quickly. For example when using `fun_prop` as discharger in + `simp`. + -/ + maxTransitionDepth := 1 + /-- Maximum number of steps `fun_prop` can take. -/ + maxSteps := 100000 +deriving Inhabited, BEq + +/-- `fun_prop` context -/ +structure Context where + /-- fun_prop config -/ + config : Config := {} /-- Name to unfold -/ constToUnfold : Batteries.RBSet Name Name.quickCmp := .ofArray defaultNamesToUnfold _ /-- Custom discharger to satisfy theorem hypotheses. -/ disch : Expr → MetaM (Option Expr) := fun _ => pure .none - /-- Maximal number of transitions between function properties - e.g. inferring differentiability from linearity -/ - maxDepth := 200 - /-- current depth -/ - depth := 0 - /-- Stack of used theorem, used to prevent trivial loops. -/ - thmStack : List Origin := [] - /-- Maximum number of steps `fun_prop` can take. -/ - maxSteps := 100000 -deriving Inhabited + /-- current transition depth -/ + transitionDepth := 0 /-- `fun_prop` state -/ structure State where - /-- Simp's cache is used as the `funProp` tactic is designed to be used inside of simp and utilize - its cache -/ + /-- Simp's cache is used as the `fun_prop` tactic is designed to be used inside of simp and + utilize its cache -/ cache : Simp.Cache := {} /-- Count the number of steps and stop when maxSteps is reached. -/ numSteps := 0 /-- Log progress and failures messages that should be displayed to the user at the end. -/ msgLog : List String := [] -/-- Log used theorem -/ -def Config.addThm (cfg : Config) (thmId : Origin) : Config := - {cfg with thmStack := thmId :: cfg.thmStack} - /-- Increase depth -/ -def Config.increaseDepth (cfg : Config) : Config := - {cfg with depth := cfg.depth + 1} - -/-- -/ -abbrev FunPropM := ReaderT FunProp.Config $ StateT FunProp.State MetaM +def Context.increaseTransitionDepth (ctx : Context) : Context := + {ctx with transitionDepth := ctx.transitionDepth + 1} +/-- Monad to run `fun_prop` tactic in. -/ +abbrev FunPropM := ReaderT FunProp.Context <| StateT FunProp.State MetaM /-- Result of `funProp`, it is a proof of function property `P f` -/ structure Result where /-- -/ proof : Expr -/-- Check if previously used theorem was `thmOrigin`. -/ -def previouslyUsedThm (thmOrigin : Origin) : FunPropM Bool := do - match (← read).thmStack.head? with - | .some thmOrigin' => return thmOrigin == thmOrigin' - | _ => return false - -/-- Puts the theorem to the stack of used theorems. -/ -def withTheorem {α} (thmOrigin : Origin) (go : FunPropM α) : FunPropM α := do - let cfg ← read - if cfg.depth > cfg.maxDepth then - throwError s!"fun_prop error, maximum depth({cfg.maxDepth}) reached!" - withReader (fun cfg => cfg.addThm thmOrigin |>.increaseDepth) do go - /-- Default names to unfold -/ def defaultUnfoldPred : Name → Bool := defaultNamesToUnfold.contains @@ -140,12 +128,37 @@ def unfoldNamePred : FunPropM (Name → Bool) := do /-- Increase heartbeat, throws error when `maxSteps` was reached -/ def increaseSteps : FunPropM Unit := do let numSteps := (← get).numSteps - let maxSteps := (← read).maxSteps + let maxSteps := (← read).config.maxSteps if numSteps > maxSteps then throwError s!"fun_prop failed, maximum number({maxSteps}) of steps exceeded" modify (fun s => {s with numSteps := s.numSteps + 1}) -/-- Log error message that will displayed to the user at the end. -/ +/-- Increase transition depth. Return `none` if maximum transition depth has been reached. -/ +def withIncreasedTransitionDepth {α} (go : FunPropM (Option α)) : FunPropM (Option α) := do + let maxDepth := (← read).config.maxTransitionDepth + let newDepth := (← read).transitionDepth + 1 + if newDepth > maxDepth then + trace[Meta.Tactic.fun_prop] + "maximum transition depth ({maxDepth}) reached + if you want `fun_prop` to continue then increase the maximum depth with \ + `fun_prop (config := \{maxTransitionDepth := {newDepth}})`" + return none + else + withReader (fun s => {s with transitionDepth := newDepth}) go + +/-- Log error message that will displayed to the user at the end. + +Messages are logged only when `transitionDepth = 0` i.e. when `fun_prop` is **not** trying to infer +function property like continuity from another property like differentiability. +The main reason is that if the user forgets to add a continuity theorem for function `foo` then +`fun_prop` should report that there is a continuity theorem for `foo` missing. If we would log +messages `transitionDepth > 0` then user will see messages saying that there is a missing theorem for +differentiability, smoothness, ... for `foo`. -/ def logError (msg : String) : FunPropM Unit := do - modify fun s => - {s with msgLog := msg::s.msgLog} + if (← read).transitionDepth = 0 then + modify fun s => + {s with msgLog := + if s.msgLog.contains msg then + s.msgLog + else + msg::s.msgLog} diff --git a/test/fun_prop2.lean b/test/fun_prop2.lean index 6e9d83406e0f8..4fcc1f27e7a48 100644 --- a/test/fun_prop2.lean +++ b/test/fun_prop2.lean @@ -29,6 +29,9 @@ example {n} (y : ℝ) (hy : y≠0) : ContDiffAt ℝ n foo y := by unfold foo; fun_prop (disch:=aesop) +example : Continuous fun ((x, _, _) : ℝ × ℝ × ℝ) ↦ x := by fun_prop +example : Continuous fun ((_, y, _) : ℝ × ℝ × ℝ) ↦ y := by fun_prop +example : Continuous fun ((_, _, z) : ℝ × ℝ × ℝ) ↦ z := by fun_prop -- This theorem is meant to work together with `measurable_of_continuousOn_compl_singleton` -- Unification of `(hf : ContinuousOn f {a}ᶜ)` with this theorem determines the point `a` to be `0` @@ -41,7 +44,7 @@ example : Measurable (fun x => x * (Real.log x) ^ 2 - Real.exp x / x) := by -- Notice that no theorems about measuability of log are used. It is infered from continuity. example : AEMeasurable (fun x => x * (Real.log x) ^ 2 - Real.exp x / x) := by - fun_prop + fun_prop (config:={maxTransitionDepth:=2}) diff --git a/test/fun_prop_dev.lean b/test/fun_prop_dev.lean index 92df53b44d108..fdcf619066278 100644 --- a/test/fun_prop_dev.lean +++ b/test/fun_prop_dev.lean @@ -44,11 +44,12 @@ variable [Obj α] [Obj β] [Obj γ] [Obj δ] [∀ x, Obj (E x)] @[fun_prop] theorem Con_apply (x : α) : Con (fun f : α → β => f x) := silentSorry @[fun_prop] theorem Con_applyDep (x : α) : Con (fun f : (x' : α) → E x' => f x) := silentSorry @[fun_prop] theorem Con_comp (f : β → γ) (g : α → β) (hf : Con f) (hg : Con g) : Con (fun x => f (g x)) := silentSorry -@[fun_prop] theorem Con_let (f : α → β → γ) (g : α → β) (hf : Con (fun (x,y) => f x y)) (hg : Con g) : Con (fun x => let y:= g x; f x y) := silentSorry +-- @[fun_prop] theorem Con_let (f : α → β → γ) (g : α → β) (hf : Con (fun (x,y) => f x y)) (hg : Con g) : Con (fun x => let y:= g x; f x y) := silentSorry @[fun_prop] theorem Con_pi (f : β → (i : α) → (E i)) (hf : ∀ i, Con (fun x => f x i)) : Con (fun x i => f x i) := silentSorry -- Lin is missing `const` theorem @[fun_prop] theorem Lin_id : Lin (fun x : α => x) := silentSorry +@[fun_prop] theorem Lin_const {β} [Obj β] [Zero β] : Lin (fun x : α => (0 : β)) := silentSorry @[fun_prop] theorem Lin_apply (x : α) : Lin (fun f : α → β => f x) := silentSorry @[fun_prop] theorem Lin_applyDep (x : α) : Lin (fun f : (x' : α) → E x' => f x) := silentSorry @[fun_prop] theorem Lin_comp (f : β → γ) (g : α → β) (hf : Lin f) (hg : Lin g) : Lin (f ∘ g) := silentSorry @@ -58,11 +59,8 @@ variable [Obj α] [Obj β] [Obj γ] [Obj δ] [∀ x, Obj (E x)] -- this is to stress test detection of loops @[fun_prop] theorem kaboom (f : α → β) (hf : Con f) : Con f := hf - --- currently only trivial loops are detected --- make it more sophisticated such that longer loops are detected --- @[fun_prop] --- theorem chabam (f : α → β) (hf : Con f) : Con f := hf +@[fun_prop] +theorem chabam (f : α → β) (hf : Con f) : Con f := hf -- transition theorem -- @@ -321,6 +319,8 @@ example (x) : Con fun (f : α ->> α) => f (f x) := by fun_prop example (x) : Con fun (f : α ->> α) => f (f (f x)) := by fun_prop +example [Zero α] : Lin (fun x : α => (0 : α) + x + (0 : α) + (0 : α) + x) := by fun_prop + noncomputable def foo : α ->> α ->> α := silentSorry noncomputable @@ -358,12 +358,14 @@ theorem iterate_con (n : Nat) (f : α → α) (hf : Con f) : Con (iterate n f) : example : let f := fun x : α => x; Con f := by fun_prop - +example : let f := fun x => x + y; ∀ y : α, ∀ z : α, Con fun x => x + f x + z := by fun_prop +example : ∀ y : α, let f := fun x => x + y; ∀ z : α, Con fun x => x + f x + z := by fun_prop +-- this is still broken +-- example : ∀ y : α, ∀ z : α, let f := fun x => x + y; Con fun x => x + f x + z := by fun_prop example (f g : α → β) (hf : Con f := by fun_prop) (hg : outParam (Con g)) : Con (fun x => f x + g x) := by fun_prop - opaque foo1 : α → α := id opaque foo2 : α → α := id @@ -383,3 +385,83 @@ def diag (f : α → α → α) (x : α) := f x x theorem diag_Con (f : α → α → α) (hf : Con (myUncurry f)) : Con (fun x => diag f x) := by fun_prop [diag,myUncurry] +namespace MultipleLambdaTheorems + +opaque A : Prop +opaque B : Prop +@[local fun_prop] theorem Con_comp' (f : β → γ) (g : α → β) (h : A) : Con (fun x => f (g x)) := silentSorry +@[local fun_prop] theorem Con_comp'' (f : β → γ) (g : α → β) (b : B) : Con (fun x => f (g x)) := silentSorry + +example (f : β → γ) (g : α → β) (h : A) : Con (fun x => f (g x)) := by fun_prop (disch := assumption) +example (f : β → γ) (g : α → β) (h : B) : Con (fun x => f (g x)) := by fun_prop (disch := assumption) + +end MultipleLambdaTheorems + + +/-- warning: `?m` is not a `fun_prop` goal! -/ +#guard_msgs in +#check_failure ((by fun_prop) : ?m) + +-- todo: warning should not have mvar id in it +-- /-- warning: `?m.71721` is not a `fun_prop` goal! -/ +-- #guard_msgs in +-- #check_failure (by exact add_Con' (by fun_prop) : Con (fun x : α => (x + x) + (x + x))) + +example : Con fun ((x, _, _) : α × α × α) => x := by fun_prop +example : Con fun ((_, x, _) : α × α × α) => x := by fun_prop +example : Con fun ((_, _, x) : α × α × α) => x := by fun_prop + +example : let f := (by exact (fun x : α => x+x)); Con f := by + intro f; + let F := fun x : α => x+x + have : Con F := by fun_prop -- this used to be problematic + fun_prop + + +def f1 (a : α) := a +def f2 (a : α) := a + +/-- +error: `fun_prop` was unable to prove `Con fun x => x + f1 x` + +Issues: + No theorems found for `f1` in order to prove `Con fun a => f1 a` +-/ +#guard_msgs in +example : Con (fun x : α => x + f1 x) := by fun_prop + +/-- +error: `fun_prop` was unable to prove `Con fun x => f1 x + f1 x` + +Issues: + No theorems found for `f1` in order to prove `Con fun a => f1 a` +-/ +#guard_msgs in +example : Con (fun x : α => f1 x + f1 x) := by fun_prop + +/-- +error: `fun_prop` was unable to prove `Con fun x => f2 x + f1 x` + +Issues: + No theorems found for `f2` in order to prove `Con fun a => f2 a` +-/ +#guard_msgs in +example : Con (fun x : α => f2 x + f1 x) := by fun_prop + + +def f3 (a : α) := a + +@[fun_prop] +theorem f3_lin : Lin (fun x : α => f3 x) := by + unfold f3; fun_prop (config:={maxTransitionDepth:=0,maxSteps:=10}) + +example : Con (fun x : α => f3 x) := by fun_prop + +/-- +error: `fun_prop` was unable to prove `Con fun x => f3 x` + +Issues: + No theorems found for `f3` in order to prove `Con fun x => f3 x` +-/ +#guard_msgs in +example : Con (fun x : α => f3 x) := by fun_prop (config:={maxTransitionDepth:=0}) From f57201f9e0f3097e72753cdcc99f9917ea055339 Mon Sep 17 00:00:00 2001 From: Jeremy Tan Jie Rui Date: Mon, 12 Aug 2024 15:04:34 +0000 Subject: [PATCH 213/359] chore: deduplicate `inv_mul_cancel`/`mul_inv_cancel`, etc. (#15717) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * `mul_left_inv`, `inv_mul_self` → `inv_mul_cancel` (for `Group`) * `mul_right_inv`, `mul_inv_self` → `mul_inv_cancel` * `add_left_neg`, `neg_add_self` → `neg_add_cancel` (for `AddGroup`) * `add_right_neg`, `add_neg_self` → `add_neg_cancel` * `inv_mul_cancel` → `inv_mul_cancel₀` (for `GroupWithZero`) * `mul_inv_cancel` → `mul_inv_cancel₀` --- Archive/Imo/Imo2008Q3.lean | 2 +- Archive/Imo/Imo2024Q2.lean | 2 +- Archive/Imo/Imo2024Q6.lean | 4 +- Archive/Wiedijk100Theorems/AreaOfACircle.lean | 6 +-- Archive/Wiedijk100Theorems/BuffonsNeedle.lean | 2 +- Counterexamples/SorgenfreyLine.lean | 4 +- Mathlib/Algebra/AddTorsor.lean | 2 +- Mathlib/Algebra/Algebra/Equiv.lean | 6 +-- Mathlib/Algebra/BigOperators/Fin.lean | 2 +- Mathlib/Algebra/BigOperators/Group/List.lean | 2 +- Mathlib/Algebra/Category/Grp/Colimits.lean | 6 +-- Mathlib/Algebra/Category/Grp/EpiMono.lean | 6 +-- .../Category/Grp/FilteredColimits.lean | 4 +- .../Algebra/Category/ModuleCat/Presheaf.lean | 2 +- Mathlib/Algebra/Category/Ring/Colimits.lean | 12 ++--- Mathlib/Algebra/CharP/Basic.lean | 4 +- Mathlib/Algebra/CharZero/Quotient.lean | 2 +- Mathlib/Algebra/DirectLimit.lean | 2 +- Mathlib/Algebra/Field/IsField.lean | 2 +- Mathlib/Algebra/Field/MinimalAxioms.lean | 4 +- Mathlib/Algebra/Field/Subfield.lean | 2 +- Mathlib/Algebra/GeomSum.lean | 10 ++--- Mathlib/Algebra/Group/Action/Defs.lean | 6 +-- Mathlib/Algebra/Group/Action/Units.lean | 4 +- Mathlib/Algebra/Group/AddChar.lean | 4 +- Mathlib/Algebra/Group/Aut.lean | 14 +++--- Mathlib/Algebra/Group/Basic.lean | 6 +-- Mathlib/Algebra/Group/Conj.lean | 2 +- Mathlib/Algebra/Group/Defs.lean | 43 +++++++++--------- Mathlib/Algebra/Group/Fin/Basic.lean | 2 +- Mathlib/Algebra/Group/Hom/Basic.lean | 4 +- Mathlib/Algebra/Group/Hom/Defs.lean | 2 +- Mathlib/Algebra/Group/Hom/Instances.lean | 2 +- Mathlib/Algebra/Group/InjSurj.lean | 4 +- Mathlib/Algebra/Group/Int.lean | 2 +- Mathlib/Algebra/Group/Invertible/Defs.lean | 4 +- Mathlib/Algebra/Group/MinimalAxioms.lean | 24 +++++----- Mathlib/Algebra/Group/Opposite.lean | 2 +- Mathlib/Algebra/Group/Pi/Basic.lean | 2 +- Mathlib/Algebra/Group/Prod.lean | 2 +- Mathlib/Algebra/Group/Semiconj/Defs.lean | 2 +- Mathlib/Algebra/Group/Subgroup/Basic.lean | 24 +++++----- Mathlib/Algebra/Group/Subgroup/Pointwise.lean | 8 ++-- .../Algebra/Group/Submonoid/Membership.lean | 2 +- Mathlib/Algebra/Group/TypeTags.lean | 4 +- Mathlib/Algebra/Group/UniqueProds.lean | 4 +- Mathlib/Algebra/Group/Units.lean | 9 ++-- Mathlib/Algebra/Group/Units/Equiv.lean | 2 +- Mathlib/Algebra/Group/Units/Hom.lean | 4 +- .../Algebra/GroupWithZero/Action/Defs.lean | 2 +- Mathlib/Algebra/GroupWithZero/Basic.lean | 26 +++++------ Mathlib/Algebra/GroupWithZero/Defs.lean | 4 +- Mathlib/Algebra/GroupWithZero/InjSurj.lean | 4 +- Mathlib/Algebra/GroupWithZero/Invertible.lean | 8 ++-- Mathlib/Algebra/GroupWithZero/NeZero.lean | 4 +- .../GroupWithZero/NonZeroDivisors.lean | 4 +- Mathlib/Algebra/GroupWithZero/Opposite.lean | 2 +- .../Algebra/GroupWithZero/Units/Basic.lean | 8 ++-- .../Algebra/GroupWithZero/Units/Lemmas.lean | 2 +- Mathlib/Algebra/GroupWithZero/WithZero.lean | 2 +- .../Algebra/Homology/DerivedCategory/Ext.lean | 2 +- Mathlib/Algebra/Homology/Homotopy.lean | 2 +- .../HomotopyCategory/DegreewiseSplit.lean | 2 +- .../Homology/HomotopyCategory/HomComplex.lean | 14 +++--- .../HomotopyCategory/MappingCone.lean | 28 ++++++------ .../HomotopyCategory/Pretriangulated.lean | 30 ++++++------- .../HomotopyCategory/Triangulated.lean | 12 ++--- .../Homology/ShortComplex/Preadditive.lean | 2 +- Mathlib/Algebra/Lie/BaseChange.lean | 2 +- Mathlib/Algebra/Lie/Weights/Cartan.lean | 6 +-- Mathlib/Algebra/Lie/Weights/Killing.lean | 6 +-- Mathlib/Algebra/Lie/Weights/RootSystem.lean | 8 ++-- Mathlib/Algebra/Module/CharacterModule.lean | 2 +- Mathlib/Algebra/Module/Defs.lean | 6 +-- Mathlib/Algebra/Module/Equiv/Basic.lean | 2 +- Mathlib/Algebra/Module/LocalizedModule.lean | 2 +- Mathlib/Algebra/Order/AddGroupWithTop.lean | 2 +- Mathlib/Algebra/Order/CauSeq/Basic.lean | 2 +- Mathlib/Algebra/Order/Field/Basic.lean | 2 +- Mathlib/Algebra/Order/Floor.lean | 2 +- Mathlib/Algebra/Order/Group/PosPart.lean | 8 ++-- .../Algebra/Order/Group/Unbundled/Abs.lean | 4 +- .../Algebra/Order/Group/Unbundled/Basic.lean | 24 +++++----- .../Order/GroupWithZero/Canonical.lean | 2 +- Mathlib/Algebra/Order/Positive/Field.lean | 2 +- Mathlib/Algebra/Order/Ring/Cast.lean | 4 +- Mathlib/Algebra/Order/ToIntervalMod.lean | 2 +- Mathlib/Algebra/PUnitInstances/Algebra.lean | 2 +- Mathlib/Algebra/Polynomial/Basic.lean | 2 +- Mathlib/Algebra/Polynomial/CancelLeads.lean | 2 +- Mathlib/Algebra/Polynomial/Degree/Lemmas.lean | 2 +- Mathlib/Algebra/Polynomial/Eval.lean | 2 +- Mathlib/Algebra/Polynomial/FieldDivision.lean | 10 ++--- Mathlib/Algebra/Polynomial/Laurent.lean | 10 ++--- Mathlib/Algebra/Polynomial/Smeval.lean | 2 +- Mathlib/Algebra/Polynomial/Taylor.lean | 2 +- Mathlib/Algebra/Quandle.lean | 8 ++-- Mathlib/Algebra/Quaternion.lean | 2 +- Mathlib/Algebra/Ring/Aut.lean | 2 +- Mathlib/Algebra/Ring/Basic.lean | 2 +- Mathlib/Algebra/Ring/BooleanRing.lean | 6 +-- Mathlib/Algebra/Ring/Defs.lean | 4 +- Mathlib/Algebra/Ring/MinimalAxioms.lean | 10 ++--- Mathlib/Algebra/Ring/Opposite.lean | 2 +- Mathlib/Algebra/Ring/Subring/Basic.lean | 4 +- Mathlib/Algebra/Ring/ULift.lean | 10 ++--- Mathlib/Algebra/RingQuot.lean | 2 +- Mathlib/Algebra/SMulWithZero.lean | 2 +- Mathlib/Algebra/Star/CHSH.lean | 2 +- Mathlib/Algebra/Star/Unitary.lean | 2 +- Mathlib/Algebra/TrivSqZeroExt.lean | 12 ++--- Mathlib/Algebra/Tropical/Basic.lean | 2 +- .../EllipticCurve/Group.lean | 14 +++--- .../EllipticCurve/Weierstrass.lean | 2 +- .../AlternatingFaceMapComplex.lean | 2 +- .../DoldKan/Degeneracies.lean | 2 +- Mathlib/AlgebraicTopology/DoldKan/Faces.lean | 4 +- .../AlgebraicTopology/ExtraDegeneracy.lean | 2 +- Mathlib/Analysis/Analytic/Basic.lean | 2 +- Mathlib/Analysis/Analytic/Meromorphic.lean | 2 +- Mathlib/Analysis/Analytic/RadiusLiminf.lean | 2 +- .../Asymptotics/SpecificAsymptotics.lean | 4 +- .../Asymptotics/SuperpolynomialDecay.lean | 12 ++--- Mathlib/Analysis/BoundedVariation.lean | 2 +- .../Calculus/BumpFunction/Normed.lean | 2 +- .../Analysis/Calculus/ContDiff/Bounds.lean | 2 +- .../Analysis/Calculus/FDeriv/Analytic.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Basic.lean | 2 +- .../Analysis/Calculus/LineDeriv/Basic.lean | 8 ++-- Mathlib/Analysis/Calculus/Rademacher.lean | 4 +- Mathlib/Analysis/Calculus/TangentCone.lean | 2 +- Mathlib/Analysis/Complex/OpenMapping.lean | 4 +- Mathlib/Analysis/Convex/Between.lean | 4 +- Mathlib/Analysis/Convex/Combination.lean | 8 ++-- Mathlib/Analysis/Convex/Cone/Basic.lean | 8 ++-- Mathlib/Analysis/Convex/Cone/Extension.lean | 8 ++-- Mathlib/Analysis/Convex/Segment.lean | 4 +- Mathlib/Analysis/Convex/Side.lean | 14 +++--- .../Convex/SpecificFunctions/Basic.lean | 8 ++-- Mathlib/Analysis/Convex/Uniform.lean | 4 +- Mathlib/Analysis/Fourier/AddCircle.lean | 2 +- .../Analysis/Fourier/FourierTransform.lean | 2 +- .../FunctionalSpaces/SobolevInequality.lean | 2 +- .../Analysis/InnerProductSpace/OfNorm.lean | 2 +- .../LocallyConvex/ContinuousOfBounded.lean | 2 +- Mathlib/Analysis/Matrix.lean | 4 +- Mathlib/Analysis/MeanInequalities.lean | 8 ++-- Mathlib/Analysis/MeanInequalitiesPow.lean | 6 +-- Mathlib/Analysis/MellinInversion.lean | 2 +- Mathlib/Analysis/MellinTransform.lean | 2 +- Mathlib/Analysis/Normed/Affine/AddTorsor.lean | 2 +- Mathlib/Analysis/Normed/Affine/Isometry.lean | 2 +- .../Analysis/Normed/Algebra/Exponential.lean | 10 ++--- Mathlib/Analysis/Normed/Field/Basic.lean | 2 +- Mathlib/Analysis/Normed/Group/AddCircle.lean | 10 ++--- Mathlib/Analysis/Normed/Group/Basic.lean | 4 +- Mathlib/Analysis/Normed/Lp/PiLp.lean | 7 +-- Mathlib/Analysis/Normed/Lp/ProdLp.lean | 6 +-- Mathlib/Analysis/Normed/Lp/lpSpace.lean | 2 +- Mathlib/Analysis/Normed/Module/Dual.lean | 2 +- Mathlib/Analysis/Normed/MulAction.lean | 2 +- Mathlib/Analysis/Normed/Operator/Banach.lean | 4 +- .../Normed/Operator/ContinuousLinearMap.lean | 2 +- .../Normed/Operator/LinearIsometry.lean | 2 +- Mathlib/Analysis/Normed/Ring/Seminorm.lean | 2 +- .../Normed/Ring/SeminormFromConst.lean | 4 +- Mathlib/Analysis/Normed/Ring/Units.lean | 2 +- Mathlib/Analysis/NormedSpace/Connected.lean | 4 +- .../NormedSpace/HahnBanach/Extension.lean | 2 +- .../HahnBanach/SeparatingDual.lean | 8 ++-- .../NormedSpace/OperatorNorm/Basic.lean | 2 +- Mathlib/Analysis/NormedSpace/Pointwise.lean | 2 +- Mathlib/Analysis/NormedSpace/RCLike.lean | 2 +- Mathlib/Analysis/NormedSpace/Real.lean | 2 +- Mathlib/Analysis/NormedSpace/RieszLemma.lean | 2 +- .../Analysis/NormedSpace/SphereNormEquiv.lean | 2 +- Mathlib/Analysis/RCLike/Basic.lean | 4 +- Mathlib/Analysis/Seminorm.lean | 2 +- .../SpecialFunctions/Complex/Arctan.lean | 8 ++-- .../SpecialFunctions/Complex/LogBounds.lean | 2 +- .../SpecialFunctions/Gamma/Basic.lean | 4 +- .../Analysis/SpecialFunctions/Gamma/Beta.lean | 2 +- .../Gaussian/GaussianIntegral.lean | 2 +- .../Analysis/SpecialFunctions/Integrals.lean | 2 +- .../Analysis/SpecialFunctions/Log/Deriv.lean | 2 +- .../SpecialFunctions/Log/NegMulLog.lean | 2 +- .../SpecialFunctions/Pow/Complex.lean | 4 +- .../Analysis/SpecialFunctions/Pow/NNReal.lean | 26 +++++------ .../Analysis/SpecialFunctions/Pow/Real.lean | 10 ++--- .../Trigonometric/Series.lean | 2 +- Mathlib/Analysis/SpecificLimits/Basic.lean | 4 +- Mathlib/Analysis/SpecificLimits/FloorPow.lean | 4 +- Mathlib/CategoryTheory/Abelian/Basic.lean | 4 +- .../Abelian/NonPreadditive.lean | 6 +-- Mathlib/CategoryTheory/Endomorphism.lean | 4 +- Mathlib/CategoryTheory/Galois/Examples.lean | 2 +- .../CategoryTheory/Groupoid/VertexGroup.lean | 2 +- .../CategoryTheory/Idempotents/Karoubi.lean | 4 +- .../CalculusOfFractions/Preadditive.lean | 6 +-- Mathlib/CategoryTheory/Preadditive/Basic.lean | 2 +- .../Preadditive/EilenbergMoore.lean | 8 ++-- .../Preadditive/EndoFunctor.lean | 8 ++-- .../Preadditive/FunctorCategory.lean | 4 +- .../CategoryTheory/Quotient/Preadditive.lean | 2 +- Mathlib/CategoryTheory/Shift/Basic.lean | 44 +++++++++---------- Mathlib/CategoryTheory/SingleObj.lean | 10 ++--- .../Sites/NonabelianCohomology/H1.lean | 6 +-- .../Triangulated/HomologicalFunctor.lean | 2 +- .../CategoryTheory/Triangulated/Opposite.lean | 6 +-- .../CategoryTheory/Triangulated/Rotate.lean | 4 +- .../Triangulated/TStructure/Basic.lean | 2 +- .../Triangulated/TriangleShift.lean | 4 +- .../Additive/AP/Three/Behrend.lean | 2 +- .../SetFamily/AhlswedeZhang.lean | 2 +- .../SetFamily/CauchyDavenport.lean | 2 +- .../Computability/AkraBazzi/AkraBazzi.lean | 8 ++-- .../AkraBazzi/GrowsPolynomially.lean | 4 +- Mathlib/Data/Complex/Basic.lean | 4 +- Mathlib/Data/Complex/Exponential.lean | 10 ++--- Mathlib/Data/DFinsupp/NeLocus.lean | 2 +- Mathlib/Data/ENNReal/Inv.lean | 4 +- Mathlib/Data/ENNReal/Operations.lean | 2 +- Mathlib/Data/Finsupp/NeLocus.lean | 2 +- Mathlib/Data/Matrix/Rank.lean | 2 +- Mathlib/Data/NNReal/Basic.lean | 10 ++--- Mathlib/Data/Nat/Choose/Sum.lean | 2 +- Mathlib/Data/Nat/Totient.lean | 2 +- Mathlib/Data/Num/Lemmas.lean | 2 +- Mathlib/Data/Rat/Cast/Defs.lean | 2 +- Mathlib/Data/Rat/Defs.lean | 4 +- Mathlib/Data/Rat/Star.lean | 2 +- Mathlib/Data/Real/Archimedean.lean | 2 +- Mathlib/Data/Real/Basic.lean | 2 +- Mathlib/Data/Real/GoldenRatio.lean | 4 +- Mathlib/Data/Real/Hyperreal.lean | 2 +- Mathlib/Data/Real/Irrational.lean | 2 +- Mathlib/Data/ZMod/Basic.lean | 4 +- Mathlib/Data/ZMod/Defs.lean | 2 +- Mathlib/Deprecated/Group.lean | 2 +- Mathlib/Deprecated/Subgroup.lean | 6 +-- Mathlib/Deprecated/Subring.lean | 2 +- Mathlib/Dynamics/Flow.lean | 4 +- Mathlib/Dynamics/Newton.lean | 2 +- Mathlib/FieldTheory/AbelRuffini.lean | 2 +- Mathlib/FieldTheory/Finite/Basic.lean | 2 +- Mathlib/FieldTheory/Minpoly/Field.lean | 4 +- Mathlib/FieldTheory/MvPolynomial.lean | 2 +- Mathlib/FieldTheory/PerfectClosure.lean | 6 +-- Mathlib/FieldTheory/RatFunc/Basic.lean | 8 ++-- .../Euclidean/Angle/Oriented/Basic.lean | 4 +- .../Euclidean/Angle/Unoriented/Affine.lean | 6 +-- Mathlib/Geometry/Euclidean/Basic.lean | 2 +- .../Manifold/Algebra/SmoothFunctions.lean | 2 +- .../Geometry/Manifold/Instances/Sphere.lean | 2 +- .../Manifold/Sheaf/LocallyRingedSpace.lean | 4 +- Mathlib/GroupTheory/Complement.lean | 8 ++-- Mathlib/GroupTheory/Congruence/Basic.lean | 2 +- Mathlib/GroupTheory/Coprod/Basic.lean | 10 ++--- Mathlib/GroupTheory/CoprodI.lean | 4 +- Mathlib/GroupTheory/Coset.lean | 4 +- Mathlib/GroupTheory/CosetCover.lean | 4 +- Mathlib/GroupTheory/DoubleCoset.lean | 6 +-- Mathlib/GroupTheory/FreeGroup/Basic.lean | 2 +- Mathlib/GroupTheory/GroupAction/Basic.lean | 6 +-- Mathlib/GroupTheory/GroupAction/ConjAct.lean | 2 +- Mathlib/GroupTheory/GroupAction/Quotient.lean | 2 +- Mathlib/GroupTheory/Index.lean | 2 +- Mathlib/GroupTheory/Perm/Basic.lean | 2 +- Mathlib/GroupTheory/Perm/Cycle/Basic.lean | 2 +- Mathlib/GroupTheory/Perm/Cycle/Type.lean | 2 +- Mathlib/GroupTheory/PushoutI.lean | 4 +- Mathlib/GroupTheory/QuotientGroup.lean | 6 +-- Mathlib/GroupTheory/Schreier.lean | 2 +- Mathlib/GroupTheory/SchurZassenhaus.lean | 4 +- Mathlib/GroupTheory/SemidirectProduct.lean | 2 +- .../GroupTheory/SpecificGroups/Cyclic.lean | 4 +- .../GroupTheory/SpecificGroups/Dihedral.lean | 4 +- .../SpecificGroups/Quaternion.lean | 4 +- Mathlib/GroupTheory/Submonoid/Inverses.lean | 8 ++-- Mathlib/GroupTheory/Torsion.lean | 2 +- Mathlib/GroupTheory/Transfer.lean | 2 +- .../AffineSpace/AffineEquiv.lean | 2 +- .../AffineSpace/AffineSubspace.lean | 2 +- .../AffineSpace/Combination.lean | 2 +- .../AffineSpace/Independent.lean | 2 +- .../LinearAlgebra/AffineSpace/Midpoint.lean | 2 +- .../LinearAlgebra/Alternating/DomCoprod.lean | 2 +- .../CliffordAlgebra/Contraction.lean | 4 +- .../LinearAlgebra/Dimension/DivisionRing.lean | 2 +- Mathlib/LinearAlgebra/Dual.lean | 2 +- Mathlib/LinearAlgebra/Lagrange.lean | 4 +- Mathlib/LinearAlgebra/LinearIndependent.lean | 2 +- .../LinearAlgebra/Matrix/Charpoly/Coeff.lean | 2 +- .../Matrix/GeneralLinearGroup/Defs.lean | 4 +- .../LinearAlgebra/Matrix/SchurComplement.lean | 6 +-- .../Matrix/SpecialLinearGroup.lean | 6 +-- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 4 +- Mathlib/LinearAlgebra/Orientation.lean | 2 +- .../LinearAlgebra/QuadraticForm/Basic.lean | 2 +- .../LinearAlgebra/QuadraticForm/Complex.lean | 2 +- Mathlib/LinearAlgebra/SymplecticGroup.lean | 2 +- .../LinearAlgebra/TensorProduct/Basic.lean | 8 ++-- Mathlib/LinearAlgebra/UnitaryGroup.lean | 4 +- .../Covering/BesicovitchVectorSpace.lean | 2 +- .../MeasureTheory/Decomposition/Lebesgue.lean | 6 +-- .../ConditionalExpectation/Basic.lean | 2 +- .../Function/ConditionalExpectation/Real.lean | 2 +- .../Function/LpSeminorm/Basic.lean | 14 +++--- .../Function/LpSeminorm/CompareExp.lean | 4 +- Mathlib/MeasureTheory/Function/LpSpace.lean | 4 +- .../Function/StronglyMeasurable/Basic.lean | 2 +- .../Function/StronglyMeasurable/Lemmas.lean | 2 +- .../Function/UniformIntegrable.lean | 2 +- .../Integral/CircleIntegral.lean | 2 +- Mathlib/MeasureTheory/Integral/Gamma.lean | 2 +- .../Integral/LebesgueNormedSpace.lean | 2 +- .../MeasureTheory/Integral/PeakFunction.lean | 4 +- .../Measure/Haar/NormedSpace.lean | 2 +- .../MeasureTheory/Measure/Haar/Unique.lean | 2 +- Mathlib/MeasureTheory/Measure/Hausdorff.lean | 4 +- .../MeasureTheory/Measure/Lebesgue/Basic.lean | 6 +-- .../Measure/Lebesgue/EqHaar.lean | 8 ++-- .../MeasureTheory/Measure/MeasureSpace.lean | 2 +- .../Measure/ProbabilityMeasure.lean | 2 +- Mathlib/MeasureTheory/Measure/Tilted.lean | 4 +- Mathlib/ModelTheory/Algebra/Field/Basic.lean | 12 ++--- Mathlib/NumberTheory/ArithmeticFunction.lean | 6 +-- Mathlib/NumberTheory/ClassNumber/Finite.lean | 2 +- Mathlib/NumberTheory/Cyclotomic/Three.lean | 2 +- Mathlib/NumberTheory/Dioph.lean | 2 +- Mathlib/NumberTheory/GaussSum.lean | 2 +- Mathlib/NumberTheory/Harmonic/GammaDeriv.lean | 4 +- .../NumberTheory/LSeries/AbstractFuncEq.lean | 4 +- .../NumberTheory/LSeries/HurwitzZetaEven.lean | 6 +-- .../LSeries/MellinEqDirichlet.lean | 2 +- .../LegendreSymbol/AddCharacter.lean | 2 +- .../NumberTheory/LegendreSymbol/Basic.lean | 2 +- .../LegendreSymbol/QuadraticChar/Basic.lean | 2 +- .../NumberTheory/Liouville/LiouvilleWith.lean | 4 +- Mathlib/NumberTheory/Modular.lean | 2 +- .../ModularForms/EisensteinSeries/Defs.lean | 6 +-- .../ModularForms/JacobiTheta/TwoVariable.lean | 2 +- .../ModularForms/SlashActions.lean | 2 +- Mathlib/NumberTheory/MulChar/Basic.lean | 4 +- .../CanonicalEmbedding/ConvexBody.lean | 6 +-- .../NumberField/Discriminant.lean | 4 +- .../NumberField/Units/DirichletTheorem.lean | 2 +- Mathlib/NumberTheory/Padics/Hensel.lean | 2 +- .../NumberTheory/Padics/PadicIntegers.lean | 6 +-- Mathlib/NumberTheory/Padics/PadicVal.lean | 2 +- Mathlib/NumberTheory/Pell.lean | 2 +- Mathlib/NumberTheory/PellMatiyasevic.lean | 2 +- Mathlib/NumberTheory/Wilson.lean | 2 +- Mathlib/NumberTheory/Zsqrtd/Basic.lean | 2 +- .../Zsqrtd/QuadraticReciprocity.lean | 2 +- Mathlib/Order/Filter/FilterProduct.lean | 2 +- Mathlib/Order/Filter/Germ/Basic.lean | 2 +- Mathlib/Order/RelIso/Group.lean | 2 +- .../Probability/Distributions/Gaussian.lean | 4 +- Mathlib/Probability/Kernel/RadonNikodym.lean | 4 +- Mathlib/Probability/Variance.lean | 4 +- .../RepresentationTheory/Action/Basic.lean | 4 +- .../RepresentationTheory/Action/Limits.lean | 2 +- .../RepresentationTheory/Action/Monoidal.lean | 4 +- .../GroupCohomology/LowDegree.lean | 12 ++--- Mathlib/RepresentationTheory/Maschke.lean | 2 +- Mathlib/RepresentationTheory/Rep.lean | 6 +-- Mathlib/RingTheory/Algebraic.lean | 2 +- Mathlib/RingTheory/Binomial.lean | 2 +- Mathlib/RingTheory/Coprime/Basic.lean | 2 +- .../RingTheory/DedekindDomain/Different.lean | 14 +++--- .../DedekindDomain/Factorization.lean | 4 +- Mathlib/RingTheory/DedekindDomain/Ideal.lean | 12 ++--- Mathlib/RingTheory/Derivation/Basic.lean | 2 +- .../FractionalIdeal/Operations.lean | 8 ++-- Mathlib/RingTheory/FreeCommRing.lean | 2 +- Mathlib/RingTheory/FreeRing.lean | 2 +- Mathlib/RingTheory/HahnSeries/Addition.lean | 4 +- Mathlib/RingTheory/HahnSeries/Summable.lean | 12 ++--- Mathlib/RingTheory/Henselian.lean | 2 +- Mathlib/RingTheory/Ideal/Maps.lean | 2 +- Mathlib/RingTheory/LaurentSeries.lean | 4 +- Mathlib/RingTheory/Localization/NumDen.lean | 2 +- Mathlib/RingTheory/MvPowerSeries/Inverse.lean | 2 +- Mathlib/RingTheory/OreLocalization/Basic.lean | 8 ++-- Mathlib/RingTheory/OreLocalization/Ring.lean | 2 +- Mathlib/RingTheory/Polynomial/Dickson.lean | 4 +- Mathlib/RingTheory/Polynomial/GaussLemma.lean | 2 +- Mathlib/RingTheory/Valuation/Integers.lean | 2 +- .../RingTheory/Valuation/ValuationRing.lean | 2 +- .../Valuation/ValuationSubring.lean | 2 +- Mathlib/SetTheory/Game/Basic.lean | 2 +- Mathlib/SetTheory/Game/Impartial.lean | 2 +- Mathlib/SetTheory/Game/PGame.lean | 44 +++++++++---------- Mathlib/SetTheory/Surreal/Basic.lean | 2 +- Mathlib/Tactic/Group.lean | 2 +- .../Topology/Algebra/ContinuousMonoidHom.lean | 2 +- Mathlib/Topology/Algebra/Field.lean | 4 +- Mathlib/Topology/Algebra/Group/Basic.lean | 4 +- Mathlib/Topology/Algebra/GroupCompletion.lean | 4 +- Mathlib/Topology/Algebra/Module/Basic.lean | 12 ++--- .../Topology/Algebra/Module/Cardinality.lean | 4 +- .../Algebra/Module/FiniteDimension.lean | 2 +- Mathlib/Topology/Algebra/UniformField.lean | 2 +- .../Topology/Algebra/Valued/ValuedField.lean | 14 +++--- Mathlib/Topology/Connected/PathConnected.lean | 2 +- .../Topology/ContinuousFunction/Ideals.lean | 2 +- Mathlib/Topology/Instances/AddCircle.lean | 2 +- .../Topology/MetricSpace/DilationEquiv.lean | 4 +- .../Topology/MetricSpace/GromovHausdorff.lean | 4 +- .../Topology/MetricSpace/IsometricSMul.lean | 2 +- Mathlib/Topology/MetricSpace/Isometry.lean | 2 +- 412 files changed, 930 insertions(+), 925 deletions(-) diff --git a/Archive/Imo/Imo2008Q3.lean b/Archive/Imo/Imo2008Q3.lean index 97f0bf9dfc6be..3a908bf008718 100644 --- a/Archive/Imo/Imo2008Q3.lean +++ b/Archive/Imo/Imo2008Q3.lean @@ -43,7 +43,7 @@ theorem p_lemma (p : ℕ) (hpp : Nat.Prime p) (hp_mod_4_eq_1 : p ≡ 1 [MOD 4]) simp only [n, Int.natAbs_sq, Int.natCast_pow, Int.ofNat_succ, Int.natCast_dvd_natCast.mp] refine (ZMod.intCast_zmod_eq_zero_iff_dvd (m ^ 2 + 1) p).mp ?_ simp only [m, Int.cast_pow, Int.cast_add, Int.cast_one, ZMod.coe_valMinAbs] - rw [pow_two, ← hy]; exact add_left_neg 1 + rw [pow_two, ← hy]; exact neg_add_cancel 1 have hnat₂ : n ≤ p / 2 := ZMod.natAbs_valMinAbs_le y have hnat₃ : p ≥ 2 * n := by linarith [Nat.div_mul_le_self p 2] set k : ℕ := p - 2 * n with hnat₄ diff --git a/Archive/Imo/Imo2024Q2.lean b/Archive/Imo/Imo2024Q2.lean index 048b384ebbf9f..cf1a6f8be4e86 100644 --- a/Archive/Imo/Imo2024Q2.lean +++ b/Archive/Imo/Imo2024Q2.lean @@ -136,7 +136,7 @@ lemma ab_add_one_dvd_a_pow_large_n_add_b : a * b + 1 ∣ a ^ h.large_n + b := by (IsUnit.mul_right_eq_zero (ZMod.unitOfCoprime _ a_coprime_ab_add_one).isUnit).1 this rw [mul_add] norm_cast - simp only [mul_right_inv, Units.val_one, ZMod.coe_unitOfCoprime] + simp only [mul_inv_cancel, Units.val_one, ZMod.coe_unitOfCoprime] norm_cast convert ZMod.natCast_self (a * b + 1) using 2 exact add_comm _ _ diff --git a/Archive/Imo/Imo2024Q6.lean b/Archive/Imo/Imo2024Q6.lean index d205d0771ba86..12e05f56b7108 100644 --- a/Archive/Imo/Imo2024Q6.lean +++ b/Archive/Imo/Imo2024Q6.lean @@ -62,9 +62,9 @@ lemma Aquaesulian.apply_zero : f 0 = 0 := by @[simp] lemma Aquaesulian.apply_neg_apply_add (x : G) : f (-(f x)) + x = 0 := by rcases h x (-(f x)) with hc | hc - · rw [add_right_neg, ← h.apply_zero] at hc + · rw [add_neg_cancel, ← h.apply_zero] at hc exact h.injective hc - · rw [add_right_neg, h.apply_zero] at hc + · rw [add_neg_cancel, h.apply_zero] at hc exact hc.symm @[simp] diff --git a/Archive/Wiedijk100Theorems/AreaOfACircle.lean b/Archive/Wiedijk100Theorems/AreaOfACircle.lean index e443becaa2f8c..74411920d93f5 100644 --- a/Archive/Wiedijk100Theorems/AreaOfACircle.lean +++ b/Archive/Wiedijk100Theorems/AreaOfACircle.lean @@ -113,12 +113,12 @@ theorem area_disc : volume (disc r) = NNReal.pi * r ^ 2 := by ring · suffices -(1 : ℝ) < (r : ℝ)⁻¹ * x by exact this.ne' calc - -(1 : ℝ) = (r : ℝ)⁻¹ * -r := by simp [inv_mul_cancel hlt.ne'] + -(1 : ℝ) = (r : ℝ)⁻¹ * -r := by simp [inv_mul_cancel₀ hlt.ne'] _ < (r : ℝ)⁻¹ * x := by nlinarith [inv_pos.mpr hlt] · suffices (r : ℝ)⁻¹ * x < 1 by exact this.ne calc (r : ℝ)⁻¹ * x < (r : ℝ)⁻¹ * r := by nlinarith [inv_pos.mpr hlt] - _ = 1 := inv_mul_cancel hlt.ne' + _ = 1 := inv_mul_cancel₀ hlt.ne' · nlinarith have hcont : ContinuousOn F (Icc (-r) r) := (by continuity : Continuous F).continuousOn calc @@ -126,6 +126,6 @@ theorem area_disc : volume (disc r) = NNReal.pi * r ^ 2 := by integral_eq_sub_of_hasDerivAt_of_le (neg_le_self r.2) hcont hderiv (continuous_const.mul hf).continuousOn.intervalIntegrable _ = NNReal.pi * (r : ℝ) ^ 2 := by - norm_num [F, inv_mul_cancel hlt.ne', ← mul_div_assoc, mul_comm π] + norm_num [F, inv_mul_cancel₀ hlt.ne', ← mul_div_assoc, mul_comm π] end Theorems100 diff --git a/Archive/Wiedijk100Theorems/BuffonsNeedle.lean b/Archive/Wiedijk100Theorems/BuffonsNeedle.lean index 3017493e369aa..062587554806b 100644 --- a/Archive/Wiedijk100Theorems/BuffonsNeedle.lean +++ b/Archive/Wiedijk100Theorems/BuffonsNeedle.lean @@ -285,7 +285,7 @@ lemma integral_min_eq_two_mul : rw [← intervalIntegral.integral_add_adjacent_intervals (b := π / 2) (c := π)] conv => lhs; arg 2; arg 1; intro θ; rw [← neg_neg θ, Real.sin_neg] · simp_rw [intervalIntegral.integral_comp_neg fun θ => min d (-θ.sin * l), ← Real.sin_add_pi, - intervalIntegral.integral_comp_add_right (fun θ => min d (θ.sin * l)), add_left_neg, + intervalIntegral.integral_comp_add_right (fun θ => min d (θ.sin * l)), neg_add_cancel, (by ring : -(π / 2) + π = π / 2), two_mul] all_goals exact intervalIntegrable_min_const_sin_mul d l _ _ diff --git a/Counterexamples/SorgenfreyLine.lean b/Counterexamples/SorgenfreyLine.lean index 53bcfb79bbf6b..6015b4d98307e 100644 --- a/Counterexamples/SorgenfreyLine.lean +++ b/Counterexamples/SorgenfreyLine.lean @@ -278,7 +278,7 @@ theorem not_separatedNhds_rat_irrational_antidiag : `Ico x (x + k⁻¹) ×ˢ Ico (-x) (-x + k⁻¹) ⊆ V`. -/ have : ∀ x : ℝₗ, Irrational (toReal x) → ∃ k : ℕ+, Ico x (x + (k : ℝₗ)⁻¹) ×ˢ Ico (-x) (-x + (k : ℝₗ)⁻¹) ⊆ V := fun x hx ↦ by - have hV : V ∈ 𝓝 (x, -x) := Vo.mem_nhds (@TV (x, -x) ⟨add_neg_self x, hx⟩) + have hV : V ∈ 𝓝 (x, -x) := Vo.mem_nhds (@TV (x, -x) ⟨add_neg_cancel x, hx⟩) exact (nhds_prod_antitone_basis_inv_pnat _ _).mem_iff.1 hV choose! k hkV using this /- Since the set of irrational numbers is a dense Gδ set in the usual topology of `ℝ`, there @@ -294,7 +294,7 @@ theorem not_separatedNhds_rat_irrational_antidiag : /- Choose a rational number `r` in the interior of the closure of `C N`, then choose `n ≥ N > 0` such that `Ico r (r + n⁻¹) × Ico (-r) (-r + n⁻¹) ⊆ U`. -/ rcases Rat.denseRange_cast.exists_mem_open isOpen_interior hN with ⟨r, hr⟩ - have hrU : ((r, -r) : ℝₗ × ℝₗ) ∈ U := @SU (r, -r) ⟨add_neg_self _, r, rfl⟩ + have hrU : ((r, -r) : ℝₗ × ℝₗ) ∈ U := @SU (r, -r) ⟨add_neg_cancel _, r, rfl⟩ obtain ⟨n, hnN, hn⟩ : ∃ n, N ≤ n ∧ Ico (r : ℝₗ) (r + (n : ℝₗ)⁻¹) ×ˢ Ico (-r : ℝₗ) (-r + (n : ℝₗ)⁻¹) ⊆ U := ((nhds_prod_antitone_basis_inv_pnat _ _).hasBasis_ge N).mem_iff.1 (Uo.mem_nhds hrU) diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index 4247a76c97d39..c0ef691f61fae 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -146,7 +146,7 @@ theorem vadd_vsub_eq_sub_vsub (g : G) (p q : P) : g +ᵥ p -ᵥ q = g - (q -ᵥ as subtracting the points and subtracting that group element. -/ theorem vsub_vadd_eq_vsub_sub (p₁ p₂ : P) (g : G) : p₁ -ᵥ (g +ᵥ p₂) = p₁ -ᵥ p₂ - g := by rw [← add_right_inj (p₂ -ᵥ p₁ : G), vsub_add_vsub_cancel, ← neg_vsub_eq_vsub_rev, vadd_vsub, ← - add_sub_assoc, ← neg_vsub_eq_vsub_rev, neg_add_self, zero_sub] + add_sub_assoc, ← neg_vsub_eq_vsub_rev, neg_add_cancel, zero_sub] /-- Cancellation subtracting the results of two subtractions. -/ @[simp] diff --git a/Mathlib/Algebra/Algebra/Equiv.lean b/Mathlib/Algebra/Algebra/Equiv.lean index 6b75f47d33700..ce14bc0b72a3c 100644 --- a/Mathlib/Algebra/Algebra/Equiv.lean +++ b/Mathlib/Algebra/Algebra/Equiv.lean @@ -604,7 +604,7 @@ instance aut : Group (A₁ ≃ₐ[R] A₁) where one_mul ϕ := ext fun x => rfl mul_one ϕ := ext fun x => rfl inv := symm - mul_left_inv ϕ := ext <| symm_apply_apply ϕ + inv_mul_cancel ϕ := ext <| symm_apply_apply ϕ theorem aut_mul (ϕ ψ : A₁ ≃ₐ[R] A₁) : ϕ * ψ = ψ.trans ϕ := rfl @@ -714,8 +714,8 @@ def algHomUnitsEquiv (R S : Type*) [CommSemiring R] [Semiring S] [Algebra R S] : toFun := fun f ↦ { (f : S →ₐ[R] S) with invFun := ↑(f⁻¹) - left_inv := (fun x ↦ show (↑(f⁻¹ * f) : S →ₐ[R] S) x = x by rw [inv_mul_self]; rfl) - right_inv := (fun x ↦ show (↑(f * f⁻¹) : S →ₐ[R] S) x = x by rw [mul_inv_self]; rfl) } + left_inv := (fun x ↦ show (↑(f⁻¹ * f) : S →ₐ[R] S) x = x by rw [inv_mul_cancel]; rfl) + right_inv := (fun x ↦ show (↑(f * f⁻¹) : S →ₐ[R] S) x = x by rw [mul_inv_cancel]; rfl) } invFun := fun f ↦ ⟨f, f.symm, f.comp_symm, f.symm_comp⟩ left_inv := fun _ ↦ rfl right_inv := fun _ ↦ rfl diff --git a/Mathlib/Algebra/BigOperators/Fin.lean b/Mathlib/Algebra/BigOperators/Fin.lean index 328d5d841314d..b544301f19d47 100644 --- a/Mathlib/Algebra/BigOperators/Fin.lean +++ b/Mathlib/Algebra/BigOperators/Fin.lean @@ -220,7 +220,7 @@ theorem partialProd_right_inv {G : Type*} [Group G] (f : Fin n → G) (i : Fin n simp only [partialProd_succ, mul_inv_rev, Fin.castSucc_mk] -- Porting note: was -- assoc_rw [hi, inv_mul_cancel_left] - rw [← mul_assoc, mul_left_eq_self, mul_assoc, hi, mul_left_inv] + rw [← mul_assoc, mul_left_eq_self, mul_assoc, hi, inv_mul_cancel] /-- Let `(g₀, g₁, ..., gₙ)` be a tuple of elements in `Gⁿ⁺¹`. Then if `k < j`, this says `(g₀g₁...gₖ₋₁)⁻¹ * g₀g₁...gₖ = gₖ`. diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 6b3d964510bb8..0251d89f06740 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -435,7 +435,7 @@ lemma prod_rotate_eq_one_of_prod_eq_one : have : n % List.length (a :: l) ≤ List.length (a :: l) := le_of_lt (Nat.mod_lt _ (by simp)) rw [← List.take_append_drop (n % List.length (a :: l)) (a :: l)] at hl rw [← rotate_mod, rotate_eq_drop_append_take this, List.prod_append, mul_eq_one_iff_inv_eq, - ← one_mul (List.prod _)⁻¹, ← hl, List.prod_append, mul_assoc, mul_inv_self, mul_one] + ← one_mul (List.prod _)⁻¹, ← hl, List.prod_append, mul_assoc, mul_inv_cancel, mul_one] end Group diff --git a/Mathlib/Algebra/Category/Grp/Colimits.lean b/Mathlib/Algebra/Category/Grp/Colimits.lean index aff94e095f625..68b7e009d27e5 100644 --- a/Mathlib/Algebra/Category/Grp/Colimits.lean +++ b/Mathlib/Algebra/Category/Grp/Colimits.lean @@ -82,7 +82,7 @@ inductive Relation : Prequotient.{w} F → Prequotient.{w} F → Prop -- And one relation per axiom | zero_add : ∀ x, Relation (add zero x) x | add_zero : ∀ x, Relation (add x zero) x - | add_left_neg : ∀ x, Relation (add (neg x) x) zero + | neg_add_cancel : ∀ x, Relation (add (neg x) x) zero | add_comm : ∀ x y, Relation (add x y) (add y x) | add_assoc : ∀ x y z, Relation (add (add x y) z) (add x (add y z)) @@ -113,7 +113,7 @@ instance : Add (ColimitType.{w} F) where instance : AddCommGroup (ColimitType.{w} F) where zero_add := Quotient.ind <| fun _ => Quotient.sound <| Relation.zero_add _ add_zero := Quotient.ind <| fun _ => Quotient.sound <| Relation.add_zero _ - add_left_neg := Quotient.ind <| fun _ => Quotient.sound <| Relation.add_left_neg _ + neg_add_cancel := Quotient.ind <| fun _ => Quotient.sound <| Relation.neg_add_cancel _ add_comm := Quotient.ind₂ <| fun _ _ => Quotient.sound <| Relation.add_comm _ _ add_assoc := Quotient.ind <| fun _ => Quotient.ind₂ <| fun _ _ => Quotient.sound <| Relation.add_assoc _ _ _ @@ -201,7 +201,7 @@ def descFun (s : Cocone F) : ColimitType.{w} F → s.pt := by | add_2 _ _ _ _ r_ih => dsimp; rw [r_ih] | zero_add => dsimp; rw [zero_add] | add_zero => dsimp; rw [add_zero] - | add_left_neg => dsimp; rw [add_left_neg] + | neg_add_cancel => dsimp; rw [neg_add_cancel] | add_comm => dsimp; rw [add_comm] | add_assoc => dsimp; rw [add_assoc] diff --git a/Mathlib/Algebra/Category/Grp/EpiMono.lean b/Mathlib/Algebra/Category/Grp/EpiMono.lean index 4ead5efd13225..31c09a8f4a5fe 100644 --- a/Mathlib/Algebra/Category/Grp/EpiMono.lean +++ b/Mathlib/Algebra/Category/Grp/EpiMono.lean @@ -186,10 +186,10 @@ def g : B →* SX' where invFun := fun x => β⁻¹ • x left_inv := fun x => by dsimp only - rw [← mul_smul, mul_left_inv, one_smul] + rw [← mul_smul, inv_mul_cancel, one_smul] right_inv := fun x => by dsimp only - rw [← mul_smul, mul_right_inv, one_smul] } + rw [← mul_smul, mul_inv_cancel, one_smul] } map_one' := by ext simp [one_smul] @@ -253,7 +253,7 @@ theorem h_apply_fromCoset_nin_range (x : B) (hx : x ∈ f.range) (b : B) (hb : b simp only [g_apply_fromCoset, leftCoset_assoc] refine Equiv.swap_apply_of_ne_of_ne (fromCoset_ne_of_nin_range _ fun r => hb ?_) (by simp) convert Subgroup.mul_mem _ (Subgroup.inv_mem _ hx) r - rw [← mul_assoc, mul_left_inv, one_mul] + rw [← mul_assoc, inv_mul_cancel, one_mul] theorem agree : f.range = { x | h x = g x } := by refine Set.ext fun b => ⟨?_, fun hb : h b = g b => by_contradiction fun r => ?_⟩ diff --git a/Mathlib/Algebra/Category/Grp/FilteredColimits.lean b/Mathlib/Algebra/Category/Grp/FilteredColimits.lean index 2fe1025945d7f..fe2f4bcd125da 100644 --- a/Mathlib/Algebra/Category/Grp/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/Grp/FilteredColimits.lean @@ -92,14 +92,14 @@ theorem colimit_inv_mk_eq (x : Σ j, F.obj j) : (G.mk.{v, u} F x)⁻¹ = G.mk F @[to_additive] noncomputable instance colimitGroup : Group (G.{v, u} F) := { colimitInv.{v, u} F, (G.{v, u} F).str with - mul_left_inv := fun x => by + inv_mul_cancel := fun x => by refine Quot.inductionOn x ?_; clear x; intro x cases' x with j x erw [colimit_inv_mk_eq, colimit_mul_mk_eq (F ⋙ forget₂ Grp MonCat.{max v u}) ⟨j, _⟩ ⟨j, _⟩ j (𝟙 j) (𝟙 j), colimit_one_eq (F ⋙ forget₂ Grp MonCat.{max v u}) j] dsimp - erw [CategoryTheory.Functor.map_id, mul_left_inv] } + erw [CategoryTheory.Functor.map_id, inv_mul_cancel] } /-- The bundled group giving the filtered colimit of a diagram. -/ @[to_additive "The bundled additive group giving the filtered colimit of a diagram."] diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean index 87fea38ec949b..622c9e6a838a6 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean @@ -177,7 +177,7 @@ lemma neg_app (f : P ⟶ Q) (X : Cᵒᵖ) : (-f).app X = -f.app X := rfl instance : AddCommGroup (P ⟶ Q) where add_assoc := by intros; ext1; simp only [add_app, add_assoc] zero_add := by intros; ext1; simp only [add_app, zero_app, zero_add] - add_left_neg := by intros; ext1; simp only [add_app, neg_app, add_left_neg, zero_app] + neg_add_cancel := by intros; ext1; simp only [add_app, neg_app, neg_add_cancel, zero_app] add_zero := by intros; ext1; simp only [add_app, zero_app, add_zero] add_comm := by intros; ext1; simp only [add_app]; apply add_comm sub_eq_add_neg := by intros; ext1; simp only [add_app, sub_app, neg_app, sub_eq_add_neg] diff --git a/Mathlib/Algebra/Category/Ring/Colimits.lean b/Mathlib/Algebra/Category/Ring/Colimits.lean index 01e11c4fde83c..5d7cc51769fd3 100644 --- a/Mathlib/Algebra/Category/Ring/Colimits.lean +++ b/Mathlib/Algebra/Category/Ring/Colimits.lean @@ -84,7 +84,7 @@ inductive Relation : Prequotient F → Prequotient F → Prop -- Make it an equi | add_zero : ∀ x, Relation (add x zero) x | one_mul : ∀ x, Relation (mul one x) x | mul_one : ∀ x, Relation (mul x one) x - | add_left_neg : ∀ x, Relation (add (neg x) x) zero + | neg_add_cancel : ∀ x, Relation (add (neg x) x) zero | add_comm : ∀ x y, Relation (add x y) (add y x) | add_assoc : ∀ x y z, Relation (add (add x y) z) (add x (add y z)) | mul_assoc : ∀ x y z, Relation (mul (mul x y) z) (mul x (mul y z)) @@ -119,7 +119,7 @@ instance ColimitType.AddGroup : AddGroup (ColimitType F) where neg := Quotient.map neg Relation.neg_1 zero_add := Quotient.ind <| fun _ => Quotient.sound <| Relation.zero_add _ add_zero := Quotient.ind <| fun _ => Quotient.sound <| Relation.add_zero _ - add_left_neg := Quotient.ind <| fun _ => Quotient.sound <| Relation.add_left_neg _ + neg_add_cancel := Quotient.ind <| fun _ => Quotient.sound <| Relation.neg_add_cancel _ add_assoc := Quotient.ind <| fun _ => Quotient.ind₂ <| fun _ _ => Quotient.sound <| Relation.add_assoc _ _ _ nsmul := nsmulRec @@ -250,7 +250,7 @@ def descFun (s : Cocone F) : ColimitType F → s.pt := by | add_zero x => dsimp; rw [add_zero] | one_mul x => dsimp; rw [one_mul] | mul_one x => dsimp; rw [mul_one] - | add_left_neg x => dsimp; rw [add_left_neg] + | neg_add_cancel x => dsimp; rw [neg_add_cancel] | add_comm x y => dsimp; rw [add_comm] | add_assoc x y z => dsimp; rw [add_assoc] | mul_assoc x y z => dsimp; rw [mul_assoc] @@ -390,7 +390,7 @@ inductive Relation : Prequotient F → Prequotient F → Prop -- Make it an equi | add_zero : ∀ x, Relation (add x zero) x | one_mul : ∀ x, Relation (mul one x) x | mul_one : ∀ x, Relation (mul x one) x - | add_left_neg : ∀ x, Relation (add (neg x) x) zero + | neg_add_cancel : ∀ x, Relation (add (neg x) x) zero | add_comm : ∀ x y, Relation (add x y) (add y x) | mul_comm : ∀ x y, Relation (mul x y) (mul y x) | add_assoc : ∀ x y z, Relation (add (add x y) z) (add x (add y z)) @@ -426,7 +426,7 @@ instance ColimitType.AddGroup : AddGroup (ColimitType F) where neg := Quotient.map neg Relation.neg_1 zero_add := Quotient.ind <| fun _ => Quotient.sound <| Relation.zero_add _ add_zero := Quotient.ind <| fun _ => Quotient.sound <| Relation.add_zero _ - add_left_neg := Quotient.ind <| fun _ => Quotient.sound <| Relation.add_left_neg _ + neg_add_cancel := Quotient.ind <| fun _ => Quotient.sound <| Relation.neg_add_cancel _ add_assoc := Quotient.ind <| fun _ => Quotient.ind₂ <| fun _ _ => Quotient.sound <| Relation.add_assoc _ _ _ nsmul := nsmulRec @@ -558,7 +558,7 @@ def descFun (s : Cocone F) : ColimitType F → s.pt := by | add_zero x => dsimp; rw [add_zero] | one_mul x => dsimp; rw [one_mul] | mul_one x => dsimp; rw [mul_one] - | add_left_neg x => dsimp; rw [add_left_neg] + | neg_add_cancel x => dsimp; rw [neg_add_cancel] | add_comm x y => dsimp; rw [add_comm] | mul_comm x y => dsimp; rw [mul_comm] | add_assoc x y z => dsimp; rw [add_assoc] diff --git a/Mathlib/Algebra/CharP/Basic.lean b/Mathlib/Algebra/CharP/Basic.lean index ce2ecc6aef10e..af8e256be0815 100644 --- a/Mathlib/Algebra/CharP/Basic.lean +++ b/Mathlib/Algebra/CharP/Basic.lean @@ -132,14 +132,14 @@ theorem CharP.neg_one_pow_char [Ring R] (p : ℕ) [CharP R p] [Fact p.Prime] : (-1 : R) ^ p = -1 := by rw [eq_neg_iff_add_eq_zero] nth_rw 2 [← one_pow p] - rw [← add_pow_char_of_commute R _ _ (Commute.one_right _), add_left_neg, + rw [← add_pow_char_of_commute R _ _ (Commute.one_right _), neg_add_cancel, zero_pow (Fact.out (p := Nat.Prime p)).ne_zero] theorem CharP.neg_one_pow_char_pow [Ring R] (p n : ℕ) [CharP R p] [Fact p.Prime] : (-1 : R) ^ p ^ n = -1 := by rw [eq_neg_iff_add_eq_zero] nth_rw 2 [← one_pow (p ^ n)] - rw [← add_pow_char_pow_of_commute R _ _ (Commute.one_right _), add_left_neg, + rw [← add_pow_char_pow_of_commute R _ _ (Commute.one_right _), neg_add_cancel, zero_pow (pow_ne_zero _ (Fact.out (p := Nat.Prime p)).ne_zero)] namespace CharP diff --git a/Mathlib/Algebra/CharZero/Quotient.lean b/Mathlib/Algebra/CharZero/Quotient.lean index 9d2cc562a162e..309850f089606 100644 --- a/Mathlib/Algebra/CharZero/Quotient.lean +++ b/Mathlib/Algebra/CharZero/Quotient.lean @@ -22,7 +22,7 @@ theorem zsmul_mem_zmultiples_iff_exists_sub_div {r : R} {z : ℤ} (hz : z ≠ 0) simp_rw [AddSubgroup.mem_zmultiples_iff, div_eq_mul_inv, ← smul_mul_assoc, eq_sub_iff_add_eq] have hz' : (z : R) ≠ 0 := Int.cast_ne_zero.mpr hz conv_rhs => simp (config := { singlePass := true }) only [← (mul_right_injective₀ hz').eq_iff] - simp_rw [← zsmul_eq_mul, smul_add, ← mul_smul_comm, zsmul_eq_mul (z : R)⁻¹, mul_inv_cancel hz', + simp_rw [← zsmul_eq_mul, smul_add, ← mul_smul_comm, zsmul_eq_mul (z : R)⁻¹, mul_inv_cancel₀ hz', mul_one, ← natCast_zsmul, smul_smul, ← add_smul] constructor · rintro ⟨k, h⟩ diff --git a/Mathlib/Algebra/DirectLimit.lean b/Mathlib/Algebra/DirectLimit.lean index 29b9817a14331..993dcaec82950 100644 --- a/Mathlib/Algebra/DirectLimit.lean +++ b/Mathlib/Algebra/DirectLimit.lean @@ -948,7 +948,7 @@ theorem exists_inv {p : Ring.DirectLimit G f} : p ≠ 0 → ∃ y, p * y = 1 := Ring.DirectLimit.induction_on p fun i x H => ⟨Ring.DirectLimit.of G f i x⁻¹, by erw [← (Ring.DirectLimit.of _ _ _).map_mul, - mul_inv_cancel fun h : x = 0 => H <| by rw [h, (Ring.DirectLimit.of _ _ _).map_zero], + mul_inv_cancel₀ fun h : x = 0 => H <| by rw [h, (Ring.DirectLimit.of _ _ _).map_zero], (Ring.DirectLimit.of _ _ _).map_one]⟩ section diff --git a/Mathlib/Algebra/Field/IsField.lean b/Mathlib/Algebra/Field/IsField.lean index 509721d425184..a3d3847cbc41f 100644 --- a/Mathlib/Algebra/Field/IsField.lean +++ b/Mathlib/Algebra/Field/IsField.lean @@ -37,7 +37,7 @@ structure IsField (R : Type u) [Semiring R] : Prop where /-- Transferring from `Semifield` to `IsField`. -/ theorem Semifield.toIsField (R : Type u) [Semifield R] : IsField R where __ := ‹Semifield R› - mul_inv_cancel {a} ha := ⟨a⁻¹, mul_inv_cancel ha⟩ + mul_inv_cancel {a} ha := ⟨a⁻¹, mul_inv_cancel₀ ha⟩ /-- Transferring from `Field` to `IsField`. -/ theorem Field.toIsField (R : Type u) [Field R] : IsField R := diff --git a/Mathlib/Algebra/Field/MinimalAxioms.lean b/Mathlib/Algebra/Field/MinimalAxioms.lean index fec439419f97c..b80f40d16d801 100644 --- a/Mathlib/Algebra/Field/MinimalAxioms.lean +++ b/Mathlib/Algebra/Field/MinimalAxioms.lean @@ -28,7 +28,7 @@ abbrev Field.ofMinimalAxioms (K : Type u) [Add K] [Mul K] [Neg K] [Inv K] [Zero K] [One K] (add_assoc : ∀ a b c : K, a + b + c = a + (b + c)) (zero_add : ∀ a : K, 0 + a = a) - (add_left_neg : ∀ a : K, -a + a = 0) + (neg_add_cancel : ∀ a : K, -a + a = 0) (mul_assoc : ∀ a b c : K, a * b * c = a * (b * c)) (mul_comm : ∀ a b : K, a * b = b * a) (one_mul : ∀ a : K, 1 * a = a) @@ -37,7 +37,7 @@ abbrev Field.ofMinimalAxioms (K : Type u) (left_distrib : ∀ a b c : K, a * (b + c) = a * b + a * c) (exists_pair_ne : ∃ x y : K, x ≠ y) : Field K := letI := CommRing.ofMinimalAxioms add_assoc zero_add - add_left_neg mul_assoc mul_comm one_mul left_distrib + neg_add_cancel mul_assoc mul_comm one_mul left_distrib { exists_pair_ne := exists_pair_ne mul_inv_cancel := mul_inv_cancel inv_zero := inv_zero diff --git a/Mathlib/Algebra/Field/Subfield.lean b/Mathlib/Algebra/Field/Subfield.lean index 49f8d0abafa74..b304f131e2f96 100644 --- a/Mathlib/Algebra/Field/Subfield.lean +++ b/Mathlib/Algebra/Field/Subfield.lean @@ -616,7 +616,7 @@ theorem closure_induction {s : Set K} {p : K → Prop} {x} (h : x ∈ closure s) (inv : ∀ x, p x → p x⁻¹) (mul : ∀ x y, p x → p y → p (x * y)) : p x := by letI : Subfield K := ⟨⟨⟨⟨⟨p, by intro _ _; exact mul _ _⟩, one⟩, - by intro _ _; exact add _ _, @add_neg_self K _ 1 ▸ add _ _ one (neg _ one)⟩, + by intro _ _; exact add _ _, @add_neg_cancel K _ 1 ▸ add _ _ one (neg _ one)⟩, by intro _; exact neg _⟩, inv⟩ exact (closure_le (t := this)).2 mem h diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 3232d89dff8f1..bce4d36d459b2 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -123,7 +123,7 @@ theorem neg_one_geom_sum [Ring α] {n : ℕ} : simp only [geom_sum_succ', Nat.even_add_one, hk] split_ifs with h · rw [h.neg_one_pow, add_zero] - · rw [(Nat.odd_iff_not_even.2 h).neg_one_pow, neg_add_self] + · rw [(Nat.odd_iff_not_even.2 h).neg_one_pow, neg_add_cancel] theorem geom_sum₂_self {α : Type*} [CommRing α] (x : α) (n : ℕ) : ∑ i ∈ range n, x ^ i * x ^ (n - 1 - i) = n * x ^ (n - 1) := @@ -355,11 +355,11 @@ theorem geom_sum_inv [DivisionRing α] {x : α} (hx1 : x ≠ 1) (hx0 : x ≠ 0) have h₃ : x - 1 ≠ 0 := mt sub_eq_zero.1 hx1 have h₄ : x * (x ^ n)⁻¹ = (x ^ n)⁻¹ * x := Nat.recOn n (by simp) fun n h => by - rw [pow_succ', mul_inv_rev, ← mul_assoc, h, mul_assoc, mul_inv_cancel hx0, mul_assoc, - inv_mul_cancel hx0] + rw [pow_succ', mul_inv_rev, ← mul_assoc, h, mul_assoc, mul_inv_cancel₀ hx0, mul_assoc, + inv_mul_cancel₀ hx0] rw [geom_sum_eq h₁, div_eq_iff_mul_eq h₂, ← mul_right_inj' h₃, ← mul_assoc, ← mul_assoc, - mul_inv_cancel h₃] - simp [mul_add, add_mul, mul_inv_cancel hx0, mul_assoc, h₄, sub_eq_add_neg, add_comm, + mul_inv_cancel₀ h₃] + simp [mul_add, add_mul, mul_inv_cancel₀ hx0, mul_assoc, h₄, sub_eq_add_neg, add_comm, add_left_comm] rw [add_comm _ (-x), add_assoc, add_assoc _ _ 1] diff --git a/Mathlib/Algebra/Group/Action/Defs.lean b/Mathlib/Algebra/Group/Action/Defs.lean index b96c5a818e280..b0519ca6ce79c 100644 --- a/Mathlib/Algebra/Group/Action/Defs.lean +++ b/Mathlib/Algebra/Group/Action/Defs.lean @@ -464,10 +464,10 @@ section Group variable [Group G] [MulAction G α] {g : G} {a b : α} @[to_additive (attr := simp)] -lemma inv_smul_smul (g : G) (a : α) : g⁻¹ • g • a = a := by rw [smul_smul, mul_left_inv, one_smul] +lemma inv_smul_smul (g : G) (a : α) : g⁻¹ • g • a = a := by rw [smul_smul, inv_mul_cancel, one_smul] @[to_additive (attr := simp)] -lemma smul_inv_smul (g : G) (a : α) : g • g⁻¹ • a = a := by rw [smul_smul, mul_right_inv, one_smul] +lemma smul_inv_smul (g : G) (a : α) : g • g⁻¹ • a = a := by rw [smul_smul, mul_inv_cancel, one_smul] @[to_additive] lemma inv_smul_eq_iff : g⁻¹ • a = b ↔ a = g • b := ⟨fun h ↦ by rw [← h, smul_inv_smul], fun h ↦ by rw [h, inv_smul_smul]⟩ @@ -489,7 +489,7 @@ end Mul variable [Group H] [MulAction G H] [SMulCommClass G H H] [IsScalarTower G H H] lemma smul_inv (g : G) (a : H) : (g • a)⁻¹ = g⁻¹ • a⁻¹ := - inv_eq_of_mul_eq_one_right $ by rw [smul_mul_smul, mul_right_inv, mul_right_inv, one_smul] + inv_eq_of_mul_eq_one_right $ by rw [smul_mul_smul, mul_inv_cancel, mul_inv_cancel, one_smul] lemma smul_zpow (g : G) (a : H) (n : ℤ) : (g • a) ^ n = g ^ n • a ^ n := by cases n <;> simp [smul_pow, smul_inv] diff --git a/Mathlib/Algebra/Group/Action/Units.lean b/Mathlib/Algebra/Group/Action/Units.lean index 9b41201135138..d7ad6e43cfdf4 100644 --- a/Mathlib/Algebra/Group/Action/Units.lean +++ b/Mathlib/Algebra/Group/Action/Units.lean @@ -69,8 +69,8 @@ instance mulAction' [Group G] [Monoid M] [MulAction G M] [SMulCommClass G M M] [IsScalarTower G M M] : MulAction G Mˣ where smul g m := ⟨g • (m : M), (g⁻¹ • ((m⁻¹ : Mˣ) : M)), - by rw [smul_mul_smul, Units.mul_inv, mul_right_inv, one_smul], - by rw [smul_mul_smul, Units.inv_mul, mul_left_inv, one_smul]⟩ + by rw [smul_mul_smul, Units.mul_inv, mul_inv_cancel, one_smul], + by rw [smul_mul_smul, Units.inv_mul, inv_mul_cancel, one_smul]⟩ one_smul m := Units.ext <| one_smul _ _ mul_smul g₁ g₂ m := Units.ext <| mul_smul _ _ _ diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index 7ac9a514409f1..a8bb7908842b0 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -271,7 +271,7 @@ inversion operation for the definition (but see `AddChar.map_neg_eq_inv` below). instance instCommGroup : CommGroup (AddChar A M) := { instCommMonoid with inv := fun ψ ↦ ψ.compAddMonoidHom negAddMonoidHom - mul_left_inv := fun ψ ↦ by ext1 x; simp [negAddMonoidHom, ← map_add_eq_mul]} + inv_mul_cancel := fun ψ ↦ by ext1 x; simp [negAddMonoidHom, ← map_add_eq_mul]} @[simp] lemma inv_apply (ψ : AddChar A M) (x : A) : ψ⁻¹ x = ψ (-x) := rfl @@ -292,7 +292,7 @@ variable {A M : Type*} [AddGroup A] [DivisionMonoid M] /-- An additive character maps negatives to inverses (when defined) -/ lemma map_neg_eq_inv (ψ : AddChar A M) (a : A) : ψ (-a) = (ψ a)⁻¹ := by apply eq_inv_of_mul_eq_one_left - simp only [← map_add_eq_mul, add_left_neg, map_zero_eq_one] + simp only [← map_add_eq_mul, neg_add_cancel, map_zero_eq_one] /-- An additive character maps integer scalar multiples to integer powers. -/ lemma map_zsmul_eq_zpow (ψ : AddChar A M) (n : ℤ) (a : A) : ψ (n • a) = (ψ a) ^ n := diff --git a/Mathlib/Algebra/Group/Aut.lean b/Mathlib/Algebra/Group/Aut.lean index ff3f00ba4f303..5f7cba0d26ab5 100644 --- a/Mathlib/Algebra/Group/Aut.lean +++ b/Mathlib/Algebra/Group/Aut.lean @@ -55,7 +55,7 @@ instance : Group (MulAut M) where mul_assoc _ _ _ := rfl one_mul _ := rfl mul_one _ := rfl - mul_left_inv := MulEquiv.self_trans_symm + inv_mul_cancel := MulEquiv.self_trans_symm instance : Inhabited (MulAut M) := ⟨1⟩ @@ -125,8 +125,8 @@ def conj [Group G] : G →* MulAut G where toFun g := { toFun := fun h => g * h * g⁻¹ invFun := fun h => g⁻¹ * h * g - left_inv := fun _ => by simp only [mul_assoc, inv_mul_cancel_left, mul_left_inv, mul_one] - right_inv := fun _ => by simp only [mul_assoc, mul_inv_cancel_left, mul_right_inv, mul_one] + left_inv := fun _ => by simp only [mul_assoc, inv_mul_cancel_left, inv_mul_cancel, mul_one] + right_inv := fun _ => by simp only [mul_assoc, mul_inv_cancel_left, mul_inv_cancel, mul_one] map_mul' := by simp only [mul_assoc, inv_mul_cancel_left, forall_const] } map_mul' g₁ g₂ := by ext h @@ -162,7 +162,7 @@ instance group : Group (AddAut A) where mul_assoc _ _ _ := rfl one_mul _ := rfl mul_one _ := rfl - mul_left_inv := AddEquiv.self_trans_symm + inv_mul_cancel := AddEquiv.self_trans_symm instance : Inhabited (AddAut A) := ⟨1⟩ @@ -233,8 +233,10 @@ def conj [AddGroup G] : G →+ Additive (AddAut G) where { toFun := fun h => g + h + -g -- this definition is chosen to match `MulAut.conj` invFun := fun h => -g + h + g - left_inv := fun _ => by simp only [add_assoc, neg_add_cancel_left, add_left_neg, add_zero] - right_inv := fun _ => by simp only [add_assoc, add_neg_cancel_left, add_right_neg, add_zero] + left_inv := fun _ => by + simp only [add_assoc, neg_add_cancel_left, neg_add_cancel, add_zero] + right_inv := fun _ => by + simp only [add_assoc, add_neg_cancel_left, add_neg_cancel, add_zero] map_add' := by simp only [add_assoc, neg_add_cancel_left, forall_const] } map_add' g₁ g₂ := by apply Additive.toMul.injective; ext h diff --git a/Mathlib/Algebra/Group/Basic.lean b/Mathlib/Algebra/Group/Basic.lean index 7f5c25f4e3dfa..9ebc3308c99d3 100644 --- a/Mathlib/Algebra/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Basic.lean @@ -636,7 +636,7 @@ theorem mul_eq_of_eq_mul_inv (h : a = c * b⁻¹) : a * b = c := by simp [h] @[to_additive] theorem mul_eq_one_iff_eq_inv : a * b = 1 ↔ a = b⁻¹ := - ⟨eq_inv_of_mul_eq_one_left, fun h ↦ by rw [h, mul_left_inv]⟩ + ⟨eq_inv_of_mul_eq_one_left, fun h ↦ by rw [h, inv_mul_cancel]⟩ @[to_additive] theorem mul_eq_one_iff_inv_eq : a * b = 1 ↔ a⁻¹ = b := by @@ -693,7 +693,7 @@ theorem div_mul_cancel (a b : G) : a / b * b = a := by rw [div_eq_mul_inv, inv_mul_cancel_right a b] @[to_additive (attr := simp) sub_self] -theorem div_self' (a : G) : a / a = 1 := by rw [div_eq_mul_inv, mul_right_inv a] +theorem div_self' (a : G) : a / a = 1 := by rw [div_eq_mul_inv, mul_inv_cancel a] @[to_additive (attr := simp)] theorem mul_div_cancel_right (a b : G) : a * b / b = a := by @@ -889,7 +889,7 @@ theorem div_eq_of_eq_mul' {a b c : G} (h : a = b * c) : a / b = c := by @[to_additive (attr := simp)] theorem mul_div_mul_left_eq_div (a b c : G) : c * a / (c * b) = a / b := by rw [div_eq_mul_inv, mul_inv_rev, mul_comm b⁻¹ c⁻¹, mul_comm c a, mul_assoc, ← mul_assoc c, - mul_right_inv, one_mul, div_eq_mul_inv] + mul_inv_cancel, one_mul, div_eq_mul_inv] @[to_additive eq_sub_of_add_eq'] theorem eq_div_of_mul_eq'' (h : c * a = b) : a = b / c := by simp [h.symm] diff --git a/Mathlib/Algebra/Group/Conj.lean b/Mathlib/Algebra/Group/Conj.lean index f73a41e75af25..0be980f48df22 100644 --- a/Mathlib/Algebra/Group/Conj.lean +++ b/Mathlib/Algebra/Group/Conj.lean @@ -81,7 +81,7 @@ variable [Group α] @[simp] theorem isConj_iff {a b : α} : IsConj a b ↔ ∃ c : α, c * a * c⁻¹ = b := ⟨fun ⟨c, hc⟩ => ⟨c, mul_inv_eq_iff_eq_mul.2 hc⟩, fun ⟨c, hc⟩ => - ⟨⟨c, c⁻¹, mul_inv_self c, inv_mul_self c⟩, mul_inv_eq_iff_eq_mul.1 hc⟩⟩ + ⟨⟨c, c⁻¹, mul_inv_cancel c, inv_mul_cancel c⟩, mul_inv_eq_iff_eq_mul.1 hc⟩⟩ -- Porting note: not in simp NF. -- @[simp] diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 5e8035c15123a..f005d732edc45 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -1032,7 +1032,7 @@ Use `Group.ofLeftAxioms` or `Group.ofRightAxioms` to define a group structure on a type with the minumum proof obligations. -/ class Group (G : Type u) extends DivInvMonoid G where - protected mul_left_inv : ∀ a : G, a⁻¹ * a = 1 + protected inv_mul_cancel : ∀ a : G, a⁻¹ * a = 1 /-- An `AddGroup` is an `AddMonoid` with a unary `-` satisfying `-a + a = 0`. @@ -1043,7 +1043,7 @@ Use `AddGroup.ofLeftAxioms` or `AddGroup.ofRightAxioms` to define an additive group structure on a type with the minumum proof obligations. -/ class AddGroup (A : Type u) extends SubNegMonoid A where - protected add_left_neg : ∀ a : A, -a + a = 0 + protected neg_add_cancel : ∀ a : A, -a + a = 0 attribute [to_additive] Group @@ -1052,46 +1052,47 @@ section Group variable [Group G] {a b c : G} @[to_additive (attr := simp)] -theorem mul_left_inv : ∀ a : G, a⁻¹ * a = 1 := - Group.mul_left_inv - -@[to_additive] -theorem inv_mul_self (a : G) : a⁻¹ * a = 1 := - mul_left_inv a +theorem inv_mul_cancel (a : G) : a⁻¹ * a = 1 := + Group.inv_mul_cancel a @[to_additive] private theorem inv_eq_of_mul (h : a * b = 1) : a⁻¹ = b := - left_inv_eq_right_inv (inv_mul_self a) h + left_inv_eq_right_inv (inv_mul_cancel a) h @[to_additive (attr := simp)] -theorem mul_right_inv (a : G) : a * a⁻¹ = 1 := by - rw [← mul_left_inv a⁻¹, inv_eq_of_mul (mul_left_inv a)] - -@[to_additive] -theorem mul_inv_self (a : G) : a * a⁻¹ = 1 := - mul_right_inv a +theorem mul_inv_cancel (a : G) : a * a⁻¹ = 1 := by + rw [← inv_mul_cancel a⁻¹, inv_eq_of_mul (inv_mul_cancel a)] + +@[deprecated (since := "2024-08-12")] alias mul_left_inv := inv_mul_cancel +@[deprecated (since := "2024-08-12")] alias mul_right_inv := mul_inv_cancel +@[deprecated (since := "2024-08-12")] alias add_left_neg := neg_add_cancel +@[deprecated (since := "2024-08-12")] alias add_right_neg := add_neg_cancel +@[deprecated (since := "2024-08-12")] alias inv_mul_self := inv_mul_cancel +@[deprecated (since := "2024-08-12")] alias mul_inv_self := mul_inv_cancel +@[deprecated (since := "2024-08-12")] alias neg_add_self := neg_add_cancel +@[deprecated (since := "2024-08-12")] alias add_right_self := add_neg_cancel @[to_additive (attr := simp)] theorem inv_mul_cancel_left (a b : G) : a⁻¹ * (a * b) = b := by - rw [← mul_assoc, mul_left_inv, one_mul] + rw [← mul_assoc, inv_mul_cancel, one_mul] @[to_additive (attr := simp)] theorem mul_inv_cancel_left (a b : G) : a * (a⁻¹ * b) = b := by - rw [← mul_assoc, mul_right_inv, one_mul] + rw [← mul_assoc, mul_inv_cancel, one_mul] @[to_additive (attr := simp)] theorem mul_inv_cancel_right (a b : G) : a * b * b⁻¹ = a := by - rw [mul_assoc, mul_right_inv, mul_one] + rw [mul_assoc, mul_inv_cancel, mul_one] @[to_additive (attr := simp)] theorem inv_mul_cancel_right (a b : G) : a * b⁻¹ * b = a := by - rw [mul_assoc, mul_left_inv, mul_one] + rw [mul_assoc, inv_mul_cancel, mul_one] @[to_additive AddGroup.toSubtractionMonoid] instance (priority := 100) Group.toDivisionMonoid : DivisionMonoid G := - { inv_inv := fun a ↦ inv_eq_of_mul (mul_left_inv a) + { inv_inv := fun a ↦ inv_eq_of_mul (inv_mul_cancel a) mul_inv_rev := - fun a b ↦ inv_eq_of_mul <| by rw [mul_assoc, mul_inv_cancel_left, mul_right_inv] + fun a b ↦ inv_eq_of_mul <| by rw [mul_assoc, mul_inv_cancel_left, mul_inv_cancel] inv_eq_of_mul := fun _ _ ↦ inv_eq_of_mul } -- see Note [lower instance priority] diff --git a/Mathlib/Algebra/Group/Fin/Basic.lean b/Mathlib/Algebra/Group/Fin/Basic.lean index ec862c8479a6a..27b58e867117f 100644 --- a/Mathlib/Algebra/Group/Fin/Basic.lean +++ b/Mathlib/Algebra/Group/Fin/Basic.lean @@ -50,7 +50,7 @@ instance instAddMonoidWithOne (n) [NeZero n] : AddMonoidWithOne (Fin n) where instance addCommGroup (n : ℕ) [NeZero n] : AddCommGroup (Fin n) where __ := addCommMonoid n __ := neg n - add_left_neg := fun ⟨a, ha⟩ ↦ + neg_add_cancel := fun ⟨a, ha⟩ ↦ Fin.ext <| (Nat.mod_add_mod _ _ _).trans <| by rw [Fin.val_zero', Nat.sub_add_cancel, Nat.mod_self] exact le_of_lt ha diff --git a/Mathlib/Algebra/Group/Hom/Basic.lean b/Mathlib/Algebra/Group/Hom/Basic.lean index de37e28179ce7..e1a97f26ed9b8 100644 --- a/Mathlib/Algebra/Group/Hom/Basic.lean +++ b/Mathlib/Algebra/Group/Hom/Basic.lean @@ -110,7 +110,7 @@ theorem _root_.injective_iff_map_eq_one {G H} [Group G] [MulOneClass H] [FunLike F G H] [MonoidHomClass F G H] (f : F) : Function.Injective f ↔ ∀ a, f a = 1 → a = 1 := ⟨fun h x => (map_eq_one_iff f h).mp, fun h x y hxy => - mul_inv_eq_one.1 <| h _ <| by rw [map_mul, hxy, ← map_mul, mul_inv_self, map_one]⟩ + mul_inv_eq_one.1 <| h _ <| by rw [map_mul, hxy, ← map_mul, mul_inv_cancel, map_one]⟩ /-- A homomorphism from a group to a monoid is injective iff its kernel is trivial, stated as an iff on the triviality of the kernel. @@ -142,7 +142,7 @@ def ofMapMulInv {H : Type*} [Group H] (f : G → H) { simp only [one_mul, inv_one, ← map_div, inv_inv] } _ = f x * f y := by { simp only [map_div] - simp only [mul_right_inv, one_mul, inv_inv] } + simp only [mul_inv_cancel, one_mul, inv_inv] } @[to_additive (attr := simp)] theorem coe_of_map_mul_inv {H : Type*} [Group H] (f : G → H) diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index b62997b508292..a64a2118ac2da 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -395,7 +395,7 @@ lemma map_comp_div' [DivInvMonoid G] [DivInvMonoid H] [MonoidHomClass F G H] (f @[to_additive (attr := simp) "Additive group homomorphisms preserve negation."] theorem map_inv [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (a : G) : f a⁻¹ = (f a)⁻¹ := - eq_inv_of_mul_eq_one_left <| map_mul_eq_one f <| inv_mul_self _ + eq_inv_of_mul_eq_one_left <| map_mul_eq_one f <| inv_mul_cancel _ @[to_additive (attr := simp)] lemma map_comp_inv [Group G] [DivisionMonoid H] [MonoidHomClass F G H] (f : F) (g : ι → G) : diff --git a/Mathlib/Algebra/Group/Hom/Instances.lean b/Mathlib/Algebra/Group/Hom/Instances.lean index 6042f9140ec71..51101f1a4791d 100644 --- a/Mathlib/Algebra/Group/Hom/Instances.lean +++ b/Mathlib/Algebra/Group/Hom/Instances.lean @@ -55,7 +55,7 @@ instance MonoidHom.commGroup {M G} [MulOneClass M] [CommGroup G] : CommGroup (M intros ext apply div_eq_mul_inv, - mul_left_inv := by intros; ext; apply mul_left_inv, + inv_mul_cancel := by intros; ext; apply inv_mul_cancel, zpow := fun n f => { toFun := fun x => f x ^ n, map_one' := by simp, diff --git a/Mathlib/Algebra/Group/InjSurj.lean b/Mathlib/Algebra/Group/InjSurj.lean index 1c5b3b20b480d..4f2c17838f4dc 100644 --- a/Mathlib/Algebra/Group/InjSurj.lean +++ b/Mathlib/Algebra/Group/InjSurj.lean @@ -291,7 +291,7 @@ protected abbrev group [Group M₂] (f : M₁ → M₂) (hf : Injective f) (one (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : Group M₁ := { hf.divInvMonoid f one mul inv div npow zpow with - mul_left_inv := fun x => hf <| by erw [mul, inv, mul_left_inv, one] } + inv_mul_cancel := fun x => hf <| by erw [mul, inv, inv_mul_cancel, one] } /-- A type endowed with `0`, `1` and `+` is an additive group with one, if it admits an injective map that preserves `0`, `1` and `+` to an additive group with one. See note @@ -476,7 +476,7 @@ protected abbrev group [Group M₁] (f : M₁ → M₂) (hf : Surjective f) (one (div : ∀ x y, f (x / y) = f x / f y) (npow : ∀ (x) (n : ℕ), f (x ^ n) = f x ^ n) (zpow : ∀ (x) (n : ℤ), f (x ^ n) = f x ^ n) : Group M₂ := { hf.divInvMonoid f one mul inv div npow zpow with - mul_left_inv := hf.forall.2 fun x => by erw [← inv, ← mul, mul_left_inv, one] } + inv_mul_cancel := hf.forall.2 fun x => by erw [← inv, ← mul, inv_mul_cancel, one] } /-- A type endowed with `0`, `1`, `+` is an additive group with one, if it admits a surjective map that preserves `0`, `1`, and `+` to an additive group with one. diff --git a/Mathlib/Algebra/Group/Int.lean b/Mathlib/Algebra/Group/Int.lean index fbdbbfb08b060..720440a2282fb 100644 --- a/Mathlib/Algebra/Group/Int.lean +++ b/Mathlib/Algebra/Group/Int.lean @@ -38,7 +38,7 @@ instance instAddCommGroup : AddCommGroup ℤ where add_assoc := Int.add_assoc add_zero := Int.add_zero zero_add := Int.zero_add - add_left_neg := Int.add_left_neg + neg_add_cancel := Int.add_left_neg nsmul := (·*·) nsmul_zero := Int.zero_mul nsmul_succ n x := diff --git a/Mathlib/Algebra/Group/Invertible/Defs.lean b/Mathlib/Algebra/Group/Invertible/Defs.lean index a5aeeb2cf162d..a44dee8152e49 100644 --- a/Mathlib/Algebra/Group/Invertible/Defs.lean +++ b/Mathlib/Algebra/Group/Invertible/Defs.lean @@ -171,11 +171,11 @@ abbrev Invertible.copy [MulOneClass α] {r : α} (hr : Invertible r) (s : α) (h /-- Each element of a group is invertible. -/ def invertibleOfGroup [Group α] (a : α) : Invertible a := - ⟨a⁻¹, inv_mul_self a, mul_inv_self a⟩ + ⟨a⁻¹, inv_mul_cancel a, mul_inv_cancel a⟩ @[simp] theorem invOf_eq_group_inv [Group α] (a : α) [Invertible a] : ⅟ a = a⁻¹ := - invOf_eq_right_inv (mul_inv_self a) + invOf_eq_right_inv (mul_inv_cancel a) /-- `1` is the inverse of itself -/ def invertibleOne [Monoid α] : Invertible (1 : α) := diff --git a/Mathlib/Algebra/Group/MinimalAxioms.lean b/Mathlib/Algebra/Group/MinimalAxioms.lean index 3cf4026f974c3..bcc623c694479 100644 --- a/Mathlib/Algebra/Group/MinimalAxioms.lean +++ b/Mathlib/Algebra/Group/MinimalAxioms.lean @@ -38,20 +38,20 @@ See note [reducible non-instances]."] abbrev Group.ofLeftAxioms {G : Type u} [Mul G] [Inv G] [One G] (assoc : ∀ a b c : G, (a * b) * c = a * (b * c)) (one_mul : ∀ a : G, 1 * a = a) - (mul_left_inv : ∀ a : G, a⁻¹ * a = 1) : Group G := + (inv_mul_cancel : ∀ a : G, a⁻¹ * a = 1) : Group G := { mul_assoc := assoc, one_mul := one_mul, - mul_left_inv := mul_left_inv, + inv_mul_cancel := inv_mul_cancel, mul_one := fun a => by - have mul_right_inv : ∀ a : G, a * a⁻¹ = 1 := fun a => + have mul_inv_cancel : ∀ a : G, a * a⁻¹ = 1 := fun a => calc a * a⁻¹ = 1 * (a * a⁻¹) := (one_mul _).symm _ = ((a * a⁻¹)⁻¹ * (a * a⁻¹)) * (a * a⁻¹) := by - rw [mul_left_inv] + rw [inv_mul_cancel] _ = (a * a⁻¹)⁻¹ * (a * ((a⁻¹ * a) * a⁻¹)) := by simp only [assoc] _ = 1 := by - rw [mul_left_inv, one_mul, mul_left_inv] - rw [← mul_left_inv a, ← assoc, mul_right_inv a, one_mul] } + rw [inv_mul_cancel, one_mul, inv_mul_cancel] + rw [← inv_mul_cancel a, ← assoc, mul_inv_cancel a, one_mul] } /-- Define a `Group` structure on a Type by proving `∀ a, a * 1 = a` and `∀ a, a * a⁻¹ = 1`. @@ -65,17 +65,17 @@ See note [reducible non-instances]."] abbrev Group.ofRightAxioms {G : Type u} [Mul G] [Inv G] [One G] (assoc : ∀ a b c : G, (a * b) * c = a * (b * c)) (mul_one : ∀ a : G, a * 1 = a) - (mul_right_inv : ∀ a : G, a * a⁻¹ = 1) : Group G := - have mul_left_inv : ∀ a : G, a⁻¹ * a = 1 := fun a => + (mul_inv_cancel : ∀ a : G, a * a⁻¹ = 1) : Group G := + have inv_mul_cancel : ∀ a : G, a⁻¹ * a = 1 := fun a => calc a⁻¹ * a = (a⁻¹ * a) * 1 := (mul_one _).symm _ = (a⁻¹ * a) * ((a⁻¹ * a) * (a⁻¹ * a)⁻¹) := by - rw [mul_right_inv] + rw [mul_inv_cancel] _ = ((a⁻¹ * (a * a⁻¹)) * a) * (a⁻¹ * a)⁻¹ := by simp only [assoc] _ = 1 := by - rw [mul_right_inv, mul_one, mul_right_inv] + rw [mul_inv_cancel, mul_one, mul_inv_cancel] { mul_assoc := assoc, mul_one := mul_one, - mul_left_inv := mul_left_inv, + inv_mul_cancel := inv_mul_cancel, one_mul := fun a => by - rw [← mul_right_inv a, assoc, mul_left_inv, mul_one] } + rw [← mul_inv_cancel a, assoc, inv_mul_cancel, mul_one] } diff --git a/Mathlib/Algebra/Group/Opposite.lean b/Mathlib/Algebra/Group/Opposite.lean index 55cb4466fa95e..1f1389337c3a4 100644 --- a/Mathlib/Algebra/Group/Opposite.lean +++ b/Mathlib/Algebra/Group/Opposite.lean @@ -180,7 +180,7 @@ instance instDivisionCommMonoid [DivisionCommMonoid α] : DivisionCommMonoid α @[to_additive] instance instGroup [Group α] : Group αᵐᵒᵖ where toDivInvMonoid := instDivInvMonoid - mul_left_inv _ := unop_injective <| mul_inv_self _ + inv_mul_cancel _ := unop_injective <| mul_inv_cancel _ @[to_additive] instance instCommGroup [CommGroup α] : CommGroup αᵐᵒᵖ where diff --git a/Mathlib/Algebra/Group/Pi/Basic.lean b/Mathlib/Algebra/Group/Pi/Basic.lean index 15492d032a04c..92f9824806c5f 100644 --- a/Mathlib/Algebra/Group/Pi/Basic.lean +++ b/Mathlib/Algebra/Group/Pi/Basic.lean @@ -214,7 +214,7 @@ instance divisionCommMonoid [∀ i, DivisionCommMonoid (f i)] : DivisionCommMono @[to_additive] instance group [∀ i, Group (f i)] : Group (∀ i, f i) where - mul_left_inv := by intros; ext; exact mul_left_inv _ + inv_mul_cancel := by intros; ext; exact inv_mul_cancel _ @[to_additive] instance commGroup [∀ i, CommGroup (f i)] : CommGroup (∀ i, f i) := { group, commMonoid with } diff --git a/Mathlib/Algebra/Group/Prod.lean b/Mathlib/Algebra/Group/Prod.lean index 8467b413ead24..d1a779b29708d 100644 --- a/Mathlib/Algebra/Group/Prod.lean +++ b/Mathlib/Algebra/Group/Prod.lean @@ -189,7 +189,7 @@ instance [DivisionCommMonoid G] [DivisionCommMonoid H] : DivisionCommMonoid (G @[to_additive] instance instGroup [Group G] [Group H] : Group (G × H) := - { mul_left_inv := fun _ => mk.inj_iff.mpr ⟨mul_left_inv _, mul_left_inv _⟩ } + { inv_mul_cancel := fun _ => mk.inj_iff.mpr ⟨inv_mul_cancel _, inv_mul_cancel _⟩ } @[to_additive] instance [Mul G] [Mul H] [IsLeftCancelMul G] [IsLeftCancelMul H] : IsLeftCancelMul (G × H) where diff --git a/Mathlib/Algebra/Group/Semiconj/Defs.lean b/Mathlib/Algebra/Group/Semiconj/Defs.lean index bcc5e310fabbb..bae3512e418c7 100644 --- a/Mathlib/Algebra/Group/Semiconj/Defs.lean +++ b/Mathlib/Algebra/Group/Semiconj/Defs.lean @@ -121,7 +121,7 @@ variable [Group G] {a x y : G} /-- `a` semiconjugates `x` to `a * x * a⁻¹`. -/ @[to_additive "`a` semiconjugates `x` to `a + x + -a`."] theorem conj_mk (a x : G) : SemiconjBy a x (a * x * a⁻¹) := by - unfold SemiconjBy; rw [mul_assoc, inv_mul_self, mul_one] + unfold SemiconjBy; rw [mul_assoc, inv_mul_cancel, mul_one] @[to_additive (attr := simp)] theorem conj_iff {a x y b : G} : diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index e7386db452760..187ab2f00201a 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -934,9 +934,9 @@ def closureCommGroupOfComm {k : Set G} (hcomm : ∀ x ∈ k, ∀ y ∈ k, x * y (fun x y z h₁ h₂ => by rw [mul_assoc, h₂, ← mul_assoc, h₁, mul_assoc]) (fun x y z h₁ h₂ => by rw [← mul_assoc, h₁, mul_assoc, h₂, ← mul_assoc]) (fun x y h => by - rw [inv_mul_eq_iff_eq_mul, ← mul_assoc, h, mul_assoc, mul_inv_self, mul_one]) + rw [inv_mul_eq_iff_eq_mul, ← mul_assoc, h, mul_assoc, mul_inv_cancel, mul_one]) fun x y h => by - rw [mul_inv_eq_iff_eq_mul, mul_assoc, h, ← mul_assoc, inv_mul_self, one_mul] } + rw [mul_inv_eq_iff_eq_mul, mul_assoc, h, ← mul_assoc, inv_mul_cancel, one_mul] } variable (G) @@ -1614,7 +1614,7 @@ def normalizer : Subgroup G where simp only [mul_assoc, mul_inv_rev] inv_mem' {a} (ha : ∀ n, n ∈ H ↔ a * n * a⁻¹ ∈ H) n := by rw [ha (a⁻¹ * n * a⁻¹⁻¹)] - simp only [inv_inv, mul_assoc, mul_inv_cancel_left, mul_right_inv, mul_one] + simp only [inv_inv, mul_assoc, mul_inv_cancel_left, mul_inv_cancel, mul_one] -- variant for sets. -- TODO should this replace `normalizer`? @@ -1630,7 +1630,7 @@ def setNormalizer (S : Set G) : Subgroup G where simp only [mul_assoc, mul_inv_rev] inv_mem' {a} (ha : ∀ n, n ∈ S ↔ a * n * a⁻¹ ∈ S) n := by rw [ha (a⁻¹ * n * a⁻¹⁻¹)] - simp only [inv_inv, mul_assoc, mul_inv_cancel_left, mul_right_inv, mul_one] + simp only [inv_inv, mul_assoc, mul_inv_cancel_left, mul_inv_cancel, mul_one] variable {H} @@ -1911,7 +1911,7 @@ theorem normalClosure_closure_eq_normalClosure {s : Set G} : as shown by `Subgroup.normalCore_eq_iSup`. -/ def normalCore (H : Subgroup G) : Subgroup G where carrier := { a : G | ∀ b : G, b * a * b⁻¹ ∈ H } - one_mem' a := by rw [mul_one, mul_inv_self]; exact H.one_mem + one_mem' a := by rw [mul_one, mul_inv_cancel]; exact H.one_mem inv_mem' {a} h b := (congr_arg (· ∈ H) conj_inv).mp (H.inv_mem (h b)) mul_mem' {a b} ha hb c := (congr_arg (· ∈ H) conj_mul).mp (H.mul_mem (ha c) (hb c)) @@ -2110,7 +2110,7 @@ def ker (f : G →* M) : Subgroup G := inv_mem' := fun {x} (hx : f x = 1) => calc f x⁻¹ = f x * f x⁻¹ := by rw [hx, one_mul] - _ = 1 := by rw [← map_mul, mul_inv_self, map_one] } + _ = 1 := by rw [← map_mul, mul_inv_cancel, map_one] } @[to_additive] theorem mem_ker (f : G →* M) {x : G} : x ∈ f.ker ↔ f x = 1 := @@ -2128,8 +2128,8 @@ theorem ker_toHomUnits {M} [Monoid M] (f : G →* M) : f.toHomUnits.ker = f.ker @[to_additive] theorem eq_iff (f : G →* M) {x y : G} : f x = f y ↔ y⁻¹ * x ∈ f.ker := by constructor <;> intro h - · rw [mem_ker, map_mul, h, ← map_mul, inv_mul_self, map_one] - · rw [← one_mul x, ← mul_inv_self y, mul_assoc, map_mul, f.mem_ker.1 h, mul_one] + · rw [mem_ker, map_mul, h, ← map_mul, inv_mul_cancel, map_one] + · rw [← one_mul x, ← mul_inv_cancel y, mul_assoc, map_mul, f.mem_ker.1 h, mul_one] @[to_additive] instance decidableMemKer [DecidableEq M] (f : G →* M) : DecidablePred (· ∈ f.ker) := fun x => @@ -2200,7 +2200,7 @@ theorem range_le_ker_iff (f : G →* G') (g : G' →* G'') : f.range ≤ g.ker @[to_additive] instance (priority := 100) normal_ker (f : G →* M) : f.ker.Normal := ⟨fun x hx y => by - rw [mem_ker, map_mul, map_mul, f.mem_ker.1 hx, mul_one, map_mul_eq_one f (mul_inv_self y)]⟩ + rw [mem_ker, map_mul, map_mul, f.mem_ker.1 hx, mul_one, map_mul_eq_one f (mul_inv_cancel y)]⟩ @[to_additive (attr := simp)] lemma ker_fst : ker (fst G G') = .prod ⊥ ⊤ := SetLike.ext fun _ => (and_true_iff _).symm @@ -2815,7 +2815,7 @@ theorem subgroupOf_sup (A A' B : Subgroup G) (hA : A ≤ B) (hA' : A' ≤ B) : theorem SubgroupNormal.mem_comm {H K : Subgroup G} (hK : H ≤ K) [hN : (H.subgroupOf K).Normal] {a b : G} (hb : b ∈ K) (h : a * b ∈ H) : b * a ∈ H := by have := (normal_subgroupOf_iff hK).mp hN (a * b) b h hb - rwa [mul_assoc, mul_assoc, mul_right_inv, mul_one] at this + rwa [mul_assoc, mul_assoc, mul_inv_cancel, mul_one] at this /-- Elements of disjoint, normal subgroups commute. -/ @[to_additive "Elements of disjoint, normal subgroups commute."] @@ -2883,8 +2883,8 @@ theorem normalClosure_eq_top_of {N : Subgroup G} [hn : N.Normal] {g g' : G} {hg · have h := hn.conj_mem _ hx c⁻¹ rwa [inv_inv] at h simp only [MonoidHom.codRestrict_apply, MulEquiv.coe_toMonoidHom, MulAut.conj_apply, coe_mk, - MonoidHom.restrict_apply, Subtype.mk_eq_mk, ← mul_assoc, mul_inv_self, one_mul] - rw [mul_assoc, mul_inv_self, mul_one] + MonoidHom.restrict_apply, Subtype.mk_eq_mk, ← mul_assoc, mul_inv_cancel, one_mul] + rw [mul_assoc, mul_inv_cancel, mul_one] rw [eq_top_iff, ← MonoidHom.range_top_of_surjective _ hs, MonoidHom.range_eq_map] refine le_trans (map_mono (eq_top_iff.1 ht)) (map_le_iff_le_comap.2 (normalClosure_le_normal ?_)) rw [Set.singleton_subset_iff, SetLike.mem_coe] diff --git a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean index b4cf683865c20..71e4f17ffc5ea 100644 --- a/Mathlib/Algebra/Group/Subgroup/Pointwise.lean +++ b/Mathlib/Algebra/Group/Subgroup/Pointwise.lean @@ -80,7 +80,7 @@ theorem closure_toSubmonoid (S : Set G) : `zsmul_induction_left`."] theorem closure_induction_left {p : (x : G) → x ∈ closure s → Prop} (one : p 1 (one_mem _)) (mul_left : ∀ x (hx : x ∈ s), ∀ (y) hy, p y hy → p (x * y) (mul_mem (subset_closure hx) hy)) - (mul_left_inv : ∀ x (hx : x ∈ s), ∀ (y) hy, p y hy → + (inv_mul_cancel : ∀ x (hx : x ∈ s), ∀ (y) hy, p y hy → p (x⁻¹ * y) (mul_mem (inv_mem (subset_closure hx)) hy)) {x : G} (h : x ∈ closure s) : p x h := by revert h @@ -91,7 +91,7 @@ theorem closure_induction_left {p : (x : G) → x ∈ closure s → Prop} (one : | mul_left x hx y hy ih => cases hx with | inl hx => exact mul_left _ hx _ hy ih - | inr hx => simpa only [inv_inv] using mul_left_inv _ hx _ hy ih + | inr hx => simpa only [inv_inv] using inv_mul_cancel _ hx _ hy ih /-- For subgroups generated by a single element, see the simpler `zpow_induction_right`. -/ @[to_additive (attr := elab_as_elim) @@ -99,14 +99,14 @@ theorem closure_induction_left {p : (x : G) → x ∈ closure s → Prop} (one : `zsmul_induction_right`."] theorem closure_induction_right {p : (x : G) → x ∈ closure s → Prop} (one : p 1 (one_mem _)) (mul_right : ∀ (x) hx, ∀ y (hy : y ∈ s), p x hx → p (x * y) (mul_mem hx (subset_closure hy))) - (mul_right_inv : ∀ (x) hx, ∀ y (hy : y ∈ s), p x hx → + (mul_inv_cancel : ∀ (x) hx, ∀ y (hy : y ∈ s), p x hx → p (x * y⁻¹) (mul_mem hx (inv_mem (subset_closure hy)))) {x : G} (h : x ∈ closure s) : p x h := closure_induction_left (s := MulOpposite.unop ⁻¹' s) (p := fun m hm => p m.unop <| by rwa [← op_closure] at hm) one (fun _x hx _y hy => mul_right _ _ _ hx) - (fun _x hx _y hy => mul_right_inv _ _ _ hx) + (fun _x hx _y hy => mul_inv_cancel _ _ _ hx) (by rwa [← op_closure]) @[to_additive (attr := simp)] diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index e3f2d53159647..8c8595e52f624 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -422,7 +422,7 @@ lemma powers_one : powers (1 : M) = ⊥ := bot_unique <| powers_le.2 <| one_mem /-- 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) - mul_left_inv y := Subtype.ext <| by + inv_mul_cancel y := Subtype.ext <| by obtain ⟨_, k, rfl⟩ := y simp only [coe_one, coe_mul, SubmonoidClass.coe_pow] rw [← pow_succ, Nat.sub_add_cancel hpos, ← pow_mul, mul_comm, pow_mul, hx, one_pow] diff --git a/Mathlib/Algebra/Group/TypeTags.lean b/Mathlib/Algebra/Group/TypeTags.lean index 2ecd644e03b28..897dc1988ea48 100644 --- a/Mathlib/Algebra/Group/TypeTags.lean +++ b/Mathlib/Algebra/Group/TypeTags.lean @@ -420,10 +420,10 @@ instance Multiplicative.divisionCommMonoid [SubtractionCommMonoid α] : { Multiplicative.divisionMonoid, Multiplicative.commSemigroup with } instance Additive.addGroup [Group α] : AddGroup (Additive α) := - { Additive.subNegMonoid with add_left_neg := @mul_left_inv α _ } + { Additive.subNegMonoid with neg_add_cancel := @inv_mul_cancel α _ } instance Multiplicative.group [AddGroup α] : Group (Multiplicative α) := - { Multiplicative.divInvMonoid with mul_left_inv := @add_left_neg α _ } + { Multiplicative.divInvMonoid with inv_mul_cancel := @neg_add_cancel α _ } instance Additive.addCommGroup [CommGroup α] : AddCommGroup (Additive α) := { Additive.addGroup, Additive.addCommMonoid with } diff --git a/Mathlib/Algebra/Group/UniqueProds.lean b/Mathlib/Algebra/Group/UniqueProds.lean index 6ac2374fea1d6..82f5f510f979c 100644 --- a/Mathlib/Algebra/Group/UniqueProds.lean +++ b/Mathlib/Algebra/Group/UniqueProds.lean @@ -376,8 +376,8 @@ open MulOpposite in let C := A.map ⟨_, mul_right_injective a⁻¹⟩ -- C = a⁻¹A let D := B.map ⟨_, mul_left_injective b⁻¹⟩ -- D = Bb⁻¹ have hcard : 1 < C.card ∨ 1 < D.card := by simp_rw [C, D, card_map]; exact hc.2.2 - have hC : 1 ∈ C := mem_map.mpr ⟨a, ha, inv_mul_self a⟩ - have hD : 1 ∈ D := mem_map.mpr ⟨b, hb, mul_inv_self b⟩ + have hC : 1 ∈ C := mem_map.mpr ⟨a, ha, inv_mul_cancel a⟩ + have hD : 1 ∈ D := mem_map.mpr ⟨b, hb, mul_inv_cancel b⟩ suffices ∃ c ∈ C, ∃ d ∈ D, (c ≠ 1 ∨ d ≠ 1) ∧ UniqueMul C D c d by simp_rw [mem_product] obtain ⟨c, hc, d, hd, hne, hu'⟩ := this diff --git a/Mathlib/Algebra/Group/Units.lean b/Mathlib/Algebra/Group/Units.lean index 3286fc67b219d..7ae26582ac369 100644 --- a/Mathlib/Algebra/Group/Units.lean +++ b/Mathlib/Algebra/Group/Units.lean @@ -323,7 +323,7 @@ instance instDivInvMonoid : DivInvMonoid αˣ where /-- Units of a monoid form a group. -/ @[to_additive "Additive units of an additive monoid form an additive group."] instance instGroup : Group αˣ where - mul_left_inv := fun u => ext u.inv_val + inv_mul_cancel := fun u => ext u.inv_val /-- Units of a commutative monoid form a commutative group. -/ @[to_additive "Additive units of an additive commutative monoid form @@ -912,7 +912,8 @@ lemma divp_eq_div [DivisionMonoid α] (a : α) (u : αˣ) : a /ₚ u = a / u := rw [div_eq_mul_inv, divp, u.val_inv_eq_inv_val] @[to_additive] -lemma Group.isUnit [Group α] (a : α) : IsUnit a := ⟨⟨a, a⁻¹, mul_inv_self _, inv_mul_self _⟩, rfl⟩ +lemma Group.isUnit [Group α] (a : α) : IsUnit a := + ⟨⟨a, a⁻¹, mul_inv_cancel _, inv_mul_cancel _⟩, rfl⟩ -- namespace end IsUnit @@ -930,7 +931,7 @@ noncomputable def invOfIsUnit [Monoid M] (h : ∀ a : M, IsUnit a) : Inv M where noncomputable def groupOfIsUnit [hM : Monoid M] (h : ∀ a : M, IsUnit a) : Group M := { hM with toInv := invOfIsUnit h, - mul_left_inv := fun a => by + inv_mul_cancel := fun a => by change ↑(h a).unit⁻¹ * a = 1 rw [Units.inv_mul_eq_iff_eq_mul, (h a).unit_spec, mul_one] } @@ -938,7 +939,7 @@ noncomputable def groupOfIsUnit [hM : Monoid M] (h : ∀ a : M, IsUnit a) : Grou noncomputable def commGroupOfIsUnit [hM : CommMonoid M] (h : ∀ a : M, IsUnit a) : CommGroup M := { hM with toInv := invOfIsUnit h, - mul_left_inv := fun a => by + inv_mul_cancel := fun a => by change ↑(h a).unit⁻¹ * a = 1 rw [Units.inv_mul_eq_iff_eq_mul, (h a).unit_spec, mul_one] } diff --git a/Mathlib/Algebra/Group/Units/Equiv.lean b/Mathlib/Algebra/Group/Units/Equiv.lean index cd95a51a8c40a..a3ae5f7904381 100644 --- a/Mathlib/Algebra/Group/Units/Equiv.lean +++ b/Mathlib/Algebra/Group/Units/Equiv.lean @@ -19,7 +19,7 @@ variable {F α β A B M N P Q G H : Type*} @[to_additive (attr := simps apply_val symm_apply) "An additive group is isomorphic to its group of additive units"] def toUnits [Group G] : G ≃* Gˣ where - toFun x := ⟨x, x⁻¹, mul_inv_self _, inv_mul_self _⟩ + toFun x := ⟨x, x⁻¹, mul_inv_cancel _, inv_mul_cancel _⟩ invFun x := x left_inv _ := rfl right_inv _ := Units.ext rfl diff --git a/Mathlib/Algebra/Group/Units/Hom.lean b/Mathlib/Algebra/Group/Units/Hom.lean index 75e7b7796e1cb..1a84f0441934c 100644 --- a/Mathlib/Algebra/Group/Units/Hom.lean +++ b/Mathlib/Algebra/Group/Units/Hom.lean @@ -140,8 +140,8 @@ and `f.toHomUnits` is the corresponding monoid homomorphism from `G` to `Mˣ`. - then its image lies in the `AddUnits` of `M`, and `f.toHomUnits` is the corresponding homomorphism from `G` to `AddUnits M`."] def toHomUnits {G M : Type*} [Group G] [Monoid M] (f : G →* M) : G →* Mˣ := - Units.liftRight f (fun g => ⟨f g, f g⁻¹, map_mul_eq_one f (mul_inv_self _), - map_mul_eq_one f (inv_mul_self _)⟩) + Units.liftRight f (fun g => ⟨f g, f g⁻¹, map_mul_eq_one f (mul_inv_cancel _), + map_mul_eq_one f (inv_mul_cancel _)⟩) fun _ => rfl @[to_additive (attr := simp)] diff --git a/Mathlib/Algebra/GroupWithZero/Action/Defs.lean b/Mathlib/Algebra/GroupWithZero/Action/Defs.lean index 5686fbc1e8818..a2e67fb1c2e12 100644 --- a/Mathlib/Algebra/GroupWithZero/Action/Defs.lean +++ b/Mathlib/Algebra/GroupWithZero/Action/Defs.lean @@ -274,7 +274,7 @@ instance AddGroup.int_smulCommClass' : SMulCommClass M ℤ A := @[simp] theorem smul_neg (r : M) (x : A) : r • -x = -(r • x) := - eq_neg_of_add_eq_zero_left <| by rw [← smul_add, neg_add_self, smul_zero] + eq_neg_of_add_eq_zero_left <| by rw [← smul_add, neg_add_cancel, smul_zero] theorem smul_sub (r : M) (x y : A) : r • (x - y) = r • x - r • y := by rw [sub_eq_add_neg, sub_eq_add_neg, smul_add, smul_neg] diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index 857d0272350c0..d5a2ad48bcf5f 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -238,11 +238,11 @@ variable [GroupWithZero G₀] {a b c g h x : G₀} theorem GroupWithZero.mul_left_injective (h : x ≠ 0) : Function.Injective fun y => x * y := fun y y' w => by - simpa only [← mul_assoc, inv_mul_cancel h, one_mul] using congr_arg (fun y => x⁻¹ * y) w + simpa only [← mul_assoc, inv_mul_cancel₀ h, one_mul] using congr_arg (fun y => x⁻¹ * y) w theorem GroupWithZero.mul_right_injective (h : x ≠ 0) : Function.Injective fun y => y * x := fun y y' w => by - simpa only [mul_assoc, mul_inv_cancel h, mul_one] using congr_arg (fun y => y * x⁻¹) w + simpa only [mul_assoc, mul_inv_cancel₀ h, mul_one] using congr_arg (fun y => y * x⁻¹) w @[simp] theorem inv_mul_cancel_right₀ (h : b ≠ 0) (a : G₀) : a * b⁻¹ * b = a := @@ -268,7 +268,7 @@ instance (priority := 100) GroupWithZero.toDivisionMonoid : DivisionMonoid G₀ inv_inv := fun a => by by_cases h : a = 0 · simp [h] - · exact left_inv_eq_right_inv (inv_mul_cancel <| inv_ne_zero h) (inv_mul_cancel h) + · exact left_inv_eq_right_inv (inv_mul_cancel₀ <| inv_ne_zero h) (inv_mul_cancel₀ h) , mul_inv_rev := fun a b => by by_cases ha : a = 0 @@ -305,16 +305,16 @@ theorem div_zero (a : G₀) : a / 0 = 0 := by rw [div_eq_mul_inv, inv_zero, mul_ theorem mul_self_mul_inv (a : G₀) : a * a * a⁻¹ = a := by by_cases h : a = 0 · rw [h, inv_zero, mul_zero] - · rw [mul_assoc, mul_inv_cancel h, mul_one] + · rw [mul_assoc, mul_inv_cancel₀ h, mul_one] /-- Multiplying `a` by its inverse and then by itself results in `a` (whether or not `a` is zero). -/ @[simp] -theorem mul_inv_mul_self (a : G₀) : a * a⁻¹ * a = a := by +theorem mul_inv_mul_cancel (a : G₀) : a * a⁻¹ * a = a := by by_cases h : a = 0 · rw [h, inv_zero, mul_zero] - · rw [mul_inv_cancel h, one_mul] + · rw [mul_inv_cancel₀ h, one_mul] /-- Multiplying `a⁻¹` by `a` twice results in `a` (whether or not `a` @@ -323,7 +323,7 @@ is zero). -/ theorem inv_mul_mul_self (a : G₀) : a⁻¹ * a * a = a := by by_cases h : a = 0 · rw [h, inv_zero, mul_zero] - · rw [inv_mul_cancel h, one_mul] + · rw [inv_mul_cancel₀ h, one_mul] /-- Multiplying `a` by itself and then dividing by itself results in `a`, whether or not `a` is @@ -334,7 +334,7 @@ theorem mul_self_div_self (a : G₀) : a * a / a = a := by rw [div_eq_mul_inv, m /-- Dividing `a` by itself and then multiplying by itself results in `a`, whether or not `a` is zero. -/ @[simp] -theorem div_self_mul_self (a : G₀) : a / a * a = a := by rw [div_eq_mul_inv, mul_inv_mul_self a] +theorem div_self_mul_self (a : G₀) : a / a * a = a := by rw [div_eq_mul_inv, mul_inv_mul_cancel a] attribute [local simp] div_eq_mul_inv mul_comm mul_assoc mul_left_comm @@ -370,10 +370,10 @@ theorem eq_zero_of_one_div_eq_zero {a : G₀} (h : 1 / a = 0) : a = 0 := Classical.byCases (fun ha => ha) fun ha => ((one_div_ne_zero ha) h).elim theorem mul_left_surjective₀ {a : G₀} (h : a ≠ 0) : Surjective fun g => a * g := fun g => - ⟨a⁻¹ * g, by simp [← mul_assoc, mul_inv_cancel h]⟩ + ⟨a⁻¹ * g, by simp [← mul_assoc, mul_inv_cancel₀ h]⟩ theorem mul_right_surjective₀ {a : G₀} (h : a ≠ 0) : Surjective fun g => g * a := fun g => - ⟨g * a⁻¹, by simp [mul_assoc, inv_mul_cancel h]⟩ + ⟨g * a⁻¹, by simp [mul_assoc, inv_mul_cancel₀ h]⟩ lemma zero_zpow : ∀ n : ℤ, n ≠ 0 → (0 : G₀) ^ n = 0 | (n : ℕ), h => by rw [zpow_natCast, zero_pow]; simpa [Int.natCast_eq_zero] using h @@ -386,15 +386,15 @@ lemma zero_zpow_eq (n : ℤ) : (0 : G₀) ^ n = if n = 0 then 1 else 0 := by lemma zpow_add_one₀ (ha : a ≠ 0) : ∀ n : ℤ, a ^ (n + 1) = a ^ n * a | (n : ℕ) => by simp only [← Int.ofNat_succ, zpow_natCast, pow_succ] - | .negSucc 0 => by erw [zpow_zero, zpow_negSucc, pow_one, inv_mul_cancel ha] + | .negSucc 0 => by erw [zpow_zero, zpow_negSucc, pow_one, inv_mul_cancel₀ ha] | .negSucc (n + 1) => by rw [Int.negSucc_eq, zpow_neg, Int.neg_add, Int.neg_add_cancel_right, zpow_neg, ← Int.ofNat_succ, - zpow_natCast, zpow_natCast, pow_succ' _ (n + 1), mul_inv_rev, mul_assoc, inv_mul_cancel ha, + zpow_natCast, zpow_natCast, pow_succ' _ (n + 1), mul_inv_rev, mul_assoc, inv_mul_cancel₀ ha, mul_one] lemma zpow_sub_one₀ (ha : a ≠ 0) (n : ℤ) : a ^ (n - 1) = a ^ n * a⁻¹ := calc - a ^ (n - 1) = a ^ (n - 1) * a * a⁻¹ := by rw [mul_assoc, mul_inv_cancel ha, mul_one] + a ^ (n - 1) = a ^ (n - 1) * a * a⁻¹ := by rw [mul_assoc, mul_inv_cancel₀ ha, mul_one] _ = a ^ n * a⁻¹ := by rw [← zpow_add_one₀ ha, Int.sub_add_cancel] lemma zpow_add₀ (ha : a ≠ 0) (m n : ℤ) : a ^ (m + n) = a ^ m * a ^ n := by diff --git a/Mathlib/Algebra/GroupWithZero/Defs.lean b/Mathlib/Algebra/GroupWithZero/Defs.lean index 55fde63c5ae05..b3d73d7ba0a35 100644 --- a/Mathlib/Algebra/GroupWithZero/Defs.lean +++ b/Mathlib/Algebra/GroupWithZero/Defs.lean @@ -192,11 +192,11 @@ attribute [simp] inv_zero section GroupWithZero variable [GroupWithZero G₀] {a : G₀} -@[simp] lemma mul_inv_cancel (h : a ≠ 0) : a * a⁻¹ = 1 := GroupWithZero.mul_inv_cancel a h +@[simp] lemma mul_inv_cancel₀ (h : a ≠ 0) : a * a⁻¹ = 1 := GroupWithZero.mul_inv_cancel a h -- See note [lower instance priority] instance (priority := 100) GroupWithZero.toMulDivCancelClass : MulDivCancelClass G₀ where - mul_div_cancel a b hb := by rw [div_eq_mul_inv, mul_assoc, mul_inv_cancel hb, mul_one] + mul_div_cancel a b hb := by rw [div_eq_mul_inv, mul_assoc, mul_inv_cancel₀ hb, mul_one] end GroupWithZero diff --git a/Mathlib/Algebra/GroupWithZero/InjSurj.lean b/Mathlib/Algebra/GroupWithZero/InjSurj.lean index c3dbdb1c4a2ac..e8de1ecfd7972 100644 --- a/Mathlib/Algebra/GroupWithZero/InjSurj.lean +++ b/Mathlib/Algebra/GroupWithZero/InjSurj.lean @@ -193,7 +193,7 @@ protected abbrev Function.Injective.groupWithZero [Zero G₀'] [Mul G₀'] [One pullback_nonzero f zero one with inv_zero := hf <| by erw [inv, zero, inv_zero], mul_inv_cancel := fun x hx => hf <| by - erw [one, mul, inv, mul_inv_cancel ((hf.ne_iff' zero).2 hx)] } + erw [one, mul, inv, mul_inv_cancel₀ ((hf.ne_iff' zero).2 hx)] } /-- Push forward a `GroupWithZero` along a surjective function. See note [reducible non-instances]. -/ @@ -206,7 +206,7 @@ protected abbrev Function.Surjective.groupWithZero [Zero G₀'] [Mul G₀'] [One { hf.monoidWithZero f zero one mul npow, hf.divInvMonoid f one mul inv div npow zpow with inv_zero := by erw [← zero, ← inv, inv_zero], mul_inv_cancel := hf.forall.2 fun x hx => by - erw [← inv, ← mul, mul_inv_cancel (mt (congr_arg f) fun h ↦ hx (h.trans zero)), one] + erw [← inv, ← mul, mul_inv_cancel₀ (mt (congr_arg f) fun h ↦ hx (h.trans zero)), one] exists_pair_ne := ⟨0, 1, h01⟩ } end GroupWithZero diff --git a/Mathlib/Algebra/GroupWithZero/Invertible.lean b/Mathlib/Algebra/GroupWithZero/Invertible.lean index 4e9d53ea9718d..36f4cf1c801f9 100644 --- a/Mathlib/Algebra/GroupWithZero/Invertible.lean +++ b/Mathlib/Algebra/GroupWithZero/Invertible.lean @@ -44,19 +44,19 @@ variable [GroupWithZero α] /-- `a⁻¹` is an inverse of `a` if `a ≠ 0` -/ def invertibleOfNonzero {a : α} (h : a ≠ 0) : Invertible a := - ⟨a⁻¹, inv_mul_cancel h, mul_inv_cancel h⟩ + ⟨a⁻¹, inv_mul_cancel₀ h, mul_inv_cancel₀ h⟩ @[simp] theorem invOf_eq_inv (a : α) [Invertible a] : ⅟ a = a⁻¹ := - invOf_eq_right_inv (mul_inv_cancel (nonzero_of_invertible a)) + invOf_eq_right_inv (mul_inv_cancel₀ (nonzero_of_invertible a)) @[simp] theorem inv_mul_cancel_of_invertible (a : α) [Invertible a] : a⁻¹ * a = 1 := - inv_mul_cancel (nonzero_of_invertible a) + inv_mul_cancel₀ (nonzero_of_invertible a) @[simp] theorem mul_inv_cancel_of_invertible (a : α) [Invertible a] : a * a⁻¹ = 1 := - mul_inv_cancel (nonzero_of_invertible a) + mul_inv_cancel₀ (nonzero_of_invertible a) /-- `a` is the inverse of `a⁻¹` -/ def invertibleInv {a : α} [Invertible a] : Invertible a⁻¹ := diff --git a/Mathlib/Algebra/GroupWithZero/NeZero.lean b/Mathlib/Algebra/GroupWithZero/NeZero.lean index 5684c8915566e..5f9e303a359d3 100644 --- a/Mathlib/Algebra/GroupWithZero/NeZero.lean +++ b/Mathlib/Algebra/GroupWithZero/NeZero.lean @@ -43,11 +43,11 @@ variable {G₀ : Type*} [GroupWithZero G₀] {a : G₀} -- Porting note: used `simpa` to prove `False` in lean3 theorem inv_ne_zero (h : a ≠ 0) : a⁻¹ ≠ 0 := fun a_eq_0 => by - have := mul_inv_cancel h + have := mul_inv_cancel₀ h simp only [a_eq_0, mul_zero, zero_ne_one] at this @[simp] -theorem inv_mul_cancel (h : a ≠ 0) : a⁻¹ * a = 1 := +theorem inv_mul_cancel₀ (h : a ≠ 0) : a⁻¹ * a = 1 := calc a⁻¹ * a = a⁻¹ * a * a⁻¹ * a⁻¹⁻¹ := by simp [inv_ne_zero h] _ = a⁻¹ * a⁻¹⁻¹ := by simp [h] diff --git a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean index cec1efcc4201e..0ed3c88eff1ab 100644 --- a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean +++ b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean @@ -182,8 +182,8 @@ theorem mul_mem_nonZeroDivisors {a b : M₁} : a * b ∈ M₁⁰ ↔ a ∈ M₁ theorem isUnit_of_mem_nonZeroDivisors {G₀ : Type*} [GroupWithZero G₀] {x : G₀} (hx : x ∈ nonZeroDivisors G₀) : IsUnit x := - ⟨⟨x, x⁻¹, mul_inv_cancel (nonZeroDivisors.ne_zero hx), - inv_mul_cancel (nonZeroDivisors.ne_zero hx)⟩, rfl⟩ + ⟨⟨x, x⁻¹, mul_inv_cancel₀ (nonZeroDivisors.ne_zero hx), + inv_mul_cancel₀ (nonZeroDivisors.ne_zero hx)⟩, rfl⟩ lemma IsUnit.mem_nonZeroDivisors {a : M} (ha : IsUnit a) : a ∈ M⁰ := fun _ h ↦ ha.mul_left_eq_zero.mp h diff --git a/Mathlib/Algebra/GroupWithZero/Opposite.lean b/Mathlib/Algebra/GroupWithZero/Opposite.lean index d051e803fed1f..5ab40612e130a 100644 --- a/Mathlib/Algebra/GroupWithZero/Opposite.lean +++ b/Mathlib/Algebra/GroupWithZero/Opposite.lean @@ -66,7 +66,7 @@ instance instGroupWithZero [GroupWithZero α] : GroupWithZero αᵃᵒᵖ where __ := instMonoidWithZero __ := instNontrivial __ := instDivInvMonoid - mul_inv_cancel _ hx := unop_injective <| mul_inv_cancel <| unop_injective.ne hx + mul_inv_cancel _ hx := unop_injective <| mul_inv_cancel₀ <| unop_injective.ne hx inv_zero := unop_injective inv_zero end AddOpposite diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index 57475724a359a..c7f47fc71a812 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -159,7 +159,7 @@ variable {a b : G₀} or the `/ₚ` operation, it is possible to write a division as a partial function with three arguments. -/ def mk0 (a : G₀) (ha : a ≠ 0) : G₀ˣ := - ⟨a, a⁻¹, mul_inv_cancel ha, inv_mul_cancel ha⟩ + ⟨a, a⁻¹, mul_inv_cancel₀ ha, inv_mul_cancel₀ ha⟩ @[simp] theorem mk0_one (h := one_ne_zero) : mk0 (1 : G₀) h = 1 := by @@ -176,11 +176,11 @@ theorem mk0_val (u : G₀ˣ) (h : (u : G₀) ≠ 0) : mk0 (u : G₀) h = u := -- Porting note: removed `simp` tag because `simpNF` says it's redundant theorem mul_inv' (u : G₀ˣ) : u * (u : G₀)⁻¹ = 1 := - mul_inv_cancel u.ne_zero + mul_inv_cancel₀ u.ne_zero -- Porting note: removed `simp` tag because `simpNF` says it's redundant theorem inv_mul' (u : G₀ˣ) : (u⁻¹ : G₀) * u = 1 := - inv_mul_cancel u.ne_zero + inv_mul_cancel₀ u.ne_zero @[simp] theorem mk0_inj {a b : G₀} (ha : a ≠ 0) (hb : b ≠ 0) : Units.mk0 a ha = Units.mk0 b hb ↔ a = b := @@ -370,7 +370,7 @@ lemma zpow_eq_zero_iff {n : ℤ} (hn : n ≠ 0) : a ^ n = 0 ↔ a = 0 := lemma zpow_ne_zero_iff {n : ℤ} (hn : n ≠ 0) : a ^ n ≠ 0 ↔ a ≠ 0 := (zpow_eq_zero_iff hn).ne lemma zpow_neg_mul_zpow_self (n : ℤ) (ha : a ≠ 0) : a ^ (-n) * a ^ n = 1 := by - rw [zpow_neg]; exact inv_mul_cancel (zpow_ne_zero n ha) + rw [zpow_neg]; exact inv_mul_cancel₀ (zpow_ne_zero n ha) theorem Ring.inverse_eq_inv (a : G₀) : Ring.inverse a = a⁻¹ := by obtain rfl | ha := eq_or_ne a 0 diff --git a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean index af2fe768fa1ac..c413ab0a165b8 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean @@ -60,7 +60,7 @@ theorem map_inv₀ : f a⁻¹ = (f a)⁻¹ := by by_cases h : a = 0 · simp [h, map_zero f] · apply eq_inv_of_mul_eq_one_left - rw [← map_mul, inv_mul_cancel h, map_one] + rw [← map_mul, inv_mul_cancel₀ h, map_one] @[simp] theorem map_div₀ : f (a / b) = f a / f b := diff --git a/Mathlib/Algebra/GroupWithZero/WithZero.lean b/Mathlib/Algebra/GroupWithZero/WithZero.lean index 61d3330c7fac0..1a8e376d48a1d 100644 --- a/Mathlib/Algebra/GroupWithZero/WithZero.lean +++ b/Mathlib/Algebra/GroupWithZero/WithZero.lean @@ -255,7 +255,7 @@ instance groupWithZero : GroupWithZero (WithZero α) where __ := divInvMonoid __ := nontrivial inv_zero := WithZero.inv_zero - mul_inv_cancel a ha := by lift a to α using ha; norm_cast; apply mul_right_inv + mul_inv_cancel a ha := by lift a to α using ha; norm_cast; apply mul_inv_cancel /-- Any group is isomorphic to the units of itself adjoined with `0`. -/ diff --git a/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean b/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean index 15e28dcd178de..cc231f472e59d 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/Ext.lean @@ -284,7 +284,7 @@ lemma add_hom (α β : Ext X Y n) : (α + β).hom = α.hom + β.hom := by biprod.lift_snd, Functor.map_id, ShiftedHom.mk₀_id_comp] lemma neg_hom (α : Ext X Y n) : (-α).hom = -α.hom := by - rw [← add_right_inj α.hom, ← add_hom, add_right_neg, add_right_neg, zero_hom] + rw [← add_right_inj α.hom, ← add_hom, add_neg_cancel, add_neg_cancel, zero_hom] /-- When an instance of `[HasDerivedCategory.{w'} C]` is available, this is the additive bijection between `Ext.{w} X Y n` and a type of morphisms in the derived category. -/ diff --git a/Mathlib/Algebra/Homology/Homotopy.lean b/Mathlib/Algebra/Homology/Homotopy.lean index 508deb6f6d08f..bc2fd8fdf31ef 100644 --- a/Mathlib/Algebra/Homology/Homotopy.lean +++ b/Mathlib/Algebra/Homology/Homotopy.lean @@ -157,7 +157,7 @@ def symm {f g : C ⟶ D} (h : Homotopy f g) : Homotopy g f where hom := -h.hom zero i j w := by rw [Pi.neg_apply, Pi.neg_apply, h.zero i j w, neg_zero] comm i := by - rw [AddMonoidHom.map_neg, AddMonoidHom.map_neg, h.comm, ← neg_add, ← add_assoc, neg_add_self, + rw [AddMonoidHom.map_neg, AddMonoidHom.map_neg, h.comm, ← neg_add, ← add_assoc, neg_add_cancel, zero_add] /-- homotopy is a transitive relation. -/ diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean index b64a9b5129608..a8ad46dc1d626 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean @@ -43,7 +43,7 @@ def cocycleOfDegreewiseSplit : Cocycle S.X₃ S.X₁ 1 := sub_comp, comp_sub, comp_sub, assoc, id_comp, d_comp_d, comp_zero, zero_sub, ← S.g.comm_assoc, reassoc_of% (s_g p), r_f (p + 2), comp_sub, comp_sub, comp_id, comp_sub, ← S.g.comm_assoc, reassoc_of% (s_g (p + 1)), d_comp_d_assoc, zero_comp, - sub_zero, add_left_neg]) + sub_zero, neg_add_cancel]) /-- The canonical morphism `S.X₃ ⟶ S.X₁⟦(1 : ℤ)⟧` attached to a degreewise split short exact sequence of cochain complexes. -/ diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean index b00732bca4501..d7de339fc2748 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean @@ -290,8 +290,8 @@ lemma comp_assoc_of_third_is_zero_cochain {n₁ n₂ n₁₂ : ℤ} lemma comp_assoc_of_second_degree_eq_neg_third_degree {n₁ n₂ n₁₂ : ℤ} (z₁ : Cochain F G n₁) (z₂ : Cochain G K (-n₂)) (z₃ : Cochain K L n₂) (h₁₂ : n₁ + (-n₂) = n₁₂) : (z₁.comp z₂ h₁₂).comp z₃ - (show n₁₂ + n₂ = n₁ by rw [← h₁₂, add_assoc, neg_add_self, add_zero]) = - z₁.comp (z₂.comp z₃ (neg_add_self n₂)) (add_zero n₁) := + (show n₁₂ + n₂ = n₁ by rw [← h₁₂, add_assoc, neg_add_cancel, add_zero]) = + z₁.comp (z₂.comp z₃ (neg_add_cancel n₂)) (add_zero n₁) := comp_assoc _ _ _ _ _ (by omega) @[simp] @@ -478,7 +478,7 @@ lemma δ_δ (n₀ n₁ n₂ : ℤ) (z : Cochain F G n₀) : δ n₁ n₂ (δ n ← h₁₂, Int.negOnePow_succ, add_comp, assoc, HomologicalComplex.d_comp_d, comp_zero, zero_add, comp_add, HomologicalComplex.d_comp_d_assoc, zero_comp, smul_zero, - add_zero, add_right_neg, Units.neg_smul, + add_zero, add_neg_cancel, Units.neg_smul, Linear.units_smul_comp, Linear.comp_units_smul] lemma δ_comp {n₁ n₂ n₁₂ : ℤ} (z₁ : Cochain F G n₁) (z₂ : Cochain G K n₂) (h : n₁ + n₂ = n₁₂) @@ -539,8 +539,8 @@ lemma δ_ofHomotopy {φ₁ φ₂ : F ⟶ G} (h : Homotopy φ₁ φ₂) : have eq := h.comm p rw [dNext_eq h.hom (show (ComplexShape.up ℤ).Rel p (p+1) by simp), prevD_eq h.hom (show (ComplexShape.up ℤ).Rel (p-1) p by simp)] at eq - rw [Cochain.ofHomotopy, δ_v (-1) 0 (neg_add_self 1) _ p p (add_zero p) (p-1) (p+1) rfl rfl] - simp only [Cochain.mk_v, add_left_neg, one_smul, Int.negOnePow_zero, + rw [Cochain.ofHomotopy, δ_v (-1) 0 (neg_add_cancel 1) _ p p (add_zero p) (p-1) (p+1) rfl rfl] + simp only [Cochain.mk_v, neg_add_cancel, one_smul, Int.negOnePow_zero, Cochain.sub_v, Cochain.ofHom_v, eq] abel @@ -548,8 +548,8 @@ lemma δ_neg_one_cochain (z : Cochain F G (-1)) : δ (-1) 0 z = Cochain.ofHom (Homotopy.nullHomotopicMap' (fun i j hij => z.v i j (by dsimp at hij; rw [← hij, add_neg_cancel_right]))) := by ext p - rw [δ_v (-1) 0 (neg_add_self 1) _ p p (add_zero p) (p-1) (p+1) rfl rfl] - simp only [neg_add_self, one_smul, Cochain.ofHom_v, Int.negOnePow_zero] + rw [δ_v (-1) 0 (neg_add_cancel 1) _ p p (add_zero p) (p-1) (p+1) rfl rfl] + simp only [neg_add_cancel, one_smul, Cochain.ofHom_v, Int.negOnePow_zero] rw [Homotopy.nullHomotopicMap'_f (show (ComplexShape.up ℤ).Rel (p-1) p by simp) (show (ComplexShape.up ℤ).Rel p (p+1) by simp)] abel diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/MappingCone.lean b/Mathlib/Algebra/Homology/HomotopyCategory/MappingCone.lean index 02c91c5e960fc..41a763809646d 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/MappingCone.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/MappingCone.lean @@ -91,9 +91,9 @@ lemma inr_f_snd_v (p : ℤ) : @[simp] lemma inl_fst : - (inl φ).comp (fst φ).1 (neg_add_self 1) = Cochain.ofHom (𝟙 F) := by + (inl φ).comp (fst φ).1 (neg_add_cancel 1) = Cochain.ofHom (𝟙 F) := by ext p - simp [Cochain.comp_v _ _ (neg_add_self 1) p (p-1) p rfl (by omega)] + simp [Cochain.comp_v _ _ (neg_add_cancel 1) p (p-1) p rfl (by omega)] @[simp] lemma inl_snd : @@ -121,7 +121,7 @@ it is also interesting to have `reassoc` variants of lemmas, like `inl_fst_assoc @[simp] lemma inl_fst_assoc {K : CochainComplex C ℤ} {d e : ℤ} (γ : Cochain F K d) (he : 1 + d = e) : (inl φ).comp ((fst φ).1.comp γ he) (by rw [← he, neg_add_cancel_left]) = γ := by - rw [← Cochain.comp_assoc _ _ _ (neg_add_self 1) (by omega) (by omega), inl_fst, + rw [← Cochain.comp_assoc _ _ _ (neg_add_cancel 1) (by omega) (by omega), inl_fst, Cochain.id_comp] @[simp] @@ -220,15 +220,15 @@ lemma ext_cochain_from_iff (i j : ℤ) (hij : i + 1 = j) exact ⟨h₁, h₂⟩ lemma id : - (fst φ).1.comp (inl φ) (add_neg_self 1) + + (fst φ).1.comp (inl φ) (add_neg_cancel 1) + (snd φ).comp (Cochain.ofHom (inr φ)) (add_zero 0) = Cochain.ofHom (𝟙 _) := by - simp [ext_cochain_from_iff φ (-1) 0 (neg_add_self 1)] + simp [ext_cochain_from_iff φ (-1) 0 (neg_add_cancel 1)] lemma id_X (p q : ℤ) (hpq : p + 1 = q) : (fst φ).1.v p q hpq ≫ (inl φ).v q p (by omega) + (snd φ).v p p (add_zero p) ≫ (inr φ).f p = 𝟙 ((mappingCone φ).X p) := by simpa only [Cochain.add_v, Cochain.comp_zero_cochain_v, Cochain.ofHom_v, id_f, - Cochain.comp_v _ _ (add_neg_self 1) p q p hpq (by omega)] + Cochain.comp_v _ _ (add_neg_cancel 1) p q p hpq (by omega)] using Cochain.congr_v (id φ) p p (add_zero p) @[reassoc] @@ -275,7 +275,7 @@ lemma d_snd_v' (n : ℤ) : lemma δ_inl : δ (-1) 0 (inl φ) = Cochain.ofHom (φ ≫ inr φ) := by ext p - simp [δ_v (-1) 0 (neg_add_self 1) (inl φ) p p (add_zero p) _ _ rfl rfl, + simp [δ_v (-1) 0 (neg_add_cancel 1) (inl φ) p p (add_zero p) _ _ rfl rfl, inl_v_d φ p (p - 1) (p + 1) (by omega) (by omega)] @[simp] @@ -353,11 +353,11 @@ variable {K : CochainComplex C ℤ} (α : Cochain F K (-1)) (β : G ⟶ K) from a cochain `α : Cochain F K (-1)` and a morphism `β : G ⟶ K` such that `δ (-1) 0 α = Cochain.ofHom (φ ≫ β)`. -/ noncomputable def desc : mappingCone φ ⟶ K := - Cocycle.homOf (descCocycle φ α (Cocycle.ofHom β) (neg_add_self 1) (by simp [eq])) + Cocycle.homOf (descCocycle φ α (Cocycle.ofHom β) (neg_add_cancel 1) (by simp [eq])) @[simp] lemma ofHom_desc : - Cochain.ofHom (desc φ α β eq) = descCochain φ α (Cochain.ofHom β) (neg_add_self 1) := by + Cochain.ofHom (desc φ α β eq) = descCochain φ α (Cochain.ofHom β) (neg_add_cancel 1) := by simp [desc] @[reassoc (attr := simp)] @@ -394,8 +394,8 @@ noncomputable def descHomotopy {K : CochainComplex C ℤ} (f₁ f₂ : mappingCo Homotopy f₁ f₂ := (Cochain.equivHomotopy f₁ f₂).symm ⟨descCochain φ γ₁ γ₂ (by norm_num), by simp only [Cochain.ofHom_comp] at h₂ - simp [ext_cochain_from_iff _ _ _ (neg_add_self 1), - δ_descCochain _ _ _ _ _ (neg_add_self 1), h₁, h₂]⟩ + simp [ext_cochain_from_iff _ _ _ (neg_add_cancel 1), + δ_descCochain _ _ _ _ _ (neg_add_cancel 1), h₁, h₂]⟩ section @@ -434,7 +434,7 @@ lemma δ_liftCochain (m' : ℤ) (hm' : m + 1 = m') : (δ n m β + α.comp (Cochain.ofHom φ) (add_zero m)).comp (Cochain.ofHom (inr φ)) (add_zero m) := by dsimp only [liftCochain] - simp only [δ_add, δ_comp α (inl φ) _ m' _ _ h hm' (neg_add_self 1), + simp only [δ_add, δ_comp α (inl φ) _ m' _ _ h hm' (neg_add_cancel 1), δ_comp_zero_cochain _ _ _ h, δ_inl, Cochain.ofHom_comp, Int.negOnePow_neg, Int.negOnePow_one, Units.neg_smul, one_smul, δ_ofHom, Cochain.comp_zero, zero_add, Cochain.add_comp, @@ -505,7 +505,7 @@ noncomputable def liftHomotopy {K : CochainComplex C ℤ} (f₁ f₂ : K ⟶ map δ (-1) 0 β + α.comp (Cochain.ofHom φ) (zero_add 0) + (Cochain.ofHom f₂).comp (snd φ) (zero_add 0)) : Homotopy f₁ f₂ := - (Cochain.equivHomotopy f₁ f₂).symm ⟨liftCochain φ α β (neg_add_self 1), by + (Cochain.equivHomotopy f₁ f₂).symm ⟨liftCochain φ α β (neg_add_cancel 1), by simp [δ_liftCochain _ _ _ _ _ (zero_add 1), ext_cochain_to_iff _ _ _ (zero_add 1), h₁, h₂]⟩ section @@ -538,7 +538,7 @@ lemma lift_desc_f {K L : CochainComplex C ℤ} (α : Cocycle K F 1) (β : Cochai (lift φ α β eq).f n ≫ (desc φ α' β' eq').f n = α.1.v n n' hnn' ≫ α'.v n' n (by omega) + β.v n n (add_zero n) ≫ β'.f n := by simp only [lift, desc, Cocycle.homOf_f, liftCocycle_coe, descCocycle_coe, Cocycle.ofHom_coe, - liftCochain_v_descCochain_v φ α.1 β α' (Cochain.ofHom β') (zero_add 1) (neg_add_self 1) 0 + liftCochain_v_descCochain_v φ α.1 β α' (Cochain.ofHom β') (zero_add 1) (neg_add_cancel 1) 0 (add_zero 0) n n n (add_zero n) (add_zero n) n' hnn', Cochain.ofHom_v] diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean index b6a0a567c16d5..0ad80b275377d 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean @@ -183,13 +183,13 @@ the mapping cone of `inr φ : L ⟶ mappingCone φ`. -/ noncomputable def rotateHomotopyEquiv : HomotopyEquiv (K⟦(1 : ℤ)⟧) (mappingCone (inr φ)) where hom := lift (inr φ) (-(Cocycle.ofHom φ).leftShift 1 1 (zero_add 1)) - (-(inl φ).leftShift 1 0 (neg_add_self 1)) (by - simp? [Cochain.δ_leftShift _ 1 0 1 (neg_add_self 1) 0 (zero_add 1)] says + (-(inl φ).leftShift 1 0 (neg_add_cancel 1)) (by + simp? [Cochain.δ_leftShift _ 1 0 1 (neg_add_cancel 1) 0 (zero_add 1)] says simp only [Int.reduceNeg, δ_neg, - Cochain.δ_leftShift _ 1 0 1 (neg_add_self 1) 0 (zero_add 1), + Cochain.δ_leftShift _ 1 0 1 (neg_add_cancel 1) 0 (zero_add 1), Int.negOnePow_one, δ_inl, Cochain.ofHom_comp, Cochain.leftShift_comp_zero_cochain, Units.neg_smul, one_smul, neg_neg, Cocycle.coe_neg, Cocycle.leftShift_coe, - Cocycle.ofHom_coe, Cochain.neg_comp, add_right_neg]) + Cocycle.ofHom_coe, Cochain.neg_comp, add_neg_cancel]) inv := desc (inr φ) 0 (triangle φ).mor₃ (by simp only [δ_zero, inr_triangleδ, Cochain.ofHom_zero]) homotopyHomInvId := Homotopy.ofEq (by @@ -208,17 +208,17 @@ noncomputable def rotateHomotopyEquiv : ⟨-(snd (inr φ)).comp ((snd φ).comp (inl (inr φ)) (zero_add (-1))) (zero_add (-1)), by ext n simp? [ext_to_iff _ _ (n + 1) rfl, ext_from_iff _ (n + 1) _ rfl, - δ_zero_cochain_comp _ _ _ (neg_add_self 1), - (inl φ).leftShift_v 1 0 (neg_add_self 1) n n (add_zero n) (n + 1) (by omega), + δ_zero_cochain_comp _ _ _ (neg_add_cancel 1), + (inl φ).leftShift_v 1 0 (neg_add_cancel 1) n n (add_zero n) (n + 1) (by omega), (Cochain.ofHom φ).leftShift_v 1 1 (zero_add 1) n (n + 1) rfl (n + 1) (by omega), - Cochain.comp_v _ _ (add_neg_self 1) n (n + 1) n rfl (by omega)] says + Cochain.comp_v _ _ (add_neg_cancel 1) n (n + 1) n rfl (by omega)] says simp only [Int.reduceNeg, Cochain.ofHom_comp, ofHom_desc, ofHom_lift, Cocycle.coe_neg, Cocycle.leftShift_coe, Cocycle.ofHom_coe, Cochain.comp_zero_cochain_v, - shiftFunctor_obj_X', δ_neg, δ_zero_cochain_comp _ _ _ (neg_add_self 1), δ_inl, + shiftFunctor_obj_X', δ_neg, δ_zero_cochain_comp _ _ _ (neg_add_cancel 1), δ_inl, Int.negOnePow_neg, Int.negOnePow_one, δ_snd, Cochain.neg_comp, Cochain.comp_assoc_of_second_is_zero_cochain, smul_neg, Units.neg_smul, one_smul, neg_neg, Cochain.comp_add, inr_snd_assoc, neg_add_rev, Cochain.add_v, Cochain.neg_v, - Cochain.comp_v _ _ (add_neg_self 1) n (n + 1) n rfl (by omega), + Cochain.comp_v _ _ (add_neg_cancel 1) n (n + 1) n rfl (by omega), Cochain.zero_cochain_comp_v, Cochain.ofHom_v, HomologicalComplex.id_f, ext_to_iff _ _ (n + 1) rfl, assoc, liftCochain_v_fst_v, (Cochain.ofHom φ).leftShift_v 1 1 (zero_add 1) n (n + 1) rfl (n + 1) (by omega), @@ -230,9 +230,9 @@ noncomputable def rotateHomotopyEquiv : inr_f_descCochain_v_assoc, inr_f_snd_v_assoc, inl_v_triangle_mor₃_f_assoc, triangle_obj₁, Iso.refl_inv, inl_v_fst_v_assoc, inr_f_triangle_mor₃_f_assoc, inr_f_fst_v_assoc, and_self, liftCochain_v_snd_v, - (inl φ).leftShift_v 1 0 (neg_add_self 1) n n (add_zero n) (n + 1) (by omega), + (inl φ).leftShift_v 1 0 (neg_add_cancel 1) n n (add_zero n) (n + 1) (by omega), Int.negOnePow_zero, inl_v_snd_v, inr_f_snd_v, zero_add, inl_v_descCochain_v, - inr_f_descCochain_v, inl_v_triangle_mor₃_f, inr_f_triangle_mor₃_f, add_left_neg]⟩ + inr_f_descCochain_v, inl_v_triangle_mor₃_f, inr_f_triangle_mor₃_f, neg_add_cancel]⟩ /-- Auxiliary definition for `rotateTrianglehIso`. -/ noncomputable def rotateHomotopyEquivComm₂Homotopy : @@ -242,9 +242,9 @@ noncomputable def rotateHomotopyEquivComm₂Homotopy : ext p dsimp [rotateHomotopyEquiv] simp [ext_from_iff _ _ _ rfl, ext_to_iff _ _ _ rfl, - (inl φ).leftShift_v 1 0 (neg_add_self 1) p p (add_zero p) (p + 1) (by omega), - δ_zero_cochain_comp _ _ _ (neg_add_self 1), - Cochain.comp_v _ _ (add_neg_self 1) p (p + 1) p rfl (by omega), + (inl φ).leftShift_v 1 0 (neg_add_cancel 1) p p (add_zero p) (p + 1) (by omega), + δ_zero_cochain_comp _ _ _ (neg_add_cancel 1), + Cochain.comp_v _ _ (add_neg_cancel 1) p (p + 1) p rfl (by omega), (Cochain.ofHom φ).leftShift_v 1 1 (zero_add 1) p (p + 1) rfl (p + 1) (by omega)]⟩ @[reassoc (attr := simp)] @@ -286,7 +286,7 @@ noncomputable def shiftIso (n : ℤ) : (mappingCone φ)⟦n⟧ ≅ mappingCone ( dsimp simp only [Cochain.δ_shift, δ_snd, Cochain.shift_neg, smul_neg, Cochain.neg_v, shiftFunctor_obj_X', Cochain.units_smul_v, Cochain.shift_v', Cochain.comp_zero_cochain_v, - Cochain.ofHom_v, Cochain.units_smul_comp, shiftFunctor_map_f', add_left_neg]) + Cochain.ofHom_v, Cochain.units_smul_comp, shiftFunctor_map_f', neg_add_cancel]) inv := desc _ (n.negOnePow • (inl φ).shift n) ((inr φ)⟦n⟧') (by ext p dsimp diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean index 9c7b600049093..b49421f3f25b8 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean @@ -58,7 +58,7 @@ the mapping cone of the morphism `mappingCone f ⟶ mappingCone (f ≫ g)`. -/ noncomputable def hom : mappingCone g ⟶ mappingCone (mappingConeCompTriangle f g).mor₁ := lift _ (descCocycle g (Cochain.ofHom (inr f)) 0 (zero_add 1) (by dsimp; simp)) - (descCochain _ 0 (Cochain.ofHom (inr (f ≫ g))) (neg_add_self 1)) (by + (descCochain _ 0 (Cochain.ofHom (inr (f ≫ g))) (neg_add_cancel 1)) (by ext p _ rfl simp [mappingConeCompTriangle, map, ext_from_iff _ _ _ rfl, inl_v_d_assoc _ (p+1) p (p+2) (by linarith) (by linarith)]) @@ -72,7 +72,7 @@ noncomputable def inv : mappingCone (mappingConeCompTriangle f g).mor₁ ⟶ map ext p rw [ext_from_iff _ (p + 1) _ rfl, ext_to_iff _ _ (p + 1) rfl] simp [map, δ_zero_cochain_comp, - Cochain.comp_v _ _ (add_neg_self 1) p (p+1) p (by linarith) (by linarith)]) + Cochain.comp_v _ _ (add_neg_cancel 1) p (p+1) p (by linarith) (by linarith)]) @[reassoc (attr := simp)] lemma hom_inv_id : hom f g ≫ inv f g = 𝟙 _ := by @@ -87,7 +87,7 @@ the morphism `mappingCone f ⟶ mappingCone (f ≫ g)`. -/ noncomputable def homotopyInvHomId : Homotopy (inv f g ≫ hom f g) (𝟙 _) := (Cochain.equivHomotopy _ _).symm ⟨-((snd _).comp ((fst (f ≫ g)).1.comp ((inl f).comp (inl _) (by linarith)) (show 1 + (-2) = -1 by linarith)) (zero_add (-1))), by - rw [δ_neg, δ_zero_cochain_comp _ _ _ (neg_add_self 1), + rw [δ_neg, δ_zero_cochain_comp _ _ _ (neg_add_cancel 1), Int.negOnePow_neg, Int.negOnePow_one, Units.neg_smul, one_smul, δ_comp _ _ (show 1 + (-2) = -1 by linarith) 2 (-1) 0 (by linarith) (by linarith) (by linarith), @@ -100,7 +100,7 @@ noncomputable def homotopyInvHomId : Homotopy (inv f g ≫ hom f g) (𝟙 _) := rw [ext_from_iff _ (n + 1) n rfl, ext_from_iff _ (n + 1) n rfl, ext_from_iff _ (n + 2) (n + 1) (by linarith)] simp? [hom, inv, ext_to_iff _ n (n + 1) rfl, map, Cochain.comp_v _ _ - (add_neg_self 1) n (n + 1) n (by linarith) (by linarith), + (add_neg_cancel 1) n (n + 1) n (by linarith) (by linarith), Cochain.comp_v _ _ (show 1 + -2 = -1 by linarith) (n + 1) (n + 2) n (by linarith) (by linarith), Cochain.comp_v _ _ (show (-1) + -1 = -2 by linarith) (n + 2) (n + 1) n @@ -113,13 +113,13 @@ noncomputable def homotopyInvHomId : Homotopy (inv f g ≫ hom f g) (𝟙 _) := Cochain.comp_assoc_of_first_is_zero_cochain, Cochain.comp_add, Cochain.comp_neg, Cochain.comp_assoc_of_second_is_zero_cochain, neg_add_rev, neg_neg, Cochain.add_v, Cochain.neg_v, - Cochain.comp_v _ _ (add_neg_self 1) n (n + 1) n (by linarith) (by linarith), + Cochain.comp_v _ _ (add_neg_cancel 1) n (n + 1) n (by linarith) (by linarith), Cochain.comp_v _ _ (show 1 + -2 = -1 by linarith) (n + 1) (n + 2) n (by linarith) (by linarith), Cochain.comp_v _ _ (show (-1) + -1 = -2 by linarith) (n + 2) (n + 1) n (by linarith) (by linarith), Cochain.ofHom_v, HomologicalComplex.id_f, Preadditive.comp_add, Preadditive.comp_neg, - inl_v_fst_v_assoc, neg_zero, add_zero, comp_id, add_left_neg, inr_f_snd_v_assoc, + inl_v_fst_v_assoc, neg_zero, add_zero, comp_id, neg_add_cancel, inr_f_snd_v_assoc, inr_f_descCochain_v_assoc, inr_f_fst_v_assoc, comp_zero, zero_add, ext_to_iff _ n (n + 1) rfl, liftCochain_v_fst_v, inl_v_descCochain_v, inl_v_fst_v, liftCochain_v_snd_v, Cochain.zero_v, inl_v_snd_v, and_self, neg_add_cancel_right, diff --git a/Mathlib/Algebra/Homology/ShortComplex/Preadditive.lean b/Mathlib/Algebra/Homology/ShortComplex/Preadditive.lean index 3c2ccb5513fef..dee25030965f5 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Preadditive.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Preadditive.lean @@ -49,7 +49,7 @@ instance : AddCommGroup (S₁ ⟶ S₂) where add_assoc := fun a b c => by ext <;> apply add_assoc add_zero := fun a => by ext <;> apply add_zero zero_add := fun a => by ext <;> apply zero_add - add_left_neg := fun a => by ext <;> apply add_left_neg + neg_add_cancel := fun a => by ext <;> apply neg_add_cancel add_comm := fun a b => by ext <;> apply add_comm sub_eq_add_neg := fun a b => by ext <;> apply sub_eq_add_neg nsmul := nsmulRec diff --git a/Mathlib/Algebra/Lie/BaseChange.lean b/Mathlib/Algebra/Lie/BaseChange.lean index 234a24fe17928..6ad77f8f9d079 100644 --- a/Mathlib/Algebra/Lie/BaseChange.lean +++ b/Mathlib/Algebra/Lie/BaseChange.lean @@ -73,7 +73,7 @@ private theorem bracket_lie_self (x : A ⊗[R] L) : ⁅x, x⁆ = 0 := by · simp only [LinearMap.map_zero, add_zero, LinearMap.zero_apply] · intro a₂ l₂ simp only [← lie_skew l₂ l₁, mul_comm a₁ a₂, TensorProduct.tmul_neg, bracket'_tmul, - add_right_neg] + add_neg_cancel] · intro y₁ y₂ hy₁ hy₂ simp only [hy₁, hy₂, add_add_add_comm, add_zero, LinearMap.add_apply, LinearMap.map_add] · intro y₁ y₂ hy₁ hy₂ diff --git a/Mathlib/Algebra/Lie/Weights/Cartan.lean b/Mathlib/Algebra/Lie/Weights/Cartan.lean index a7869de313652..a78eaf08e3214 100644 --- a/Mathlib/Algebra/Lie/Weights/Cartan.lean +++ b/Mathlib/Algebra/Lie/Weights/Cartan.lean @@ -258,9 +258,9 @@ Note that the name "coroot space" is not standard as this space does not seem to informal literature. -/ def corootSpace : LieIdeal R H := LieModuleHom.range <| ((rootSpace H 0).incl.comp <| - rootSpaceProduct R L H α (-α) 0 (add_neg_self α)).codRestrict H.toLieSubmodule (by + rootSpaceProduct R L H α (-α) 0 (add_neg_cancel α)).codRestrict H.toLieSubmodule (by rw [← rootSpace_zero_eq] - exact fun p ↦ (rootSpaceProduct R L H α (-α) 0 (add_neg_self α) p).property) + exact fun p ↦ (rootSpaceProduct R L H α (-α) 0 (add_neg_cancel α) p).property) lemma mem_corootSpace {x : H} : x ∈ corootSpace α ↔ @@ -291,7 +291,7 @@ lemma mem_corootSpace' {x : H} : exists_and_right, exists_eq_right, mem_setOf_eq, s] refine ⟨fun ⟨_, y, hy, z, hz, hyz⟩ ↦ ⟨y, hy, z, hz, hyz⟩, fun ⟨y, hy, z, hz, hyz⟩ ↦ ⟨?_, y, hy, z, hz, hyz⟩⟩ - convert (rootSpaceProduct R L H α (-α) 0 (add_neg_self α) (⟨y, hy⟩ ⊗ₜ[R] ⟨z, hz⟩)).property + convert (rootSpaceProduct R L H α (-α) 0 (add_neg_cancel α) (⟨y, hy⟩ ⊗ₜ[R] ⟨z, hz⟩)).property simp [hyz] end LieAlgebra diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean index 76f58fc396815..cee61eb9c3bc1 100644 --- a/Mathlib/Algebra/Lie/Weights/Killing.lean +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -382,7 +382,7 @@ lemma root_apply_cartanEquivDual_symm_ne_zero {α : Weight K H L} (hα : α.IsNo lemma root_apply_coroot {α : Weight K H L} (hα : α.IsNonZero) : α (coroot α) = 2 := by rw [← Weight.coe_coe] - simpa [coroot] using inv_mul_cancel (root_apply_cartanEquivDual_symm_ne_zero hα) + simpa [coroot] using inv_mul_cancel₀ (root_apply_cartanEquivDual_symm_ne_zero hα) @[simp] lemma coroot_eq_zero_iff {α : Weight K H L} : coroot α = 0 ↔ α.IsZero := by @@ -490,7 +490,7 @@ lemma exists_isSl2Triple_of_weight_isNonZero {α : Weight K H L} (hα : α.IsNon let f := (2 * (α h)⁻¹) • f' replace hef : ⁅⁅e, f⁆, e⁆ = 2 • e := by have : ⁅⁅e, f'⁆, e⁆ = α h • e := lie_eq_smul_of_mem_rootSpace heα h - rw [lie_smul, smul_lie, this, ← smul_assoc, smul_eq_mul, mul_assoc, inv_mul_cancel hh, + rw [lie_smul, smul_lie, this, ← smul_assoc, smul_eq_mul, mul_assoc, inv_mul_cancel₀ hh, mul_one, two_smul, two_smul] refine ⟨⁅e, f⁆, e, f, ⟨fun contra ↦ ?_, rfl, hef, ?_⟩, heα, Submodule.smul_mem _ _ hfα⟩ · rw [contra] at hef @@ -523,7 +523,7 @@ lemma _root_.IsSl2Triple.h_eq_coroot {α : Weight K H L} (hα : α.IsNonZero) have := ht.h_ne_zero contrapose! this simpa [this] using h_eq - rw [h_eq, smul_smul, mul_assoc, inv_mul_cancel hef₀, mul_one, smul_assoc, coroot] + rw [h_eq, smul_smul, mul_assoc, inv_mul_cancel₀ hef₀, mul_one, smul_assoc, coroot] lemma finrank_rootSpace_eq_one (α : Weight K H L) (hα : α.IsNonZero) : finrank K (rootSpace H α) = 1 := by diff --git a/Mathlib/Algebra/Lie/Weights/RootSystem.lean b/Mathlib/Algebra/Lie/Weights/RootSystem.lean index b38901b4a096c..e4cecabbee665 100644 --- a/Mathlib/Algebra/Lie/Weights/RootSystem.lean +++ b/Mathlib/Algebra/Lie/Weights/RootSystem.lean @@ -158,7 +158,7 @@ lemma chainBotCoeff_add_chainTopCoeff : intro e apply rootSpace_neg_nsmul_add_chainTop_of_le α β e rw [← Nat.succ_add, ← Nat.cast_smul_eq_nsmul ℤ, ← neg_smul, coe_chainTop, ← add_assoc, - ← add_smul, Nat.cast_add, neg_add, add_assoc, neg_add_self, add_zero, neg_smul, ← smul_neg, + ← add_smul, Nat.cast_add, neg_add, add_assoc, neg_add_cancel, add_zero, neg_smul, ← smul_neg, Nat.cast_smul_eq_nsmul] exact weightSpace_chainTopCoeff_add_one_nsmul_add (-α) β (Weight.IsNonZero.neg hα) @@ -194,7 +194,7 @@ lemma le_chainBotCoeff_of_rootSpace_ne_top chainBotCoeff_add_chainTopCoeff] at hn have := rootSpace_neg_nsmul_add_chainTop_of_lt α β hα hn rwa [← Nat.cast_smul_eq_nsmul ℤ, ← neg_smul, coe_chainTop, ← add_assoc, - ← add_smul, Nat.cast_add, neg_add, add_assoc, neg_add_self, add_zero] at this + ← add_smul, Nat.cast_add, neg_add, add_assoc, neg_add_cancel, add_zero] at this /-- Members of the `α`-chain through `β` are the only roots of the form `β - kα`. -/ lemma rootSpace_zsmul_add_ne_bot_iff (hα : α.IsNonZero) (n : ℤ) : @@ -245,7 +245,7 @@ lemma chainLength_of_eq_zsmul_add (β' : Weight K H L) (n : ℤ) (hβ' : (β' : rw [← chainTopCoeff_add_chainBotCoeff, ← chainTopCoeff_add_chainBotCoeff, Nat.cast_add, Nat.cast_add, chainTopCoeff_of_eq_zsmul_add α β hα β' n hβ', chainBotCoeff_of_eq_zsmul_add α β hα β' n hβ', sub_eq_add_neg, add_add_add_comm, - neg_add_self, add_zero] + neg_add_cancel, add_zero] lemma chainTopCoeff_zero_right [Nontrivial L] (hα : α.IsNonZero) : chainTopCoeff α (0 : Weight K H L) = 1 := by @@ -266,7 +266,7 @@ lemma chainTopCoeff_zero_right [Nontrivial L] (hα : α.IsNonZero) : have : (toEnd K L L f ^ (chainTopCoeff α (0 : Weight K H L) + 1)) x ∈ rootSpace H (-α) := by convert toEnd_pow_apply_mem hf hx (chainTopCoeff α (0 : Weight K H L) + 1) using 2 rw [coe_chainTop', Weight.coe_zero, add_zero, succ_nsmul', - add_assoc, smul_neg, neg_add_self, add_zero] + add_assoc, smul_neg, neg_add_cancel, add_zero] simpa using (finrank_eq_one_iff_of_nonzero' ⟨f, hf⟩ (by simpa using isSl2.f_ne_zero)).mp (finrank_rootSpace_eq_one _ hα.neg) ⟨_, this⟩ apply_fun (⁅f, ·⁆) at hk diff --git a/Mathlib/Algebra/Module/CharacterModule.lean b/Mathlib/Algebra/Module/CharacterModule.lean index bd0b99a07a4fe..ec58e2bbe758b 100644 --- a/Mathlib/Algebra/Module/CharacterModule.lean +++ b/Mathlib/Algebra/Module/CharacterModule.lean @@ -140,7 +140,7 @@ protected lemma int.divByNat_self (n : ℕ) : obtain rfl | h0 := eq_or_ne n 0 · apply map_zero exact (AddCircle.coe_eq_zero_iff _).mpr - ⟨1, by simp [mul_inv_cancel (Nat.cast_ne_zero (R := ℚ).mpr h0)]⟩ + ⟨1, by simp [mul_inv_cancel₀ (Nat.cast_ne_zero (R := ℚ).mpr h0)]⟩ variable {A} diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index 058a63cdb7344..53aee9653932d 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -226,7 +226,7 @@ variable [Ring R] [AddCommGroup M] [Module R M] (r s : R) (x y : M) @[simp] theorem neg_smul : -r • x = -(r • x) := - eq_neg_of_add_eq_zero_left <| by rw [← add_smul, add_left_neg, zero_smul] + eq_neg_of_add_eq_zero_left <| by rw [← add_smul, neg_add_cancel, zero_smul] -- Porting note (#10618): simp can prove this --@[simp] @@ -256,10 +256,10 @@ abbrev Module.addCommMonoidToAddCommGroup [Ring R] [AddCommMonoid M] [Module R M] : AddCommGroup M := { (inferInstance : AddCommMonoid M) with neg := fun a => (-1 : R) • a - add_left_neg := fun a => + neg_add_cancel := fun a => show (-1 : R) • a + a = 0 by nth_rw 2 [← one_smul R a] - rw [← add_smul, add_left_neg, zero_smul] + rw [← add_smul, neg_add_cancel, zero_smul] zsmul := fun z a => (z : R) • a zsmul_zero' := fun a => by simpa only [Int.cast_zero] using zero_smul R a zsmul_succ' := fun z a => by simp [add_comm, add_smul] diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index 41c09bf110bf0..fcb550321d2d3 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -81,7 +81,7 @@ instance automorphismGroup : Group (M ≃ₗ[R] M) where mul_assoc f g h := rfl mul_one f := ext fun x ↦ rfl one_mul f := ext fun x ↦ rfl - mul_left_inv f := ext <| f.left_inv + inv_mul_cancel f := ext <| f.left_inv @[simp] lemma coe_one : ↑(1 : M ≃ₗ[R] M) = id := rfl diff --git a/Mathlib/Algebra/Module/LocalizedModule.lean b/Mathlib/Algebra/Module/LocalizedModule.lean index fbd7276dd85e0..ff78e7064cfa4 100644 --- a/Mathlib/Algebra/Module/LocalizedModule.lean +++ b/Mathlib/Algebra/Module/LocalizedModule.lean @@ -208,7 +208,7 @@ instance {M : Type*} [AddCommGroup M] [Module R M] : Neg (LocalizedModule S M) w instance {M : Type*} [AddCommGroup M] [Module R M] : AddCommGroup (LocalizedModule S M) := { show AddCommMonoid (LocalizedModule S M) by infer_instance with - add_left_neg := by + neg_add_cancel := by rintro ⟨m, s⟩ change (liftOn (mk m s) (fun x => mk (-x.1) x.2) fun ⟨m1, s1⟩ ⟨m2, s2⟩ ⟨u, hu⟩ => by diff --git a/Mathlib/Algebra/Order/AddGroupWithTop.lean b/Mathlib/Algebra/Order/AddGroupWithTop.lean index 2ebd3144d382b..a977f5a486e68 100644 --- a/Mathlib/Algebra/Order/AddGroupWithTop.lean +++ b/Mathlib/Algebra/Order/AddGroupWithTop.lean @@ -111,7 +111,7 @@ instance : LinearOrderedAddCommGroupWithTop (WithTop α) where add_neg_cancel := by rintro (a | a) ha · exact (ha rfl).elim - · exact (WithTop.coe_add ..).symm.trans (WithTop.coe_eq_coe.2 (add_neg_self a)) + · exact (WithTop.coe_add ..).symm.trans (WithTop.coe_eq_coe.2 (add_neg_cancel a)) end LinearOrderedAddCommGroup diff --git a/Mathlib/Algebra/Order/CauSeq/Basic.lean b/Mathlib/Algebra/Order/CauSeq/Basic.lean index 3688f98cebaca..9523087b69005 100644 --- a/Mathlib/Algebra/Order/CauSeq/Basic.lean +++ b/Mathlib/Algebra/Order/CauSeq/Basic.lean @@ -75,7 +75,7 @@ theorem rat_inv_continuous_lemma {β : Type*} [DivisionRing β] (abv : β → α rw [inv_sub_inv' ((abv_pos abv).1 a0) ((abv_pos abv).1 b0), abv_mul abv, abv_mul abv, abv_inv abv, abv_inv abv, abv_sub abv] refine lt_of_mul_lt_mul_left (lt_of_mul_lt_mul_right ?_ b0.le) a0.le - rw [mul_assoc, inv_mul_cancel_right₀ b0.ne', ← mul_assoc, mul_inv_cancel a0.ne', one_mul] + rw [mul_assoc, inv_mul_cancel_right₀ b0.ne', ← mul_assoc, mul_inv_cancel₀ a0.ne', one_mul] refine h.trans_le ?_ gcongr diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 61dd3d24f00a7..2dee94a512441 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -760,7 +760,7 @@ theorem sub_one_div_inv_le_two (a2 : 2 ≤ a) : (1 - 1 / a)⁻¹ ≤ 2 := by -- take inverses on both sides and use the assumption `2 ≤ a`. convert (one_div a).le.trans (inv_le_inv_of_le zero_lt_two a2) using 1 -- show `1 - 1 / 2 = 1 / 2`. - rw [sub_eq_iff_eq_add, ← two_mul, mul_inv_cancel two_ne_zero] + rw [sub_eq_iff_eq_add, ← two_mul, mul_inv_cancel₀ two_ne_zero] /-! ### Results about `IsLUB` -/ diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 6e2e7f168de95..517130d8d4219 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -1375,7 +1375,7 @@ theorem round_two_inv : round (2⁻¹ : α) = 1 := by @[simp] theorem round_neg_two_inv : round (-2⁻¹ : α) = 0 := by - simp only [round_eq, ← one_div, add_left_neg, floor_zero] + simp only [round_eq, ← one_div, neg_add_cancel, floor_zero] @[simp] theorem round_eq_zero_iff {x : α} : round x = 0 ↔ x ∈ Ico (-(1 / 2)) ((1 : α) / 2) := by diff --git a/Mathlib/Algebra/Order/Group/PosPart.lean b/Mathlib/Algebra/Order/Group/PosPart.lean index 6f473a0f6d627..8afd1f0c6ea35 100644 --- a/Mathlib/Algebra/Order/Group/PosPart.lean +++ b/Mathlib/Algebra/Order/Group/PosPart.lean @@ -124,7 +124,7 @@ lemma leOnePart_eq_one : a⁻ᵐ = 1 ↔ 1 ≤ a := by simp [leOnePart_eq_one'] -- Bourbaki A.VI.12 Prop 9 a) @[to_additive (attr := simp)] lemma oneLePart_div_leOnePart (a : α) : a⁺ᵐ / a⁻ᵐ = a := by - rw [div_eq_mul_inv, mul_inv_eq_iff_eq_mul, leOnePart, mul_sup, mul_one, mul_right_inv, sup_comm, + rw [div_eq_mul_inv, mul_inv_eq_iff_eq_mul, leOnePart, mul_sup, mul_one, mul_inv_cancel, sup_comm, oneLePart] @[to_additive (attr := simp)] lemma leOnePart_div_oneLePart (a : α) : a⁻ᵐ / a⁺ᵐ = a⁻¹ := by @@ -142,19 +142,19 @@ lemma leOnePart_eq_inv_inf_one (a : α) : a⁻ᵐ = (a ⊓ 1)⁻¹ := by -- Bourbaki A.VI.12 Prop 9 d) @[to_additive] lemma oneLePart_mul_leOnePart (a : α) : a⁺ᵐ * a⁻ᵐ = |a|ₘ := by - rw [oneLePart, sup_mul, one_mul, leOnePart, mul_sup, mul_one, mul_inv_self, sup_assoc, + rw [oneLePart, sup_mul, one_mul, leOnePart, mul_sup, mul_one, mul_inv_cancel, sup_assoc, ← sup_assoc a, sup_eq_right.2 le_sup_right] exact sup_eq_left.2 <| one_le_mabs a @[to_additive] lemma leOnePart_mul_oneLePart (a : α) : a⁻ᵐ * a⁺ᵐ = |a|ₘ := by - rw [oneLePart, mul_sup, mul_one, leOnePart, sup_mul, one_mul, inv_mul_self, sup_assoc, + rw [oneLePart, mul_sup, mul_one, leOnePart, sup_mul, one_mul, inv_mul_cancel, sup_assoc, ← @sup_assoc _ _ a, sup_eq_right.2 le_sup_right] exact sup_eq_left.2 <| one_le_mabs a -- Bourbaki A.VI.12 Prop 9 a) -- a⁺ᵐ ⊓ a⁻ᵐ = 0 (`a⁺` and `a⁻` are co-prime, and, since they are positive, disjoint) @[to_additive] lemma oneLePart_inf_leOnePart_eq_one (a : α) : a⁺ᵐ ⊓ a⁻ᵐ = 1 := by - rw [← mul_left_inj a⁻ᵐ⁻¹, inf_mul, one_mul, mul_right_inv, ← div_eq_mul_inv, + rw [← mul_left_inj a⁻ᵐ⁻¹, inf_mul, one_mul, mul_inv_cancel, ← div_eq_mul_inv, oneLePart_div_leOnePart, leOnePart_eq_inv_inf_one, inv_inv] end covariantmulop diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean index 270be2e5a791c..5ad660f7f4f37 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Abs.lean @@ -93,7 +93,7 @@ variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] @[to_additive (attr := simp) abs_nonneg] lemma one_le_mabs (a : α) : 1 ≤ |a|ₘ := by apply pow_two_semiclosed _ - rw [mabs, pow_two, mul_sup, sup_mul, ← pow_two, mul_left_inv, sup_comm, ← sup_assoc] + rw [mabs, pow_two, mul_sup, sup_mul, ← pow_two, inv_mul_cancel, sup_comm, ← sup_assoc] apply le_sup_right @[to_additive (attr := simp)] lemma mabs_mabs (a : α) : |(|a|ₘ)|ₘ = |a|ₘ := @@ -224,7 +224,7 @@ variable [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} · simp [mabs_of_le_one h] @[to_additive add_abs_nonneg] lemma one_le_mul_mabs (a : α) : 1 ≤ a * |a|ₘ := by - rw [← mul_right_inv a]; exact mul_le_mul_left' (inv_le_mabs a) _ + rw [← mul_inv_cancel a]; exact mul_le_mul_left' (inv_le_mabs a) _ @[to_additive] lemma inv_mabs_le_inv (a : α) : |a|ₘ⁻¹ ≤ a⁻¹ := by simpa using inv_mabs_le a⁻¹ diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean index 2014aa988c223..f0e95176dbb3d 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Basic.lean @@ -58,11 +58,11 @@ theorem inv_mul_le_iff_le_mul : b⁻¹ * a ≤ c ↔ a ≤ b * c := by @[to_additive neg_le_iff_add_nonneg'] theorem inv_le_iff_one_le_mul' : a⁻¹ ≤ b ↔ 1 ≤ a * b := - (mul_le_mul_iff_left a).symm.trans <| by rw [mul_inv_self] + (mul_le_mul_iff_left a).symm.trans <| by rw [mul_inv_cancel] @[to_additive] theorem le_inv_iff_mul_le_one_left : a ≤ b⁻¹ ↔ b * a ≤ 1 := - (mul_le_mul_iff_left b).symm.trans <| by rw [mul_inv_self] + (mul_le_mul_iff_left b).symm.trans <| by rw [mul_inv_cancel] @[to_additive] theorem le_inv_mul_iff_le : 1 ≤ b⁻¹ * a ↔ b ≤ a := by @@ -82,12 +82,12 @@ variable [LT α] [CovariantClass α α (· * ·) (· < ·)] {a b c : α} /-- Uses `left` co(ntra)variant. -/ @[to_additive (attr := simp) Left.neg_pos_iff "Uses `left` co(ntra)variant."] theorem Left.one_lt_inv_iff : 1 < a⁻¹ ↔ a < 1 := by - rw [← mul_lt_mul_iff_left a, mul_inv_self, mul_one] + rw [← mul_lt_mul_iff_left a, mul_inv_cancel, mul_one] /-- Uses `left` co(ntra)variant. -/ @[to_additive (attr := simp) "Uses `left` co(ntra)variant."] theorem Left.inv_lt_one_iff : a⁻¹ < 1 ↔ 1 < a := by - rw [← mul_lt_mul_iff_left a, mul_inv_self, mul_one] + rw [← mul_lt_mul_iff_left a, mul_inv_cancel, mul_one] @[to_additive (attr := simp)] theorem lt_inv_mul_iff_mul_lt : b < a⁻¹ * c ↔ a * b < c := by @@ -100,11 +100,11 @@ theorem inv_mul_lt_iff_lt_mul : b⁻¹ * a < c ↔ a < b * c := by @[to_additive] theorem inv_lt_iff_one_lt_mul' : a⁻¹ < b ↔ 1 < a * b := - (mul_lt_mul_iff_left a).symm.trans <| by rw [mul_inv_self] + (mul_lt_mul_iff_left a).symm.trans <| by rw [mul_inv_cancel] @[to_additive] theorem lt_inv_iff_mul_lt_one' : a < b⁻¹ ↔ b * a < 1 := - (mul_lt_mul_iff_left b).symm.trans <| by rw [mul_inv_self] + (mul_lt_mul_iff_left b).symm.trans <| by rw [mul_inv_cancel] @[to_additive] theorem lt_inv_mul_iff_lt : 1 < b⁻¹ * a ↔ b < a := by @@ -134,11 +134,11 @@ theorem Right.one_le_inv_iff : 1 ≤ a⁻¹ ↔ a ≤ 1 := by @[to_additive neg_le_iff_add_nonneg] theorem inv_le_iff_one_le_mul : a⁻¹ ≤ b ↔ 1 ≤ b * a := - (mul_le_mul_iff_right a).symm.trans <| by rw [inv_mul_self] + (mul_le_mul_iff_right a).symm.trans <| by rw [inv_mul_cancel] @[to_additive] theorem le_inv_iff_mul_le_one_right : a ≤ b⁻¹ ↔ a * b ≤ 1 := - (mul_le_mul_iff_right b).symm.trans <| by rw [inv_mul_self] + (mul_le_mul_iff_right b).symm.trans <| by rw [inv_mul_cancel] @[to_additive (attr := simp)] theorem mul_inv_le_iff_le_mul : a * b⁻¹ ≤ c ↔ a ≤ c * b := @@ -170,20 +170,20 @@ variable [LT α] [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} /-- Uses `right` co(ntra)variant. -/ @[to_additive (attr := simp) "Uses `right` co(ntra)variant."] theorem Right.inv_lt_one_iff : a⁻¹ < 1 ↔ 1 < a := by - rw [← mul_lt_mul_iff_right a, inv_mul_self, one_mul] + rw [← mul_lt_mul_iff_right a, inv_mul_cancel, one_mul] /-- Uses `right` co(ntra)variant. -/ @[to_additive (attr := simp) Right.neg_pos_iff "Uses `right` co(ntra)variant."] theorem Right.one_lt_inv_iff : 1 < a⁻¹ ↔ a < 1 := by - rw [← mul_lt_mul_iff_right a, inv_mul_self, one_mul] + rw [← mul_lt_mul_iff_right a, inv_mul_cancel, one_mul] @[to_additive] theorem inv_lt_iff_one_lt_mul : a⁻¹ < b ↔ 1 < b * a := - (mul_lt_mul_iff_right a).symm.trans <| by rw [inv_mul_self] + (mul_lt_mul_iff_right a).symm.trans <| by rw [inv_mul_cancel] @[to_additive] theorem lt_inv_iff_mul_lt_one : a < b⁻¹ ↔ a * b < 1 := - (mul_lt_mul_iff_right b).symm.trans <| by rw [inv_mul_self] + (mul_lt_mul_iff_right b).symm.trans <| by rw [inv_mul_cancel] @[to_additive (attr := simp)] theorem mul_inv_lt_iff_lt_mul : a * b⁻¹ < c ↔ a < c * b := by diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index fa19028c7999f..3984c64e8056f 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -241,7 +241,7 @@ instance : LinearOrderedAddCommGroupWithTop (Additive αᵒᵈ) := { Additive.subNegMonoid, instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual, Additive.instNontrivial with neg_top := set_option backward.isDefEq.lazyProjDelta false in @inv_zero _ (_) - 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 : Additive.toMul a ≠ 0) } lemma pow_lt_pow_succ (ha : 1 < a) : a ^ n < a ^ n.succ := by rw [← one_mul (a ^ n), pow_succ'] diff --git a/Mathlib/Algebra/Order/Positive/Field.lean b/Mathlib/Algebra/Order/Positive/Field.lean index 8b843e5dc8c3b..8178078ac60be 100644 --- a/Mathlib/Algebra/Order/Positive/Field.lean +++ b/Mathlib/Algebra/Order/Positive/Field.lean @@ -35,6 +35,6 @@ theorem coe_zpow (x : { x : K // 0 < x }) (n : ℤ) : ↑(x ^ n) = (x : K) ^ n : instance : LinearOrderedCommGroup { x : K // 0 < x } := { Positive.Subtype.inv, Positive.linearOrderedCancelCommMonoid with - mul_left_inv := fun a => Subtype.ext <| inv_mul_cancel a.2.ne' } + inv_mul_cancel := fun a => Subtype.ext <| inv_mul_cancel₀ a.2.ne' } end Positive diff --git a/Mathlib/Algebra/Order/Ring/Cast.lean b/Mathlib/Algebra/Order/Ring/Cast.lean index 0b74081308b9e..98e9a362cfbdd 100644 --- a/Mathlib/Algebra/Order/Ring/Cast.lean +++ b/Mathlib/Algebra/Order/Ring/Cast.lean @@ -86,10 +86,10 @@ lemma cast_le_neg_one_or_one_le_cast_of_ne_zero (hn : n ≠ 0) : (n : R) ≤ -1 lemma nneg_mul_add_sq_of_abs_le_one (n : ℤ) (hx : |x| ≤ 1) : (0 : R) ≤ n * x + n * n := by have hnx : 0 < n → 0 ≤ x + n := fun hn => by have := _root_.add_le_add (neg_le_of_abs_le hx) (cast_one_le_of_pos hn) - rwa [add_left_neg] at this + rwa [neg_add_cancel] at this have hnx' : n < 0 → x + n ≤ 0 := fun hn => by have := _root_.add_le_add (le_of_abs_le hx) (cast_le_neg_one_of_neg hn) - rwa [add_right_neg] at this + rwa [add_neg_cancel] at this rw [← mul_add, mul_nonneg_iff] rcases lt_trichotomy n 0 with (h | rfl | h) · exact Or.inr ⟨mod_cast h.le, hnx' h⟩ diff --git a/Mathlib/Algebra/Order/ToIntervalMod.lean b/Mathlib/Algebra/Order/ToIntervalMod.lean index 706b51d23eb81..749e1b6de5802 100644 --- a/Mathlib/Algebra/Order/ToIntervalMod.lean +++ b/Mathlib/Algebra/Order/ToIntervalMod.lean @@ -528,7 +528,7 @@ theorem tfae_modEq : · intro h rw [h, eq_comm, toIocMod_eq_iff, Set.right_mem_Ioc] refine ⟨lt_add_of_pos_right a hp, toIcoDiv hp a b - 1, ?_⟩ - rw [sub_one_zsmul, add_add_add_comm, add_right_neg, add_zero] + rw [sub_one_zsmul, add_add_add_comm, add_neg_cancel, add_zero] conv_lhs => rw [← toIcoMod_add_toIcoDiv_zsmul hp a b, h] tfae_have 2 → 1 · rw [← not_exists, not_imp_comm] diff --git a/Mathlib/Algebra/PUnitInstances/Algebra.lean b/Mathlib/Algebra/PUnitInstances/Algebra.lean index 21ebff8072370..a9860604f2c27 100644 --- a/Mathlib/Algebra/PUnitInstances/Algebra.lean +++ b/Mathlib/Algebra/PUnitInstances/Algebra.lean @@ -25,7 +25,7 @@ instance commGroup : CommGroup PUnit where mul_assoc := by intros; rfl one_mul := by intros; rfl mul_one := by intros; rfl - mul_left_inv := by intros; rfl + inv_mul_cancel := by intros; rfl mul_comm := by intros; rfl -- shortcut instances diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index fa151eea0e863..fdb174e68d4da 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -1019,7 +1019,7 @@ theorem coeff_sub (p q : R[X]) (n : ℕ) : coeff (p - q) n = coeff p n - coeff q -- @[simp] -- Porting note (#10618): simp can prove this theorem monomial_neg (n : ℕ) (a : R) : monomial n (-a) = -monomial n a := by - rw [eq_neg_iff_add_eq_zero, ← monomial_add, neg_add_self, monomial_zero_right] + rw [eq_neg_iff_add_eq_zero, ← monomial_add, neg_add_cancel, monomial_zero_right] theorem monomial_sub (n : ℕ) : monomial n (a - b) = monomial n a - monomial n b := by rw [sub_eq_add_neg, monomial_add, monomial_neg] diff --git a/Mathlib/Algebra/Polynomial/CancelLeads.lean b/Mathlib/Algebra/Polynomial/CancelLeads.lean index f92bee28b5c77..9d4b754822053 100644 --- a/Mathlib/Algebra/Polynomial/CancelLeads.lean +++ b/Mathlib/Algebra/Polynomial/CancelLeads.lean @@ -64,7 +64,7 @@ theorem natDegree_cancelLeads_lt_of_natDegree_le_natDegree_of_comm add_comm _ p.natDegree] simp only [coeff_mul_X_pow, coeff_neg, coeff_C_mul, add_tsub_cancel_left, coeff_add] rw [add_comm p.natDegree, tsub_add_cancel_of_le h, ← leadingCoeff, ← leadingCoeff, comm, - add_right_neg] + add_neg_cancel] end Ring diff --git a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean index f49898f66a5bd..9683ab124fc69 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Lemmas.lean @@ -382,7 +382,7 @@ theorem irreducible_mul_leadingCoeff_inv {p : K[X]} : theorem monic_mul_leadingCoeff_inv {p : K[X]} (h : p ≠ 0) : Monic (p * C (leadingCoeff p)⁻¹) := by rw [Monic, leadingCoeff_mul, leadingCoeff_C, - mul_inv_cancel (show leadingCoeff p ≠ 0 from mt leadingCoeff_eq_zero.1 h)] + mul_inv_cancel₀ (show leadingCoeff p ≠ 0 from mt leadingCoeff_eq_zero.1 h)] -- `simp` normal form of `degree_mul_leadingCoeff_inv` @[simp] lemma degree_leadingCoeff_inv {p : K[X]} (hp0 : p ≠ 0) : diff --git a/Mathlib/Algebra/Polynomial/Eval.lean b/Mathlib/Algebra/Polynomial/Eval.lean index 026dd4a013229..47cf7b8a57706 100644 --- a/Mathlib/Algebra/Polynomial/Eval.lean +++ b/Mathlib/Algebra/Polynomial/Eval.lean @@ -1103,7 +1103,7 @@ alias eval_int_cast := eval_intCast @[simp] theorem eval₂_neg {S} [Ring S] (f : R →+* S) {x : S} : (-p).eval₂ f x = -p.eval₂ f x := by - rw [eq_neg_iff_add_eq_zero, ← eval₂_add, add_left_neg, eval₂_zero] + rw [eq_neg_iff_add_eq_zero, ← eval₂_add, neg_add_cancel, eval₂_zero] @[simp] theorem eval₂_sub {S} [Ring S] (f : R →+* S) {x : S} : diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index f77638479e29a..955c34ceffb85 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -267,7 +267,7 @@ theorem isUnit_iff_degree_eq_zero : IsUnit p ↔ degree p = 0 := isUnit_iff_dvd_one.2 ⟨C (coeff p 0)⁻¹, by conv in p => rw [eq_C_of_degree_le_zero this] - rw [← C_mul, _root_.mul_inv_cancel hc, C_1]⟩⟩ + rw [← C_mul, mul_inv_cancel₀ hc, C_1]⟩⟩ /-- Division of polynomials. See `Polynomial.divByMonic` for more details. -/ def div (p q : R[X]) := @@ -466,7 +466,7 @@ theorem coeff_inv_units (u : R[X]ˣ) (n : ℕ) : ((↑u : R[X]).coeff n)⁻¹ = coeff_C, coeff_C, inv_eq_one_div] split_ifs · rw [div_eq_iff_mul_eq (coeff_coe_units_zero_ne_zero u), coeff_zero_eq_eval_zero, - coeff_zero_eq_eval_zero, ← eval_mul, ← Units.val_mul, inv_mul_self] + coeff_zero_eq_eval_zero, ← eval_mul, ← Units.val_mul, inv_mul_cancel] simp · simp @@ -489,18 +489,18 @@ theorem div_C_mul : p / (C a * q) = C a⁻¹ * (p / q) := by · simp [ha] simp only [div_def, leadingCoeff_mul, mul_inv, leadingCoeff_C, C.map_mul, mul_assoc] congr 3 - rw [mul_left_comm q, ← mul_assoc, ← C.map_mul, mul_inv_cancel ha, C.map_one, one_mul] + rw [mul_left_comm q, ← mul_assoc, ← C.map_mul, mul_inv_cancel₀ ha, C.map_one, one_mul] theorem C_mul_dvd (ha : a ≠ 0) : C a * p ∣ q ↔ p ∣ q := ⟨fun h => dvd_trans (dvd_mul_left _ _) h, fun ⟨r, hr⟩ => ⟨C a⁻¹ * r, by - rw [mul_assoc, mul_left_comm p, ← mul_assoc, ← C.map_mul, _root_.mul_inv_cancel ha, C.map_one, + rw [mul_assoc, mul_left_comm p, ← mul_assoc, ← C.map_mul, mul_inv_cancel₀ ha, C.map_one, one_mul, hr]⟩⟩ theorem dvd_C_mul (ha : a ≠ 0) : p ∣ Polynomial.C a * q ↔ p ∣ q := ⟨fun ⟨r, hr⟩ => ⟨C a⁻¹ * r, by - rw [mul_left_comm p, ← hr, ← mul_assoc, ← C.map_mul, _root_.inv_mul_cancel ha, C.map_one, + rw [mul_left_comm p, ← hr, ← mul_assoc, ← C.map_mul, inv_mul_cancel₀ ha, C.map_one, one_mul]⟩, fun h => dvd_trans h (dvd_mul_left _ _)⟩ diff --git a/Mathlib/Algebra/Polynomial/Laurent.lean b/Mathlib/Algebra/Polynomial/Laurent.lean index ca06a13294c58..99365bbe5de8b 100644 --- a/Mathlib/Algebra/Polynomial/Laurent.lean +++ b/Mathlib/Algebra/Polynomial/Laurent.lean @@ -226,8 +226,8 @@ theorem _root_.Polynomial.toLaurent_C_mul_X_pow (n : ℕ) (r : R) : instance invertibleT (n : ℤ) : Invertible (T n : R[T;T⁻¹]) where invOf := T (-n) - invOf_mul_self := by rw [← T_add, add_left_neg, T_zero] - mul_invOf_self := by rw [← T_add, add_right_neg, T_zero] + invOf_mul_self := by rw [← T_add, neg_add_cancel, T_zero] + mul_invOf_self := by rw [← T_add, add_neg_cancel, T_zero] @[simp] theorem invOf_T (n : ℤ) : ⅟ (T n : R[T;T⁻¹]) = T (-n) := @@ -346,8 +346,8 @@ theorem exists_T_pow (f : R[T;T⁻¹]) : ∃ (n : ℕ) (f' : R[X]), toLaurent f' · cases' n with n n · exact ⟨0, Polynomial.C a * X ^ n, by simp⟩ · refine ⟨n + 1, Polynomial.C a, ?_⟩ - simp only [Int.negSucc_eq, Polynomial.toLaurent_C, Int.ofNat_succ, mul_T_assoc, add_left_neg, - T_zero, mul_one] + simp only [Int.negSucc_eq, Polynomial.toLaurent_C, Int.ofNat_succ, mul_T_assoc, + neg_add_cancel, T_zero, mul_one] /-- This is a version of `exists_T_pow` stated as an induction principle. -/ @[elab_as_elim] @@ -511,7 +511,7 @@ theorem isLocalization : IsLocalization (Submonoid.closure ({X} : Set R[X])) R[T have := (Submonoid.closure ({X} : Set R[X])).pow_mem Submonoid.mem_closure_singleton_self n refine ⟨(f, ⟨_, this⟩), ?_⟩ simp only [algebraMap_eq_toLaurent, Polynomial.toLaurent_X_pow, mul_T_assoc, - add_left_neg, T_zero, mul_one] + neg_add_cancel, T_zero, mul_one] exists_of_eq := fun {f g} => by rw [algebraMap_eq_toLaurent, algebraMap_eq_toLaurent, Polynomial.toLaurent_inj] rintro rfl diff --git a/Mathlib/Algebra/Polynomial/Smeval.lean b/Mathlib/Algebra/Polynomial/Smeval.lean index 6db5f23f5b21e..657e2d1b19d52 100644 --- a/Mathlib/Algebra/Polynomial/Smeval.lean +++ b/Mathlib/Algebra/Polynomial/Smeval.lean @@ -153,7 +153,7 @@ variable (R : Type*) [Ring R] {S : Type*} [AddCommGroup S] [Pow S ℕ] [Module R @[simp] theorem smeval_neg : (-p).smeval x = - p.smeval x := by - rw [← add_eq_zero_iff_eq_neg, ← smeval_add, add_left_neg, smeval_zero] + rw [← add_eq_zero_iff_eq_neg, ← smeval_add, neg_add_cancel, smeval_zero] @[simp] theorem smeval_sub : (p - q).smeval x = p.smeval x - q.smeval x := by diff --git a/Mathlib/Algebra/Polynomial/Taylor.lean b/Mathlib/Algebra/Polynomial/Taylor.lean index 3df1513d4b612..6f9f0e82544ef 100644 --- a/Mathlib/Algebra/Polynomial/Taylor.lean +++ b/Mathlib/Algebra/Polynomial/Taylor.lean @@ -122,7 +122,7 @@ theorem eq_zero_of_hasseDeriv_eq_zero {R} [CommRing R] (f : R[X]) (r : R) /-- Taylor's formula. -/ theorem sum_taylor_eq {R} [CommRing R] (f : R[X]) (r : R) : ((taylor r f).sum fun i a => C a * (X - C r) ^ i) = f := by - rw [← comp_eq_sum_left, sub_eq_add_neg, ← C_neg, ← taylor_apply, taylor_taylor, neg_add_self, + rw [← comp_eq_sum_left, sub_eq_add_neg, ← C_neg, ← taylor_apply, taylor_taylor, neg_add_cancel, taylor_zero] end Polynomial diff --git a/Mathlib/Algebra/Quandle.lean b/Mathlib/Algebra/Quandle.lean index 080f659613491..aa9218bc51e55 100644 --- a/Mathlib/Algebra/Quandle.lean +++ b/Mathlib/Algebra/Quandle.lean @@ -549,7 +549,7 @@ inductive PreEnvelGroupRel' (R : Type u) [Rack R] : PreEnvelGroup R → PreEnvel | assoc (a b c : PreEnvelGroup R) : PreEnvelGroupRel' R (mul (mul a b) c) (mul a (mul b c)) | one_mul (a : PreEnvelGroup R) : PreEnvelGroupRel' R (mul unit a) a | mul_one (a : PreEnvelGroup R) : PreEnvelGroupRel' R (mul a unit) a - | mul_left_inv (a : PreEnvelGroup R) : PreEnvelGroupRel' R (mul (inv a) a) unit + | inv_mul_cancel (a : PreEnvelGroup R) : PreEnvelGroupRel' R (mul (inv a) a) unit | act_incl (x y : R) : PreEnvelGroupRel' R (mul (mul (incl x) (incl y)) (inv (incl x))) (incl (x ◃ y)) @@ -611,8 +611,8 @@ instance (R : Type*) [Rack R] : DivInvMonoid (EnvelGroup R) where mul_one a := Quotient.inductionOn a fun a => Quotient.sound (PreEnvelGroupRel'.mul_one a).rel instance (R : Type*) [Rack R] : Group (EnvelGroup R) := - { mul_left_inv := fun a => - Quotient.inductionOn a fun a => Quotient.sound (PreEnvelGroupRel'.mul_left_inv a).rel } + { inv_mul_cancel := fun a => + Quotient.inductionOn a fun a => Quotient.sound (PreEnvelGroupRel'.inv_mul_cancel a).rel } instance EnvelGroup.inhabited (R : Type*) [Rack R] : Inhabited (EnvelGroup R) := ⟨1⟩ @@ -652,7 +652,7 @@ theorem well_def {R : Type*} [Rack R] {G : Type*} [Group G] (f : R →◃ Quandl | _, _, assoc a b c => by apply mul_assoc | _, _, PreEnvelGroupRel'.one_mul a => by simp [toEnvelGroup.mapAux] | _, _, PreEnvelGroupRel'.mul_one a => by simp [toEnvelGroup.mapAux] - | _, _, PreEnvelGroupRel'.mul_left_inv a => by simp [toEnvelGroup.mapAux] + | _, _, PreEnvelGroupRel'.inv_mul_cancel a => by simp [toEnvelGroup.mapAux] | _, _, act_incl x y => by simp [toEnvelGroup.mapAux] end toEnvelGroup.mapAux diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 86c899ed938f5..97126ffd2f19d 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -1225,7 +1225,7 @@ instance instGroupWithZero : GroupWithZero ℍ[R] := -- Porting note: the aliased definition confuse TC search letI : Semiring ℍ[R] := inferInstanceAs (Semiring ℍ[R,-1,-1]) rw [instInv_inv, Algebra.mul_smul_comm (normSq a)⁻¹ a (star a), self_mul_star, smul_coe, - inv_mul_cancel (normSq_ne_zero.2 ha), coe_one] } + inv_mul_cancel₀ (normSq_ne_zero.2 ha), coe_one] } @[norm_cast, simp] theorem coe_inv (x : R) : ((x⁻¹ : R) : ℍ[R]) = (↑x)⁻¹ := diff --git a/Mathlib/Algebra/Ring/Aut.lean b/Mathlib/Algebra/Ring/Aut.lean index 1b3d80415eb6a..881e53ba2eaeb 100644 --- a/Mathlib/Algebra/Ring/Aut.lean +++ b/Mathlib/Algebra/Ring/Aut.lean @@ -47,7 +47,7 @@ instance : Group (RingAut R) where mul_assoc _ _ _ := rfl one_mul _ := rfl mul_one _ := rfl - mul_left_inv := RingEquiv.self_trans_symm + inv_mul_cancel := RingEquiv.self_trans_symm instance : Inhabited (RingAut R) := ⟨1⟩ diff --git a/Mathlib/Algebra/Ring/Basic.lean b/Mathlib/Algebra/Ring/Basic.lean index 8c18a0c09898a..4ae5170c85619 100644 --- a/Mathlib/Algebra/Ring/Basic.lean +++ b/Mathlib/Algebra/Ring/Basic.lean @@ -97,7 +97,7 @@ variable {α : Type*} [Group α] [HasDistribNeg α] @[simp] theorem inv_neg' (a : α) : (-a)⁻¹ = -a⁻¹ := by - rw [eq_comm, eq_inv_iff_mul_eq_one, neg_mul, mul_neg, neg_neg, mul_left_inv] + rw [eq_comm, eq_inv_iff_mul_eq_one, neg_mul, mul_neg, neg_neg, inv_mul_cancel] end Group diff --git a/Mathlib/Algebra/Ring/BooleanRing.lean b/Mathlib/Algebra/Ring/BooleanRing.lean index 761b053c9354f..f04b26d3b23f6 100644 --- a/Mathlib/Algebra/Ring/BooleanRing.lean +++ b/Mathlib/Algebra/Ring/BooleanRing.lean @@ -70,7 +70,7 @@ theorem add_self : a + a = 0 := by theorem neg_eq : -a = a := calc -a = -a + 0 := by rw [add_zero] - _ = -a + -a + a := by rw [← neg_add_self, add_assoc] + _ = -a + -a + a := by rw [← neg_add_cancel, add_assoc] _ = a := by rw [add_self, zero_add] theorem add_eq_zero' : a + b = 0 ↔ a = b := @@ -385,7 +385,7 @@ abbrev GeneralizedBooleanAlgebra.toNonUnitalCommRing [GeneralizedBooleanAlgebra zero_mul := bot_inf_eq mul_zero := inf_bot_eq neg := id - add_left_neg := symmDiff_self + neg_add_cancel := symmDiff_self add_comm := symmDiff_comm mul := (· ⊓ ·) mul_assoc := inf_assoc @@ -526,7 +526,7 @@ instance : BooleanRing Bool where zero_add := Bool.false_xor add_zero := Bool.xor_false sub_eq_add_neg _ _ := rfl - add_left_neg := Bool.xor_self + neg_add_cancel := Bool.xor_self add_comm := xor_comm mul_assoc := and_assoc one_mul := Bool.true_and diff --git a/Mathlib/Algebra/Ring/Defs.lean b/Mathlib/Algebra/Ring/Defs.lean index 21d4663cec14a..3cc194109119e 100644 --- a/Mathlib/Algebra/Ring/Defs.lean +++ b/Mathlib/Algebra/Ring/Defs.lean @@ -323,8 +323,8 @@ variable [NonUnitalNonAssocRing α] instance (priority := 100) NonUnitalNonAssocRing.toHasDistribNeg : HasDistribNeg α where neg := Neg.neg neg_neg := neg_neg - neg_mul a b := eq_neg_of_add_eq_zero_left <| by rw [← right_distrib, add_left_neg, zero_mul] - mul_neg a b := eq_neg_of_add_eq_zero_left <| by rw [← left_distrib, add_left_neg, mul_zero] + neg_mul a b := eq_neg_of_add_eq_zero_left <| by rw [← right_distrib, neg_add_cancel, zero_mul] + mul_neg a b := eq_neg_of_add_eq_zero_left <| by rw [← left_distrib, neg_add_cancel, mul_zero] theorem mul_sub_left_distrib (a b c : α) : a * (b - c) = a * b - a * c := by simpa only [sub_eq_add_neg, neg_mul_eq_mul_neg] using mul_add a b (-c) diff --git a/Mathlib/Algebra/Ring/MinimalAxioms.lean b/Mathlib/Algebra/Ring/MinimalAxioms.lean index a5b2c44bbca96..de81f68a2e3b8 100644 --- a/Mathlib/Algebra/Ring/MinimalAxioms.lean +++ b/Mathlib/Algebra/Ring/MinimalAxioms.lean @@ -31,13 +31,13 @@ abbrev Ring.ofMinimalAxioms {R : Type u} [Add R] [Mul R] [Neg R] [Zero R] [One R] (add_assoc : ∀ a b c : R, a + b + c = a + (b + c)) (zero_add : ∀ a : R, 0 + a = a) - (add_left_neg : ∀ a : R, -a + a = 0) + (neg_add_cancel : ∀ a : R, -a + a = 0) (mul_assoc : ∀ a b c : R, a * b * c = a * (b * c)) (one_mul : ∀ a : R, 1 * a = a) (mul_one : ∀ a : R, a * 1 = a) (left_distrib : ∀ a b c : R, a * (b + c) = a * b + a * c) (right_distrib : ∀ a b c : R, (a + b) * c = a * c + b * c) : Ring R := - letI := AddGroup.ofLeftAxioms add_assoc zero_add add_left_neg + letI := AddGroup.ofLeftAxioms add_assoc zero_add neg_add_cancel haveI add_comm : ∀ a b, a + b = b + a := by intro a b have h₁ : (1 + 1 : R) * (a + b) = a + (a + b) + b := by @@ -66,7 +66,7 @@ abbrev Ring.ofMinimalAxioms {R : Type u} mul_assoc := mul_assoc one_mul := one_mul mul_one := mul_one - add_left_neg := add_left_neg + neg_add_cancel := neg_add_cancel zsmul := (· • ·) } /-- Define a `CommRing` structure on a Type by proving a minimized set of axioms. @@ -76,7 +76,7 @@ abbrev CommRing.ofMinimalAxioms {R : Type u} [Add R] [Mul R] [Neg R] [Zero R] [One R] (add_assoc : ∀ a b c : R, a + b + c = a + (b + c)) (zero_add : ∀ a : R, 0 + a = a) - (add_left_neg : ∀ a : R, -a + a = 0) + (neg_add_cancel : ∀ a : R, -a + a = 0) (mul_assoc : ∀ a b c : R, a * b * c = a * (b * c)) (mul_comm : ∀ a b : R, a * b = b * a) (one_mul : ∀ a : R, 1 * a = a) @@ -85,6 +85,6 @@ abbrev CommRing.ofMinimalAxioms {R : Type u} rw [mul_comm, one_mul] haveI right_distrib : ∀ a b c : R, (a + b) * c = a * c + b * c := fun a b c => by rw [mul_comm, left_distrib, mul_comm, mul_comm b c] - letI := Ring.ofMinimalAxioms add_assoc zero_add add_left_neg mul_assoc + letI := Ring.ofMinimalAxioms add_assoc zero_add neg_add_cancel mul_assoc one_mul mul_one left_distrib right_distrib { mul_comm := mul_comm } diff --git a/Mathlib/Algebra/Ring/Opposite.lean b/Mathlib/Algebra/Ring/Opposite.lean index 42f58ea315695..6f3e9818f4b83 100644 --- a/Mathlib/Algebra/Ring/Opposite.lean +++ b/Mathlib/Algebra/Ring/Opposite.lean @@ -78,7 +78,7 @@ instance instGroupWithZero [GroupWithZero α] : GroupWithZero αᵐᵒᵖ where __ := instMonoidWithZero __ := instNontrivial __ := instDivInvMonoid - mul_inv_cancel _ hx := unop_injective <| inv_mul_cancel <| unop_injective.ne hx + mul_inv_cancel _ hx := unop_injective <| inv_mul_cancel₀ <| unop_injective.ne hx inv_zero := unop_injective inv_zero end MulOpposite diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index c4878d4541ca9..365272a227add 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -659,7 +659,7 @@ variable {K : Type u} [DivisionRing K] instance instField : Field (center K) where inv a := ⟨a⁻¹, Set.inv_mem_center a.prop⟩ - mul_inv_cancel a ha := Subtype.ext <| mul_inv_cancel <| Subtype.coe_injective.ne ha + mul_inv_cancel a ha := Subtype.ext <| mul_inv_cancel₀ <| Subtype.coe_injective.ne ha div a b := ⟨a / b, Set.div_mem_center a.prop b.prop⟩ div_eq_mul_inv a b := Subtype.ext <| div_eq_mul_inv _ _ inv_zero := Subtype.ext inv_zero @@ -1145,7 +1145,7 @@ variable {s : Set R} protected theorem InClosure.recOn {C : R → Prop} {x : R} (hx : x ∈ closure s) (h1 : C 1) (hneg1 : C (-1)) (hs : ∀ z ∈ s, ∀ n, C n → C (z * n)) (ha : ∀ {x y}, C x → C y → C (x + y)) : C x := by - have h0 : C 0 := add_neg_self (1 : R) ▸ ha h1 hneg1 + have h0 : C 0 := add_neg_cancel (1 : R) ▸ ha h1 hneg1 rcases exists_list_of_mem_closure hx with ⟨L, HL, rfl⟩ clear hx induction' L with hd tl ih diff --git a/Mathlib/Algebra/Ring/ULift.lean b/Mathlib/Algebra/Ring/ULift.lean index 0c7f9261dbc9e..d22da7ec4c22e 100644 --- a/Mathlib/Algebra/Ring/ULift.lean +++ b/Mathlib/Algebra/Ring/ULift.lean @@ -79,7 +79,7 @@ instance commSemiring [CommSemiring α] : CommSemiring (ULift α) := instance nonUnitalNonAssocRing [NonUnitalNonAssocRing α] : NonUnitalNonAssocRing (ULift α) := { zero := (0 : ULift α), add := (· + ·), mul := (· * ·), sub := Sub.sub, neg := Neg.neg, nsmul := AddMonoid.nsmul, zsmul := SubNegMonoid.zsmul, add_assoc, zero_add, add_zero, - add_left_neg, add_comm, left_distrib, right_distrib, zero_mul, mul_zero, sub_eq_add_neg, + neg_add_cancel, add_comm, left_distrib, right_distrib, zero_mul, mul_zero, sub_eq_add_neg, nsmul_zero := fun _ => AddMonoid.nsmul_zero _, nsmul_succ := fun _ _ => AddMonoid.nsmul_succ _ _, zsmul_zero' := SubNegMonoid.zsmul_zero', zsmul_succ' := SubNegMonoid.zsmul_succ', @@ -88,7 +88,7 @@ instance nonUnitalNonAssocRing [NonUnitalNonAssocRing α] : NonUnitalNonAssocRin instance nonUnitalRing [NonUnitalRing α] : NonUnitalRing (ULift α) := { zero := (0 : ULift α), add := (· + ·), mul := (· * ·), sub := Sub.sub, neg := Neg.neg, nsmul := AddMonoid.nsmul, zsmul := SubNegMonoid.zsmul, add_assoc, zero_add, add_zero, add_comm, - add_left_neg, left_distrib, right_distrib, zero_mul, mul_zero, mul_assoc, sub_eq_add_neg + neg_add_cancel, left_distrib, right_distrib, zero_mul, mul_zero, mul_assoc, sub_eq_add_neg nsmul_zero := fun _ => AddMonoid.nsmul_zero _, nsmul_succ := fun _ _ => AddMonoid.nsmul_succ _ _, zsmul_zero' := SubNegMonoid.zsmul_zero', zsmul_succ' := SubNegMonoid.zsmul_succ', @@ -99,7 +99,7 @@ instance nonAssocRing [NonAssocRing α] : NonAssocRing (ULift α) := neg := Neg.neg, nsmul := AddMonoid.nsmul, natCast := fun n => ULift.up n, intCast := fun n => ULift.up n, zsmul := SubNegMonoid.zsmul, intCast_ofNat := addGroupWithOne.intCast_ofNat, add_assoc, zero_add, - add_zero, add_left_neg, add_comm, left_distrib, right_distrib, zero_mul, mul_zero, one_mul, + add_zero, neg_add_cancel, add_comm, left_distrib, right_distrib, zero_mul, mul_zero, one_mul, mul_one, sub_eq_add_neg, nsmul_zero := fun _ => AddMonoid.nsmul_zero _, nsmul_succ := fun _ _ => AddMonoid.nsmul_succ _ _, zsmul_zero' := SubNegMonoid.zsmul_zero', zsmul_succ' := SubNegMonoid.zsmul_succ', @@ -112,7 +112,7 @@ instance ring [Ring α] : Ring (ULift α) := neg := Neg.neg, nsmul := AddMonoid.nsmul, npow := Monoid.npow, zsmul := SubNegMonoid.zsmul, intCast_ofNat := addGroupWithOne.intCast_ofNat, add_assoc, zero_add, add_zero, add_comm, left_distrib, right_distrib, zero_mul, mul_zero, mul_assoc, one_mul, mul_one, sub_eq_add_neg, - add_left_neg, nsmul_zero := fun _ => AddMonoid.nsmul_zero _, natCast := fun n => ULift.up n, + neg_add_cancel, nsmul_zero := fun _ => AddMonoid.nsmul_zero _, natCast := fun n => ULift.up n, intCast := fun n => ULift.up n, nsmul_succ := fun _ _ => AddMonoid.nsmul_succ _ _, natCast_zero := AddMonoidWithOne.natCast_zero, natCast_succ := AddMonoidWithOne.natCast_succ, npow_zero := fun _ => Monoid.npow_zero _, npow_succ := fun _ _ => Monoid.npow_succ _ _, @@ -123,7 +123,7 @@ instance nonUnitalCommRing [NonUnitalCommRing α] : NonUnitalCommRing (ULift α) { zero := (0 : ULift α), add := (· + ·), mul := (· * ·), sub := Sub.sub, neg := Neg.neg, nsmul := AddMonoid.nsmul, zsmul := SubNegMonoid.zsmul, zero_mul, add_assoc, zero_add, add_zero, mul_zero, left_distrib, right_distrib, add_comm, mul_assoc, mul_comm, - nsmul_zero := fun _ => AddMonoid.nsmul_zero _, add_left_neg, + nsmul_zero := fun _ => AddMonoid.nsmul_zero _, neg_add_cancel, nsmul_succ := fun _ _ => AddMonoid.nsmul_succ _ _, sub_eq_add_neg, zsmul_zero' := SubNegMonoid.zsmul_zero', zsmul_succ' := SubNegMonoid.zsmul_succ', diff --git a/Mathlib/Algebra/RingQuot.lean b/Mathlib/Algebra/RingQuot.lean index af06d01d6c194..f0e0f3bcd667b 100644 --- a/Mathlib/Algebra/RingQuot.lean +++ b/Mathlib/Algebra/RingQuot.lean @@ -337,7 +337,7 @@ private def intCast {R : Type uR} [Ring R] (r : R → R → Prop) (z : ℤ) : Ri instance instRing {R : Type uR} [Ring R] (r : R → R → Prop) : Ring (RingQuot r) := { RingQuot.instSemiring r with neg := Neg.neg - add_left_neg := by + neg_add_cancel := by rintro ⟨⟨⟩⟩ simp [neg_quot, add_quot, ← zero_quot] sub := Sub.sub diff --git a/Mathlib/Algebra/SMulWithZero.lean b/Mathlib/Algebra/SMulWithZero.lean index df3c43ceda5fc..bba8bf4abb805 100644 --- a/Mathlib/Algebra/SMulWithZero.lean +++ b/Mathlib/Algebra/SMulWithZero.lean @@ -197,7 +197,7 @@ theorem smul_inv₀ [SMulCommClass α β β] [IsScalarTower α β β] (c : α) ( obtain rfl | hx := eq_or_ne x 0 · simp only [inv_zero, smul_zero] · refine inv_eq_of_mul_eq_one_left ?_ - rw [smul_mul_smul, inv_mul_cancel hc, inv_mul_cancel hx, one_smul] + rw [smul_mul_smul, inv_mul_cancel₀ hc, inv_mul_cancel₀ hx, one_smul] end GroupWithZero diff --git a/Mathlib/Algebra/Star/CHSH.lean b/Mathlib/Algebra/Star/CHSH.lean index 58d2771af2470..72b30f5ea9329 100644 --- a/Mathlib/Algebra/Star/CHSH.lean +++ b/Mathlib/Algebra/Star/CHSH.lean @@ -150,7 +150,7 @@ we prepare some easy lemmas about √2. -- defeated me. Thanks for the rescue from Shing Tak Lam! theorem tsirelson_inequality_aux : √2 * √2 ^ 3 = √2 * (2 * (√2)⁻¹ + 4 * ((√2)⁻¹ * 2⁻¹)) := by ring_nf - rw [mul_inv_cancel (ne_of_gt (Real.sqrt_pos.2 (show (2 : ℝ) > 0 by norm_num)))] + rw [mul_inv_cancel₀ (ne_of_gt (Real.sqrt_pos.2 (show (2 : ℝ) > 0 by norm_num)))] convert congr_arg (· ^ 2) (@Real.sq_sqrt 2 (by norm_num)) using 1 <;> (try simp only [← pow_mul]) <;> norm_num diff --git a/Mathlib/Algebra/Star/Unitary.lean b/Mathlib/Algebra/Star/Unitary.lean index 4c41792f28ab4..4b73bc722f118 100644 --- a/Mathlib/Algebra/Star/Unitary.lean +++ b/Mathlib/Algebra/Star/Unitary.lean @@ -89,7 +89,7 @@ theorem mul_star_self (U : unitary R) : U * star U = 1 := instance : Group (unitary R) := { Submonoid.toMonoid _ with inv := star - mul_left_inv := star_mul_self } + inv_mul_cancel := star_mul_self } instance : InvolutiveStar (unitary R) := ⟨by diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index 56213cf92612d..ed6277acc897f 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -727,17 +727,17 @@ protected theorem inv_one : (1 : tsze R M)⁻¹ = (1 : tsze R M) := by protected theorem inv_mul_cancel {x : tsze R M} (hx : fst x ≠ 0) : x⁻¹ * x = 1 := by ext - · rw [fst_mul, fst_inv, inv_mul_cancel hx, fst_one] - · rw [snd_mul, snd_inv, snd_one, smul_neg, op_smul_op_smul, inv_mul_cancel hx, op_one, one_smul, - fst_inv, add_right_neg] + · rw [fst_mul, fst_inv, inv_mul_cancel₀ hx, fst_one] + · rw [snd_mul, snd_inv, snd_one, smul_neg, op_smul_op_smul, inv_mul_cancel₀ hx, op_one, one_smul, + fst_inv, add_neg_cancel] variable [SMulCommClass R Rᵐᵒᵖ M] protected theorem mul_inv_cancel {x : tsze R M} (hx : fst x ≠ 0) : x * x⁻¹ = 1 := by ext - · rw [fst_mul, fst_inv, fst_one, mul_inv_cancel hx] - · rw [snd_mul, snd_inv, snd_one, smul_neg, smul_comm, smul_smul, mul_inv_cancel hx, one_smul, - fst_inv, add_left_neg] + · rw [fst_mul, fst_inv, fst_one, mul_inv_cancel₀ hx] + · rw [snd_mul, snd_inv, snd_one, smul_neg, smul_comm, smul_smul, mul_inv_cancel₀ hx, one_smul, + fst_inv, neg_add_cancel] protected theorem mul_inv_rev (a b : tsze R M) : (a * b)⁻¹ = b⁻¹ * a⁻¹ := by diff --git a/Mathlib/Algebra/Tropical/Basic.lean b/Mathlib/Algebra/Tropical/Basic.lean index ac29efb162ca7..7d620bea68f04 100644 --- a/Mathlib/Algebra/Tropical/Basic.lean +++ b/Mathlib/Algebra/Tropical/Basic.lean @@ -410,7 +410,7 @@ instance instGroupTropical [AddGroup R] : Group (Tropical R) := { instMonoidTropical with inv := Inv.inv div_eq_mul_inv := fun _ _ => untrop_injective <| by simp [sub_eq_add_neg] - mul_left_inv := fun _ => untrop_injective <| add_left_neg _ + inv_mul_cancel := fun _ => untrop_injective <| neg_add_cancel _ zpow := fun n x => trop <| n • untrop x zpow_zero' := fun _ => untrop_injective <| zero_zsmul _ zpow_succ' := fun _ _ => untrop_injective <| SubNegMonoid.zsmul_succ' _ _ diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean index c8ca6cc20d486..c87f5ce1c0540 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean @@ -328,13 +328,13 @@ lemma XYIdeal_neg_mul {x y : F} (h : W.Nonsingular x y) : refine ⟨C <| C W_X⁻¹ * -(X + C (2 * x + W.a₂)), C <| C <| W_X⁻¹ * W.a₁, 0, C <| C <| W_X⁻¹ * -1, ?_⟩ rw [← mul_right_inj' <| C_ne_zero.mpr <| C_ne_zero.mpr hx] - simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel hx] + simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel₀ hx] C_simp ring1 · let W_Y := 2 * y + W.a₁ * x + W.a₃ refine ⟨0, C <| C W_Y⁻¹, C <| C <| W_Y⁻¹ * -1, 0, ?_⟩ rw [negY, ← mul_right_inj' <| C_ne_zero.mpr <| C_ne_zero.mpr hy] - simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel hy] + simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel₀ hy] C_simp ring1 @@ -376,13 +376,13 @@ lemma XYIdeal_mul_XYIdeal {x₁ x₂ y₁ y₂ : F} (h₁ : W.Equation x₁ y₁ 0, C (C y⁻¹) * (Y - W.negPolynomial), ?_⟩, by rw [map_add, map_one, _root_.map_mul <| mk W, AdjoinRoot.mk_self, mul_zero, add_zero]⟩ rw [polynomial, negPolynomial, ← mul_right_inj' <| C_ne_zero.mpr <| C_ne_zero.mpr hxy] - simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel hxy] + simp only [mul_add, ← mul_assoc, ← C_mul, mul_inv_cancel₀ hxy] linear_combination (norm := (rw [b₂, b₄, negY]; C_simp; ring1)) -4 * congr_arg C (congr_arg C <| (equation_iff ..).mp h₁) · replace hx := sub_ne_zero_of_ne hx refine ⟨_, ⟨⟨C <| C (x₁ - x₂)⁻¹, C <| C <| (x₁ - x₂)⁻¹ * -1, 0, ?_⟩, map_one _⟩⟩ rw [← mul_right_inj' <| C_ne_zero.mpr <| C_ne_zero.mpr hx] - simp only [← mul_assoc, mul_add, ← C_mul, mul_inv_cancel hx] + simp only [← mul_assoc, mul_add, ← C_mul, mul_inv_cancel₀ hx] C_simp ring1 @@ -558,7 +558,7 @@ noncomputable instance : AddCommGroup W.Point where zsmul := zsmulRec zero_add := zero_add add_zero := add_zero - add_left_neg _ := by rw [add_eq_zero] + neg_add_cancel _ := by rw [add_eq_zero] add_comm _ _ := toClass_injective <| by simp only [map_add, add_comm] add_assoc _ _ _ := toClass_injective <| by simp only [map_add, add_assoc] @@ -579,8 +579,8 @@ noncomputable instance : AddCommGroup W.Point where simp only [map_add, toAffineAddEquiv_apply, toAffineLift_zero, zero_add] add_zero _ := (toAffineAddEquiv W).injective <| by simp only [map_add, toAffineAddEquiv_apply, toAffineLift_zero, add_zero] - add_left_neg P := (toAffineAddEquiv W).injective <| by - simp only [map_add, toAffineAddEquiv_apply, toAffineLift_neg, add_left_neg, toAffineLift_zero] + neg_add_cancel P := (toAffineAddEquiv W).injective <| by + simp only [map_add, toAffineAddEquiv_apply, toAffineLift_neg, neg_add_cancel, toAffineLift_zero] add_comm _ _ := (toAffineAddEquiv W).injective <| by simp only [map_add, add_comm] add_assoc _ _ _ := (toAffineAddEquiv W).injective <| by simp only [map_add, add_assoc] diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean index a4e832ba1c175..5bb09227528a8 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean @@ -218,7 +218,7 @@ instance instGroup : Group (VariableChange R) where mul := comp one_mul := id_comp mul_one := comp_id - mul_left_inv := comp_left_inv + inv_mul_cancel := comp_left_inv mul_assoc := comp_assoc end VariableChange diff --git a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean index e025d70905a16..4a3ffcf186da8 100644 --- a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean +++ b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean @@ -214,7 +214,7 @@ def ε [Limits.HasZeroObject C] : rw [alternatingFaceMapComplex_obj_d, objD, Fin.sum_univ_two, Fin.val_zero, pow_zero, one_smul, Fin.val_one, pow_one, neg_smul, one_smul, add_comp, neg_comp, SimplicialObject.δ_naturality, SimplicialObject.δ_naturality] - apply add_right_neg + apply add_neg_cancel naturality X Y f := by apply HomologicalComplex.to_single_hom_ext dsimp diff --git a/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean b/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean index 2ff8defbbf428..dd949cbd7cb71 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Degeneracies.lean @@ -86,7 +86,7 @@ theorem σ_comp_P_eq_zero (X : SimplicialObject C) {n q : ℕ} (i : Fin (n + 1)) SimplicialObject.δ_comp_σ_of_le X (show (0 : Fin 2) ≤ Fin.castSucc 0 by rw [Fin.castSucc_zero]), SimplicialObject.δ_comp_σ_self_assoc, SimplicialObject.δ_comp_σ_succ_assoc] - simp only [add_right_neg, add_zero, zero_add] + simp only [add_neg_cancel, add_zero, zero_add] · rw [← id_comp (X.σ i), ← (P_add_Q_f q n.succ : _ = 𝟙 (X.obj _)), add_comp, add_comp, P_succ] have v : HigherFacesVanish q ((P q).f n.succ ≫ X.σ i) := diff --git a/Mathlib/AlgebraicTopology/DoldKan/Faces.lean b/Mathlib/AlgebraicTopology/DoldKan/Faces.lean index 9bd3942d31cbc..a232a8b12caaf 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Faces.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Faces.lean @@ -117,7 +117,7 @@ theorem comp_Hσ_eq {Y : C} {n a q : ℕ} {φ : Y ⟶ X _[n + 1]} (v : HigherFac rw [X.δ_comp_σ_self' (Fin.castSucc_mk _ _ _).symm, X.δ_comp_σ_succ' (Fin.succ_mk _ _ _).symm] simp only [comp_id, pow_add _ (a + 1) 1, pow_one, mul_neg, mul_one, neg_mul, neg_smul, - add_right_neg] + add_neg_cancel] · -- c + a = 0 rw [← Finset.sum_add_distrib] apply Finset.sum_eq_zero @@ -146,7 +146,7 @@ theorem comp_Hσ_eq_zero {Y : C} {n q : ℕ} {φ : Y ⟶ X _[n + 1]} (v : Higher Fin.cast_mk, Fin.castSucc_mk] simp only [Fin.mk_zero, Fin.val_zero, pow_zero, one_zsmul, Fin.mk_one, Fin.val_one, pow_one, neg_smul, comp_neg] - erw [δ_comp_σ_self, δ_comp_σ_succ, add_right_neg] + erw [δ_comp_σ_self, δ_comp_σ_succ, add_neg_cancel] · intro j dsimp [Fin.cast, Fin.castLE, Fin.castLT] rw [comp_zsmul, comp_zsmul, δ_comp_σ_of_gt', v.comp_δ_eq_zero_assoc, zero_comp, zsmul_zero] diff --git a/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean b/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean index ea16d454ce082..4c1705a39ed6d 100644 --- a/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean +++ b/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean @@ -407,7 +407,7 @@ noncomputable def homotopyEquiv {C : Type*} [Category C] [Preadditive C] [HasZer Preadditive.comp_add, drop_obj, s_comp_δ₀, Preadditive.sum_comp, Preadditive.zsmul_comp, Preadditive.comp_sum, Preadditive.comp_zsmul, zsmul_neg, s_comp_δ, pow_add, pow_one, mul_neg, mul_one, neg_zsmul, neg_neg, - neg_add_cancel_comm_assoc, add_left_neg, zero_comp] } + neg_add_cancel_comm_assoc, neg_add_cancel, zero_comp] } end ExtraDegeneracy diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index e81266a4b101c..cce923522de40 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -936,7 +936,7 @@ theorem Asymptotics.IsBigO.continuousMultilinearMap_apply_eq_zero {n : ℕ} {p : rw [← mul_assoc] simp [norm_mul, mul_pow] _ ≤ 0 + ε := by - rw [inv_mul_cancel (norm_pos_iff.mp k_pos)] + rw [inv_mul_cancel₀ (norm_pos_iff.mp k_pos)] simpa using h₃.le /-- If a formal multilinear series `p` represents the zero function at `x : E`, then the diff --git a/Mathlib/Analysis/Analytic/Meromorphic.lean b/Mathlib/Analysis/Analytic/Meromorphic.lean index 1c8988bb6dd45..5fc9bcc71b8be 100644 --- a/Mathlib/Analysis/Analytic/Meromorphic.lean +++ b/Mathlib/Analysis/Analytic/Meromorphic.lean @@ -210,7 +210,7 @@ lemma iff_eventuallyEq_zpow_smul_analyticAt {f : 𝕜 → E} {x : 𝕜} : Meromo ∃ (n : ℤ) (g : 𝕜 → E), AnalyticAt 𝕜 g x ∧ ∀ᶠ z in 𝓝[≠] x, f z = (z - x) ^ n • g z := by refine ⟨fun ⟨n, hn⟩ ↦ ⟨-n, _, ⟨hn, eventually_nhdsWithin_iff.mpr ?_⟩⟩, ?_⟩ · filter_upwards with z hz - rw [← mul_smul, ← zpow_natCast, ← zpow_add₀ (sub_ne_zero.mpr hz), add_left_neg, + rw [← mul_smul, ← zpow_natCast, ← zpow_add₀ (sub_ne_zero.mpr hz), neg_add_cancel, zpow_zero, one_smul] · refine fun ⟨n, g, hg_an, hg_eq⟩ ↦ MeromorphicAt.congr ?_ (EventuallyEq.symm hg_eq) exact (((MeromorphicAt.id x).sub (.const _ x)).zpow _).smul hg_an.meromorphicAt diff --git a/Mathlib/Analysis/Analytic/RadiusLiminf.lean b/Mathlib/Analysis/Analytic/RadiusLiminf.lean index 0bee5b674904b..175d6acca7461 100644 --- a/Mathlib/Analysis/Analytic/RadiusLiminf.lean +++ b/Mathlib/Analysis/Analytic/RadiusLiminf.lean @@ -40,7 +40,7 @@ theorem radius_eq_liminf : have : 0 < (n : ℝ) := Nat.cast_pos.2 hn conv_lhs => rw [one_div, ENNReal.le_inv_iff_mul_le, ← ENNReal.coe_mul, ENNReal.coe_le_one_iff, one_div, ← - NNReal.rpow_one r, ← mul_inv_cancel this.ne', NNReal.rpow_mul, ← NNReal.mul_rpow, ← + NNReal.rpow_one r, ← mul_inv_cancel₀ this.ne', NNReal.rpow_mul, ← NNReal.mul_rpow, ← NNReal.one_rpow n⁻¹, NNReal.rpow_le_rpow_iff (inv_pos.2 this), mul_comm, NNReal.rpow_natCast] apply le_antisymm <;> refine ENNReal.le_of_forall_nnreal_lt fun r hr => ?_ diff --git a/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean b/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean index e3f7304c5b371..a0281a8808ecb 100644 --- a/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/SpecificAsymptotics.lean @@ -134,10 +134,10 @@ theorem Filter.Tendsto.cesaro_smul {E : Type*} [NormedAddCommGroup E] [NormedSpa · filter_upwards [Ici_mem_atTop 1] with n npos have nposℝ : (0 : ℝ) < n := Nat.cast_pos.2 npos simp only [smul_sub, sum_sub_distrib, sum_const, card_range, sub_right_inj] - rw [← Nat.cast_smul_eq_nsmul ℝ, smul_smul, inv_mul_cancel nposℝ.ne', one_smul] + rw [← Nat.cast_smul_eq_nsmul ℝ, smul_smul, inv_mul_cancel₀ nposℝ.ne', one_smul] · filter_upwards [Ici_mem_atTop 1] with n npos have nposℝ : (0 : ℝ) < n := Nat.cast_pos.2 npos - rw [Algebra.id.smul_eq_mul, inv_mul_cancel nposℝ.ne'] + rw [Algebra.id.smul_eq_mul, inv_mul_cancel₀ nposℝ.ne'] /-- The Cesaro average of a converging sequence converges to the same limit. -/ theorem Filter.Tendsto.cesaro {u : ℕ → ℝ} {l : ℝ} (h : Tendsto u atTop (𝓝 l)) : diff --git a/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean b/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean index 29a5817582655..45672ad5d7b49 100644 --- a/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean +++ b/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean @@ -177,12 +177,12 @@ variable [TopologicalSpace β] [Field β] (l k f) theorem superpolynomialDecay_mul_const_iff [ContinuousMul β] {c : β} (hc0 : c ≠ 0) : (SuperpolynomialDecay l k fun n => f n * c) ↔ SuperpolynomialDecay l k f := - ⟨fun h => (h.mul_const c⁻¹).congr fun x => by simp [mul_assoc, mul_inv_cancel hc0], fun h => + ⟨fun h => (h.mul_const c⁻¹).congr fun x => by simp [mul_assoc, mul_inv_cancel₀ hc0], fun h => h.mul_const c⟩ theorem superpolynomialDecay_const_mul_iff [ContinuousMul β] {c : β} (hc0 : c ≠ 0) : (SuperpolynomialDecay l k fun n => c * f n) ↔ SuperpolynomialDecay l k f := - ⟨fun h => (h.const_mul c⁻¹).congr fun x => by simp [← mul_assoc, inv_mul_cancel hc0], fun h => + ⟨fun h => (h.const_mul c⁻¹).congr fun x => by simp [← mul_assoc, inv_mul_cancel₀ hc0], fun h => h.const_mul c⟩ variable {l k f} @@ -210,7 +210,7 @@ theorem superpolynomialDecay_iff_abs_isBoundedUnder (hk : Tendsto k l atTop) : ((eventually_map.1 hm).mp ?_) refine (hk.eventually_ne_atTop 0).mono fun x hk0 hx => ?_ refine Eq.trans_le ?_ (mul_le_mul_of_nonneg_left hx <| abs_nonneg (k x)⁻¹) - rw [← abs_mul, ← mul_assoc, pow_succ', ← mul_assoc, inv_mul_cancel hk0, one_mul] + rw [← abs_mul, ← mul_assoc, pow_succ', ← mul_assoc, inv_mul_cancel₀ hk0, one_mul] theorem superpolynomialDecay_iff_zpow_tendsto_zero (hk : Tendsto k l atTop) : SuperpolynomialDecay l k f ↔ ∀ z : ℤ, Tendsto (fun a : α => k a ^ z * f a) l (𝓝 0) := by @@ -251,7 +251,7 @@ theorem superpolynomialDecay_param_mul_iff (hk : Tendsto k l atTop) : SuperpolynomialDecay l k (k * f) ↔ SuperpolynomialDecay l k f := ⟨fun h => (h.inv_param_mul hk).congr' - ((hk.eventually_ne_atTop 0).mono fun x hx => by simp [← mul_assoc, inv_mul_cancel hx]), + ((hk.eventually_ne_atTop 0).mono fun x hx => by simp [← mul_assoc, inv_mul_cancel₀ hx]), fun h => h.param_mul⟩ theorem superpolynomialDecay_mul_param_iff (hk : Tendsto k l atTop) : @@ -306,7 +306,7 @@ theorem superpolynomialDecay_iff_isBigO (hk : Tendsto k l atTop) : ((isBigO_refl (fun a => k a ^ z) l).mul (h (-(z + 1)))).trans (IsBigO.of_bound 1 <| hk0.mono fun a ha0 => ?_) simp only [one_mul, neg_add z 1, zpow_add₀ ha0, ← mul_assoc, zpow_neg, - mul_inv_cancel (zpow_ne_zero z ha0), zpow_one] + mul_inv_cancel₀ (zpow_ne_zero z ha0), zpow_one] rfl theorem superpolynomialDecay_iff_isLittleO (hk : Tendsto k l atTop) : @@ -319,7 +319,7 @@ theorem superpolynomialDecay_iff_isLittleO (hk : Tendsto k l atTop) : have : f =o[l] fun x : α => k x * k x ^ (z - 1) := by simpa using this.mul_isBigO ((superpolynomialDecay_iff_isBigO f hk).1 h <| z - 1) refine this.trans_isBigO (IsBigO.of_bound 1 (hk0.mono fun x hkx => le_of_eq ?_)) - rw [one_mul, zpow_sub_one₀ hkx, mul_comm (k x), mul_assoc, inv_mul_cancel hkx, mul_one] + rw [one_mul, zpow_sub_one₀ hkx, mul_comm (k x), mul_assoc, inv_mul_cancel₀ hkx, mul_one] end NormedLinearOrderedField diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index d902df8fafe19..2e4ffbbc89320 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -633,7 +633,7 @@ protected theorem add {f : α → E} {s : Set α} (hf : LocallyBoundedVariationO symm refine additive_of_isTotal ((· : α) ≤ ·) (variationOnFromTo f s) (· ∈ s) ?_ ?_ ha hb hc · rintro x y _xs _ys - simp only [variationOnFromTo.eq_neg_swap f s y x, Subtype.coe_mk, add_right_neg, + simp only [variationOnFromTo.eq_neg_swap f s y x, Subtype.coe_mk, add_neg_cancel, forall_true_left] · rintro x y z xy yz xs ys zs rw [variationOnFromTo.eq_of_le f s xy, variationOnFromTo.eq_of_le f s yz, diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean index 207a397d2c912..81bf4037d473e 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean @@ -63,7 +63,7 @@ theorem integral_pos : 0 < ∫ x, f x ∂μ := by theorem integral_normed : ∫ x, f.normed μ x ∂μ = 1 := by simp_rw [ContDiffBump.normed, div_eq_mul_inv, mul_comm (f _), ← smul_eq_mul, integral_smul] - exact inv_mul_cancel f.integral_pos.ne' + exact inv_mul_cancel₀ f.integral_pos.ne' theorem support_normed_eq : Function.support (f.normed μ) = Metric.ball c f.rOut := by unfold ContDiffBump.normed diff --git a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean index ba117c50a4db6..9d0be42f520d6 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean @@ -441,7 +441,7 @@ theorem norm_iteratedFDerivWithin_comp_le_aux {Fu Gu : Type u} [NormedAddCommGro ring _ = ∑ i ∈ Finset.range (n + 1), (n ! : ℝ) * 1 * C * D ^ (n + 1) * ((n - i)! : ℝ)⁻¹ := by congr! with i hi - · apply inv_mul_cancel + · apply inv_mul_cancel₀ simpa only [Ne, Nat.cast_eq_zero] using i.factorial_ne_zero · rw [← pow_add] congr 1 diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index acb1cb8f0b2cc..b488789c0194b 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -478,6 +478,6 @@ theorem hasSum_iteratedFDeriv [CharZero 𝕜] {y : E} (hy : y ∈ EMetric.ball 0 HasSum (fun n ↦ (n ! : 𝕜)⁻¹ • iteratedFDeriv 𝕜 n f x fun _ ↦ y) (f (x + y)) := by convert h.hasSum hy with n rw [← h.factorial_smul y n, smul_comm, ← smul_assoc, nsmul_eq_mul, - mul_inv_cancel <| cast_ne_zero.mpr n.factorial_ne_zero, one_smul] + mul_inv_cancel₀ <| cast_ne_zero.mpr n.factorial_ne_zero, one_smul] end HasFPowerSeriesOnBall diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index 7c30027db61b2..031c14bda50c7 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -444,7 +444,7 @@ theorem HasFDerivAt.lim (hf : HasFDerivAt f f' x) (v : E) {α : Type*} {c : α refine (eventually_ne_of_tendsto_norm_atTop hc (0 : 𝕜)).mono fun y hy => ?_ convert mem_of_mem_nhds hU dsimp only - rw [← mul_smul, mul_inv_cancel hy, one_smul] + rw [← mul_smul, mul_inv_cancel₀ hy, one_smul] theorem HasFDerivAt.unique (h₀ : HasFDerivAt f f₀' x) (h₁ : HasFDerivAt f f₁' x) : f₀' = f₁' := by rw [← hasFDerivWithinAt_univ] at h₀ h₁ diff --git a/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean b/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean index 7c8970d5cbecb..eb08a5c169706 100644 --- a/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/LineDeriv/Basic.lean @@ -506,7 +506,7 @@ theorem HasLineDerivWithinAt.smul (h : HasLineDerivWithinAt 𝕜 f f' s x v) (c theorem hasLineDerivWithinAt_smul_iff {c : 𝕜} (hc : c ≠ 0) : HasLineDerivWithinAt 𝕜 f (c • f') s x (c • v) ↔ HasLineDerivWithinAt 𝕜 f f' s x v := - ⟨fun h ↦ by simpa [smul_smul, inv_mul_cancel hc] using h.smul (c ⁻¹), fun h ↦ h.smul c⟩ + ⟨fun h ↦ by simpa [smul_smul, inv_mul_cancel₀ hc] using h.smul (c ⁻¹), fun h ↦ h.smul c⟩ theorem HasLineDerivAt.smul (h : HasLineDerivAt 𝕜 f f' x v) (c : 𝕜) : HasLineDerivAt 𝕜 f (c • f') x (c • v) := by @@ -515,7 +515,7 @@ theorem HasLineDerivAt.smul (h : HasLineDerivAt 𝕜 f f' x v) (c : 𝕜) : theorem hasLineDerivAt_smul_iff {c : 𝕜} (hc : c ≠ 0) : HasLineDerivAt 𝕜 f (c • f') x (c • v) ↔ HasLineDerivAt 𝕜 f f' x v := - ⟨fun h ↦ by simpa [smul_smul, inv_mul_cancel hc] using h.smul (c ⁻¹), fun h ↦ h.smul c⟩ + ⟨fun h ↦ by simpa [smul_smul, inv_mul_cancel₀ hc] using h.smul (c ⁻¹), fun h ↦ h.smul c⟩ theorem LineDifferentiableWithinAt.smul (h : LineDifferentiableWithinAt 𝕜 f s x v) (c : 𝕜) : LineDifferentiableWithinAt 𝕜 f s x (c • v) := @@ -523,7 +523,7 @@ theorem LineDifferentiableWithinAt.smul (h : LineDifferentiableWithinAt 𝕜 f s theorem lineDifferentiableWithinAt_smul_iff {c : 𝕜} (hc : c ≠ 0) : LineDifferentiableWithinAt 𝕜 f s x (c • v) ↔ LineDifferentiableWithinAt 𝕜 f s x v := - ⟨fun h ↦ by simpa [smul_smul, inv_mul_cancel hc] using h.smul (c ⁻¹), fun h ↦ h.smul c⟩ + ⟨fun h ↦ by simpa [smul_smul, inv_mul_cancel₀ hc] using h.smul (c ⁻¹), fun h ↦ h.smul c⟩ theorem LineDifferentiableAt.smul (h : LineDifferentiableAt 𝕜 f x v) (c : 𝕜) : LineDifferentiableAt 𝕜 f x (c • v) := @@ -531,7 +531,7 @@ theorem LineDifferentiableAt.smul (h : LineDifferentiableAt 𝕜 f x v) (c : theorem lineDifferentiableAt_smul_iff {c : 𝕜} (hc : c ≠ 0) : LineDifferentiableAt 𝕜 f x (c • v) ↔ LineDifferentiableAt 𝕜 f x v := - ⟨fun h ↦ by simpa [smul_smul, inv_mul_cancel hc] using h.smul (c ⁻¹), fun h ↦ h.smul c⟩ + ⟨fun h ↦ by simpa [smul_smul, inv_mul_cancel₀ hc] using h.smul (c ⁻¹), fun h ↦ h.smul c⟩ theorem lineDeriv_smul {c : 𝕜} : lineDeriv 𝕜 f x (c • v) = c • lineDeriv 𝕜 f x v := by rcases eq_or_ne c 0 with rfl|hc diff --git a/Mathlib/Analysis/Calculus/Rademacher.lean b/Mathlib/Analysis/Calculus/Rademacher.lean index 46a730bff90c4..5d94dc5f85717 100644 --- a/Mathlib/Analysis/Calculus/Rademacher.lean +++ b/Mathlib/Analysis/Calculus/Rademacher.lean @@ -280,8 +280,8 @@ theorem hasFderivAt_of_hasLineDerivAt_of_closure {f : E → F} rcases eq_or_ne v 0 with rfl|v_ne · simp obtain ⟨w, ρ, w_mem, hvw, hρ⟩ : ∃ w ρ, w ∈ sphere 0 1 ∧ v = ρ • w ∧ ρ = ‖v‖ := by - refine ⟨‖v‖⁻¹ • v, ‖v‖, by simp [norm_smul, inv_mul_cancel (norm_ne_zero_iff.2 v_ne)], ?_, rfl⟩ - simp [smul_smul, mul_inv_cancel (norm_ne_zero_iff.2 v_ne)] + refine ⟨‖v‖⁻¹ • v, ‖v‖, by simp [norm_smul, inv_mul_cancel₀ (norm_ne_zero_iff.2 v_ne)], ?_, rfl⟩ + simp [smul_smul, mul_inv_cancel₀ (norm_ne_zero_iff.2 v_ne)] have norm_rho : ‖ρ‖ = ρ := by rw [hρ, norm_norm] have rho_pos : 0 ≤ ρ := by simp [hρ] obtain ⟨y, yq, hy⟩ : ∃ y ∈ q, ‖w - y‖ < δ := by simpa [← dist_eq_norm] using hq w_mem diff --git a/Mathlib/Analysis/Calculus/TangentCone.lean b/Mathlib/Analysis/Calculus/TangentCone.lean index d75973e9bd895..e4c5b68d694c1 100644 --- a/Mathlib/Analysis/Calculus/TangentCone.lean +++ b/Mathlib/Analysis/Calculus/TangentCone.lean @@ -104,7 +104,7 @@ theorem tangentConeAt.lim_zero {α : Type*} (l : Filter α) {c : α → 𝕜} {d rw [zero_mul] at C have : ∀ᶠ n in l, ‖c n‖⁻¹ * ‖c n • d n‖ = ‖d n‖ := by refine (eventually_ne_of_tendsto_norm_atTop hc 0).mono fun n hn => ?_ - rw [norm_smul, ← mul_assoc, inv_mul_cancel, one_mul] + rw [norm_smul, ← mul_assoc, inv_mul_cancel₀, one_mul] rwa [Ne, norm_eq_zero] have D : Tendsto (fun n => ‖d n‖) l (𝓝 0) := Tendsto.congr' this C rw [tendsto_zero_iff_norm_tendsto_zero] diff --git a/Mathlib/Analysis/Complex/OpenMapping.lean b/Mathlib/Analysis/Complex/OpenMapping.lean index c126494745575..9ce4d01945319 100644 --- a/Mathlib/Analysis/Complex/OpenMapping.lean +++ b/Mathlib/Analysis/Complex/OpenMapping.lean @@ -132,14 +132,14 @@ theorem AnalyticAt.eventually_constant_or_nhds_le_map_nhds {z₀ : E} (hg : Anal let w : E := ‖z - z₀‖⁻¹ • (z - z₀) have h3 : ∀ t ∈ ball (0 : ℂ) r, gray w t = g z₀ := by have e1 : IsPreconnected (ball (0 : ℂ) r) := (convex_ball 0 r).isPreconnected - have e2 : w ∈ sphere (0 : E) 1 := by simp [w, norm_smul, inv_mul_cancel h'] + have e2 : w ∈ sphere (0 : E) 1 := by simp [w, norm_smul, inv_mul_cancel₀ h'] specialize h1 w e2 apply h1.eqOn_of_preconnected_of_eventuallyEq analyticOn_const e1 (mem_ball_self hr) simpa [ray, gray] using h w e2 have h4 : ‖z - z₀‖ < r := by simpa [dist_eq_norm] using mem_ball.mp hz replace h4 : ↑‖z - z₀‖ ∈ ball (0 : ℂ) r := by simpa only [mem_ball_zero_iff, norm_eq_abs, abs_ofReal, abs_norm] - simpa only [ray, gray, w, smul_smul, mul_inv_cancel h', one_smul, add_sub_cancel, + simpa only [ray, gray, w, smul_smul, mul_inv_cancel₀ h', one_smul, add_sub_cancel, Function.comp_apply, coe_smul] using h3 (↑‖z - z₀‖) h4 · right -- Otherwise, it is open along at least one direction and that implies the result diff --git a/Mathlib/Analysis/Convex/Between.lean b/Mathlib/Analysis/Convex/Between.lean index b15ff896c87f7..ccbae3b13d6c8 100644 --- a/Mathlib/Analysis/Convex/Between.lean +++ b/Mathlib/Analysis/Convex/Between.lean @@ -585,14 +585,14 @@ theorem wbtw_iff_left_eq_or_right_mem_image_Ici {x y z : P} : · rw [Set.mem_image] refine Or.inr ⟨r⁻¹, one_le_inv hr0' hr1, ?_⟩ simp only [lineMap_apply, smul_smul, vadd_vsub] - rw [inv_mul_cancel hr0'.ne', one_smul, vsub_vadd] + rw [inv_mul_cancel₀ hr0'.ne', one_smul, vsub_vadd] · simp · rcases h with (rfl | ⟨r, ⟨hr, rfl⟩⟩) · exact wbtw_self_left _ _ _ · rw [Set.mem_Ici] at hr refine ⟨r⁻¹, ⟨inv_nonneg.2 (zero_le_one.trans hr), inv_le_one hr⟩, ?_⟩ simp only [lineMap_apply, smul_smul, vadd_vsub] - rw [inv_mul_cancel (one_pos.trans_le hr).ne', one_smul, vsub_vadd] + rw [inv_mul_cancel₀ (one_pos.trans_le hr).ne', one_smul, vsub_vadd] theorem Wbtw.right_mem_image_Ici_of_left_ne {x y z : P} (h : Wbtw R x y z) (hne : x ≠ y) : z ∈ lineMap x y '' Set.Ici (1 : R) := diff --git a/Mathlib/Analysis/Convex/Combination.lean b/Mathlib/Analysis/Convex/Combination.lean index 58926a2e62815..a1b480de2d161 100644 --- a/Mathlib/Analysis/Convex/Combination.lean +++ b/Mathlib/Analysis/Convex/Combination.lean @@ -60,10 +60,10 @@ theorem Finset.centerMass_insert (ha : i ∉ t) (hw : ∑ j ∈ t, w j ≠ 0) : ((∑ j ∈ t, w j) / (w i + ∑ j ∈ t, w j)) • t.centerMass w z := by simp only [centerMass, sum_insert ha, smul_add, (mul_smul _ _ _).symm, ← div_eq_inv_mul] congr 2 - rw [div_mul_eq_mul_div, mul_inv_cancel hw, one_div] + rw [div_mul_eq_mul_div, mul_inv_cancel₀ hw, one_div] theorem Finset.centerMass_singleton (hw : w i ≠ 0) : ({i} : Finset ι).centerMass w z = z i := by - rw [centerMass, sum_singleton, sum_singleton, ← mul_smul, inv_mul_cancel hw, one_smul] + rw [centerMass, sum_singleton, sum_singleton, ← mul_smul, inv_mul_cancel₀ hw, one_smul] @[simp] lemma Finset.centerMass_neg_left : t.centerMass (-w) z = t.centerMass w z := by simp [centerMass, inv_neg] @@ -163,7 +163,7 @@ theorem Convex.centerMass_mem (hs : Convex R s) : have wz : ∑ j ∈ t, w j • z j = 0 := sum_eq_zero fun i hi => by simp [ws i hi] simp only [centerMass, sum_insert hi, wz, hsum_t, add_zero] simp only [hsum_t, add_zero] at hpos - rw [← mul_smul, inv_mul_cancel (ne_of_gt hpos), one_smul] + rw [← mul_smul, inv_mul_cancel₀ (ne_of_gt hpos), one_smul] exact zi · rw [Finset.centerMass_insert _ _ _ hi hsum_t] refine convex_iff_div.1 hs zi (ht hs₀ ?_ ?_) ?_ (sum_nonneg hs₀) hpos @@ -262,7 +262,7 @@ theorem Finset.centroid_mem_convexHull (s : Finset E) (hs : s.Nonempty) : apply s.centerMass_id_mem_convexHull · simp only [inv_nonneg, imp_true_iff, Nat.cast_nonneg, Finset.centroidWeights_apply] · have hs_card : (s.card : R) ≠ 0 := by simp [Finset.nonempty_iff_ne_empty.mp hs] - simp only [hs_card, Finset.sum_const, nsmul_eq_mul, mul_inv_cancel, Ne, not_false_iff, + simp only [hs_card, Finset.sum_const, nsmul_eq_mul, mul_inv_cancel₀, Ne, not_false_iff, Finset.centroidWeights_apply, zero_lt_one] theorem convexHull_range_eq_exists_affineCombination (v : ι → E) : convexHull R (range v) = diff --git a/Mathlib/Analysis/Convex/Cone/Basic.lean b/Mathlib/Analysis/Convex/Cone/Basic.lean index 095dee8252a90..fbc04fd0cfde0 100644 --- a/Mathlib/Analysis/Convex/Cone/Basic.lean +++ b/Mathlib/Analysis/Convex/Cone/Basic.lean @@ -339,7 +339,7 @@ theorem Salient.anti {S T : ConvexCone 𝕜 E} (h : T ≤ S) : S.Salient → T.S /-- A flat cone is always pointed (contains `0`). -/ theorem Flat.pointed {S : ConvexCone 𝕜 E} (hS : S.Flat) : S.Pointed := by obtain ⟨x, hx, _, hxneg⟩ := hS - rw [Pointed, ← add_neg_self x] + rw [Pointed, ← add_neg_cancel x] exact add_mem S hx hxneg /-- A blunt cone (one not containing `0`) is always salient. -/ @@ -509,7 +509,7 @@ theorem salient_positive : Salient (positive 𝕜 E) := fun x xs hx hx' => (calc 0 < x := lt_of_le_of_ne xs hx.symm _ ≤ x + -x := le_add_of_nonneg_right hx' - _ = 0 := add_neg_self x + _ = 0 := add_neg_cancel x ) /-- The positive cone of an ordered module is always pointed. -/ @@ -574,9 +574,9 @@ theorem mem_toCone : x ∈ hs.toCone s ↔ ∃ c : 𝕜, 0 < c ∧ ∃ y ∈ s, theorem mem_toCone' : x ∈ hs.toCone s ↔ ∃ c : 𝕜, 0 < c ∧ c • x ∈ s := by refine hs.mem_toCone.trans ⟨?_, ?_⟩ · rintro ⟨c, hc, y, hy, rfl⟩ - exact ⟨c⁻¹, inv_pos.2 hc, by rwa [smul_smul, inv_mul_cancel hc.ne', one_smul]⟩ + exact ⟨c⁻¹, inv_pos.2 hc, by rwa [smul_smul, inv_mul_cancel₀ hc.ne', one_smul]⟩ · rintro ⟨c, hc, hcx⟩ - exact ⟨c⁻¹, inv_pos.2 hc, _, hcx, by rw [smul_smul, inv_mul_cancel hc.ne', one_smul]⟩ + exact ⟨c⁻¹, inv_pos.2 hc, _, hcx, by rw [smul_smul, inv_mul_cancel₀ hc.ne', one_smul]⟩ theorem subset_toCone : s ⊆ hs.toCone s := fun x hx => hs.mem_toCone'.2 ⟨1, zero_lt_one, by rwa [one_smul]⟩ diff --git a/Mathlib/Analysis/Convex/Cone/Extension.lean b/Mathlib/Analysis/Convex/Cone/Extension.lean index 8b2c4f3992353..0e1685f92d6f4 100644 --- a/Mathlib/Analysis/Convex/Cone/Extension.lean +++ b/Mathlib/Analysis/Convex/Cone/Extension.lean @@ -92,20 +92,20 @@ theorem step (nonneg : ∀ x : f.domain, (x : E) ∈ s → 0 ≤ f x) rcases lt_trichotomy r 0 with (hr | hr | hr) · have : -(r⁻¹ • x) - y ∈ s := by rwa [← s.smul_mem_iff (neg_pos.2 hr), smul_sub, smul_neg, neg_smul, neg_neg, smul_smul, - mul_inv_cancel hr.ne, one_smul, sub_eq_add_neg, neg_smul, neg_neg] + mul_inv_cancel₀ hr.ne, one_smul, sub_eq_add_neg, neg_smul, neg_neg] -- Porting note: added type annotation and `by exact` replace : f (r⁻¹ • ⟨x, hx⟩) ≤ c := le_c (r⁻¹ • ⟨x, hx⟩) (by exact this) rwa [← mul_le_mul_left (neg_pos.2 hr), neg_mul, neg_mul, neg_le_neg_iff, f.map_smul, - smul_eq_mul, ← mul_assoc, mul_inv_cancel hr.ne, one_mul] at this + smul_eq_mul, ← mul_assoc, mul_inv_cancel₀ hr.ne, one_mul] at this · subst r simp only [zero_smul, add_zero] at hzs ⊢ apply nonneg exact hzs · have : r⁻¹ • x + y ∈ s := by - rwa [← s.smul_mem_iff hr, smul_add, smul_smul, mul_inv_cancel hr.ne', one_smul] + rwa [← s.smul_mem_iff hr, smul_add, smul_smul, mul_inv_cancel₀ hr.ne', one_smul] -- Porting note: added type annotation and `by exact` replace : c ≤ f (r⁻¹ • ⟨x, hx⟩) := c_le (r⁻¹ • ⟨x, hx⟩) (by exact this) - rwa [← mul_le_mul_left hr, f.map_smul, smul_eq_mul, ← mul_assoc, mul_inv_cancel hr.ne', + rwa [← mul_le_mul_left hr, f.map_smul, smul_eq_mul, ← mul_assoc, mul_inv_cancel₀ hr.ne', one_mul] at this theorem exists_top (p : E →ₗ.[ℝ] ℝ) (hp_nonneg : ∀ x : p.domain, (x : E) ∈ s → 0 ≤ p x) diff --git a/Mathlib/Analysis/Convex/Segment.lean b/Mathlib/Analysis/Convex/Segment.lean index 6816d169f39c0..87adc0a70ac2e 100644 --- a/Mathlib/Analysis/Convex/Segment.lean +++ b/Mathlib/Analysis/Convex/Segment.lean @@ -360,9 +360,9 @@ theorem mem_segment_iff_sameRay : x ∈ [y -[𝕜] z] ↔ SameRay 𝕜 (x - y) ( refine ⟨sameRay_of_mem_segment, fun h => ?_⟩ rcases h.exists_eq_smul_add with ⟨a, b, ha, hb, hab, hxy, hzx⟩ rw [add_comm, sub_add_sub_cancel] at hxy hzx - rw [← mem_segment_translate _ (-x), neg_add_self] + rw [← mem_segment_translate _ (-x), neg_add_cancel] refine ⟨b, a, hb, ha, add_comm a b ▸ hab, ?_⟩ - rw [← sub_eq_neg_add, ← neg_sub, hxy, ← sub_eq_neg_add, hzx, smul_neg, smul_comm, neg_add_self] + rw [← sub_eq_neg_add, ← neg_sub, hxy, ← sub_eq_neg_add, hzx, smul_neg, smul_comm, neg_add_cancel] open AffineMap diff --git a/Mathlib/Analysis/Convex/Side.lean b/Mathlib/Analysis/Convex/Side.lean index ad3b21ef618f1..13933585a56e0 100644 --- a/Mathlib/Analysis/Convex/Side.lean +++ b/Mathlib/Analysis/Convex/Side.lean @@ -587,7 +587,7 @@ theorem wOppSide_iff_exists_wbtw {s : AffineSubspace R P} {x y : P} : rw [add_comm (y -ᵥ p₂), smul_sub, smul_add, add_sub_assoc, add_assoc, add_right_eq_self, div_eq_inv_mul, ← neg_vsub_eq_vsub_rev, smul_neg, ← smul_smul, ← h, smul_smul, ← neg_smul, ← sub_smul, ← div_eq_inv_mul, ← div_eq_inv_mul, ← neg_div, ← sub_div, sub_eq_add_neg, - ← neg_add, neg_div, div_self (Left.add_pos hr₁ hr₂).ne.symm, neg_one_smul, neg_add_self] + ← neg_add, neg_div, div_self (Left.add_pos hr₁ hr₂).ne.symm, neg_one_smul, neg_add_cancel] 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₁ @@ -666,7 +666,7 @@ theorem setOf_wSameSide_eq_image2 {s : AffineSubspace R P} {x p : P} (hx : x ∉ refine ⟨0, le_rfl, p₂, hp₂, ?_⟩ simp [h] · refine ⟨r₁ / r₂, (div_pos hr₁ hr₂).le, p₂, hp₂, ?_⟩ - rw [div_eq_inv_mul, ← smul_smul, h, smul_smul, inv_mul_cancel hr₂.ne.symm, one_smul, + rw [div_eq_inv_mul, ← smul_smul, h, smul_smul, inv_mul_cancel₀ hr₂.ne.symm, one_smul, vsub_vadd] · rintro ⟨t, ht, p', hp', rfl⟩ exact wSameSide_smul_vsub_vadd_right x hp hp' ht @@ -683,7 +683,7 @@ theorem setOf_sSameSide_eq_image2 {s : AffineSubspace R P} {x p : P} (hx : x ∉ · rw [vsub_eq_zero_iff_eq] at h exact False.elim (hy (h.symm ▸ hp₂)) · refine ⟨r₁ / r₂, div_pos hr₁ hr₂, p₂, hp₂, ?_⟩ - rw [div_eq_inv_mul, ← smul_smul, h, smul_smul, inv_mul_cancel hr₂.ne.symm, one_smul, + rw [div_eq_inv_mul, ← smul_smul, h, smul_smul, inv_mul_cancel₀ hr₂.ne.symm, one_smul, vsub_vadd] · rintro ⟨t, ht, p', hp', rfl⟩ exact sSameSide_smul_vsub_vadd_right hx hp hp' ht @@ -701,8 +701,8 @@ theorem setOf_wOppSide_eq_image2 {s : AffineSubspace R P} {x p : P} (hx : x ∉ refine ⟨0, le_rfl, p₂, hp₂, ?_⟩ simp [h] · refine ⟨-r₁ / r₂, (div_neg_of_neg_of_pos (Left.neg_neg_iff.2 hr₁) hr₂).le, p₂, hp₂, ?_⟩ - rw [div_eq_inv_mul, ← smul_smul, neg_smul, h, smul_neg, smul_smul, inv_mul_cancel hr₂.ne.symm, - one_smul, neg_vsub_eq_vsub_rev, vsub_vadd] + rw [div_eq_inv_mul, ← smul_smul, neg_smul, h, smul_neg, smul_smul, + inv_mul_cancel₀ hr₂.ne.symm, one_smul, neg_vsub_eq_vsub_rev, vsub_vadd] · rintro ⟨t, ht, p', hp', rfl⟩ exact wOppSide_smul_vsub_vadd_right x hp hp' ht @@ -718,8 +718,8 @@ theorem setOf_sOppSide_eq_image2 {s : AffineSubspace R P} {x p : P} (hx : x ∉ · rw [vsub_eq_zero_iff_eq] at h exact False.elim (hy (h ▸ hp₂)) · refine ⟨-r₁ / r₂, div_neg_of_neg_of_pos (Left.neg_neg_iff.2 hr₁) hr₂, p₂, hp₂, ?_⟩ - rw [div_eq_inv_mul, ← smul_smul, neg_smul, h, smul_neg, smul_smul, inv_mul_cancel hr₂.ne.symm, - one_smul, neg_vsub_eq_vsub_rev, vsub_vadd] + rw [div_eq_inv_mul, ← smul_smul, neg_smul, h, smul_neg, smul_smul, + inv_mul_cancel₀ hr₂.ne.symm, one_smul, neg_vsub_eq_vsub_rev, vsub_vadd] · rintro ⟨t, ht, p', hp', rfl⟩ exact sOppSide_smul_vsub_vadd_right hx hp hp' ht diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean index 728bd424c1242..922b18543ffe0 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean @@ -95,7 +95,7 @@ theorem one_add_mul_self_lt_rpow_one_add {s : ℝ} (hs : -1 ≤ s) (hs' : s ≠ 1 + p * s < (1 + s) ^ p := by have hp' : 0 < p := zero_lt_one.trans hp rcases eq_or_lt_of_le hs with rfl | hs - · rwa [add_right_neg, zero_rpow hp'.ne', mul_neg_one, add_neg_lt_iff_lt_add, zero_add] + · rwa [add_neg_cancel, zero_rpow hp'.ne', mul_neg_one, add_neg_lt_iff_lt_add, zero_add] have hs1 : 0 < 1 + s := neg_lt_iff_pos_add'.mp hs rcases le_or_lt (1 + p * s) 0 with hs2 | hs2 · exact hs2.trans_lt (rpow_pos_of_pos hs1 _) @@ -131,7 +131,7 @@ with `s ≠ 0`, we have `(1 + s) ^ p < 1 + p * s`. -/ theorem rpow_one_add_lt_one_add_mul_self {s : ℝ} (hs : -1 ≤ s) (hs' : s ≠ 0) {p : ℝ} (hp1 : 0 < p) (hp2 : p < 1) : (1 + s) ^ p < 1 + p * s := by rcases eq_or_lt_of_le hs with rfl | hs - · rwa [add_right_neg, zero_rpow hp1.ne', mul_neg_one, lt_add_neg_iff_add_lt, zero_add] + · rwa [add_neg_cancel, zero_rpow hp1.ne', mul_neg_one, lt_add_neg_iff_add_lt, zero_add] have hs1 : 0 < 1 + s := neg_lt_iff_pos_add'.mp hs have hs2 : 0 < 1 + p * s := by rw [← neg_lt_iff_pos_add'] @@ -180,7 +180,7 @@ theorem strictConvexOn_rpow {p : ℝ} (hp : 1 < p) : StrictConvexOn ℝ (Ici 0) mul_div_assoc, ← rpow_sub hy, sub_sub_cancel_left, rpow_neg_one, mul_assoc, ← div_eq_inv_mul, sub_eq_add_neg, ← mul_neg, ← neg_div, neg_sub, _root_.sub_div, div_self hy.ne'] apply one_add_mul_self_lt_rpow_one_add _ _ hp - · rw [le_sub_iff_add_le, add_left_neg, div_nonneg_iff] + · rw [le_sub_iff_add_le, neg_add_cancel, div_nonneg_iff] exact Or.inl ⟨hx, hy.le⟩ · rw [sub_ne_zero] exact ((div_lt_one hy).mpr hxy).ne @@ -190,7 +190,7 @@ theorem strictConvexOn_rpow {p : ℝ} (hp : 1 < p) : StrictConvexOn ℝ (Ici 0) ← div_mul_eq_mul_div, mul_div_assoc, ← rpow_sub hy, sub_sub_cancel_left, rpow_neg_one, mul_assoc, ← div_eq_inv_mul, _root_.sub_div, div_self hy.ne'] apply one_add_mul_self_lt_rpow_one_add _ _ hp - · rw [le_sub_iff_add_le, add_left_neg, div_nonneg_iff] + · rw [le_sub_iff_add_le, neg_add_cancel, div_nonneg_iff] exact Or.inl ⟨hz, hy.le⟩ · rw [sub_ne_zero] exact ((one_lt_div hy).mpr hyz).ne' diff --git a/Mathlib/Analysis/Convex/Uniform.lean b/Mathlib/Analysis/Convex/Uniform.lean index 0b7f265c69a97..a5ad07edba1ae 100644 --- a/Mathlib/Analysis/Convex/Uniform.lean +++ b/Mathlib/Analysis/Convex/Uniform.lean @@ -68,12 +68,12 @@ theorem exists_forall_closed_ball_dist_add_le_two_sub (hε : 0 < ε) : have hδ' : 0 < 1 - δ' := sub_pos_of_lt (min_lt_of_left_lt one_half_lt_one) have h₁ : ∀ z : E, 1 - δ' < ‖z‖ → ‖‖z‖⁻¹ • z‖ = 1 := by rintro z hz - rw [norm_smul_of_nonneg (inv_nonneg.2 <| norm_nonneg _), inv_mul_cancel (hδ'.trans hz).ne'] + rw [norm_smul_of_nonneg (inv_nonneg.2 <| norm_nonneg _), inv_mul_cancel₀ (hδ'.trans hz).ne'] have h₂ : ∀ z : E, ‖z‖ ≤ 1 → 1 - δ' ≤ ‖z‖ → ‖‖z‖⁻¹ • z - z‖ ≤ δ' := by rintro z hz hδz nth_rw 3 [← one_smul ℝ z] rwa [← sub_smul, norm_smul_of_nonneg (sub_nonneg_of_le <| one_le_inv (hδ'.trans_le hδz) hz), - sub_mul, inv_mul_cancel (hδ'.trans_le hδz).ne', one_mul, sub_le_comm] + sub_mul, inv_mul_cancel₀ (hδ'.trans_le hδz).ne', one_mul, sub_le_comm] set x' := ‖x‖⁻¹ • x set y' := ‖y‖⁻¹ • y have hxy' : ε / 3 ≤ ‖x' - y'‖ := diff --git a/Mathlib/Analysis/Fourier/AddCircle.lean b/Mathlib/Analysis/Fourier/AddCircle.lean index 7ddc90e99180e..45761b766817c 100644 --- a/Mathlib/Analysis/Fourier/AddCircle.lean +++ b/Mathlib/Analysis/Fourier/AddCircle.lean @@ -255,7 +255,7 @@ theorem orthonormal_fourier : Orthonormal ℂ (@fourierLp T _ 2 _) := by rw [ContinuousMap.inner_toLp (@haarAddCircle T hT) (fourier i) (fourier j)] simp_rw [← fourier_neg, ← fourier_add] split_ifs with h - · simp_rw [h, neg_add_self] + · simp_rw [h, neg_add_cancel] have : ⇑(@fourier T 0) = (fun _ => 1 : AddCircle T → ℂ) := by ext1; exact fourier_zero rw [this, integral_const, measure_univ, ENNReal.one_toReal, Complex.real_smul, Complex.ofReal_one, mul_one] diff --git a/Mathlib/Analysis/Fourier/FourierTransform.lean b/Mathlib/Analysis/Fourier/FourierTransform.lean index 7580cc618b9f9..57b7a7debd124 100644 --- a/Mathlib/Analysis/Fourier/FourierTransform.lean +++ b/Mathlib/Analysis/Fourier/FourierTransform.lean @@ -136,7 +136,7 @@ theorem fourierIntegral_convergent_iff (he : Continuous e) -- then use it for both directions refine ⟨fun hf ↦ ?_, fun hf ↦ aux hf w⟩ have := aux hf (-w) - simp_rw [← mul_smul (e _) (e _) (f _), ← e.map_add_eq_mul, LinearMap.map_neg, neg_add_self, + simp_rw [← mul_smul (e _) (e _) (f _), ← e.map_add_eq_mul, LinearMap.map_neg, neg_add_cancel, e.map_zero_eq_one, one_smul] at this -- the `(e _)` speeds up elaboration considerably exact this diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 6ef37580057c0..82122b6dce018 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -450,7 +450,7 @@ theorem eLpNorm_le_eLpNorm_fderiv_one {u : E → F} (hu : ContDiff ℝ 1 u) (h2 ← ENNReal.rpow_le_rpow_iff h0p, ENNReal.mul_rpow_of_nonneg _ _ h0p.le, ENNReal.coe_rpow_of_nonneg _ h0p.le, eLpNormLESNormFDerivOneConst, ← NNReal.rpow_mul, eLpNorm_nnreal_pow_eq_lintegral hp.symm.pos.ne', - inv_mul_cancel h0p.ne', NNReal.rpow_one] + inv_mul_cancel₀ h0p.ne', NNReal.rpow_one] exact lintegral_pow_le_pow_lintegral_fderiv μ hu h2u hp.coe @[deprecated (since := "2024-07-27")] diff --git a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean index 542e4fe0e33dc..cc637d03ff57a 100644 --- a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean +++ b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean @@ -278,7 +278,7 @@ private theorem I_prop : innerProp' E (I : 𝕜) := by by_cases hI : (I : 𝕜) = 0 · rw [hI, ← Nat.cast_zero]; exact nat_prop _ intro x y - have hI' : (-I : 𝕜) * I = 1 := by rw [← inv_I, inv_mul_cancel hI] + have hI' : (-I : 𝕜) * I = 1 := by rw [← inv_I, inv_mul_cancel₀ hI] rw [conj_I, inner_, inner_, mul_left_comm] congr 1 rw [smul_smul, I_mul_I_of_nonzero hI, neg_one_smul] diff --git a/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean b/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean index 9e29639ef61d0..a58db852ed74d 100644 --- a/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean +++ b/Mathlib/Analysis/LocallyConvex/ContinuousOfBounded.lean @@ -136,7 +136,7 @@ theorem LinearMap.continuousAt_zero_of_locally_bounded (f : E →ₛₗ[σ] F) rcases hu n h with ⟨y, hy, hu1⟩ convert hy rw [← hu1, ← mul_smul] - simp only [h, mul_inv_cancel, Ne, Nat.cast_eq_zero, not_false_iff, one_smul] + simp only [h, mul_inv_cancel₀, Ne, Nat.cast_eq_zero, not_false_iff, one_smul] -- The image `(fun n ↦ n • u n)` is von Neumann bounded: have h_bounded : IsVonNBounded 𝕜 (Set.range fun n : ℕ => (n : 𝕜) • u n) := h_tendsto.cauchySeq.totallyBounded_range.isVonNBounded 𝕜 diff --git a/Mathlib/Analysis/Matrix.lean b/Mathlib/Analysis/Matrix.lean index d692efa7722f7..e691c6be69c1d 100644 --- a/Mathlib/Analysis/Matrix.lean +++ b/Mathlib/Analysis/Matrix.lean @@ -399,13 +399,13 @@ private theorem norm_unitOf (a : α) : ‖unitOf a‖₊ = 1 := by split_ifs with h · simp · rw [← nnnorm_eq_zero] at h - rw [nnnorm_smul, nnnorm_inv, nnnorm_norm, mul_inv_cancel h] + rw [nnnorm_smul, nnnorm_inv, nnnorm_norm, mul_inv_cancel₀ h] private theorem mul_unitOf (a : α) : a * unitOf a = algebraMap _ _ (‖a‖₊ : ℝ) := by simp only [unitOf, coe_nnnorm] split_ifs with h · simp [h] - · rw [mul_smul_comm, mul_inv_cancel h, Algebra.algebraMap_eq_smul_one] + · rw [mul_smul_comm, mul_inv_cancel₀ h, Algebra.algebraMap_eq_smul_one] end diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index faa0d4a8ba9e3..7dc403aec3cca 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -153,7 +153,7 @@ theorem geom_mean_le_arith_mean {ι : Type*} (s : Finset ι) (w : ι → ℝ) (z · simp_rw [div_eq_mul_inv, mul_assoc, mul_comm, ← mul_assoc, ← Finset.sum_mul, mul_comm] · exact fun _ hi => div_nonneg (hw _ hi) (le_of_lt hw') · simp_rw [div_eq_mul_inv, ← Finset.sum_mul] - exact mul_inv_cancel (by linarith) + exact mul_inv_cancel₀ (by linarith) theorem geom_mean_weighted_of_constant (w z : ι → ℝ) (x : ℝ) (hw : ∀ i ∈ s, 0 ≤ w i) (hw' : ∑ i ∈ s, w i = 1) (hz : ∀ i ∈ s, 0 ≤ z i) (hx : ∀ i ∈ s, w i ≠ 0 → z i = x) : @@ -295,7 +295,7 @@ theorem harm_mean_le_geom_mean {ι : Type*} (s : Finset ι) (hs : s.Nonempty) (w rw [← Real.rpow_mul (le_of_lt <| hz i hi) (w _) n⁻¹, div_eq_mul_inv (w _) n] · exact fun i hi ↦ div_pos (hw i hi) hw' · simp_rw [div_eq_mul_inv (w _) (∑ i in s, w i), ← Finset.sum_mul _ _ (∑ i in s, w i)⁻¹] - exact mul_inv_cancel hw'.ne' + exact mul_inv_cancel₀ hw'.ne' end Real @@ -422,9 +422,9 @@ theorem inner_le_Lp_mul_Lq (f g : ι → ℝ≥0) {p q : ℝ} (hpq : p.IsConjExp · rw [Ne, rpow_eq_zero_iff, not_and_or] exact Or.inl hG_zero refine inner_le_Lp_mul_Lp_of_norm_le_one s f' g' hpq (le_of_eq ?_) (le_of_eq ?_) - · simp_rw [f', div_rpow, ← sum_div, ← rpow_mul, one_div, inv_mul_cancel hpq.ne_zero, rpow_one, + · simp_rw [f', div_rpow, ← sum_div, ← rpow_mul, one_div, inv_mul_cancel₀ hpq.ne_zero, rpow_one, div_self hF_zero] - · simp_rw [g', div_rpow, ← sum_div, ← rpow_mul, one_div, inv_mul_cancel hpq.symm.ne_zero, + · simp_rw [g', div_rpow, ← sum_div, ← rpow_mul, one_div, inv_mul_cancel₀ hpq.symm.ne_zero, rpow_one, div_self hG_zero] /-- **Weighted Hölder inequality**. -/ diff --git a/Mathlib/Analysis/MeanInequalitiesPow.lean b/Mathlib/Analysis/MeanInequalitiesPow.lean index 1eb4445cc0366..bd0833e79f889 100644 --- a/Mathlib/Analysis/MeanInequalitiesPow.lean +++ b/Mathlib/Analysis/MeanInequalitiesPow.lean @@ -76,7 +76,7 @@ theorem pow_sum_div_card_le_sum_pow {f : ι → ℝ} (n : ℕ) (hf : ∀ a ∈ s Set.mem_Ici.2 (hf i hi) · simpa only [inv_mul_eq_div, one_div, Algebra.id.smul_eq_mul] using this · simp only [one_div, inv_nonneg, Nat.cast_nonneg, imp_true_iff] - · simpa only [one_div, Finset.sum_const, nsmul_eq_mul] using mul_inv_cancel hs0.ne' + · simpa only [one_div, Finset.sum_const, nsmul_eq_mul] using mul_inv_cancel₀ hs0.ne' theorem zpow_arith_mean_le_arith_mean_zpow (w z : ι → ℝ) (hw : ∀ i ∈ s, 0 ≤ w i) (hw' : ∑ i ∈ s, w i = 1) (hz : ∀ i ∈ s, 0 < z i) (m : ℤ) : @@ -173,7 +173,7 @@ theorem add_rpow_le_rpow_add {p : ℝ} (a b : ℝ≥0) (hp1 : 1 ≤ p) : a ^ p + nth_rw 4 [← mul_one ((a + b) ^ p)] exact (mul_le_mul_left hab_0').mpr h rwa [div_eq_mul_inv, div_eq_mul_inv, mul_add, mul_comm (a ^ p), mul_comm (b ^ p), ← mul_assoc, ← - mul_assoc, mul_inv_cancel hab_0, one_mul, one_mul] at h_mul + mul_assoc, mul_inv_cancel₀ hab_0, one_mul, one_mul] at h_mul theorem rpow_add_rpow_le_add {p : ℝ} (a b : ℝ≥0) (hp1 : 1 ≤ p) : (a ^ p + b ^ p) ^ (1 / p) ≤ a + b := by @@ -185,7 +185,7 @@ theorem rpow_add_rpow_le_add {p : ℝ} (a b : ℝ≥0) (hp1 : 1 ≤ p) : theorem rpow_add_rpow_le {p q : ℝ} (a b : ℝ≥0) (hp_pos : 0 < p) (hpq : p ≤ q) : (a ^ q + b ^ q) ^ (1 / q) ≤ (a ^ p + b ^ p) ^ (1 / p) := by have h_rpow : ∀ a : ℝ≥0, a ^ q = (a ^ p) ^ (q / p) := fun a => by - rw [← NNReal.rpow_mul, div_eq_inv_mul, ← mul_assoc, _root_.mul_inv_cancel hp_pos.ne.symm, + rw [← NNReal.rpow_mul, div_eq_inv_mul, ← mul_assoc, mul_inv_cancel₀ hp_pos.ne.symm, one_mul] have h_rpow_add_rpow_le_add : ((a ^ p) ^ (q / p) + (b ^ p) ^ (q / p)) ^ (1 / (q / p)) ≤ a ^ p + b ^ p := by diff --git a/Mathlib/Analysis/MellinInversion.lean b/Mathlib/Analysis/MellinInversion.lean index 24939d01cea82..c9680254f0421 100644 --- a/Mathlib/Analysis/MellinInversion.lean +++ b/Mathlib/Analysis/MellinInversion.lean @@ -118,4 +118,4 @@ theorem mellin_inversion (σ : ℝ) (f : ℝ → E) {x : ℝ} (hx : 0 < x) (hf : rw [mul_comm σ, ← rpow_def_of_pos hx, Real.exp_log hx, ← Complex.ofReal_cpow hx.le] norm_cast rw [← smul_assoc, smul_eq_mul, Real.rpow_neg hx.le, - inv_mul_cancel (ne_of_gt (rpow_pos_of_pos hx σ)), one_smul] + inv_mul_cancel₀ (ne_of_gt (rpow_pos_of_pos hx σ)), one_smul] diff --git a/Mathlib/Analysis/MellinTransform.lean b/Mathlib/Analysis/MellinTransform.lean index 19eccb18c49fb..81136cb0e8e3c 100644 --- a/Mathlib/Analysis/MellinTransform.lean +++ b/Mathlib/Analysis/MellinTransform.lean @@ -118,7 +118,7 @@ theorem mellin_comp_rpow (f : ℝ → E) (s : ℂ) (a : ℝ) : conv_rhs => rw [← integral_comp_rpow_Ioi _ ha, ← integral_smul] refine setIntegral_congr measurableSet_Ioi fun t ht => ?_ dsimp only - rw [← mul_smul, ← mul_assoc, inv_mul_cancel (mt abs_eq_zero.1 ha), one_mul, ← smul_assoc, + rw [← mul_smul, ← mul_assoc, inv_mul_cancel₀ (mt abs_eq_zero.1 ha), one_mul, ← smul_assoc, real_smul] rw [ofReal_cpow (le_of_lt ht), ← cpow_mul_ofReal_nonneg (le_of_lt ht), ← cpow_add _ _ (ofReal_ne_zero.mpr <| ne_of_gt ht), ofReal_sub, ofReal_one, mul_sub, diff --git a/Mathlib/Analysis/Normed/Affine/AddTorsor.lean b/Mathlib/Analysis/Normed/Affine/AddTorsor.lean index 070c79f4057a1..caa4e080fffdf 100644 --- a/Mathlib/Analysis/Normed/Affine/AddTorsor.lean +++ b/Mathlib/Analysis/Normed/Affine/AddTorsor.lean @@ -206,7 +206,7 @@ theorem antilipschitzWith_lineMap {p₁ p₂ : Q} (h : p₁ ≠ p₂) : AntilipschitzWith (nndist p₁ p₂)⁻¹ (lineMap p₁ p₂ : 𝕜 → Q) := AntilipschitzWith.of_le_mul_dist fun c₁ c₂ => by rw [dist_lineMap_lineMap, NNReal.coe_inv, ← dist_nndist, mul_left_comm, - inv_mul_cancel (dist_ne_zero.2 h), mul_one] + inv_mul_cancel₀ (dist_ne_zero.2 h), mul_one] variable (𝕜) diff --git a/Mathlib/Analysis/Normed/Affine/Isometry.lean b/Mathlib/Analysis/Normed/Affine/Isometry.lean index 7b4aca414185d..8fac64e2bb06c 100644 --- a/Mathlib/Analysis/Normed/Affine/Isometry.lean +++ b/Mathlib/Analysis/Normed/Affine/Isometry.lean @@ -503,7 +503,7 @@ instance instGroup : Group (P ≃ᵃⁱ[𝕜] P) where one_mul := trans_refl mul_one := refl_trans mul_assoc _ _ _ := trans_assoc _ _ _ - mul_left_inv := self_trans_symm + inv_mul_cancel := self_trans_symm @[simp] theorem coe_one : ⇑(1 : P ≃ᵃⁱ[𝕜] P) = id := diff --git a/Mathlib/Analysis/Normed/Algebra/Exponential.lean b/Mathlib/Analysis/Normed/Algebra/Exponential.lean index 273d045957370..dbed897cdb5f6 100644 --- a/Mathlib/Analysis/Normed/Algebra/Exponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/Exponential.lean @@ -283,14 +283,14 @@ noncomputable def invertibleExpOfMemBall [CharZero 𝕂] {x : 𝔸} have hnx : -x ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius := by rw [EMetric.mem_ball, ← neg_zero, edist_neg_neg] exact hx - rw [← exp_add_of_commute_of_mem_ball (Commute.neg_left <| Commute.refl x) hnx hx, neg_add_self, - exp_zero] + rw [← exp_add_of_commute_of_mem_ball (Commute.neg_left <| Commute.refl x) hnx hx, + neg_add_cancel, exp_zero] mul_invOf_self := by have hnx : -x ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius := by rw [EMetric.mem_ball, ← neg_zero, edist_neg_neg] exact hx - rw [← exp_add_of_commute_of_mem_ball (Commute.neg_right <| Commute.refl x) hx hnx, add_neg_self, - exp_zero] + rw [← exp_add_of_commute_of_mem_ball (Commute.neg_right <| Commute.refl x) hx hnx, + add_neg_cancel, exp_zero] theorem isUnit_exp_of_mem_ball [CharZero 𝕂] {x : 𝔸} (hx : x ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius) : IsUnit (exp 𝕂 x) := @@ -463,7 +463,7 @@ theorem exp_mem_unitary_of_mem_skewAdjoint [StarRing 𝔸] [ContinuousStar 𝔸] (h : x ∈ skewAdjoint 𝔸) : exp 𝕂 x ∈ unitary 𝔸 := by rw [unitary.mem_iff, star_exp, skewAdjoint.mem_iff.mp h, ← exp_add_of_commute (Commute.refl x).neg_left, ← exp_add_of_commute (Commute.refl x).neg_right, - add_left_neg, add_right_neg, exp_zero, and_self_iff] + neg_add_cancel, add_neg_cancel, exp_zero, and_self_iff] end diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index d2126e4827d15..a8160837b5391 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -826,7 +826,7 @@ instance (priority := 100) NormedDivisionRing.to_hasContinuousInv₀ : HasContin calc ‖e⁻¹ - r⁻¹‖ = ‖r‖⁻¹ * ‖r - e‖ * ‖e‖⁻¹ := by rw [← norm_inv, ← norm_inv, ← norm_mul, ← norm_mul, _root_.mul_sub, _root_.sub_mul, - mul_assoc _ e, inv_mul_cancel r0, mul_inv_cancel e0, one_mul, mul_one] + mul_assoc _ e, inv_mul_cancel₀ r0, mul_inv_cancel₀ e0, one_mul, mul_one] -- Porting note: `ENNReal.{mul_sub, sub_mul}` should be `protected` _ = ‖r - e‖ / ‖r‖ / ‖e‖ := by field_simp [mul_comm] _ ≤ ‖r - e‖ / ‖r‖ / ε := by gcongr diff --git a/Mathlib/Analysis/Normed/Group/AddCircle.lean b/Mathlib/Analysis/Normed/Group/AddCircle.lean index 1f41936177370..8cd44a44c9c69 100644 --- a/Mathlib/Analysis/Normed/Group/AddCircle.lean +++ b/Mathlib/Analysis/Normed/Group/AddCircle.lean @@ -84,7 +84,7 @@ theorem norm_eq {x : ℝ} : ‖(x : AddCircle p)‖ = |x - round (p⁻¹ * x) * · simp have hx := norm_coe_mul p x p⁻¹ rw [abs_inv, eq_inv_mul_iff_mul_eq₀ ((not_congr abs_eq_zero).mpr hp)] at hx - rw [← hx, inv_mul_cancel hp, this, ← abs_mul, mul_sub, mul_inv_cancel_left₀ hp, mul_comm p] + rw [← hx, inv_mul_cancel₀ hp, this, ← abs_mul, mul_sub, mul_inv_cancel_left₀ hp, mul_comm p] clear! x p intros x rw [quotient_norm_eq, abs_sub_round_eq_min] @@ -121,13 +121,13 @@ theorem norm_le_half_period {x : AddCircle p} (hp : p ≠ 0) : ‖x‖ ≤ |p| / obtain ⟨x⟩ := x change ‖(x : AddCircle p)‖ ≤ |p| / 2 rw [norm_eq, ← mul_le_mul_left (abs_pos.mpr (inv_ne_zero hp)), ← abs_mul, mul_sub, mul_left_comm, - ← mul_div_assoc, ← abs_mul, inv_mul_cancel hp, mul_one, abs_one] + ← mul_div_assoc, ← abs_mul, inv_mul_cancel₀ hp, mul_one, abs_one] exact abs_sub_round (p⁻¹ * x) @[simp] theorem norm_half_period_eq : ‖(↑(p / 2) : AddCircle p)‖ = |p| / 2 := by rcases eq_or_ne p 0 with (rfl | hp); · simp - rw [norm_eq, ← mul_div_assoc, inv_mul_cancel hp, one_div, round_two_inv, Int.cast_one, + rw [norm_eq, ← mul_div_assoc, inv_mul_cancel₀ hp, one_div, round_two_inv, Int.cast_one, one_mul, (by linarith : p / 2 - p = -(p / 2)), abs_neg, abs_div, abs_two] theorem norm_coe_eq_abs_iff {x : ℝ} (hp : p ≠ 0) : ‖(x : AddCircle p)‖ = |x| ↔ |x| ≤ |p| / 2 := by @@ -149,9 +149,9 @@ theorem norm_coe_eq_abs_iff {x : ℝ} (hp : p ≠ 0) : ‖(x : AddCircle p)‖ = obtain ⟨hx₁, hx₂⟩ := abs_le.mp hx replace hx₂ := Ne.lt_of_le hx' hx₂ constructor - · rwa [← mul_le_mul_left hp, ← mul_assoc, mul_inv_cancel hp.ne.symm, one_mul, mul_neg, ← + · rwa [← mul_le_mul_left hp, ← mul_assoc, mul_inv_cancel₀ hp.ne.symm, one_mul, mul_neg, ← mul_div_assoc, mul_one] - · rwa [← mul_lt_mul_left hp, ← mul_assoc, mul_inv_cancel hp.ne.symm, one_mul, ← mul_div_assoc, + · rwa [← mul_lt_mul_left hp, ← mul_assoc, mul_inv_cancel₀ hp.ne.symm, one_mul, ← mul_div_assoc, mul_one] open Metric diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 5f3c32b121bf3..8c619280197ef 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -199,7 +199,7 @@ abbrev SeminormedGroup.ofMulDist [Norm E] [Group E] [PseudoMetricSpace E] SeminormedGroup E where dist_eq x y := by rw [h₁]; apply le_antisymm - · simpa only [div_eq_mul_inv, ← mul_right_inv y] using h₂ _ _ _ + · simpa only [div_eq_mul_inv, ← mul_inv_cancel y] using h₂ _ _ _ · simpa only [div_mul_cancel, one_mul] using h₂ (x / y) 1 y -- See note [reducible non-instances] @@ -212,7 +212,7 @@ abbrev SeminormedGroup.ofMulDist' [Norm E] [Group E] [PseudoMetricSpace E] dist_eq x y := by rw [h₁]; apply le_antisymm · simpa only [div_mul_cancel, one_mul] using h₂ (x / y) 1 y - · simpa only [div_eq_mul_inv, ← mul_right_inv y] using h₂ _ _ _ + · simpa only [div_eq_mul_inv, ← mul_inv_cancel y] using h₂ _ _ _ -- See note [reducible non-instances] /-- Construct a seminormed group from a multiplication-invariant pseudodistance. -/ diff --git a/Mathlib/Analysis/Normed/Lp/PiLp.lean b/Mathlib/Analysis/Normed/Lp/PiLp.lean index e1c8c3ba748a1..a1f0fe08603c7 100644 --- a/Mathlib/Analysis/Normed/Lp/PiLp.lean +++ b/Mathlib/Analysis/Normed/Lp/PiLp.lean @@ -614,7 +614,8 @@ instance instBoundedSMul [SeminormedRing 𝕜] [∀ i, SeminormedAddCommGroup ( · have hp0 : 0 < p.toReal := zero_lt_one.trans_le hp have hpt : p ≠ ⊤ := p.toReal_pos_iff_ne_top.mp hp0 rw [nnnorm_eq_sum hpt, nnnorm_eq_sum hpt, one_div, NNReal.rpow_inv_le_iff hp0, - NNReal.mul_rpow, ← NNReal.rpow_mul, inv_mul_cancel hp0.ne', NNReal.rpow_one, Finset.mul_sum] + NNReal.mul_rpow, ← NNReal.rpow_mul, inv_mul_cancel₀ hp0.ne', NNReal.rpow_one, + Finset.mul_sum] simp_rw [← NNReal.mul_rpow, smul_apply] exact Finset.sum_le_sum fun i _ => NNReal.rpow_le_rpow (nnnorm_smul_le _ _) hp0.le @@ -799,8 +800,8 @@ theorem nnnorm_equiv_symm_single (i : ι) (b : β i) : have hp0 : (p : ℝ) ≠ 0 := mod_cast (zero_lt_one.trans_le <| Fact.out (p := 1 ≤ (p : ℝ≥0∞))).ne' rw [nnnorm_eq_sum ENNReal.coe_ne_top, ENNReal.coe_toReal, Fintype.sum_eq_single i, - WithLp.equiv_symm_pi_apply, Pi.single_eq_same, ← NNReal.rpow_mul, one_div, mul_inv_cancel hp0, - NNReal.rpow_one] + WithLp.equiv_symm_pi_apply, Pi.single_eq_same, ← NNReal.rpow_mul, one_div, + mul_inv_cancel₀ hp0, NNReal.rpow_one] intro j hij rw [WithLp.equiv_symm_pi_apply, Pi.single_eq_of_ne hij, nnnorm_zero, NNReal.zero_rpow hp0] diff --git a/Mathlib/Analysis/Normed/Lp/ProdLp.lean b/Mathlib/Analysis/Normed/Lp/ProdLp.lean index 76b7315100d34..6d23a952e2391 100644 --- a/Mathlib/Analysis/Normed/Lp/ProdLp.lean +++ b/Mathlib/Analysis/Normed/Lp/ProdLp.lean @@ -673,7 +673,7 @@ theorem nnnorm_equiv_symm_fst (x : α) : simp [prod_nnnorm_eq_sup] | coe p => have hp0 : (p : ℝ) ≠ 0 := mod_cast (zero_lt_one.trans_le <| Fact.out (p := 1 ≤ (p : ℝ≥0∞))).ne' - simp [prod_nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] + simp [prod_nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel₀ hp0] @[simp] theorem nnnorm_equiv_symm_snd (y : β) : @@ -683,7 +683,7 @@ theorem nnnorm_equiv_symm_snd (y : β) : simp [prod_nnnorm_eq_sup] | coe p => have hp0 : (p : ℝ) ≠ 0 := mod_cast (zero_lt_one.trans_le <| Fact.out (p := 1 ≤ (p : ℝ≥0∞))).ne' - simp [prod_nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel hp0] + simp [prod_nnnorm_eq_add, NNReal.zero_rpow hp0, ← NNReal.rpow_mul, mul_inv_cancel₀ hp0] @[simp] theorem norm_equiv_symm_fst (x : α) : ‖(WithLp.equiv p (α × β)).symm (x, 0)‖ = ‖x‖ := @@ -744,7 +744,7 @@ instance instProdBoundedSMul : BoundedSMul 𝕜 (WithLp p (α × β)) := · have hp0 : 0 < p.toReal := zero_lt_one.trans_le hp have hpt : p ≠ ⊤ := p.toReal_pos_iff_ne_top.mp hp0 rw [prod_nnnorm_eq_add hpt, prod_nnnorm_eq_add hpt, one_div, NNReal.rpow_inv_le_iff hp0, - NNReal.mul_rpow, ← NNReal.rpow_mul, inv_mul_cancel hp0.ne', NNReal.rpow_one, mul_add, + NNReal.mul_rpow, ← NNReal.rpow_mul, inv_mul_cancel₀ hp0.ne', NNReal.rpow_one, mul_add, ← NNReal.mul_rpow, ← NNReal.mul_rpow] exact add_le_add (NNReal.rpow_le_rpow (nnnorm_smul_le _ _) hp0.le) diff --git a/Mathlib/Analysis/Normed/Lp/lpSpace.lean b/Mathlib/Analysis/Normed/Lp/lpSpace.lean index 636faeded0845..23bb44f6901d5 100644 --- a/Mathlib/Analysis/Normed/Lp/lpSpace.lean +++ b/Mathlib/Analysis/Normed/Lp/lpSpace.lean @@ -177,7 +177,7 @@ theorem of_exponent_ge {p q : ℝ≥0∞} {f : ∀ i, E i} (hfq : Memℓp f q) ( use A ^ q.toReal⁻¹ rintro x ⟨i, rfl⟩ have : 0 ≤ ‖f i‖ ^ q.toReal := by positivity - simpa [← Real.rpow_mul, mul_inv_cancel hq.ne'] using + simpa [← Real.rpow_mul, mul_inv_cancel₀ hq.ne'] using Real.rpow_le_rpow this (hA ⟨i, rfl⟩) (inv_nonneg.mpr hq.le) · apply memℓp_gen have hf' := hfq.summable hq diff --git a/Mathlib/Analysis/Normed/Module/Dual.lean b/Mathlib/Analysis/Normed/Module/Dual.lean index 1ba59eddf1452..de23a0957c44e 100644 --- a/Mathlib/Analysis/Normed/Module/Dual.lean +++ b/Mathlib/Analysis/Normed/Module/Dual.lean @@ -188,7 +188,7 @@ theorem smul_mem_polar {s : Set E} {x' : Dual 𝕜 E} {c : 𝕜} (hc : ∀ z, z rw [eq z] apply mul_le_mul (le_of_eq rfl) (hc z hzs) (norm_nonneg _) (norm_nonneg _) have cancel : ‖c⁻¹‖ * ‖c‖ = 1 := by - simp only [c_zero, norm_eq_zero, Ne, not_false_iff, inv_mul_cancel, norm_inv] + simp only [c_zero, norm_eq_zero, Ne, not_false_iff, inv_mul_cancel₀, norm_inv] rwa [cancel] at le theorem polar_ball_subset_closedBall_div {c : 𝕜} (hc : 1 < ‖c‖) {r : ℝ} (hr : 0 < r) : diff --git a/Mathlib/Analysis/Normed/MulAction.lean b/Mathlib/Analysis/Normed/MulAction.lean index a7acc637d6fa0..618fab80dff3f 100644 --- a/Mathlib/Analysis/Normed/MulAction.lean +++ b/Mathlib/Analysis/Normed/MulAction.lean @@ -85,7 +85,7 @@ theorem norm_smul (r : α) (x : β) : ‖r • x‖ = ‖r‖ * ‖x‖ := by calc ‖r‖ * ‖x‖ = ‖r‖ * ‖r⁻¹ • r • x‖ := by rw [inv_smul_smul₀ h] _ ≤ ‖r‖ * (‖r⁻¹‖ * ‖r • x‖) := by gcongr; apply norm_smul_le - _ = ‖r • x‖ := by rw [norm_inv, ← mul_assoc, mul_inv_cancel (mt norm_eq_zero.1 h), one_mul] + _ = ‖r • x‖ := by rw [norm_inv, ← mul_assoc, mul_inv_cancel₀ (mt norm_eq_zero.1 h), one_mul] theorem nnnorm_smul (r : α) (x : β) : ‖r • x‖₊ = ‖r‖₊ * ‖x‖₊ := NNReal.eq <| norm_smul r x diff --git a/Mathlib/Analysis/Normed/Operator/Banach.lean b/Mathlib/Analysis/Normed/Operator/Banach.lean index cacdacc7a65c6..bfc054cac483e 100644 --- a/Mathlib/Analysis/Normed/Operator/Banach.lean +++ b/Mathlib/Analysis/Normed/Operator/Banach.lean @@ -127,7 +127,7 @@ theorem exists_approx_preimage_norm_le (surj : Surjective f) : have J : ‖f (σ' d⁻¹ • x) - y‖ ≤ 1 / 2 * ‖y‖ := calc ‖f (σ' d⁻¹ • x) - y‖ = ‖d⁻¹ • f x - (d⁻¹ * d) • y‖ := by - rwa [f.map_smulₛₗ _, inv_mul_cancel, one_smul, map_inv₀, map_inv₀, + rwa [f.map_smulₛₗ _, inv_mul_cancel₀, one_smul, map_inv₀, map_inv₀, RingHomCompTriple.comp_apply, RingHom.id_apply] _ = ‖d⁻¹ • (f x - d • y)‖ := by rw [mul_smul, smul_sub] _ = ‖d‖⁻¹ * ‖f x - d • y‖ := by rw [norm_smul, norm_inv] @@ -136,7 +136,7 @@ theorem exists_approx_preimage_norm_le (surj : Surjective f) : simp only [δ] ring _ = ‖y‖ / 2 := by - rw [inv_mul_cancel, one_mul] + rw [inv_mul_cancel₀, one_mul] simp [norm_eq_zero, hd] _ = 1 / 2 * ‖y‖ := by ring rw [← dist_eq_norm] at J diff --git a/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean b/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean index 04b14c0e398ad..da7b7a93c50c0 100644 --- a/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean +++ b/Mathlib/Analysis/Normed/Operator/ContinuousLinearMap.lean @@ -182,7 +182,7 @@ theorem ContinuousLinearEquiv.homothety_inverse (a : ℝ) (ha : 0 < a) (f : E intro hf y calc ‖f.symm y‖ = a⁻¹ * (a * ‖f.symm y‖) := by - rw [← mul_assoc, inv_mul_cancel (ne_of_lt ha).symm, one_mul] + rw [← mul_assoc, inv_mul_cancel₀ (ne_of_lt ha).symm, one_mul] _ = a⁻¹ * ‖f (f.symm y)‖ := by rw [hf] _ = a⁻¹ * ‖y‖ := by simp diff --git a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean index e6b3a063b3a61..dcba3cf3c238e 100644 --- a/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean +++ b/Mathlib/Analysis/Normed/Operator/LinearIsometry.lean @@ -733,7 +733,7 @@ instance instGroup : Group (E ≃ₗᵢ[R] E) where one_mul := trans_refl mul_one := refl_trans mul_assoc _ _ _ := trans_assoc _ _ _ - mul_left_inv := self_trans_symm + inv_mul_cancel := self_trans_symm @[simp] theorem coe_one : ⇑(1 : E ≃ₗᵢ[R] E) = id := diff --git a/Mathlib/Analysis/Normed/Ring/Seminorm.lean b/Mathlib/Analysis/Normed/Ring/Seminorm.lean index e75189e80f622..9d88d4bb64e68 100644 --- a/Mathlib/Analysis/Normed/Ring/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Ring/Seminorm.lean @@ -335,7 +335,7 @@ def RingSeminorm.toRingNorm {K : Type*} [Field K] (f : RingSeminorm K) (hnt : f obtain ⟨c, hc⟩ := RingSeminorm.ne_zero_iff.mp hnt by_contra hn0 have hc0 : f c = 0 := by - rw [← mul_one c, ← mul_inv_cancel hn0, ← mul_assoc, mul_comm c, mul_assoc] + rw [← mul_one c, ← mul_inv_cancel₀ hn0, ← mul_assoc, mul_comm c, mul_assoc] exact le_antisymm (le_trans (map_mul_le_mul f _ _) diff --git a/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean b/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean index e0a99d58dedb7..bf8d8abdf6efe 100644 --- a/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean +++ b/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean @@ -96,7 +96,7 @@ theorem seminormFromConst_seq_antitone (x : R) : Antitone (seminormFromConst_seq rw [Nat.one_le_iff_ne_zero, ne_eq, Nat.sub_eq_zero_iff_le, not_le] exact lt_of_le_of_ne hmn heq rw [hpm c h1, mul_div_assoc, div_eq_mul_inv, pow_sub₀ _ hc hmn, mul_assoc, mul_comm (f c ^ m)⁻¹, - ← mul_assoc (f c ^ n), mul_inv_cancel (pow_ne_zero n hc), one_mul, div_eq_mul_inv] + ← mul_assoc (f c ^ n), mul_inv_cancel₀ (pow_ne_zero n hc), one_mul, div_eq_mul_inv] /-- The real-valued function sending `x ∈ R` to the limit of `(f (x * c^n))/((f c)^n)`. -/ def seminormFromConst' (x : R) : ℝ := @@ -252,7 +252,7 @@ theorem seminormFromConst_const_mul (x : R) : simp only [seminormFromConst_seq_def] ext n ring_nf - rw [mul_assoc _ (f c), mul_inv_cancel hc, mul_one] + rw [mul_assoc _ (f c), mul_inv_cancel₀ hc, mul_one] simpa [hterm] using tendsto_const_nhds.mul hlim end Ring diff --git a/Mathlib/Analysis/Normed/Ring/Units.lean b/Mathlib/Analysis/Normed/Ring/Units.lean index b55db66c82a51..c46b35e61a4d4 100644 --- a/Mathlib/Analysis/Normed/Ring/Units.lean +++ b/Mathlib/Analysis/Normed/Ring/Units.lean @@ -56,7 +56,7 @@ def add (x : Rˣ) (t : R) (h : ‖t‖ < ‖(↑x⁻¹ : R)‖⁻¹) : Rˣ := ‖-(↑x⁻¹ * t)‖ = ‖↑x⁻¹ * t‖ := by rw [norm_neg] _ ≤ ‖(↑x⁻¹ : R)‖ * ‖t‖ := norm_mul_le (x⁻¹).1 _ _ < ‖(↑x⁻¹ : R)‖ * ‖(↑x⁻¹ : R)‖⁻¹ := by nlinarith only [h, hpos] - _ = 1 := mul_inv_cancel (ne_of_gt hpos))) + _ = 1 := mul_inv_cancel₀ (ne_of_gt hpos))) (x + t) (by simp [mul_add]) _ rfl /-- In a complete normed ring, an element `y` of distance less than `‖x⁻¹‖⁻¹` from `x` is a unit. diff --git a/Mathlib/Analysis/NormedSpace/Connected.lean b/Mathlib/Analysis/NormedSpace/Connected.lean index ded791be4e517..64d98223dda5e 100644 --- a/Mathlib/Analysis/NormedSpace/Connected.lean +++ b/Mathlib/Analysis/NormedSpace/Connected.lean @@ -145,14 +145,14 @@ theorem isPathConnected_sphere (h : 1 < Module.rank ℝ E) (x : E) {r : ℝ} (hr apply Subset.antisymm · rintro - ⟨y, hy, rfl⟩ have : ‖y‖ ≠ 0 := by simpa using hy - simp [f, norm_smul, abs_of_nonneg hr, mul_assoc, inv_mul_cancel this] + simp [f, norm_smul, abs_of_nonneg hr, mul_assoc, inv_mul_cancel₀ this] · intro y hy refine ⟨y - x, ?_, ?_⟩ · intro H simp only [mem_singleton_iff, sub_eq_zero] at H simp only [H, mem_sphere_iff_norm, sub_self, norm_zero] at hy exact rpos.ne hy - · simp [f, mem_sphere_iff_norm.1 hy, mul_inv_cancel rpos.ne'] + · simp [f, mem_sphere_iff_norm.1 hy, mul_inv_cancel₀ rpos.ne'] rwa [this] at C /-- In a real vector space of dimension `> 1`, any sphere of nonnegative radius is connected. -/ diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean index a92b67bff8d6a..b759259be866a 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/Extension.lean @@ -156,7 +156,7 @@ theorem coord_norm' {x : E} (h : x ≠ 0) : ‖(‖x‖ : 𝕜) • coord 𝕜 x -/ set_option maxSynthPendingDepth 2 in rw [norm_smul (α := 𝕜) (x := coord 𝕜 x h), RCLike.norm_coe_norm, coord_norm, - mul_inv_cancel (mt norm_eq_zero.mp h)] + mul_inv_cancel₀ (mt norm_eq_zero.mp h)] /-- Corollary of Hahn-Banach. Given a nonzero element `x` of a normed space, there exists an element of the dual space, of norm `1`, whose value on `x` is `‖x‖`. -/ diff --git a/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean b/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean index 740851e2901ba..af94c78f743e0 100644 --- a/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean +++ b/Mathlib/Analysis/NormedSpace/HahnBanach/SeparatingDual.lean @@ -101,7 +101,7 @@ theorem dualMap_surjective_iff {W} [AddCommGroup W] [Module R W] [FiniteDimensio lemma exists_eq_one {x : V} (hx : x ≠ 0) : ∃ f : V →L[R] R, f x = 1 := by rcases exists_ne_zero (R := R) hx with ⟨f, hf⟩ - exact ⟨(f x)⁻¹ • f, inv_mul_cancel hf⟩ + exact ⟨(f x)⁻¹ • f, inv_mul_cancel₀ hf⟩ theorem exists_eq_one_ne_zero_of_ne_zero_pair {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : ∃ f : V →L[R] R, f x = 1 ∧ f y ≠ 0 := by @@ -110,7 +110,7 @@ theorem exists_eq_one_ne_zero_of_ne_zero_pair {x y : V} (hx : x ≠ 0) (hy : y · exact ⟨u, ux, uy⟩ obtain ⟨v, vy⟩ : ∃ v : V →L[R] R, v y = 1 := exists_eq_one hy rcases ne_or_eq (v x) 0 with vx|vx - · exact ⟨(v x)⁻¹ • v, inv_mul_cancel vx, show (v x)⁻¹ * v y ≠ 0 by simp [vx, vy]⟩ + · exact ⟨(v x)⁻¹ • v, inv_mul_cancel₀ vx, show (v x)⁻¹ * v y ≠ 0 by simp [vx, vy]⟩ · exact ⟨u + v, by simp [ux, vx], by simp [uy, vy]⟩ variable [TopologicalAddGroup V] @@ -132,14 +132,14 @@ theorem exists_continuousLinearEquiv_apply_eq [ContinuousSMul R V] simp only [id_eq, eq_mpr_eq_cast, RingHom.id_apply, smul_eq_mul, AddHom.toFun_eq_coe, -- Note: #8386 had to change `map_smulₛₗ` into `map_smulₛₗ _` AddHom.coe_mk, map_add, map_smulₛₗ _, map_sub, Gx, mul_sub, mul_one, add_sub_cancel] - rw [mul_comm (G z), ← mul_assoc, inv_mul_cancel Gy] + rw [mul_comm (G z), ← mul_assoc, inv_mul_cancel₀ Gy] simp only [smul_sub, one_mul] abel right_inv := fun z ↦ by -- Note: #8386 had to change `map_smulₛₗ` into `map_smulₛₗ _` simp only [map_add, map_smulₛₗ _, map_mul, map_inv₀, RingHom.id_apply, map_sub, Gx, smul_eq_mul, mul_sub, mul_one] - rw [mul_comm _ (G y), ← mul_assoc, mul_inv_cancel Gy] + rw [mul_comm _ (G y), ← mul_assoc, mul_inv_cancel₀ Gy] simp only [smul_sub, one_mul, add_sub_cancel] abel continuous_toFun := continuous_id.add (G.continuous.smul continuous_const) diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean index c69364cab074f..639f15c04f7ce 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean @@ -275,7 +275,7 @@ one controls the norm of `f`. -/ theorem opNorm_le_of_unit_norm [NormedSpace ℝ E] [NormedSpace ℝ F] {f : E →L[ℝ] F} {C : ℝ} (hC : 0 ≤ C) (hf : ∀ x, ‖x‖ = 1 → ‖f x‖ ≤ C) : ‖f‖ ≤ C := by refine opNorm_le_bound' f hC fun x hx => ?_ - have H₁ : ‖‖x‖⁻¹ • x‖ = 1 := by rw [norm_smul, norm_inv, norm_norm, inv_mul_cancel hx] + have H₁ : ‖‖x‖⁻¹ • x‖ = 1 := by rw [norm_smul, norm_inv, norm_norm, inv_mul_cancel₀ hx] have H₂ := hf _ H₁ rwa [map_smul, norm_smul, norm_inv, norm_norm, ← div_eq_inv_mul, _root_.div_le_iff] at H₂ exact (norm_nonneg x).lt_of_ne' hx diff --git a/Mathlib/Analysis/NormedSpace/Pointwise.lean b/Mathlib/Analysis/NormedSpace/Pointwise.lean index a524841a74047..bb061dfe9e186 100644 --- a/Mathlib/Analysis/NormedSpace/Pointwise.lean +++ b/Mathlib/Analysis/NormedSpace/Pointwise.lean @@ -384,7 +384,7 @@ theorem NormedSpace.sphere_nonempty [Nontrivial E] {x : E} {r : ℝ} : simp only [mem_sphere_iff_norm, add_sub_cancel_right, norm_smul, Real.norm_eq_abs, norm_inv, norm_norm, ne_eq, norm_eq_zero] simp only [abs_norm, ne_eq, norm_eq_zero] - rw [inv_mul_cancel this, mul_one, abs_eq_self.mpr hr] + rw [inv_mul_cancel₀ this, mul_one, abs_eq_self.mpr hr] theorem smul_sphere [Nontrivial E] (c : 𝕜) (x : E) {r : ℝ} (hr : 0 ≤ r) : c • sphere x r = sphere (c • x) (‖c‖ * r) := by diff --git a/Mathlib/Analysis/NormedSpace/RCLike.lean b/Mathlib/Analysis/NormedSpace/RCLike.lean index 4ecdab1b44a2a..2237337083a4d 100644 --- a/Mathlib/Analysis/NormedSpace/RCLike.lean +++ b/Mathlib/Analysis/NormedSpace/RCLike.lean @@ -61,7 +61,7 @@ theorem LinearMap.bound_of_sphere_bound {r : ℝ} (r_pos : 0 < r) (c : ℝ) (f : have r_ne_zero : (r : 𝕜) ≠ 0 := RCLike.ofReal_ne_zero.mpr r_pos.ne' have eq : f z = ‖z‖ / r * f z₁ := by rw [hz₁, LinearMap.map_smul, smul_eq_mul] - rw [← mul_assoc, ← mul_assoc, div_mul_cancel₀ _ r_ne_zero, mul_inv_cancel, one_mul] + rw [← mul_assoc, ← mul_assoc, div_mul_cancel₀ _ r_ne_zero, mul_inv_cancel₀, one_mul] simp only [z_zero, RCLike.ofReal_eq_zero, norm_eq_zero, Ne, not_false_iff] rw [eq, norm_mul, norm_div, RCLike.norm_coe_norm, RCLike.norm_of_nonneg r_pos.le, div_mul_eq_mul_div, div_mul_eq_mul_div, mul_comm] diff --git a/Mathlib/Analysis/NormedSpace/Real.lean b/Mathlib/Analysis/NormedSpace/Real.lean index a71f4ae0f993f..3106fe40f746a 100644 --- a/Mathlib/Analysis/NormedSpace/Real.lean +++ b/Mathlib/Analysis/NormedSpace/Real.lean @@ -114,7 +114,7 @@ theorem exists_norm_eq {c : ℝ} (hc : 0 ≤ c) : ∃ x : E, ‖x‖ = c := by rcases exists_ne (0 : E) with ⟨x, hx⟩ rw [← norm_ne_zero_iff] at hx use c • ‖x‖⁻¹ • x - simp [norm_smul, Real.norm_of_nonneg hc, abs_of_nonneg hc, inv_mul_cancel hx] + simp [norm_smul, Real.norm_of_nonneg hc, abs_of_nonneg hc, inv_mul_cancel₀ hx] @[simp] theorem range_norm : range (norm : E → ℝ) = Ici 0 := diff --git a/Mathlib/Analysis/NormedSpace/RieszLemma.lean b/Mathlib/Analysis/NormedSpace/RieszLemma.lean index 6900ae79b49a9..415a55f95a3f4 100644 --- a/Mathlib/Analysis/NormedSpace/RieszLemma.lean +++ b/Mathlib/Analysis/NormedSpace/RieszLemma.lean @@ -91,7 +91,7 @@ theorem riesz_lemma_of_norm_lt {c : 𝕜} (hc : 1 < ‖c‖) {R : ℝ} (hR : ‖ rescale_to_shell hc Rpos x0 refine ⟨d • x, dxlt.le, fun y hy => ?_⟩ set y' := d⁻¹ • y - have yy' : y = d • y' := by simp [y', smul_smul, mul_inv_cancel d0] + have yy' : y = d • y' := by simp [y', smul_smul, mul_inv_cancel₀ d0] calc 1 = ‖c‖ / R * (R / ‖c‖) := by field_simp [Rpos.ne', (zero_lt_one.trans hc).ne'] _ ≤ ‖c‖ / R * ‖d • x‖ := by gcongr diff --git a/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean b/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean index d40f285042ae8..fd8a5748ac567 100644 --- a/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean +++ b/Mathlib/Analysis/NormedSpace/SphereNormEquiv.lean @@ -29,7 +29,7 @@ noncomputable def homeomorphUnitSphereProd : ({0}ᶜ : Set E) ≃ₜ (sphere (0 : E) 1 × Ioi (0 : ℝ)) where toFun x := (⟨‖x.1‖⁻¹ • x.1, by rw [mem_sphere_zero_iff_norm, norm_smul, norm_inv, norm_norm, - inv_mul_cancel (norm_ne_zero_iff.2 x.2)]⟩, ⟨‖x.1‖, norm_pos_iff.2 x.2⟩) + inv_mul_cancel₀ (norm_ne_zero_iff.2 x.2)]⟩, ⟨‖x.1‖, norm_pos_iff.2 x.2⟩) invFun x := ⟨x.2.1 • x.1.1, smul_ne_zero x.2.2.out.ne' (ne_of_mem_sphere x.1.2 one_ne_zero)⟩ left_inv x := Subtype.eq <| by simp [smul_inv_smul₀ (norm_ne_zero_iff.2 x.2)] right_inv diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 5a44f14f45a95..5fc1882b59565 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -436,7 +436,7 @@ theorem inv_def (z : K) : z⁻¹ = conj z * ((‖z‖ ^ 2)⁻¹ : ℝ) := by rcases eq_or_ne z 0 with (rfl | h₀) · simp · apply inv_eq_of_mul_eq_one_right - rw [← mul_assoc, mul_conj, ofReal_inv, ofReal_pow, mul_inv_cancel] + rw [← mul_assoc, mul_conj, ofReal_inv, ofReal_pow, mul_inv_cancel₀] simpa @[simp, rclike_simps] @@ -811,7 +811,7 @@ lemma _root_.StarModule.instOrderedSMul {A : Type*} [NonUnitalRing A] [StarRing lt_of_smul_lt_smul_of_pos {x y c} hxy hc := by have : c⁻¹ • c • x < c⁻¹ • c • y := StarModule.smul_lt_smul_of_pos hxy (RCLike.inv_pos_of_pos hc) - simpa [smul_smul, inv_mul_cancel hc.ne'] using this + simpa [smul_smul, inv_mul_cancel₀ hc.ne'] using this instance {A : Type*} [NonUnitalRing A] [StarRing A] [PartialOrder A] [StarOrderedRing A] [Module ℝ A] [StarModule ℝ A] [IsScalarTower ℝ A A] [SMulCommClass ℝ A A] : diff --git a/Mathlib/Analysis/Seminorm.lean b/Mathlib/Analysis/Seminorm.lean index 44d821de11bfc..ed51cb1a81624 100644 --- a/Mathlib/Analysis/Seminorm.lean +++ b/Mathlib/Analysis/Seminorm.lean @@ -1187,7 +1187,7 @@ lemma rescale_to_shell_zpow (p : Seminorm 𝕜 E) {c : 𝕜} (hc : 1 < ‖c‖) exact (div_lt_iff εpos).1 (hn.2) · show ε / ‖c‖ ≤ p (c ^ (-(n + 1)) • x) rw [zpow_neg, div_le_iff cpos, map_smul_eq_mul, norm_inv, norm_zpow, zpow_add₀ (ne_of_gt cpos), - zpow_one, mul_inv_rev, mul_comm, ← mul_assoc, ← mul_assoc, mul_inv_cancel (ne_of_gt cpos), + zpow_one, mul_inv_rev, mul_comm, ← mul_assoc, ← mul_assoc, mul_inv_cancel₀ (ne_of_gt cpos), one_mul, ← div_eq_inv_mul, le_div_iff (zpow_pos_of_pos cpos _), mul_comm] exact (le_div_iff εpos).1 hn.1 · show ‖(c ^ (-(n + 1)))‖⁻¹ ≤ ε⁻¹ * ‖c‖ * p x diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Arctan.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Arctan.lean index 9843be4b015da..352fd5d5ee08d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Arctan.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Arctan.lean @@ -28,7 +28,7 @@ theorem tan_arctan {z : ℂ} (h₁ : z ≠ I) (h₂ : z ≠ -I) : tan (arctan z) rw [div_div_eq_mul_div, div_mul_cancel₀ _ two_ne_zero, ← div_mul_eq_mul_div, -- multiply top and bottom by `exp (arctan z * I)` ← mul_div_mul_right _ _ (exp_ne_zero (arctan z * I)), sub_mul, add_mul, - ← exp_add, neg_mul, add_left_neg, exp_zero, ← exp_add, ← two_mul] + ← exp_add, neg_mul, neg_add_cancel, exp_zero, ← exp_add, ← two_mul] have z₁ : 1 + z * I ≠ 0 := by contrapose! h₁ rw [add_eq_zero_iff_neg_eq, ← div_eq_iff I_ne_zero, div_I, neg_one_mul, neg_neg] at h₁ @@ -124,9 +124,9 @@ theorem hasSum_arctan {z : ℂ} (hz : ‖z‖ < 1) : dsimp only convert hasSum_fintype (_ : Fin 2 → ℂ) using 1 rw [Fin.sum_univ_two, Fin.val_zero, Fin.val_one, Odd.neg_one_pow (n := 2 * k + 0 + 1) (by simp), - add_left_neg, zero_mul, zero_div, mul_zero, zero_add, show 2 * k + 1 + 1 = 2 * (k + 1) by ring, - Even.neg_one_pow (n := 2 * (k + 1)) (by simp), ← mul_div_assoc (_ / _), ← mul_assoc, - show -I / 2 * (1 + 1) = -I by ring] + neg_add_cancel, zero_mul, zero_div, mul_zero, zero_add, + show 2 * k + 1 + 1 = 2 * (k + 1) by ring, Even.neg_one_pow (n := 2 * (k + 1)) (by simp), + ← mul_div_assoc (_ / _), ← mul_assoc, show -I / 2 * (1 + 1) = -I by ring] congr 1 rw [mul_pow, pow_succ' I, pow_mul, I_sq, show -I * _ = -(I * I) * (-1) ^ k * z ^ (2 * k + 1) by ring, I_mul_I, neg_neg, one_mul] diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean b/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean index 111f476afb799..695b5f7943f0f 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean @@ -100,7 +100,7 @@ lemma hasDerivAt_log_sub_logTaylor (n : ℕ) {z : ℂ} (hz : 1 + z ∈ slitPlane have hz' : -z ≠ 1 := by intro H rw [neg_eq_iff_eq_neg] at H - simp only [H, add_right_neg] at hz + simp only [H, add_neg_cancel] at hz exact slitPlane_ne_zero hz rfl simp_rw [← mul_pow, neg_one_mul, geom_sum_eq hz', ← neg_add', div_neg, add_comm z] field_simp [slitPlane_ne_zero hz] diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index 4ba4ca7492589..10561370fb1a4 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -378,8 +378,8 @@ lemma integral_cpow_mul_exp_neg_mul_Ioi {a : ℂ} {r : ℝ} (ha : 0 < a.re) (hr refine MeasureTheory.setIntegral_congr measurableSet_Ioi (fun x hx ↦ ?_) rw [mem_Ioi] at hx rw [mul_cpow_ofReal_nonneg hr.le hx.le, ← mul_assoc, one_div, ← ofReal_inv, - ← mul_cpow_ofReal_nonneg (inv_pos.mpr hr).le hr.le, ← ofReal_mul r⁻¹, inv_mul_cancel hr.ne', - ofReal_one, one_cpow, one_mul] + ← mul_cpow_ofReal_nonneg (inv_pos.mpr hr).le hr.le, ← ofReal_mul r⁻¹, + inv_mul_cancel₀ hr.ne', ofReal_one, one_cpow, one_mul] _ = 1 / r * ∫ (t : ℝ) in Ioi 0, (1 / r) ^ (a - 1) * t ^ (a - 1) * exp (-t) := by simp_rw [← ofReal_mul] rw [integral_comp_mul_left_Ioi (fun x ↦ _ * x ^ (a - 1) * exp (-x)) _ hr, mul_zero, diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean index 33440aea19dfd..486e57156cc16 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean @@ -91,7 +91,7 @@ theorem betaIntegral_symm (u v : ℂ) : betaIntegral v u = betaIntegral u v := b rw [inv_neg, inv_one, neg_one_smul, ← intervalIntegral.integral_symm] at this simp? at this says simp only [neg_mul, one_mul, ofReal_add, ofReal_neg, ofReal_one, sub_add_cancel_right, neg_neg, - mul_one, add_left_neg, mul_zero, zero_add] at this + mul_one, neg_add_cancel, mul_zero, zero_add] at this conv_lhs at this => arg 1; intro x; rw [add_comm, ← sub_eq_add_neg, mul_comm] exact this diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean index 1d8f5a101e74d..e2a7aed0ba9c4 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean @@ -90,7 +90,7 @@ theorem integrableOn_rpow_mul_exp_neg_mul_rpow {p s b : ℝ} (hs : -1 < s) (hp : suffices IntegrableOn (fun x ↦ (b ^ (-p⁻¹)) ^ s * (x ^ s * exp (-x ^ p))) (Ioi 0) by rw [show 0 = b ^ (-p⁻¹) * 0 by rw [mul_zero], ← integrableOn_Ioi_comp_mul_left_iff _ _ hib] refine this.congr_fun (fun _ hx => ?_) measurableSet_Ioi - rw [← mul_assoc, mul_rpow, mul_rpow, ← rpow_mul (z := p), neg_mul, neg_mul, inv_mul_cancel, + rw [← mul_assoc, mul_rpow, mul_rpow, ← rpow_mul (z := p), neg_mul, neg_mul, inv_mul_cancel₀, rpow_neg_one, mul_inv_cancel_left₀] all_goals linarith [mem_Ioi.mp hx] refine Integrable.const_mul ?_ _ diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index 982e2cb9557cc..ca08532bc9736 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -462,7 +462,7 @@ theorem integral_exp_mul_complex {c : ℂ} (hc : c ≠ 0) : theorem integral_log (h : (0 : ℝ) ∉ [[a, b]]) : ∫ x in a..b, log x = b * log b - a * log a - b + a := by have h' := fun x (hx : x ∈ [[a, b]]) => ne_of_mem_of_not_mem hx h - have heq := fun x hx => mul_inv_cancel (h' x hx) + have heq := fun x hx => mul_inv_cancel₀ (h' x hx) convert integral_mul_deriv_eq_deriv_mul (fun x hx => hasDerivAt_log (h' x hx)) (fun x _ => hasDerivAt_id x) (continuousOn_inv₀.mono <| subset_compl_singleton_iff.mpr h).intervalIntegrable diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean index 52f3f9747cae0..7817f8158e5eb 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean @@ -287,7 +287,7 @@ theorem hasSum_log_sub_log_of_abs_lt_one {x : ℝ} (h : |x| < 1) : · intro m hm rw [range_two_mul, Set.mem_setOf_eq, ← Nat.even_add_one] at hm dsimp [term] - rw [Even.neg_pow hm, neg_one_mul, neg_add_self] + rw [Even.neg_pow hm, neg_one_mul, neg_add_cancel] @[deprecated (since := "2024-01-31")] alias hasSum_log_sub_log_of_abs_lt_1 := hasSum_log_sub_log_of_abs_lt_one diff --git a/Mathlib/Analysis/SpecialFunctions/Log/NegMulLog.lean b/Mathlib/Analysis/SpecialFunctions/Log/NegMulLog.lean index 87070fe6bbc9d..66773834ae8aa 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/NegMulLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/NegMulLog.lean @@ -45,7 +45,7 @@ lemma differentiableOn_mul_log : DifferentiableOn ℝ (fun x ↦ x * log x) {0} lemma deriv_mul_log {x : ℝ} (hx : x ≠ 0) : deriv (fun x ↦ x * log x) x = log x + 1 := by rw [deriv_mul differentiableAt_id' (differentiableAt_log hx)] simp only [deriv_id'', one_mul, deriv_log', ne_eq, add_right_inj] - exact mul_inv_cancel hx + exact mul_inv_cancel₀ hx lemma hasDerivAt_mul_log {x : ℝ} (hx : x ≠ 0) : HasDerivAt (fun x ↦ x * log x) (log x + 1) x := by rw [← deriv_mul_log hx, hasDerivAt_deriv_iff] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean index 2afd6ee97c767..0560cfa53a183 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean @@ -141,7 +141,7 @@ alias cpow_int_cast := cpow_intCast @[simp] theorem cpow_nat_inv_pow (x : ℂ) {n : ℕ} (hn : n ≠ 0) : (x ^ (n⁻¹ : ℂ)) ^ n = x := by - rw [← cpow_nat_mul, mul_inv_cancel, cpow_one] + rw [← cpow_nat_mul, mul_inv_cancel₀, cpow_one] assumption_mod_cast /-- See Note [no_index around OfNat.ofNat] -/ @@ -174,7 +174,7 @@ lemma cpow_ofNat_mul' {x : ℂ} {n : ℕ} [n.AtLeastTwo] (hlt : -π < OfNat.ofNa lemma pow_cpow_nat_inv {x : ℂ} {n : ℕ} (h₀ : n ≠ 0) (hlt : -(π / n) < x.arg) (hle : x.arg ≤ π / n) : (x ^ n) ^ (n⁻¹ : ℂ) = x := by - rw [← cpow_nat_mul', mul_inv_cancel (Nat.cast_ne_zero.2 h₀), cpow_one] + rw [← cpow_nat_mul', mul_inv_cancel₀ (Nat.cast_ne_zero.2 h₀), cpow_one] · rwa [← div_lt_iff' (Nat.cast_pos.2 h₀.bot_lt), neg_div] · rwa [← le_div_iff' (Nat.cast_pos.2 h₀.bot_lt)] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index 331b5e0ea8d63..957d08e87a3bb 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -286,7 +286,7 @@ theorem rpow_eq_rpow_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x ^ z = y ^ (rpow_left_injective hz).eq_iff theorem rpow_left_surjective {x : ℝ} (hx : x ≠ 0) : Function.Surjective fun y : ℝ≥0 => y ^ x := - fun y => ⟨y ^ x⁻¹, by simp_rw [← rpow_mul, _root_.inv_mul_cancel hx, rpow_one]⟩ + fun y => ⟨y ^ x⁻¹, by simp_rw [← rpow_mul, inv_mul_cancel₀ hx, rpow_one]⟩ theorem rpow_left_bijective {x : ℝ} (hx : x ≠ 0) : Function.Bijective fun y : ℝ≥0 => y ^ x := ⟨rpow_left_injective hx, rpow_left_surjective hx⟩ @@ -306,10 +306,10 @@ theorem rpow_one_div_eq_iff {x y : ℝ≥0} {z : ℝ} (hz : z ≠ 0) : x ^ (1 / rw [← rpow_eq_rpow_iff hz, rpow_self_rpow_inv hz] @[simp] lemma rpow_rpow_inv {y : ℝ} (hy : y ≠ 0) (x : ℝ≥0) : (x ^ y) ^ y⁻¹ = x := by - rw [← rpow_mul, mul_inv_cancel hy, rpow_one] + rw [← rpow_mul, mul_inv_cancel₀ hy, rpow_one] @[simp] lemma rpow_inv_rpow {y : ℝ} (hy : y ≠ 0) (x : ℝ≥0) : (x ^ y⁻¹) ^ y = x := by - rw [← rpow_mul, inv_mul_cancel hy, rpow_one] + rw [← rpow_mul, inv_mul_cancel₀ hy, rpow_one] theorem pow_rpow_inv_natCast (x : ℝ≥0) {n : ℕ} (hn : n ≠ 0) : (x ^ n) ^ (n⁻¹ : ℝ) = x := by rw [← NNReal.coe_inj, coe_rpow, NNReal.coe_pow] @@ -655,13 +655,13 @@ theorem rpow_lt_rpow_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z < y ^ theorem le_rpow_inv_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ≤ y ^ z⁻¹ ↔ x ^ z ≤ y := by nth_rw 1 [← rpow_one x] - nth_rw 1 [← @_root_.mul_inv_cancel _ _ z hz.ne'] + nth_rw 1 [← @mul_inv_cancel₀ _ _ z hz.ne'] rw [rpow_mul, @rpow_le_rpow_iff _ _ z⁻¹ (by simp [hz])] @[deprecated le_rpow_inv_iff (since := "2024-07-10")] theorem le_rpow_one_div_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ≤ y ^ (1 / z) ↔ x ^ z ≤ y := by nth_rw 1 [← rpow_one x] - nth_rw 1 [← @_root_.mul_inv_cancel _ _ z hz.ne'] + nth_rw 1 [← @mul_inv_cancel₀ _ _ z hz.ne'] rw [rpow_mul, ← one_div, @rpow_le_rpow_iff _ _ (1 / z) (by simp [hz])] theorem rpow_inv_lt_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z⁻¹ < y ↔ x < y ^ z := by @@ -669,24 +669,24 @@ theorem rpow_inv_lt_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z⁻¹ < theorem lt_rpow_inv_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x < y ^ z⁻¹ ↔ x ^ z < y := by nth_rw 1 [← rpow_one x] - nth_rw 1 [← @_root_.mul_inv_cancel _ _ z (ne_of_lt hz).symm] + nth_rw 1 [← @mul_inv_cancel₀ _ _ z (ne_of_lt hz).symm] rw [rpow_mul, @rpow_lt_rpow_iff _ _ z⁻¹ (by simp [hz])] @[deprecated lt_rpow_inv_iff (since := "2024-07-10")] theorem lt_rpow_one_div_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x < y ^ (1 / z) ↔ x ^ z < y := by nth_rw 1 [← rpow_one x] - nth_rw 1 [← @_root_.mul_inv_cancel _ _ z (ne_of_lt hz).symm] + nth_rw 1 [← @mul_inv_cancel₀ _ _ z (ne_of_lt hz).symm] rw [rpow_mul, ← one_div, @rpow_lt_rpow_iff _ _ (1 / z) (by simp [hz])] theorem rpow_inv_le_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ z⁻¹ ≤ y ↔ x ≤ y ^ z := by nth_rw 1 [← ENNReal.rpow_one y] - nth_rw 1 [← @_root_.mul_inv_cancel _ _ z hz.ne.symm] + nth_rw 1 [← @mul_inv_cancel₀ _ _ z hz.ne.symm] rw [ENNReal.rpow_mul, ENNReal.rpow_le_rpow_iff (inv_pos.2 hz)] @[deprecated rpow_inv_le_iff (since := "2024-07-10")] theorem rpow_one_div_le_iff {x y : ℝ≥0∞} {z : ℝ} (hz : 0 < z) : x ^ (1 / z) ≤ y ↔ x ≤ y ^ z := by nth_rw 1 [← ENNReal.rpow_one y] - nth_rw 2 [← @_root_.mul_inv_cancel _ _ z hz.ne.symm] + nth_rw 2 [← @mul_inv_cancel₀ _ _ z hz.ne.symm] rw [ENNReal.rpow_mul, ← one_div, ENNReal.rpow_le_rpow_iff (one_div_pos.2 hz)] theorem rpow_lt_rpow_of_exponent_lt {x : ℝ≥0∞} {y z : ℝ} (hx : 1 < x) (hx' : x ≠ ⊤) (hyz : y < z) : @@ -829,16 +829,16 @@ theorem ofReal_rpow_of_nonneg {x p : ℝ} (hx_nonneg : 0 ≤ x) (hp_nonneg : 0 exact ofReal_rpow_of_pos (hx_nonneg.lt_of_ne hx0.symm) @[simp] lemma rpow_rpow_inv {y : ℝ} (hy : y ≠ 0) (x : ℝ≥0∞) : (x ^ y) ^ y⁻¹ = x := by - rw [← rpow_mul, mul_inv_cancel hy, rpow_one] + rw [← rpow_mul, mul_inv_cancel₀ hy, rpow_one] @[simp] lemma rpow_inv_rpow {y : ℝ} (hy : y ≠ 0) (x : ℝ≥0∞) : (x ^ y⁻¹) ^ y = x := by - rw [← rpow_mul, inv_mul_cancel hy, rpow_one] + rw [← rpow_mul, inv_mul_cancel₀ hy, rpow_one] lemma pow_rpow_inv_natCast {n : ℕ} (hn : n ≠ 0) (x : ℝ≥0∞) : (x ^ n) ^ (n⁻¹ : ℝ) = x := by - rw [← rpow_natCast, ← rpow_mul, mul_inv_cancel (by positivity), rpow_one] + rw [← rpow_natCast, ← rpow_mul, mul_inv_cancel₀ (by positivity), rpow_one] lemma rpow_inv_natCast_pow {n : ℕ} (hn : n ≠ 0) (x : ℝ≥0∞) : (x ^ (n⁻¹ : ℝ)) ^ n = x := by - rw [← rpow_natCast, ← rpow_mul, inv_mul_cancel (by positivity), rpow_one] + rw [← rpow_natCast, ← rpow_mul, inv_mul_cancel₀ (by positivity), rpow_one] lemma rpow_natCast_mul (x : ℝ≥0∞) (n : ℕ) (z : ℝ) : x ^ (n * z) = (x ^ n) ^ z := by rw [rpow_mul, rpow_natCast] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index 05be81afdde37..f36e3a7733dc3 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -446,18 +446,18 @@ theorem mul_log_eq_log_iff {x y z : ℝ} (hx : 0 < x) (hz : 0 < z) : by rintro rfl; rw [log_rpow hx]⟩ @[simp] lemma rpow_rpow_inv (hx : 0 ≤ x) (hy : y ≠ 0) : (x ^ y) ^ y⁻¹ = x := by - rw [← rpow_mul hx, mul_inv_cancel hy, rpow_one] + rw [← rpow_mul hx, mul_inv_cancel₀ hy, rpow_one] @[simp] lemma rpow_inv_rpow (hx : 0 ≤ x) (hy : y ≠ 0) : (x ^ y⁻¹) ^ y = x := by - rw [← rpow_mul hx, inv_mul_cancel hy, rpow_one] + rw [← rpow_mul hx, inv_mul_cancel₀ hy, rpow_one] theorem pow_rpow_inv_natCast (hx : 0 ≤ x) (hn : n ≠ 0) : (x ^ n) ^ (n⁻¹ : ℝ) = x := by have hn0 : (n : ℝ) ≠ 0 := Nat.cast_ne_zero.2 hn - rw [← rpow_natCast, ← rpow_mul hx, mul_inv_cancel hn0, rpow_one] + rw [← rpow_natCast, ← rpow_mul hx, mul_inv_cancel₀ hn0, rpow_one] theorem rpow_inv_natCast_pow (hx : 0 ≤ x) (hn : n ≠ 0) : (x ^ (n⁻¹ : ℝ)) ^ n = x := by have hn0 : (n : ℝ) ≠ 0 := Nat.cast_ne_zero.2 hn - rw [← rpow_natCast, ← rpow_mul hx, inv_mul_cancel hn0, rpow_one] + rw [← rpow_natCast, ← rpow_mul hx, inv_mul_cancel₀ hn0, rpow_one] lemma rpow_natCast_mul (hx : 0 ≤ x) (n : ℕ) (z : ℝ) : x ^ (n * z) = (x ^ n) ^ z := by rw [rpow_mul hx, rpow_natCast] @@ -683,7 +683,7 @@ theorem rpow_le_rpow_of_exponent_ge' (hx0 : 0 ≤ x) (hx1 : x ≤ 1) (hz : 0 ≤ theorem rpow_left_injOn {x : ℝ} (hx : x ≠ 0) : InjOn (fun y : ℝ => y ^ x) { y : ℝ | 0 ≤ y } := by rintro y hy z hz (hyz : y ^ x = z ^ x) - rw [← rpow_one y, ← rpow_one z, ← _root_.mul_inv_cancel hx, rpow_mul hy, rpow_mul hz, hyz] + rw [← rpow_one y, ← rpow_one z, ← mul_inv_cancel₀ hx, rpow_mul hy, rpow_mul hz, hyz] lemma rpow_left_inj (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : z ≠ 0) : x ^ z = y ^ z ↔ x = y := (rpow_left_injOn hz).eq_iff hx hy diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Series.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Series.lean index 0401b158fb2b4..dc4a20da971af 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Series.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Series.lean @@ -40,7 +40,7 @@ theorem Complex.hasSum_cos' (z : ℂ) : convert hasSum_fintype (_ : Fin 2 → ℂ) using 1 rw [Fin.sum_univ_two] simp_rw [Fin.val_zero, Fin.val_one, add_zero, pow_succ, pow_mul, mul_pow, neg_sq, ← two_mul, - neg_mul, mul_neg, neg_div, add_right_neg, zero_div, add_zero, + neg_mul, mul_neg, neg_div, add_neg_cancel, zero_div, add_zero, mul_div_cancel_left₀ _ (two_ne_zero : (2 : ℂ) ≠ 0)] theorem Complex.hasSum_sin' (z : ℂ) : diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index cc06f030ef173..293a306a56774 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -642,7 +642,7 @@ theorem tendsto_nat_floor_mul_div_atTop {a : R} (ha : 0 ≤ a) : apply tendsto_of_tendsto_of_tendsto_of_le_of_le' A tendsto_const_nhds · refine eventually_atTop.2 ⟨1, fun x hx ↦ ?_⟩ simp only [le_div_iff (zero_lt_one.trans_le hx), _root_.sub_mul, - inv_mul_cancel (zero_lt_one.trans_le hx).ne'] + inv_mul_cancel₀ (zero_lt_one.trans_le hx).ne'] have := Nat.lt_floor_add_one (a * x) linarith · refine eventually_atTop.2 ⟨1, fun x hx ↦ ?_⟩ @@ -662,7 +662,7 @@ theorem tendsto_nat_ceil_mul_div_atTop {a : R} (ha : 0 ≤ a) : rw [le_div_iff (zero_lt_one.trans_le hx)] exact Nat.le_ceil _ · refine eventually_atTop.2 ⟨1, fun x hx ↦ ?_⟩ - simp [div_le_iff (zero_lt_one.trans_le hx), inv_mul_cancel (zero_lt_one.trans_le hx).ne', + simp [div_le_iff (zero_lt_one.trans_le hx), inv_mul_cancel₀ (zero_lt_one.trans_le hx).ne', (Nat.ceil_lt_add_one (mul_nonneg ha (zero_le_one.trans hx))).le, add_mul] theorem tendsto_nat_ceil_div_atTop : Tendsto (fun x ↦ (⌈x⌉₊ : R) / x) atTop (𝓝 1) := by diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index 475f25eb1ace8..1000c3b72484d 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -160,7 +160,7 @@ theorem tendsto_div_of_monotone_of_exists_subseq_tendsto_div (u : ℕ → ℝ) ( simp_rw [div_eq_inv_mul] calc d < (n : ℝ)⁻¹ * n * (l - ε * (1 + l)) := by - rw [inv_mul_cancel, one_mul] + rw [inv_mul_cancel₀, one_mul] · linarith only [hε] · exact Nat.cast_ne_zero.2 (ne_of_gt npos) _ = (n : ℝ)⁻¹ * (n * l - ε * (1 + l) * n) := by ring @@ -226,7 +226,7 @@ theorem sum_div_pow_sq_le_div_sq (N : ℕ) {j : ℝ} (hj : 0 < j) {c : ℝ} (hc · exact sub_pos.2 (pow_lt_one (inv_nonneg.2 cpos.le) (inv_lt_one hc) two_ne_zero) have : c ^ 3 = c ^ 2 * c := by ring simp only [mul_sub, this, mul_one, inv_pow, sub_le_sub_iff_left] - rw [mul_assoc, mul_comm c, ← mul_assoc, mul_inv_cancel (sq_pos_of_pos cpos).ne', one_mul] + rw [mul_assoc, mul_comm c, ← mul_assoc, mul_inv_cancel₀ (sq_pos_of_pos cpos).ne', one_mul] simpa using pow_le_pow_right hc.le one_le_two have C : c⁻¹ ^ 2 < 1 := pow_lt_one (inv_nonneg.2 cpos.le) (inv_lt_one hc) two_ne_zero calc diff --git a/Mathlib/CategoryTheory/Abelian/Basic.lean b/Mathlib/CategoryTheory/Abelian/Basic.lean index c972721a7fa74..b7e4dbcf98889 100644 --- a/Mathlib/CategoryTheory/Abelian/Basic.lean +++ b/Mathlib/CategoryTheory/Abelian/Basic.lean @@ -550,7 +550,7 @@ abbrev pullbackToBiproduct : pullback f g ⟶ X ⊞ Y := `(0, g)`. -/ abbrev pullbackToBiproductFork : KernelFork (biprod.desc f (-g)) := KernelFork.ofι (pullbackToBiproduct f g) <| by - rw [biprod.lift_desc, comp_neg, pullback.condition, add_right_neg] + rw [biprod.lift_desc, comp_neg, pullback.condition, add_neg_cancel] /-- The canonical map `pullback f g ⟶ X ⊞ Y` is a kernel of the map induced by `(f, -g)`. -/ @@ -581,7 +581,7 @@ abbrev biproductToPushout : Y ⊞ Z ⟶ pushout f g := `X ⟶ Y ⊞ Z` induced by `f` and `-g`. -/ abbrev biproductToPushoutCofork : CokernelCofork (biprod.lift f (-g)) := CokernelCofork.ofπ (biproductToPushout f g) <| by - rw [biprod.lift_desc, neg_comp, pushout.condition, add_right_neg] + rw [biprod.lift_desc, neg_comp, pushout.condition, add_neg_cancel] /-- The cofork induced by the canonical map `Y ⊞ Z ⟶ pushout f g` is in fact a colimit cokernel cofork. -/ diff --git a/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean b/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean index b1d5be478eea8..c8c9e63e83d6f 100644 --- a/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean +++ b/Mathlib/CategoryTheory/Abelian/NonPreadditive.lean @@ -364,9 +364,9 @@ theorem add_comm {X Y : C} (a b : X ⟶ Y) : a + b = b + a := by theorem add_neg {X Y : C} (a b : X ⟶ Y) : a + -b = a - b := by rw [add_def, neg_neg] -theorem add_neg_self {X Y : C} (a : X ⟶ Y) : a + -a = 0 := by rw [add_neg, sub_self] +theorem add_neg_cancel {X Y : C} (a : X ⟶ Y) : a + -a = 0 := by rw [add_neg, sub_self] -theorem neg_add_self {X Y : C} (a : X ⟶ Y) : -a + a = 0 := by rw [add_comm, add_neg_self] +theorem neg_add_cancel {X Y : C} (a : X ⟶ Y) : -a + a = 0 := by rw [add_comm, add_neg_cancel] theorem neg_sub' {X Y : C} (a b : X ⟶ Y) : -(a - b) = -a + b := by rw [neg_def, neg_def] @@ -406,7 +406,7 @@ def preadditive : Preadditive C where zero_add := neg_neg add_zero := add_zero neg := fun f => -f - add_left_neg := neg_add_self + neg_add_cancel := neg_add_cancel sub_eq_add_neg := fun f g => (add_neg f g).symm -- Porting note: autoParam failed add_comm := add_comm nsmul := nsmulRec diff --git a/Mathlib/CategoryTheory/Endomorphism.lean b/Mathlib/CategoryTheory/Endomorphism.lean index f1a262d510296..30f789b4ad257 100644 --- a/Mathlib/CategoryTheory/Endomorphism.lean +++ b/Mathlib/CategoryTheory/Endomorphism.lean @@ -93,7 +93,7 @@ end MulAction /-- In a groupoid, endomorphisms form a group -/ instance group {C : Type u} [Groupoid.{v} C] (X : C) : Group (End X) where - mul_left_inv := Groupoid.comp_inv + inv_mul_cancel := Groupoid.comp_inv inv := Groupoid.inv end End @@ -128,7 +128,7 @@ instance : Group (Aut X) where mul_assoc _ _ _ := (Iso.trans_assoc _ _ _).symm one_mul := Iso.trans_refl mul_one := Iso.refl_trans - mul_left_inv := Iso.self_symm_id + inv_mul_cancel := Iso.self_symm_id theorem Aut_mul_def (f g : Aut X) : f * g = g.trans f := rfl diff --git a/Mathlib/CategoryTheory/Galois/Examples.lean b/Mathlib/CategoryTheory/Galois/Examples.lean index 55178c71b48ec..f7b722ebca408 100644 --- a/Mathlib/CategoryTheory/Galois/Examples.lean +++ b/Mathlib/CategoryTheory/Galois/Examples.lean @@ -54,7 +54,7 @@ noncomputable def Action.imageComplement {X Y : Action FintypeCat (MonCat.of G)} use X.ρ g⁻¹ x calc (X.ρ g⁻¹ ≫ f.hom) x = (Y.ρ g⁻¹ * Y.ρ g) y.val := by rw [f.comm, FintypeCat.comp_apply, h]; rfl - _ = y.val := by rw [← map_mul, mul_left_inv, Action.ρ_one, FintypeCat.id_apply] + _ = y.val := by rw [← map_mul, inv_mul_cancel, Action.ρ_one, FintypeCat.id_apply] map_one' := by simp only [Action.ρ_one]; rfl map_mul' := fun g h ↦ FintypeCat.hom_ext _ _ <| fun y ↦ Subtype.ext <| by exact congrFun (MonoidHom.map_mul Y.ρ g h) y.val diff --git a/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean b/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean index 03b0e8ff4431f..31aaca1e47820 100644 --- a/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean +++ b/Mathlib/CategoryTheory/Groupoid/VertexGroup.lean @@ -43,7 +43,7 @@ instance vertexGroup (c : C) : Group (c ⟶ c) where one_mul := Category.id_comp mul_one := Category.comp_id inv := Groupoid.inv - mul_left_inv := inv_comp + inv_mul_cancel := inv_comp -- dsimp loops when applying this lemma to its LHS, -- probably https://github.com/leanprover/lean4/pull/2867 diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index 395386e40d32b..121e9a8862368 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -177,9 +177,9 @@ instance instAddCommGroupHom [Preadditive C] {P Q : Karoubi C} : AddCommGroup (P add_comm f g := by ext apply add_comm - add_left_neg f := by + neg_add_cancel f := by ext - apply add_left_neg + apply neg_add_cancel zsmul := zsmulRec nsmul := nsmulRec diff --git a/Mathlib/CategoryTheory/Localization/CalculusOfFractions/Preadditive.lean b/Mathlib/CategoryTheory/Localization/CalculusOfFractions/Preadditive.lean index 217e3e9b2940e..d7bc5297388e4 100644 --- a/Mathlib/CategoryTheory/Localization/CalculusOfFractions/Preadditive.lean +++ b/Mathlib/CategoryTheory/Localization/CalculusOfFractions/Preadditive.lean @@ -158,13 +158,13 @@ lemma zero_add' (f : L.obj X ⟶ L.obj Y) : add' W (L.map 0) f = f := by rw [add'_comm, add'_zero] -lemma add'_left_neg' (f : L.obj X ⟶ L.obj Y) : +lemma neg'_add'_self (f : L.obj X ⟶ L.obj Y) : add' W (neg' W f) f = L.map 0 := by obtain ⟨α, rfl⟩ := exists_leftFraction L W f have := inverts L W _ α.hs rw [add'_eq W _ _ (LeftFraction₂.mk (-α.f) α.f α.s α.hs) (neg'_eq W _ _ rfl) rfl] simp only [← cancel_mono (L.map α.s), LeftFraction.map_comp_map_s, ← L.map_comp, - Limits.zero_comp, add_left_neg] + Limits.zero_comp, neg_add_cancel] lemma add'_assoc (f₁ f₂ f₃ : L.obj X ⟶ L.obj Y) : add' W (add' W f₁ f₂) f₃ = add' W f₁ (add' W f₂ f₃) := by @@ -229,7 +229,7 @@ noncomputable def addCommGroup' : AddCommGroup (L.obj X ⟶ L.obj Y) := by add_zero := add'_zero _ add_comm := add'_comm _ zero_add := zero_add' _ - add_left_neg := add'_left_neg' _ + neg_add_cancel := neg'_add'_self _ nsmul := nsmulRec zsmul := zsmulRec } diff --git a/Mathlib/CategoryTheory/Preadditive/Basic.lean b/Mathlib/CategoryTheory/Preadditive/Basic.lean index 3e19591f3a522..8df1e560ccfc4 100644 --- a/Mathlib/CategoryTheory/Preadditive/Basic.lean +++ b/Mathlib/CategoryTheory/Preadditive/Basic.lean @@ -194,7 +194,7 @@ Lean 4 so the previous instance needed modification. Was following my nose here. instance {X : C} : Ring (End X) := { (inferInstance : Semiring (End X)), (inferInstance : AddCommGroup (End X)) with - add_left_neg := add_left_neg } + neg_add_cancel := neg_add_cancel } instance moduleEndRight {X Y : C} : Module (End Y) (X ⟶ Y) where smul_add _ _ _ := add_comp _ _ _ _ _ _ diff --git a/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean b/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean index 61a44cda69610..6770c026e762f 100644 --- a/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean +++ b/Mathlib/CategoryTheory/Preadditive/EilenbergMoore.lean @@ -85,10 +85,10 @@ instance Monad.algebraPreadditive : Preadditive (Monad.Algebra T) where intros ext simp only [negSucc_zsmul, neg_inj, ← Nat.cast_smul_eq_nsmul ℤ] - add_left_neg := by + neg_add_cancel := by intros ext - apply add_left_neg + apply neg_add_cancel add_comm := by intros ext @@ -166,10 +166,10 @@ instance Comonad.coalgebraPreadditive : Preadditive (Comonad.Coalgebra U) where intros ext simp only [negSucc_zsmul, neg_inj, ← Nat.cast_smul_eq_nsmul ℤ] - add_left_neg := by + neg_add_cancel := by intros ext - apply add_left_neg + apply neg_add_cancel add_comm := by intros ext diff --git a/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean b/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean index c9856b25bf323..006286e5be91c 100644 --- a/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean +++ b/Mathlib/CategoryTheory/Preadditive/EndoFunctor.lean @@ -85,10 +85,10 @@ instance Endofunctor.algebraPreadditive : Preadditive (Endofunctor.Algebra F) wh intros apply Algebra.Hom.ext simp only [negSucc_zsmul, neg_inj, ← Nat.cast_smul_eq_nsmul ℤ] - add_left_neg := by + neg_add_cancel := by intros apply Algebra.Hom.ext - apply add_left_neg + apply neg_add_cancel add_comm := by intros apply Algebra.Hom.ext @@ -163,10 +163,10 @@ instance Endofunctor.coalgebraPreadditive : Preadditive (Endofunctor.Coalgebra F intros apply Coalgebra.Hom.ext simp only [negSucc_zsmul, neg_inj, ← Nat.cast_smul_eq_nsmul ℤ] - add_left_neg := by + neg_add_cancel := by intros apply Coalgebra.Hom.ext - apply add_left_neg + apply neg_add_cancel add_comm := by intros apply Coalgebra.Hom.ext diff --git a/Mathlib/CategoryTheory/Preadditive/FunctorCategory.lean b/Mathlib/CategoryTheory/Preadditive/FunctorCategory.lean index 739c65736d1e6..f1786b2cbca39 100644 --- a/Mathlib/CategoryTheory/Preadditive/FunctorCategory.lean +++ b/Mathlib/CategoryTheory/Preadditive/FunctorCategory.lean @@ -57,11 +57,11 @@ instance functorCategoryPreadditive : Preadditive (C ⥤ D) where dsimp ext apply sub_eq_add_neg - add_left_neg := by + neg_add_cancel := by intros dsimp ext - apply add_left_neg } + apply neg_add_cancel } add_comp := by intros dsimp diff --git a/Mathlib/CategoryTheory/Quotient/Preadditive.lean b/Mathlib/CategoryTheory/Quotient/Preadditive.lean index 80e1204fe4404..db607b900b1d4 100644 --- a/Mathlib/CategoryTheory/Quotient/Preadditive.lean +++ b/Mathlib/CategoryTheory/Quotient/Preadditive.lean @@ -67,7 +67,7 @@ def preadditive zero_add := by rintro ⟨_⟩; exact congr_arg (functor r).map (zero_add _) add_zero := by rintro ⟨_⟩; exact congr_arg (functor r).map (add_zero _) add_comm := by rintro ⟨_⟩ ⟨_⟩; exact congr_arg (functor r).map (add_comm _ _) - add_left_neg := by rintro ⟨_⟩; exact congr_arg (functor r).map (add_left_neg _) + neg_add_cancel := by rintro ⟨_⟩; exact congr_arg (functor r).map (neg_add_cancel _) -- todo: use a better defeq nsmul := nsmulRec zsmul := zsmulRec } diff --git a/Mathlib/CategoryTheory/Shift/Basic.lean b/Mathlib/CategoryTheory/Shift/Basic.lean index 53c86f1c6708a..d830b0b79c97d 100644 --- a/Mathlib/CategoryTheory/Shift/Basic.lean +++ b/Mathlib/CategoryTheory/Shift/Basic.lean @@ -409,7 +409,7 @@ def shiftEquiv' (i j : A) (h : i + j = 0) : C ≌ C where rfl /-- Shifting by `n` and shifting by `-n` forms an equivalence. -/ -abbrev shiftEquiv (n : A) : C ≌ C := shiftEquiv' C n (-n) (add_neg_self n) +abbrev shiftEquiv (n : A) : C ≌ C := shiftEquiv' C n (-n) (add_neg_cancel n) variable (X Y : C) (f : X ⟶ Y) @@ -431,14 +431,14 @@ abbrev shiftNegShift (i : A) : X⟦-i⟧⟦i⟧ ≅ X := variable {X Y} theorem shift_shift_neg' (i : A) : - f⟦i⟧'⟦-i⟧' = (shiftFunctorCompIsoId C i (-i) (add_neg_self i)).hom.app X ≫ - f ≫ (shiftFunctorCompIsoId C i (-i) (add_neg_self i)).inv.app Y := - (NatIso.naturality_2 (shiftFunctorCompIsoId C i (-i) (add_neg_self i)) f).symm + f⟦i⟧'⟦-i⟧' = (shiftFunctorCompIsoId C i (-i) (add_neg_cancel i)).hom.app X ≫ + f ≫ (shiftFunctorCompIsoId C i (-i) (add_neg_cancel i)).inv.app Y := + (NatIso.naturality_2 (shiftFunctorCompIsoId C i (-i) (add_neg_cancel i)) f).symm theorem shift_neg_shift' (i : A) : - f⟦-i⟧'⟦i⟧' = (shiftFunctorCompIsoId C (-i) i (neg_add_self i)).hom.app X ≫ f ≫ - (shiftFunctorCompIsoId C (-i) i (neg_add_self i)).inv.app Y := - (NatIso.naturality_2 (shiftFunctorCompIsoId C (-i) i (neg_add_self i)) f).symm + f⟦-i⟧'⟦i⟧' = (shiftFunctorCompIsoId C (-i) i (neg_add_cancel i)).hom.app X ≫ f ≫ + (shiftFunctorCompIsoId C (-i) i (neg_add_cancel i)).inv.app Y := + (NatIso.naturality_2 (shiftFunctorCompIsoId C (-i) i (neg_add_cancel i)) f).symm theorem shift_equiv_triangle (n : A) (X : C) : (shiftShiftNeg X n).inv⟦n⟧' ≫ (shiftNegShift (X⟦n⟧) n).hom = 𝟙 (X⟦n⟧) := @@ -449,40 +449,40 @@ section theorem shift_shiftFunctorCompIsoId_hom_app (n m : A) (h : n + m = 0) (X : C) : ((shiftFunctorCompIsoId C n m h).hom.app X)⟦n⟧' = (shiftFunctorCompIsoId C m n - (by rw [← neg_eq_of_add_eq_zero_left h, add_right_neg])).hom.app (X⟦n⟧) := by + (by rw [← neg_eq_of_add_eq_zero_left h, add_neg_cancel])).hom.app (X⟦n⟧) := by dsimp [shiftFunctorCompIsoId] simpa only [Functor.map_comp, ← shiftFunctorAdd'_zero_add_inv_app n X, ← shiftFunctorAdd'_add_zero_inv_app n X] using shiftFunctorAdd'_assoc_inv_app n m n 0 0 n h - (by rw [← neg_eq_of_add_eq_zero_left h, add_right_neg]) (by rw [h, zero_add]) X + (by rw [← neg_eq_of_add_eq_zero_left h, add_neg_cancel]) (by rw [h, zero_add]) X theorem shift_shiftFunctorCompIsoId_inv_app (n m : A) (h : n + m = 0) (X : C) : ((shiftFunctorCompIsoId C n m h).inv.app X)⟦n⟧' = ((shiftFunctorCompIsoId C m n - (by rw [← neg_eq_of_add_eq_zero_left h, add_right_neg])).inv.app (X⟦n⟧)) := by + (by rw [← neg_eq_of_add_eq_zero_left h, add_neg_cancel])).inv.app (X⟦n⟧)) := by rw [← cancel_mono (((shiftFunctorCompIsoId C n m h).hom.app X)⟦n⟧'), ← Functor.map_comp, Iso.inv_hom_id_app, Functor.map_id, shift_shiftFunctorCompIsoId_hom_app, Iso.inv_hom_id_app] rfl -theorem shift_shiftFunctorCompIsoId_add_neg_self_hom_app (n : A) (X : C) : - ((shiftFunctorCompIsoId C n (-n) (add_neg_self n)).hom.app X)⟦n⟧' = - (shiftFunctorCompIsoId C (-n) n (neg_add_self n)).hom.app (X⟦n⟧) := by +theorem shift_shiftFunctorCompIsoId_add_neg_cancel_hom_app (n : A) (X : C) : + ((shiftFunctorCompIsoId C n (-n) (add_neg_cancel n)).hom.app X)⟦n⟧' = + (shiftFunctorCompIsoId C (-n) n (neg_add_cancel n)).hom.app (X⟦n⟧) := by apply shift_shiftFunctorCompIsoId_hom_app -theorem shift_shiftFunctorCompIsoId_add_neg_self_inv_app (n : A) (X : C) : - ((shiftFunctorCompIsoId C n (-n) (add_neg_self n)).inv.app X)⟦n⟧' = - (shiftFunctorCompIsoId C (-n) n (neg_add_self n)).inv.app (X⟦n⟧) := by +theorem shift_shiftFunctorCompIsoId_add_neg_cancel_inv_app (n : A) (X : C) : + ((shiftFunctorCompIsoId C n (-n) (add_neg_cancel n)).inv.app X)⟦n⟧' = + (shiftFunctorCompIsoId C (-n) n (neg_add_cancel n)).inv.app (X⟦n⟧) := by apply shift_shiftFunctorCompIsoId_inv_app -theorem shift_shiftFunctorCompIsoId_neg_add_self_hom_app (n : A) (X : C) : - ((shiftFunctorCompIsoId C (-n) n (neg_add_self n)).hom.app X)⟦-n⟧' = - (shiftFunctorCompIsoId C n (-n) (add_neg_self n)).hom.app (X⟦-n⟧) := by +theorem shift_shiftFunctorCompIsoId_neg_add_cancel_hom_app (n : A) (X : C) : + ((shiftFunctorCompIsoId C (-n) n (neg_add_cancel n)).hom.app X)⟦-n⟧' = + (shiftFunctorCompIsoId C n (-n) (add_neg_cancel n)).hom.app (X⟦-n⟧) := by apply shift_shiftFunctorCompIsoId_hom_app -theorem shift_shiftFunctorCompIsoId_neg_add_self_inv_app (n : A) (X : C) : - ((shiftFunctorCompIsoId C (-n) n (neg_add_self n)).inv.app X)⟦-n⟧' = - (shiftFunctorCompIsoId C n (-n) (add_neg_self n)).inv.app (X⟦-n⟧) := by +theorem shift_shiftFunctorCompIsoId_neg_add_cancel_inv_app (n : A) (X : C) : + ((shiftFunctorCompIsoId C (-n) n (neg_add_cancel n)).inv.app X)⟦-n⟧' = + (shiftFunctorCompIsoId C n (-n) (add_neg_cancel n)).inv.app (X⟦-n⟧) := by apply shift_shiftFunctorCompIsoId_inv_app end diff --git a/Mathlib/CategoryTheory/SingleObj.lean b/Mathlib/CategoryTheory/SingleObj.lean index 252c040732722..b7e8e14ddffd6 100644 --- a/Mathlib/CategoryTheory/SingleObj.lean +++ b/Mathlib/CategoryTheory/SingleObj.lean @@ -80,12 +80,12 @@ See . -/ instance groupoid : Groupoid (SingleObj G) where inv x := x⁻¹ - inv_comp := mul_right_inv - comp_inv := mul_left_inv + inv_comp := mul_inv_cancel + comp_inv := inv_mul_cancel theorem inv_as_inv {x y : SingleObj G} (f : x ⟶ y) : inv f = f⁻¹ := by apply IsIso.inv_eq_of_hom_inv_id - rw [comp_as_mul, inv_mul_self, id_as_one] + rw [comp_as_mul, inv_mul_cancel, id_as_one] /-- Abbreviation that allows writing `CategoryTheory.SingleObj.star` rather than `Quiver.SingleObj.star`. @@ -142,11 +142,11 @@ def differenceFunctor (f : C → G) : C ⥤ SingleObj G where map {x y} _ := f y * (f x)⁻¹ map_id := by intro - simp only [SingleObj.id_as_one, mul_right_inv] + simp only [SingleObj.id_as_one, mul_inv_cancel] map_comp := by intros dsimp - rw [SingleObj.comp_as_mul, ← mul_assoc, mul_left_inj, mul_assoc, inv_mul_self, mul_one] + rw [SingleObj.comp_as_mul, ← mul_assoc, mul_left_inj, mul_assoc, inv_mul_cancel, mul_one] /-- A monoid homomorphism `f: M → End X` into the endomorphisms of an object `X` of a category `C` induces a functor `SingleObj M ⥤ C`. -/ diff --git a/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean b/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean index ea0013a826848..2f3b34ebb98a9 100644 --- a/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean +++ b/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean @@ -116,7 +116,7 @@ instance : Group (OneCochain G U) where mul_assoc _ _ _ := by ext; apply mul_assoc one_mul _ := by ext; apply one_mul mul_one _ := by ext; apply mul_one - mul_left_inv _ := by ext; apply mul_left_inv + inv_mul_cancel _ := by ext; apply inv_mul_cancel end OneCochain @@ -141,7 +141,7 @@ lemma ev_refl (γ : OneCocycle G U) (i : I) ⦃T : C⦄ (a : T ⟶ U i) : lemma ev_symm (γ : OneCocycle G U) (i j : I) ⦃T : C⦄ (a : T ⟶ U i) (b : T ⟶ U j) : γ.ev i j a b = (γ.ev j i b a)⁻¹ := by rw [← mul_left_inj (γ.ev j i b a), γ.ev_trans i j i a b a, - ev_refl, mul_left_inv] + ev_refl, inv_mul_cancel] end OneCocycle @@ -161,7 +161,7 @@ lemma symm {γ₁ γ₂ : OneCochain G U} {α : ZeroCochain G U} (h : OneCohomol OneCohomologyRelation γ₂ γ₁ α⁻¹ := fun i j T a b ↦ by rw [← mul_left_inj (G.map b.op (α j)), mul_assoc, ← h i j a b, mul_assoc, Cochain₀.inv_apply, map_inv, inv_mul_cancel_left, - Cochain₀.inv_apply, map_inv, mul_left_inv, mul_one] + Cochain₀.inv_apply, map_inv, inv_mul_cancel, mul_one] lemma trans {γ₁ γ₂ γ₃ : OneCochain G U} {α β : ZeroCochain G U} (h₁₂ : OneCohomologyRelation γ₁ γ₂ α) (h₂₃ : OneCohomologyRelation γ₂ γ₃ β) : diff --git a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean index 3c73fda709c13..f5584c6bfe3e0 100644 --- a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean +++ b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean @@ -221,7 +221,7 @@ lemma homologySequence_exact₁ : (Iso.refl _) (Iso.refl _) ?_ (by simp) dsimp simp only [homologySequenceδ, neg_comp, map_neg, comp_id, - F.shiftIso_hom_app_comp_shiftMap_of_add_eq_zero T.mor₃ (-1) (neg_add_self 1) n₀ n₁ (by omega)] + F.shiftIso_hom_app_comp_shiftMap_of_add_eq_zero T.mor₃ (-1) (neg_add_cancel 1) n₀ n₁ (by omega)] lemma homologySequence_epi_shift_map_mor₁_iff : Epi ((F.shift n₀).map T.mor₁) ↔ (F.shift n₀).map T.mor₂ = 0 := diff --git a/Mathlib/CategoryTheory/Triangulated/Opposite.lean b/Mathlib/CategoryTheory/Triangulated/Opposite.lean index 43a838e1e4767..4de06f2ff9fb9 100644 --- a/Mathlib/CategoryTheory/Triangulated/Opposite.lean +++ b/Mathlib/CategoryTheory/Triangulated/Opposite.lean @@ -146,9 +146,9 @@ noncomputable def opShiftFunctorEquivalence (n : ℤ) : Cᵒᵖ ≌ Cᵒᵖ wher functor_unitIso_comp X := Quiver.Hom.unop_inj (by dsimp [shiftFunctorOpIso] erw [comp_id, Functor.map_id, comp_id] - change (shiftFunctorCompIsoId C n (-n) (add_neg_self n)).inv.app (X.unop⟦-n⟧) ≫ - ((shiftFunctorCompIsoId C (-n) n (neg_add_self n)).hom.app X.unop)⟦-n⟧' = 𝟙 _ - rw [shift_shiftFunctorCompIsoId_neg_add_self_hom_app n X.unop, Iso.inv_hom_id_app]) + change (shiftFunctorCompIsoId C n (-n) (add_neg_cancel n)).inv.app (X.unop⟦-n⟧) ≫ + ((shiftFunctorCompIsoId C (-n) n (neg_add_cancel n)).hom.app X.unop)⟦-n⟧' = 𝟙 _ + rw [shift_shiftFunctorCompIsoId_neg_add_cancel_hom_app n X.unop, Iso.inv_hom_id_app]) /-! The naturality of the unit and counit isomorphisms are restated in the following lemmas so as to mitigate the need for `erw`. -/ diff --git a/Mathlib/CategoryTheory/Triangulated/Rotate.lean b/Mathlib/CategoryTheory/Triangulated/Rotate.lean index e72194eedf1f4..4ef56167fa648 100644 --- a/Mathlib/CategoryTheory/Triangulated/Rotate.lean +++ b/Mathlib/CategoryTheory/Triangulated/Rotate.lean @@ -72,8 +72,8 @@ def Triangle.invRotate (T : Triangle C) : Triangle C := end attribute [local simp] shift_shift_neg' shift_neg_shift' - shift_shiftFunctorCompIsoId_add_neg_self_inv_app - shift_shiftFunctorCompIsoId_add_neg_self_hom_app + shift_shiftFunctorCompIsoId_add_neg_cancel_inv_app + shift_shiftFunctorCompIsoId_add_neg_cancel_hom_app variable (C) diff --git a/Mathlib/CategoryTheory/Triangulated/TStructure/Basic.lean b/Mathlib/CategoryTheory/Triangulated/TStructure/Basic.lean index 9c9794b186478..99dd35d10775c 100644 --- a/Mathlib/CategoryTheory/Triangulated/TStructure/Basic.lean +++ b/Mathlib/CategoryTheory/Triangulated/TStructure/Basic.lean @@ -82,7 +82,7 @@ lemma exists_triangle (A : C) (n₀ n₁ : ℤ) (h : n₀ + 1 = n₁) : refine isomorphic_distinguished _ (Triangle.shift_distinguished _ mem (-n₀)) _ ?_ refine Triangle.isoMk _ _ (Iso.refl _) e.symm (Iso.refl _) ?_ ?_ ?_ all_goals dsimp; simp [T] - exact ⟨_, _, t.LE_shift _ _ _ (neg_add_self n₀) _ hX, + exact ⟨_, _, t.LE_shift _ _ _ (neg_add_cancel n₀) _ hX, t.GE_shift _ _ _ (by omega) _ hY, _, _, _, hT'⟩ lemma predicateShift_LE (a n n' : ℤ) (hn' : a + n = n') : diff --git a/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean b/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean index 53109ea474327..b811588238be2 100644 --- a/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean +++ b/Mathlib/CategoryTheory/Triangulated/TriangleShift.lean @@ -121,7 +121,7 @@ noncomputable def invRotateInvRotateInvRotateIso : (by aesop_cat) (by dsimp [shiftFunctorCompIsoId] - simp [shiftFunctorComm_eq C _ _ _ (add_neg_self (1 : ℤ))])) + simp [shiftFunctorComm_eq C _ _ _ (add_neg_cancel (1 : ℤ))])) (by aesop_cat) /-- The inverse of the rotation of triangles can be expressed using a double @@ -133,7 +133,7 @@ noncomputable def invRotateIsoRotateRotateShiftFunctorNegOne : _ ≅ invRotate C ⋙ Triangle.shiftFunctor C 0 := isoWhiskerLeft _ (Triangle.shiftFunctorZero C).symm _ ≅ invRotate C ⋙ Triangle.shiftFunctor C 1 ⋙ Triangle.shiftFunctor C (-1) := - isoWhiskerLeft _ (Triangle.shiftFunctorAdd' C 1 (-1) 0 (add_neg_self 1)) + isoWhiskerLeft _ (Triangle.shiftFunctorAdd' C 1 (-1) 0 (add_neg_cancel 1)) _ ≅ invRotate C ⋙ (rotate C ⋙ rotate C ⋙ rotate C) ⋙ Triangle.shiftFunctor C (-1) := isoWhiskerLeft _ (isoWhiskerRight (rotateRotateRotateIso C).symm _) _ ≅ (invRotate C ⋙ rotate C) ⋙ rotate C ⋙ rotate C ⋙ Triangle.shiftFunctor C (-1) := diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean index 22b62ad1ed1aa..27372279d8fdb 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean @@ -389,7 +389,7 @@ theorem le_N (hN : 2 ≤ N) : (2 * dValue N - 1) ^ nValue N ≤ N := by suffices i : (2 * dValue N : ℝ) ≤ (N : ℝ) ^ (nValue N : ℝ)⁻¹ by rw [← rpow_natCast] apply (rpow_le_rpow (mul_nonneg zero_le_two (cast_nonneg _)) i (cast_nonneg _)).trans - rw [← rpow_mul (cast_nonneg _), inv_mul_cancel, rpow_one] + rw [← rpow_mul (cast_nonneg _), inv_mul_cancel₀, rpow_one] rw [cast_ne_zero] apply (nValue_pos hN).ne' rw [← le_div_iff'] diff --git a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean index 5d91141edd867..d7ff51b86ec91 100644 --- a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean +++ b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean @@ -81,7 +81,7 @@ private lemma Fintype.sum_div_mul_card_choose_card : rw [mem_powersetCard_univ.1 hs] simp_rw [sum_congr rfl this, sum_const, card_powersetCard, card_univ, nsmul_eq_mul, mul_div, mul_comm, ← mul_div] - rw [← mul_sum, ← mul_inv_cancel (cast_ne_zero.mpr card_ne_zero : (card α : ℚ) ≠ 0), ← mul_add, + rw [← mul_sum, ← mul_inv_cancel₀ (cast_ne_zero.mpr card_ne_zero : (card α : ℚ) ≠ 0), ← mul_add, add_comm _ ((card α)⁻¹ : ℚ), ← sum_insert (f := fun x : ℕ ↦ (x⁻¹ : ℚ)) not_mem_range_self, ← range_succ] have (n) (hn : n ∈ range (card α + 1)) : diff --git a/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean b/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean index 6ac3a3c304ad4..aed46e2c9d07d 100644 --- a/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean +++ b/Mathlib/Combinatorics/SetFamily/CauchyDavenport.lean @@ -145,7 +145,7 @@ lemma Finset.min_le_card_mul (hs : s.Nonempty) (ht : t.Nonempty) : obtain hsg | hsg := eq_or_ne (op g • s) s · have hS : (zpowers g : Set α) ⊆ a⁻¹ • (s : Set α) := by refine forall_mem_zpowers.2 <| @zpow_induction_right _ _ _ (· ∈ a⁻¹ • (s : Set α)) - ⟨_, ha, inv_mul_self _⟩ (fun c hc ↦ ?_) fun c hc ↦ ?_ + ⟨_, ha, inv_mul_cancel _⟩ (fun c hc ↦ ?_) fun c hc ↦ ?_ · rw [← hsg, coe_smul_finset, smul_comm] exact Set.smul_mem_smul_set hc · simp only diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index 6f0e19fdcf5c8..f324a7c9a4781 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -1031,7 +1031,7 @@ lemma rpow_p_mul_one_sub_smoothingFn_le : case n_gt_one => rwa [Set.mem_Ioi, Nat.one_lt_cast] case bn_gt_one => - calc 1 = b i * (b i)⁻¹ := by rw [mul_inv_cancel (by positivity)] + calc 1 = b i * (b i)⁻¹ := by rw [mul_inv_cancel₀ (by positivity)] _ ≤ b i * ⌈(b i)⁻¹⌉₊ := by gcongr; exact Nat.le_ceil _ _ < b i * n := by gcongr; rw [Nat.cast_lt]; exact hn case le => calc b i * n ≤ 1 * n := by have := R.b_lt_one i; gcongr @@ -1127,7 +1127,7 @@ lemma rpow_p_mul_one_add_smoothingFn_ge : rw [Nat.one_lt_cast] exact hn' case bn_gt_one => - calc 1 = b i * (b i)⁻¹ := by rw [mul_inv_cancel (by positivity)] + calc 1 = b i * (b i)⁻¹ := by rw [mul_inv_cancel₀ (by positivity)] _ ≤ b i * ⌈(b i)⁻¹⌉₊ := by gcongr; exact Nat.le_ceil _ _ < b i * n := by gcongr; rw [Nat.cast_lt]; exact hn case le => calc b i * n ≤ 1 * n := by have := R.b_lt_one i; gcongr @@ -1284,7 +1284,7 @@ lemma T_isBigO_smoothingFn_mul_asympBound : refine mul_nonpos_of_nonpos_of_nonneg ?_ g_pos rw [sub_nonpos] calc 1 ≤ 2 * (c₁⁻¹ * c₁) * (1/2) := by - rw [inv_mul_cancel (by positivity : c₁ ≠ 0)]; norm_num + rw [inv_mul_cancel₀ (by positivity : c₁ ≠ 0)]; norm_num _ = (2 * c₁⁻¹) * c₁ * (1/2) := by ring _ ≤ C * c₁ * (1 - ε n) := by gcongr · rw [hC]; exact le_max_left _ _ @@ -1435,7 +1435,7 @@ lemma smoothingFn_mul_asympBound_isBigO_T : _ = C * (2 * c₁) := by ring _ ≤ (2 * c₁)⁻¹ * (2 * c₁) := by gcongr; exact min_le_left _ _ _ = c₁⁻¹ * c₁ := by ring - _ = 1 := inv_mul_cancel (by positivity) + _ = 1 := inv_mul_cancel₀ (by positivity) _ = C * ((1 + ε n) * asympBound g a b n) := by ring /-- The **Akra-Bazzi theorem**: `T ∈ O(n^p (1 + ∑_u^n g(u) / u^{p+1}))` -/ diff --git a/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean b/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean index 1d84e2de7fdfb..8c7637a070683 100644 --- a/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean +++ b/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean @@ -675,8 +675,8 @@ lemma GrowsPolynomially.of_isTheta {f g : ℝ → ℝ} (hg : GrowsPolynomially g have h_ub_pos : 0 < c₂ * c₄ * c₁⁻¹ := by positivity refine ⟨c₁ * c₂⁻¹ * c₃, h_lb_pos, ?_⟩ refine ⟨c₂ * c₄ * c₁⁻¹, h_ub_pos, ?_⟩ - have c₂_cancel : c₂⁻¹ * c₂ = 1 := inv_mul_cancel (by positivity) - have c₁_cancel : c₁⁻¹ * c₁ = 1 := inv_mul_cancel (by positivity) + have c₂_cancel : c₂⁻¹ * c₂ = 1 := inv_mul_cancel₀ (by positivity) + have c₁_cancel : c₁⁻¹ * c₁ = 1 := inv_mul_cancel₀ (by positivity) filter_upwards [(tendsto_id.const_mul_atTop hb_pos).eventually_forall_ge_atTop hf', (tendsto_id.const_mul_atTop hb_pos).eventually_forall_ge_atTop hf_lb, (tendsto_id.const_mul_atTop hb_pos).eventually_forall_ge_atTop hf_ub, diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index a24c1a1d55375..2b1a91b9ab87f 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -334,7 +334,7 @@ instance addCommGroup : AddCommGroup ℂ := zero_add := by intros; ext <;> simp add_zero := by intros; ext <;> simp add_comm := by intros; ext <;> simp [add_comm] - add_left_neg := by intros; ext <;> simp } + neg_add_cancel := by intros; ext <;> simp } instance addGroupWithOne : AddGroupWithOne ℂ := @@ -678,7 +678,7 @@ protected theorem inv_zero : (0⁻¹ : ℂ) = 0 := by rw [← ofReal_zero, ← ofReal_inv, inv_zero] protected theorem mul_inv_cancel {z : ℂ} (h : z ≠ 0) : z * z⁻¹ = 1 := by - rw [inv_def, ← mul_assoc, mul_conj, ← ofReal_mul, mul_inv_cancel (mt normSq_eq_zero.1 h), + rw [inv_def, ← mul_assoc, mul_conj, ← ofReal_mul, mul_inv_cancel₀ (mt normSq_eq_zero.1 h), ofReal_one] noncomputable instance instDivInvMonoid : DivInvMonoid ℂ where diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 7f358613f520b..9e9f187d59df4 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -177,7 +177,7 @@ theorem exp_add : exp (x + y) = exp x * exp y := by rw [← h₂, Nat.cast_mul, Nat.cast_mul, mul_inv, mul_inv] simp only [mul_left_comm (m.choose I : ℂ), mul_assoc, mul_left_comm (m.choose I : ℂ)⁻¹, mul_comm (m.choose I : ℂ)] - rw [inv_mul_cancel h₁] + rw [inv_mul_cancel₀ h₁] simp [div_eq_mul_inv, mul_comm, mul_assoc, mul_left_comm] simp_rw [exp, exp', lim_mul_lim] apply (lim_eq_lim_of_equiv _).symm @@ -210,10 +210,10 @@ theorem exp_nat_mul (x : ℂ) : ∀ n : ℕ, exp (n * x) = exp x ^ n @[simp] theorem exp_ne_zero : exp x ≠ 0 := fun h => - zero_ne_one <| by rw [← exp_zero, ← add_neg_self x, exp_add, h]; simp + zero_ne_one <| by rw [← exp_zero, ← add_neg_cancel x, exp_add, h]; simp theorem exp_neg : exp (-x) = (exp x)⁻¹ := by - rw [← mul_right_inj' (exp_ne_zero x), ← exp_add]; simp [mul_inv_cancel (exp_ne_zero x)] + rw [← mul_right_inj' (exp_ne_zero x), ← exp_add]; simp [mul_inv_cancel₀ (exp_ne_zero x)] theorem exp_sub : exp (x - y) = exp x / exp y := by simp [sub_eq_add_neg, exp_add, exp_neg, div_eq_mul_inv] @@ -377,7 +377,7 @@ theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_ @[simp] theorem cosh_sq_sub_sinh_sq : cosh x ^ 2 - sinh x ^ 2 = 1 := by - rw [sq_sub_sq, cosh_add_sinh, cosh_sub_sinh, ← exp_add, add_neg_self, exp_zero] + rw [sq_sub_sq, cosh_add_sinh, cosh_sub_sinh, ← exp_add, add_neg_cancel, exp_zero] theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by rw [← cosh_sq_sub_sinh_sq x] @@ -1081,7 +1081,7 @@ theorem sum_div_factorial_le {α : Type*} [LinearOrderedField α] (n j : ℕ) (h have h₄ : (n.succ - 1 : α) = n := by simp rw [geom_sum_inv h₁ h₂, eq_div_iff_mul_eq h₃, mul_comm _ (n.factorial * n : α), ← mul_assoc (n.factorial⁻¹ : α), ← mul_inv_rev, h₄, ← mul_assoc (n.factorial * n : α), - mul_comm (n : α) n.factorial, mul_inv_cancel h₃, one_mul, mul_comm] + mul_comm (n : α) n.factorial, mul_inv_cancel₀ h₃, one_mul, mul_comm] _ ≤ n.succ / (n.factorial * n : α) := by gcongr; apply sub_le_self; positivity theorem exp_bound {x : ℂ} (hx : abs x ≤ 1) {n : ℕ} (hn : 0 < n) : diff --git a/Mathlib/Data/DFinsupp/NeLocus.lean b/Mathlib/Data/DFinsupp/NeLocus.lean index ff0984c562220..7764ffc564385 100644 --- a/Mathlib/Data/DFinsupp/NeLocus.lean +++ b/Mathlib/Data/DFinsupp/NeLocus.lean @@ -124,7 +124,7 @@ theorem neLocus_neg_neg : neLocus (-f) (-g) = f.neLocus g := theorem neLocus_neg : neLocus (-f) g = f.neLocus (-g) := by rw [← neLocus_neg_neg, neg_neg] theorem neLocus_eq_support_sub : f.neLocus g = (f - g).support := by - rw [← @neLocus_add_right α N _ _ _ _ _ (-g), add_right_neg, neLocus_zero_right, sub_eq_add_neg] + rw [← @neLocus_add_right α N _ _ _ _ _ (-g), add_neg_cancel, neLocus_zero_right, sub_eq_add_neg] @[simp] theorem neLocus_sub_left : neLocus (f - g₁) (f - g₂) = neLocus g₁ g₂ := by diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 4115d63979f9a..f32953b91a158 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -55,7 +55,7 @@ theorem coe_inv_le : (↑r⁻¹ : ℝ≥0∞) ≤ (↑r)⁻¹ := @[simp, norm_cast] theorem coe_inv (hr : r ≠ 0) : (↑r⁻¹ : ℝ≥0∞) = (↑r)⁻¹ := - coe_inv_le.antisymm <| sInf_le <| mem_setOf.2 <| by rw [← coe_mul, mul_inv_cancel hr, coe_one] + coe_inv_le.antisymm <| sInf_le <| mem_setOf.2 <| by rw [← coe_mul, mul_inv_cancel₀ hr, coe_one] @[norm_cast] theorem coe_inv_two : ((2⁻¹ : ℝ≥0) : ℝ≥0∞) = 2⁻¹ := by rw [coe_inv _root_.two_ne_zero, coe_two] @@ -86,7 +86,7 @@ protected theorem inv_pow : ∀ {a : ℝ≥0∞} {n : ℕ}, (a ^ n)⁻¹ = a⁻ protected theorem mul_inv_cancel (h0 : a ≠ 0) (ht : a ≠ ∞) : a * a⁻¹ = 1 := by lift a to ℝ≥0 using ht norm_cast at h0; norm_cast - exact mul_inv_cancel h0 + exact mul_inv_cancel₀ h0 protected theorem inv_mul_cancel (h0 : a ≠ 0) (ht : a ≠ ∞) : a⁻¹ * a = 1 := mul_comm a a⁻¹ ▸ ENNReal.mul_inv_cancel h0 ht diff --git a/Mathlib/Data/ENNReal/Operations.lean b/Mathlib/Data/ENNReal/Operations.lean index bcc3076cd29db..9f5df7b047487 100644 --- a/Mathlib/Data/ENNReal/Operations.lean +++ b/Mathlib/Data/ENNReal/Operations.lean @@ -65,7 +65,7 @@ theorem mul_left_strictMono (h0 : a ≠ 0) (hinf : a ≠ ∞) : StrictMono (a * rw [coe_ne_zero] at h0 intro x y h contrapose! h - simpa only [← mul_assoc, ← coe_mul, inv_mul_cancel h0, coe_one, one_mul] + simpa only [← mul_assoc, ← coe_mul, inv_mul_cancel₀ h0, coe_one, one_mul] using mul_le_mul_left' h (↑a⁻¹) @[gcongr] protected theorem mul_lt_mul_left' (h0 : a ≠ 0) (hinf : a ≠ ⊤) (bc : b < c) : diff --git a/Mathlib/Data/Finsupp/NeLocus.lean b/Mathlib/Data/Finsupp/NeLocus.lean index 0bdacdc4fd4d0..99397666dda77 100644 --- a/Mathlib/Data/Finsupp/NeLocus.lean +++ b/Mathlib/Data/Finsupp/NeLocus.lean @@ -123,7 +123,7 @@ theorem neLocus_neg_neg : neLocus (-f) (-g) = f.neLocus g := theorem neLocus_neg : neLocus (-f) g = f.neLocus (-g) := by rw [← neLocus_neg_neg, neg_neg] theorem neLocus_eq_support_sub : f.neLocus g = (f - g).support := by - rw [← neLocus_add_right _ _ (-g), add_right_neg, neLocus_zero_right, sub_eq_add_neg] + rw [← neLocus_add_right _ _ (-g), add_neg_cancel, neLocus_zero_right, sub_eq_add_neg] @[simp] theorem neLocus_sub_left : neLocus (f - g₁) (f - g₂) = neLocus g₁ g₂ := by diff --git a/Mathlib/Data/Matrix/Rank.lean b/Mathlib/Data/Matrix/Rank.lean index 9c78955d6091f..8ae0a21d0de4a 100644 --- a/Mathlib/Data/Matrix/Rank.lean +++ b/Mathlib/Data/Matrix/Rank.lean @@ -76,7 +76,7 @@ theorem rank_unit [StrongRankCondition R] [DecidableEq n] (A : (Matrix n n R)ˣ) (A : Matrix n n R).rank = Fintype.card n := by apply le_antisymm (rank_le_card_width (A : Matrix n n R)) _ have := rank_mul_le_left (A : Matrix n n R) (↑A⁻¹ : Matrix n n R) - rwa [← Units.val_mul, mul_inv_self, Units.val_one, rank_one] at this + rwa [← Units.val_mul, mul_inv_cancel, Units.val_one, rank_one] at this theorem rank_of_isUnit [StrongRankCondition R] [DecidableEq n] (A : Matrix n n R) (h : IsUnit A) : A.rank = Fintype.card n := by diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index 26fe515738bcf..18de3f526e468 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -798,22 +798,22 @@ section Inv @[simp] theorem inv_le {r p : ℝ≥0} (h : r ≠ 0) : r⁻¹ ≤ p ↔ 1 ≤ r * p := by - rw [← mul_le_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel h] + rw [← mul_le_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel₀ h] theorem inv_le_of_le_mul {r p : ℝ≥0} (h : 1 ≤ r * p) : r⁻¹ ≤ p := by by_cases r = 0 <;> simp [*, inv_le] @[simp] theorem le_inv_iff_mul_le {r p : ℝ≥0} (h : p ≠ 0) : r ≤ p⁻¹ ↔ r * p ≤ 1 := by - rw [← mul_le_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel h, mul_comm] + rw [← mul_le_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel₀ h, mul_comm] @[simp] theorem lt_inv_iff_mul_lt {r p : ℝ≥0} (h : p ≠ 0) : r < p⁻¹ ↔ r * p < 1 := by - rw [← mul_lt_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel h, mul_comm] + rw [← mul_lt_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel₀ h, mul_comm] theorem mul_le_iff_le_inv {a b r : ℝ≥0} (hr : r ≠ 0) : r * a ≤ b ↔ a ≤ r⁻¹ * b := by have : 0 < r := lt_of_le_of_ne (zero_le r) hr.symm - rw [← mul_le_mul_left (inv_pos.mpr this), ← mul_assoc, inv_mul_cancel hr, one_mul] + rw [← mul_le_mul_left (inv_pos.mpr this), ← mul_assoc, inv_mul_cancel₀ hr, one_mul] theorem le_div_iff_mul_le {a b r : ℝ≥0} (hr : r ≠ 0) : a ≤ b / r ↔ a * r ≤ b := le_div_iff₀ hr @@ -865,7 +865,7 @@ theorem le_of_forall_lt_one_mul_le {x y : ℝ≥0} (h : ∀ a < 1, a * x ≤ y) have hx' : x⁻¹ ≠ 0 := by rwa [Ne, inv_eq_zero] have : a * x⁻¹ < 1 := by rwa [← lt_inv_iff_mul_lt hx', inv_inv] have : a * x⁻¹ * x ≤ y := h _ this - rwa [mul_assoc, inv_mul_cancel hx, mul_one] at this + rwa [mul_assoc, inv_mul_cancel₀ hx, mul_one] at this nonrec theorem half_le_self (a : ℝ≥0) : a / 2 ≤ a := half_le_self bot_le diff --git a/Mathlib/Data/Nat/Choose/Sum.lean b/Mathlib/Data/Nat/Choose/Sum.lean index a04f3438b28c7..ac70377042b48 100644 --- a/Mathlib/Data/Nat/Choose/Sum.lean +++ b/Mathlib/Data/Nat/Choose/Sum.lean @@ -136,7 +136,7 @@ theorem Int.alternating_sum_range_choose {n : ℕ} : | zero => simp | succ n => have h := add_pow (-1 : ℤ) 1 n.succ - simp only [one_pow, mul_one, add_left_neg] at h + simp only [one_pow, mul_one, neg_add_cancel] at h rw [← h, zero_pow n.succ_ne_zero, if_neg (Nat.succ_ne_zero n)] theorem Int.alternating_sum_range_choose_of_ne {n : ℕ} (h0 : n ≠ 0) : diff --git a/Mathlib/Data/Nat/Totient.lean b/Mathlib/Data/Nat/Totient.lean index 4353ad9cbb284..3bfd3d40e7e5c 100644 --- a/Mathlib/Data/Nat/Totient.lean +++ b/Mathlib/Data/Nat/Totient.lean @@ -296,7 +296,7 @@ theorem totient_eq_mul_prod_factors (n : ℕ) : refine prod_congr rfl fun p hp => ?_ have hp := pos_of_mem_primeFactorsList (List.mem_toFinset.mp hp) have hp' : (p : ℚ) ≠ 0 := cast_ne_zero.mpr hp.ne.symm - rw [sub_mul, one_mul, mul_comm, mul_inv_cancel hp', cast_pred hp] + rw [sub_mul, one_mul, mul_comm, mul_inv_cancel₀ hp', cast_pred hp] theorem totient_gcd_mul_totient_mul (a b : ℕ) : φ (a.gcd b) * φ (a * b) = φ a * φ b * a.gcd b := by have shuffle : diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index ac22db399e145..54d7f68a2f02a 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -1277,7 +1277,7 @@ instance addCommGroup : AddCommGroup ZNum := add_comm := by transfer neg := Neg.neg zsmul := zsmulRec - add_left_neg := by transfer } + neg_add_cancel := by transfer } instance addMonoidWithOne : AddMonoidWithOne ZNum := { ZNum.addMonoid with diff --git a/Mathlib/Data/Rat/Cast/Defs.lean b/Mathlib/Data/Rat/Cast/Defs.lean index b9599b8be7cc3..e9e86730c7e10 100644 --- a/Mathlib/Data/Rat/Cast/Defs.lean +++ b/Mathlib/Data/Rat/Cast/Defs.lean @@ -159,7 +159,7 @@ lemma cast_divInt_of_ne_zero (a : ℤ) {b : ℤ} (b0 : (b : α) ≠ 0) : (a /. b ((divInt_eq_iff b0' <| ne_of_gt <| Int.natCast_pos.2 h.bot_lt).1 e) rw [Int.cast_mul, Int.cast_mul, Int.cast_natCast] at this rw [eq_comm, cast_def, div_eq_mul_inv, eq_div_iff_mul_eq d0, mul_assoc, (d.commute_cast _).eq, - ← mul_assoc, this, mul_assoc, mul_inv_cancel b0, mul_one] + ← mul_assoc, this, mul_assoc, mul_inv_cancel₀ b0, mul_one] @[norm_cast] lemma cast_mkRat_of_ne_zero (a : ℤ) {b : ℕ} (hb : (b : α) ≠ 0) : (mkRat a b : α) = a / b := by diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index 1057e0b57a1f5..196de7dc71c26 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -268,7 +268,7 @@ protected theorem add_assoc : a + b + c = a + (b + c) := congr 2 ac_rfl -protected lemma add_left_neg : -a + a = 0 := by +protected lemma neg_add_cancel : -a + a = 0 := by simp [add_def, normalize_eq_mkRat, Int.neg_mul, Int.add_comm, ← Int.sub_eq_add_neg] @[deprecated zero_divInt (since := "2024-03-18")] @@ -327,7 +327,7 @@ instance addCommGroup : AddCommGroup ℚ where add_zero := Rat.add_zero add_comm := Rat.add_comm add_assoc := Rat.add_assoc - add_left_neg := Rat.add_left_neg + neg_add_cancel := Rat.neg_add_cancel sub_eq_add_neg := Rat.sub_eq_add_neg nsmul := nsmulRec zsmul := zsmulRec diff --git a/Mathlib/Data/Rat/Star.lean b/Mathlib/Data/Rat/Star.lean index 69ea7bbd03e38..a09ca020eab6d 100644 --- a/Mathlib/Data/Rat/Star.lean +++ b/Mathlib/Data/Rat/Star.lean @@ -30,7 +30,7 @@ namespace NNRat exact nsmul_mem (subset_closure <| mem_range_self _) _ rw [nsmul_eq_mul] push_cast - rw [mul_assoc, pow_sub₀, pow_one, mul_right_comm, ← mul_pow, mul_inv_cancel, one_pow, one_mul, + rw [mul_assoc, pow_sub₀, pow_one, mul_right_comm, ← mul_pow, mul_inv_cancel₀, one_pow, one_mul, ← div_eq_mul_inv, num_div_den] all_goals simp [x.den_pos.ne', Nat.one_le_iff_ne_zero, *] diff --git a/Mathlib/Data/Real/Archimedean.lean b/Mathlib/Data/Real/Archimedean.lean index 1d868592b239f..9c2e636a07e34 100644 --- a/Mathlib/Data/Real/Archimedean.lean +++ b/Mathlib/Data/Real/Archimedean.lean @@ -67,7 +67,7 @@ theorem exists_isLUB {S : Set ℝ} (hne : S.Nonempty) (hbdd : BddAbove S) : ∃ intro n n0 y yS have := (Int.sub_one_lt_floor _).trans_le (Int.cast_le.2 <| (hf n).2 _ ⟨y, yS, Int.floor_le _⟩) simp only [Rat.cast_div, Rat.cast_intCast, Rat.cast_natCast, gt_iff_lt] - rwa [lt_div_iff (Nat.cast_pos.2 n0 : (_ : ℝ) < _), sub_mul, _root_.inv_mul_cancel] + rwa [lt_div_iff (Nat.cast_pos.2 n0 : (_ : ℝ) < _), sub_mul, inv_mul_cancel₀] exact ne_of_gt (Nat.cast_pos.2 n0) have hg : IsCauSeq abs (fun n => f n / n : ℕ → ℚ) := by intro ε ε0 diff --git a/Mathlib/Data/Real/Basic.lean b/Mathlib/Data/Real/Basic.lean index 09ea95dad271d..a78b3dc25b856 100644 --- a/Mathlib/Data/Real/Basic.lean +++ b/Mathlib/Data/Real/Basic.lean @@ -189,7 +189,7 @@ instance commRing : CommRing ℝ where mul_assoc a b c := by apply ext_cauchy; simp only [cauchy_mul, mul_assoc] left_distrib a b c := by apply ext_cauchy; simp only [cauchy_add, cauchy_mul, mul_add] right_distrib a b c := by apply ext_cauchy; simp only [cauchy_add, cauchy_mul, add_mul] - add_left_neg a := by apply ext_cauchy; simp [cauchy_add, cauchy_neg, cauchy_zero] + neg_add_cancel a := by apply ext_cauchy; simp [cauchy_add, cauchy_neg, cauchy_zero] natCast_zero := by apply ext_cauchy; simp [cauchy_zero] natCast_succ n := by apply ext_cauchy; simp [cauchy_one, cauchy_add] intCast_negSucc z := by apply ext_cauchy; simp [cauchy_neg, cauchy_natCast] diff --git a/Mathlib/Data/Real/GoldenRatio.lean b/Mathlib/Data/Real/GoldenRatio.lean index 6c52a2323af45..63905e911911d 100644 --- a/Mathlib/Data/Real/GoldenRatio.lean +++ b/Mathlib/Data/Real/GoldenRatio.lean @@ -189,7 +189,7 @@ theorem Real.coe_fib_eq' : · simp · simp only [goldenRatio, goldenConj] ring_nf - rw [mul_inv_cancel]; norm_num + rw [mul_inv_cancel₀]; norm_num · exact fib_isSol_fibRec · -- Porting note: Rewrote this proof suffices LinearRecurrence.IsSolution fibRec @@ -211,7 +211,7 @@ theorem fib_golden_conj_exp (n : ℕ) : Nat.fib (n + 1) - φ * Nat.fib n = ψ ^ rw [mul_div, div_sub_div_same, mul_sub, ← pow_succ'] ring_nf have nz : sqrt 5 ≠ 0 := by norm_num - rw [← (mul_inv_cancel nz).symm, one_mul] + rw [← (mul_inv_cancel₀ nz).symm, one_mul] /-- Relationship between the Fibonacci Sequence, Golden Ratio and its exponents --/ theorem fib_golden_exp' (n : ℕ) : φ * Nat.fib (n + 1) + Nat.fib n = φ ^ (n + 1) := by diff --git a/Mathlib/Data/Real/Hyperreal.lean b/Mathlib/Data/Real/Hyperreal.lean index 65ccac0d2bc60..3ec6b4835c02e 100644 --- a/Mathlib/Data/Real/Hyperreal.lean +++ b/Mathlib/Data/Real/Hyperreal.lean @@ -159,7 +159,7 @@ theorem omega_ne_zero : ω ≠ 0 := omega_pos.ne' theorem epsilon_mul_omega : ε * ω = 1 := - @inv_mul_cancel _ _ ω omega_ne_zero + @inv_mul_cancel₀ _ _ ω omega_ne_zero theorem lt_of_tendsto_zero_of_pos {f : ℕ → ℝ} (hf : Tendsto f atTop (𝓝 0)) : ∀ {r : ℝ}, 0 < r → ofSeq f < (r : ℝ*) := fun hr ↦ diff --git a/Mathlib/Data/Real/Irrational.lean b/Mathlib/Data/Real/Irrational.lean index f056a12a6ab0c..3a8c9b0473110 100644 --- a/Mathlib/Data/Real/Irrational.lean +++ b/Mathlib/Data/Real/Irrational.lean @@ -328,7 +328,7 @@ theorem of_mul_rat (h : Irrational (x * q)) : Irrational x := h.mul_cases.resolve_right q.not_irrational theorem mul_rat (h : Irrational x) {q : ℚ} (hq : q ≠ 0) : Irrational (x * q) := - of_mul_rat q⁻¹ <| by rwa [mul_assoc, ← cast_mul, mul_inv_cancel hq, cast_one, mul_one] + of_mul_rat q⁻¹ <| by rwa [mul_assoc, ← cast_mul, mul_inv_cancel₀ hq, cast_one, mul_one] theorem of_rat_mul : Irrational (q * x) → Irrational x := mul_comm x q ▸ of_mul_rat q diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index dfdf5727a1b17..6c696dc0794a1 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -843,7 +843,7 @@ theorem val_coe_unit_coprime {n : ℕ} (u : (ZMod n)ˣ) : Nat.Coprime (u : ZMod cases' n with n · rcases Int.units_eq_one_or u with (rfl | rfl) <;> simp apply Nat.coprime_of_mul_modEq_one ((u⁻¹ : Units (ZMod (n + 1))) : ZMod (n + 1)).val - have := Units.ext_iff.1 (mul_right_inv u) + have := Units.ext_iff.1 (mul_inv_cancel u) rw [Units.val_one] at this rw [← eq_iff_modEq_nat, Nat.cast_one, ← this]; clear this rw [← natCast_zmod_val ((u * u⁻¹ : Units (ZMod (n + 1))) : ZMod (n + 1))] @@ -1019,7 +1019,7 @@ theorem neg_val' {n : ℕ} [NeZero n] (a : ZMod n) : (-a).val = (n - a.val) % n _ = (n - val a) % n := Nat.ModEq.add_right_cancel' (val a) (by - rw [Nat.ModEq, ← val_add, add_left_neg, tsub_add_cancel_of_le a.val_le, Nat.mod_self, + rw [Nat.ModEq, ← val_add, neg_add_cancel, tsub_add_cancel_of_le a.val_le, Nat.mod_self, val_zero]) theorem neg_val {n : ℕ} [NeZero n] (a : ZMod n) : (-a).val = if a = 0 then 0 else n - a.val := by diff --git a/Mathlib/Data/ZMod/Defs.lean b/Mathlib/Data/ZMod/Defs.lean index 9ad1cacc2fb16..324dec3252219 100644 --- a/Mathlib/Data/ZMod/Defs.lean +++ b/Mathlib/Data/ZMod/Defs.lean @@ -147,7 +147,7 @@ instance commRing (n : ℕ) : CommRing (ZMod n) where (inferInstanceAs (CommRing ℤ)).nsmul_succ fun n => (inferInstanceAs (CommRing (Fin n.succ))).nsmul_succ -- Porting note: `match` didn't work here - add_left_neg := Nat.casesOn n (@add_left_neg Int _) fun n => @add_left_neg (Fin n.succ) _ + neg_add_cancel := Nat.casesOn n (@neg_add_cancel Int _) fun n => @neg_add_cancel (Fin n.succ) _ add_comm := Nat.casesOn n (@add_comm Int _) fun n => @add_comm (Fin n.succ) _ mul := Nat.casesOn n (@Mul.mul Int _) fun n => @Mul.mul (Fin n.succ) _ mul_assoc := Nat.casesOn n (@mul_assoc Int _) fun n => @mul_assoc (Fin n.succ) _ diff --git a/Mathlib/Deprecated/Group.lean b/Mathlib/Deprecated/Group.lean index 20326690cf9f3..bd990fbce0326 100644 --- a/Mathlib/Deprecated/Group.lean +++ b/Mathlib/Deprecated/Group.lean @@ -246,7 +246,7 @@ theorem map_one : f 1 = 1 := /-- A group homomorphism sends inverses to inverses. -/ @[to_additive "An additive group homomorphism sends negations to negations."] theorem map_inv (hf : IsGroupHom f) (a : α) : f a⁻¹ = (f a)⁻¹ := - eq_inv_of_mul_eq_one_left <| by rw [← hf.map_mul, inv_mul_self, hf.map_one] + eq_inv_of_mul_eq_one_left <| by rw [← hf.map_mul, inv_mul_cancel, hf.map_one] @[to_additive] theorem map_div (hf : IsGroupHom f) (a b : α) : f (a / b) = f a / f b := by diff --git a/Mathlib/Deprecated/Subgroup.lean b/Mathlib/Deprecated/Subgroup.lean index 5d229f6e86fd0..566783c7bb5b5 100644 --- a/Mathlib/Deprecated/Subgroup.lean +++ b/Mathlib/Deprecated/Subgroup.lean @@ -150,7 +150,7 @@ structure IsNormalSubgroup [Group G] (s : Set G) extends IsSubgroup s : Prop whe @[to_additive] theorem isNormalSubgroup_of_commGroup [CommGroup G] {s : Set G} (hs : IsSubgroup s) : IsNormalSubgroup s := - { hs with normal := fun n hn g => by rwa [mul_right_comm, mul_right_inv, one_mul] } + { hs with normal := fun n hn g => by rwa [mul_right_comm, mul_inv_cancel, one_mul] } theorem Additive.isNormalAddSubgroup [Group G] {s : Set G} (hs : IsNormalSubgroup s) : @IsNormalAddSubgroup (Additive G) _ s := @@ -290,13 +290,13 @@ theorem one_ker_inv' {f : G → H} (hf : IsGroupHom f) {a b : G} (h : f (a⁻¹ @[to_additive] theorem inv_ker_one {f : G → H} (hf : IsGroupHom f) {a b : G} (h : f a = f b) : f (a * b⁻¹) = 1 := by - have : f a * (f b)⁻¹ = 1 := by rw [h, mul_right_inv] + have : f a * (f b)⁻¹ = 1 := by rw [h, mul_inv_cancel] rwa [← hf.map_inv, ← hf.map_mul] at this @[to_additive] theorem inv_ker_one' {f : G → H} (hf : IsGroupHom f) {a b : G} (h : f a = f b) : f (a⁻¹ * b) = 1 := by - have : (f a)⁻¹ * f b = 1 := by rw [h, mul_left_inv] + have : (f a)⁻¹ * f b = 1 := by rw [h, inv_mul_cancel] rwa [← hf.map_inv, ← hf.map_mul] at this @[to_additive] diff --git a/Mathlib/Deprecated/Subring.lean b/Mathlib/Deprecated/Subring.lean index 4663dfe6fd00b..6a652a4c798a3 100644 --- a/Mathlib/Deprecated/Subring.lean +++ b/Mathlib/Deprecated/Subring.lean @@ -115,7 +115,7 @@ theorem exists_list_of_mem_closure {a : R} (h : a ∈ closure s) : protected theorem InClosure.recOn {C : R → Prop} {x : R} (hx : x ∈ closure s) (h1 : C 1) (hneg1 : C (-1)) (hs : ∀ z ∈ s, ∀ n, C n → C (z * n)) (ha : ∀ {x y}, C x → C y → C (x + y)) : C x := by - have h0 : C 0 := add_neg_self (1 : R) ▸ ha h1 hneg1 + have h0 : C 0 := add_neg_cancel (1 : R) ▸ ha h1 hneg1 rcases exists_list_of_mem_closure hx with ⟨L, HL, rfl⟩ clear hx induction' L with hd tl ih diff --git a/Mathlib/Dynamics/Flow.lean b/Mathlib/Dynamics/Flow.lean index 2a75ae3171516..bc7748dcf6949 100644 --- a/Mathlib/Dynamics/Flow.lean +++ b/Mathlib/Dynamics/Flow.lean @@ -165,8 +165,8 @@ theorem continuous_toFun (t : τ) : Continuous (ϕ.toFun t) := by def toHomeomorph (t : τ) : (α ≃ₜ α) where toFun := ϕ t invFun := ϕ (-t) - left_inv x := by rw [← map_add, neg_add_self, map_zero_apply] - right_inv x := by rw [← map_add, add_neg_self, map_zero_apply] + left_inv x := by rw [← map_add, neg_add_cancel, map_zero_apply] + right_inv x := by rw [← map_add, add_neg_cancel, map_zero_apply] theorem image_eq_preimage (t : τ) (s : Set α) : ϕ t '' s = ϕ (-t) ⁻¹' s := (ϕ.toHomeomorph t).toEquiv.image_eq_preimage s diff --git a/Mathlib/Dynamics/Newton.lean b/Mathlib/Dynamics/Newton.lean index 7783cb5d36d2c..94030f064065f 100644 --- a/Mathlib/Dynamics/Newton.lean +++ b/Mathlib/Dynamics/Newton.lean @@ -94,7 +94,7 @@ theorem aeval_pow_two_pow_dvd_aeval_iterate_newtonMap isUnit_aeval_of_isUnit_aeval_of_isNilpotent_sub h' <| isNilpotent_iterate_newtonMap_sub_of_isNilpotent h n rw [derivative_map, eval_map_algebraMap, ← mul_assoc, mul_neg, Ring.mul_inverse_cancel _ this, - neg_mul, one_mul, add_right_neg] + neg_mul, one_mul, add_neg_cancel] · rw [neg_mul, even_two.neg_pow, mul_pow, pow_succ, pow_mul] exact dvd_mul_of_dvd_right (pow_dvd_pow_of_dvd ih 2) _ diff --git a/Mathlib/FieldTheory/AbelRuffini.lean b/Mathlib/FieldTheory/AbelRuffini.lean index ab3b3dba2d9ba..4df8eb8e0f91f 100644 --- a/Mathlib/FieldTheory/AbelRuffini.lean +++ b/Mathlib/FieldTheory/AbelRuffini.lean @@ -154,7 +154,7 @@ theorem splits_X_pow_sub_one_of_X_pow_sub_C {F : Type*} [Field F] {E : Type*} [F apply @splits_of_exists_multiset F E _ _ i (X ^ n - 1) (s.map fun c : E => c / b) rw [leadingCoeff_X_pow_sub_one hn', RingHom.map_one, C_1, one_mul, Multiset.map_map] have C_mul_C : C (i a⁻¹) * C (i a) = 1 := by - rw [← C_mul, ← i.map_mul, inv_mul_cancel ha, i.map_one, C_1] + rw [← C_mul, ← i.map_mul, inv_mul_cancel₀ ha, i.map_one, C_1] have key1 : (X ^ n - 1 : F[X]).map i = C (i a⁻¹) * ((X ^ n - C a).map i).comp (C b * X) := by rw [Polynomial.map_sub, Polynomial.map_sub, Polynomial.map_pow, map_X, map_C, Polynomial.map_one, sub_comp, pow_comp, X_comp, C_comp, mul_pow, ← C_pow, hb, mul_sub, ← diff --git a/Mathlib/FieldTheory/Finite/Basic.lean b/Mathlib/FieldTheory/Finite/Basic.lean index a99d07385ebeb..dad50e120b86c 100644 --- a/Mathlib/FieldTheory/Finite/Basic.lean +++ b/Mathlib/FieldTheory/Finite/Basic.lean @@ -78,7 +78,7 @@ theorem exists_root_sum_quadratic [Fintype R] {f g : R[X]} (hf2 : degree f = 2) simp only [disjoint_left, mem_image] at this push_neg at this rcases this with ⟨x, ⟨a, _, ha⟩, ⟨b, _, hb⟩⟩ - exact ⟨a, b, by rw [ha, ← hb, eval_neg, neg_add_self]⟩ + exact ⟨a, b, by rw [ha, ← hb, eval_neg, neg_add_cancel]⟩ fun hd : Disjoint _ _ => lt_irrefl (2 * ((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)).card) <| calc 2 * ((univ.image fun x : R => eval x f) ∪ univ.image fun x : R => eval x (-g)).card diff --git a/Mathlib/FieldTheory/Minpoly/Field.lean b/Mathlib/FieldTheory/Minpoly/Field.lean index d467329227d96..32d3a549998d0 100644 --- a/Mathlib/FieldTheory/Minpoly/Field.lean +++ b/Mathlib/FieldTheory/Minpoly/Field.lean @@ -123,9 +123,9 @@ theorem eq_of_irreducible [Nontrivial B] {p : A[X]} (hp1 : Irreducible p) have : p.leadingCoeff ≠ 0 := leadingCoeff_ne_zero.mpr hp1.ne_zero apply eq_of_irreducible_of_monic · exact Associated.irreducible ⟨⟨C p.leadingCoeff⁻¹, C p.leadingCoeff, - by rwa [← C_mul, inv_mul_cancel, C_1], by rwa [← C_mul, mul_inv_cancel, C_1]⟩, rfl⟩ hp1 + by rwa [← C_mul, inv_mul_cancel₀, C_1], by rwa [← C_mul, mul_inv_cancel₀, C_1]⟩, rfl⟩ hp1 · rw [aeval_mul, hp2, zero_mul] - · rwa [Polynomial.Monic, leadingCoeff_mul, leadingCoeff_C, mul_inv_cancel] + · rwa [Polynomial.Monic, leadingCoeff_mul, leadingCoeff_C, mul_inv_cancel₀] theorem add_algebraMap {B : Type*} [CommRing B] [Algebra A B] {x : B} (hx : IsIntegral A x) (a : A) : minpoly A (x + algebraMap A B a) = (minpoly A x).comp (X - C a) := by diff --git a/Mathlib/FieldTheory/MvPolynomial.lean b/Mathlib/FieldTheory/MvPolynomial.lean index fb3b4bacc303b..06bf5382869b6 100644 --- a/Mathlib/FieldTheory/MvPolynomial.lean +++ b/Mathlib/FieldTheory/MvPolynomial.lean @@ -33,7 +33,7 @@ theorem quotient_mk_comp_C_injective (I : Ideal (MvPolynomial σ K)) (hI : I ≠ rw [RingHom.comp_apply, Ideal.Quotient.eq_zero_iff_mem] at hx refine _root_.by_contradiction fun hx0 => absurd (I.eq_top_iff_one.2 ?_) hI have := I.mul_mem_left (MvPolynomial.C x⁻¹) hx - rwa [← MvPolynomial.C.map_mul, inv_mul_cancel hx0, MvPolynomial.C_1] at this + rwa [← MvPolynomial.C.map_mul, inv_mul_cancel₀ hx0, MvPolynomial.C_1] at this end MvPolynomial diff --git a/Mathlib/FieldTheory/PerfectClosure.lean b/Mathlib/FieldTheory/PerfectClosure.lean index 2a8d5e6a3662b..8572b59904f89 100644 --- a/Mathlib/FieldTheory/PerfectClosure.lean +++ b/Mathlib/FieldTheory/PerfectClosure.lean @@ -268,9 +268,9 @@ instance instAddCommGroup : AddCommGroup (PerfectClosure K p) := congr_arg (Quot.mk _) <| by simp only [iterate_map_zero, iterate_zero_apply, add_zero] sub_eq_add_neg := fun a b => rfl - add_left_neg := fun e => + neg_add_cancel := fun e => Quot.inductionOn e fun ⟨n, x⟩ => by - simp only [quot_mk_eq_mk, neg_mk, mk_add_mk, iterate_map_neg, add_left_neg, mk_zero] + simp only [quot_mk_eq_mk, neg_mk, mk_add_mk, iterate_map_neg, neg_add_cancel, mk_zero] add_comm := fun e f => Quot.inductionOn e fun ⟨m, x⟩ => Quot.inductionOn f fun ⟨n, y⟩ => congr_arg (Quot.mk _) <| by simp only [add_comm] @@ -488,7 +488,7 @@ instance instDivisionRing : DivisionRing (PerfectClosure K p) where rw [mk_inv, mk_mul_mk] refine (eq_iff K p _ _).2 ?_ simp only [iterate_map_one, iterate_map_zero, iterate_zero_apply, ← iterate_map_mul] at this ⊢ - rw [mul_inv_cancel this, iterate_map_one] + rw [mul_inv_cancel₀ this, iterate_map_one] inv_zero := congr_arg (Quot.mk (R K p)) (by rw [inv_zero]) nnqsmul := _ qsmul := _ diff --git a/Mathlib/FieldTheory/RatFunc/Basic.lean b/Mathlib/FieldTheory/RatFunc/Basic.lean index 884dd87c10a5c..323ae1661d472 100644 --- a/Mathlib/FieldTheory/RatFunc/Basic.lean +++ b/Mathlib/FieldTheory/RatFunc/Basic.lean @@ -167,7 +167,7 @@ theorem mul_inv_cancel : ∀ {p : RatFunc K}, p ≠ 0 → p * p⁻¹ = 1 have : p ≠ 0 := fun hp => h <| by rw [hp, ofFractionRing_zero] simpa only [← ofFractionRing_inv, ← ofFractionRing_mul, ← ofFractionRing_one, ofFractionRing.injEq] using -- Porting note: `ofFractionRing.injEq` was not present - _root_.mul_inv_cancel this + mul_inv_cancel₀ this end IsDomain @@ -258,7 +258,7 @@ macro "frac_tac" : tactic => `(tactic| repeat (rintro (⟨⟩ : RatFunc _)) <;> ← ofFractionRing_inv, add_assoc, zero_add, add_zero, mul_assoc, mul_zero, mul_one, mul_add, inv_zero, add_comm, add_left_comm, mul_comm, mul_left_comm, sub_eq_add_neg, div_eq_mul_inv, - add_mul, zero_mul, one_mul, neg_mul, mul_neg, add_right_neg]) + add_mul, zero_mul, one_mul, neg_mul, mul_neg, add_neg_cancel]) /-- Solve equations for `RatFunc K` by applying `RatFunc.induction_on`. -/ macro "smul_tac" : tactic => `(tactic| @@ -308,7 +308,7 @@ def instAddCommGroup : AddCommGroup (RatFunc K) where zero_add := by frac_tac add_zero := by frac_tac neg := Neg.neg - add_left_neg := by frac_tac + neg_add_cancel := by frac_tac sub := Sub.sub sub_eq_add_neg := by frac_tac nsmul := (· • ·) @@ -826,7 +826,7 @@ def numDenom (x : RatFunc K) : K[X] × K[X] := Polynomial.leadingCoeff_C, div_C_mul, div_C_mul, ← mul_assoc, ← Polynomial.C_mul, ← mul_assoc, ← Polynomial.C_mul] constructor <;> congr <;> - rw [inv_div, mul_comm, mul_div_assoc, ← mul_assoc, inv_inv, _root_.mul_inv_cancel ha', + rw [inv_div, mul_comm, mul_div_assoc, ← mul_assoc, inv_inv, mul_inv_cancel₀ ha', one_mul, inv_div]) @[simp] diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean index 611f7d373fa2f..0e765f0f53f1a 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean @@ -236,13 +236,13 @@ theorem two_zsmul_oangle_neg_self_right (x : V) : (2 : ℤ) • o.oangle x (-x) negated, results in 0. -/ @[simp] theorem oangle_add_oangle_rev_neg_left (x y : V) : o.oangle (-x) y + o.oangle (-y) x = 0 := by - rw [oangle_neg_left_eq_neg_right, oangle_rev, add_left_neg] + rw [oangle_neg_left_eq_neg_right, oangle_rev, neg_add_cancel] /-- Adding the angles between two vectors in each order, with the second vector in each angle negated, results in 0. -/ @[simp] theorem oangle_add_oangle_rev_neg_right (x y : V) : o.oangle x (-y) + o.oangle y (-x) = 0 := by - rw [o.oangle_rev (-x), oangle_neg_left_eq_neg_right, add_neg_self] + rw [o.oangle_rev (-x), oangle_neg_left_eq_neg_right, add_neg_cancel] /-- Multiplying the first vector passed to `oangle` by a positive real does not change the angle. -/ diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean index 9b74a140ec8b5..1bdfc60311837 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean @@ -302,8 +302,8 @@ theorem _root_.Wbtw.angle₂₁₃_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : Wbtw simp at hp₂p₁ replace hr0 := hr0.lt_of_ne hr0'.symm refine ⟨vsub_ne_zero.2 hp₂p₁, r⁻¹, inv_pos.2 hr0, ?_⟩ - rw [AffineMap.lineMap_apply, vadd_vsub_assoc, vsub_self, add_zero, smul_smul, inv_mul_cancel hr0', - one_smul] + rw [AffineMap.lineMap_apply, vadd_vsub_assoc, vsub_self, add_zero, smul_smul, + inv_mul_cancel₀ hr0', one_smul] /-- If the second of three points is strictly between the other two, the angle at the first point is zero. -/ @@ -351,7 +351,7 @@ theorem angle_eq_zero_iff_ne_and_wbtw {p₁ p₂ p₃ : P} : rintro ⟨hp₁p₂, r, hr0, hp₃p₂⟩ rcases le_or_lt 1 r with (hr1 | hr1) · refine Or.inl ⟨vsub_ne_zero.1 hp₁p₂, r⁻¹, ⟨(inv_pos.2 hr0).le, inv_le_one hr1⟩, ?_⟩ - rw [AffineMap.lineMap_apply, hp₃p₂, smul_smul, inv_mul_cancel hr0.ne.symm, one_smul, + rw [AffineMap.lineMap_apply, hp₃p₂, smul_smul, inv_mul_cancel₀ hr0.ne.symm, one_smul, vsub_vadd] · refine Or.inr ⟨?_, r, ⟨hr0.le, hr1.le⟩, ?_⟩ · rw [← @vsub_ne_zero V, hp₃p₂, smul_ne_zero_iff] diff --git a/Mathlib/Geometry/Euclidean/Basic.lean b/Mathlib/Geometry/Euclidean/Basic.lean index 666165ddc7558..61e9e1a2ede92 100644 --- a/Mathlib/Geometry/Euclidean/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Basic.lean @@ -122,7 +122,7 @@ theorem dist_smul_vadd_eq_dist {v : V} (p₁ p₂ : P) (hv : v ≠ 0) (r : ℝ) have hd : discrim ⟪v, v⟫ (2 * ⟪v, p₁ -ᵥ p₂⟫) 0 = 2 * ⟪v, p₁ -ᵥ p₂⟫ * (2 * ⟪v, p₁ -ᵥ p₂⟫) := by rw [discrim] ring - rw [quadratic_eq_zero_iff hvi hd, add_left_neg, zero_div, neg_mul_eq_neg_mul, ← + rw [quadratic_eq_zero_iff hvi hd, neg_add_cancel, zero_div, neg_mul_eq_neg_mul, ← mul_sub_right_distrib, sub_eq_add_neg, ← mul_two, mul_assoc, mul_div_assoc, mul_div_mul_left, mul_div_assoc] norm_num diff --git a/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean b/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean index 2df8aa07d4881..594cfdfdde353 100644 --- a/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean +++ b/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean @@ -135,7 +135,7 @@ instance group {G : Type*} [Group G] [TopologicalSpace G] [ChartedSpace H' G] [L Group C^∞⟮I, N; I', G⟯ := { SmoothMap.monoid with inv := fun f => ⟨fun x => (f x)⁻¹, f.smooth.inv⟩ - mul_left_inv := fun a => by ext; exact mul_left_inv _ + inv_mul_cancel := fun a => by ext; exact inv_mul_cancel _ div := fun f g => ⟨f / g, f.smooth.div g.smooth⟩ div_eq_mul_inv := fun f g => by ext; exact div_eq_mul_inv _ _ } diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index 8329c6f2f2b24..cd4fb9f824108 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -125,7 +125,7 @@ theorem stereoInvFunAux_mem (hv : ‖v‖ = 1) {w : E} (hw : w ∈ (ℝ ∙ v) have h₁ : (0 : ℝ) < ‖w‖ ^ 2 + 4 := by positivity suffices ‖(4 : ℝ) • w + (‖w‖ ^ 2 - 4) • v‖ = ‖w‖ ^ 2 + 4 by simp only [mem_sphere_zero_iff_norm, norm_smul, Real.norm_eq_abs, abs_inv, this, - abs_of_pos h₁, stereoInvFunAux_apply, inv_mul_cancel h₁.ne'] + abs_of_pos h₁, stereoInvFunAux_apply, inv_mul_cancel₀ h₁.ne'] suffices ‖(4 : ℝ) • w + (‖w‖ ^ 2 - 4) • v‖ ^ 2 = (‖w‖ ^ 2 + 4) ^ 2 by simpa [sq_eq_sq_iff_abs_eq_abs, abs_of_pos h₁] using this rw [Submodule.mem_orthogonal_singleton_iff_inner_left] at hw diff --git a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean index be3a04e05b8cf..fcd7714d974ef 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean @@ -84,14 +84,14 @@ theorem smoothSheafCommRing.isUnit_stalk_iff {x : M} convert RingHom.map_one _ apply Subtype.ext ext y - apply mul_inv_cancel + apply mul_inv_cancel₀ exact hVf y · rw [← map_mul] -- Qualified the name to avoid Lean not finding a `OneHomClass` #8386 convert RingHom.map_one _ apply Subtype.ext ext y - apply inv_mul_cancel + apply inv_mul_cancel₀ exact hVf y · intro y exact ((contDiffAt_inv _ (hVf y)).contMDiffAt).comp y diff --git a/Mathlib/GroupTheory/Complement.lean b/Mathlib/GroupTheory/Complement.lean index d7820cbb29123..b0b6b1c0ae21e 100644 --- a/Mathlib/GroupTheory/Complement.lean +++ b/Mathlib/GroupTheory/Complement.lean @@ -119,7 +119,7 @@ theorem isComplement_univ_left : IsComplement univ S ↔ ∃ g : G, S = {g} := b · obtain ⟨a, _⟩ := h.2 1 exact ⟨a.2.1, a.2.2⟩ · have : (⟨⟨_, mem_top a⁻¹⟩, ⟨a, ha⟩⟩ : (⊤ : Set G) × S) = ⟨⟨_, mem_top b⁻¹⟩, ⟨b, hb⟩⟩ := - h.1 ((inv_mul_self a).trans (inv_mul_self b).symm) + h.1 ((inv_mul_cancel a).trans (inv_mul_cancel b).symm) exact Subtype.ext_iff.mp (Prod.ext_iff.mp this).2 · rintro ⟨g, rfl⟩ exact isComplement_univ_singleton @@ -131,7 +131,7 @@ theorem isComplement_univ_right : IsComplement S univ ↔ ∃ g : G, S = {g} := · obtain ⟨a, _⟩ := h.2 1 exact ⟨a.1.1, a.1.2⟩ · have : (⟨⟨a, ha⟩, ⟨_, mem_top a⁻¹⟩⟩ : S × (⊤ : Set G)) = ⟨⟨b, hb⟩, ⟨_, mem_top b⁻¹⟩⟩ := - h.1 ((mul_inv_self a).trans (mul_inv_self b).symm) + h.1 ((mul_inv_cancel a).trans (mul_inv_cancel b).symm) exact Subtype.ext_iff.mp (Prod.ext_iff.mp this).1 · rintro ⟨g, rfl⟩ exact isComplement_singleton_univ @@ -385,12 +385,12 @@ theorem equiv_snd_eq_self_of_mem_of_one_mem {g : G} (h1 : 1 ∈ S) (hg : g ∈ T theorem equiv_snd_eq_one_of_mem_of_one_mem {g : G} (h1 : 1 ∈ T) (hg : g ∈ S) : (hST.equiv g).snd = ⟨1, h1⟩ := by ext - rw [equiv_snd_eq_inv_mul, equiv_fst_eq_self_of_mem_of_one_mem _ h1 hg, inv_mul_self] + rw [equiv_snd_eq_inv_mul, equiv_fst_eq_self_of_mem_of_one_mem _ h1 hg, inv_mul_cancel] theorem equiv_fst_eq_one_of_mem_of_one_mem {g : G} (h1 : 1 ∈ S) (hg : g ∈ T) : (hST.equiv g).fst = ⟨1, h1⟩ := by ext - rw [equiv_fst_eq_mul_inv, equiv_snd_eq_self_of_mem_of_one_mem _ h1 hg, mul_inv_self] + rw [equiv_fst_eq_mul_inv, equiv_snd_eq_self_of_mem_of_one_mem _ h1 hg, mul_inv_cancel] -- This lemma has always been bad, but the linter only noticed after lean4#2644. @[simp, nolint simpNF] diff --git a/Mathlib/GroupTheory/Congruence/Basic.lean b/Mathlib/GroupTheory/Congruence/Basic.lean index 2c1b0c3b3405a..f9137be6f7410 100644 --- a/Mathlib/GroupTheory/Congruence/Basic.lean +++ b/Mathlib/GroupTheory/Congruence/Basic.lean @@ -1040,7 +1040,7 @@ variable [Group M] [Group N] [Group P] (c : Con M) /-- Multiplicative congruence relations preserve inversion. -/ @[to_additive "Additive congruence relations preserve negation."] protected theorem inv {x y} (h : c x y) : c x⁻¹ y⁻¹ := - c.map_of_mul_left_rel_one Inv.inv (fun x => by simp only [mul_left_inv, c.refl 1]) h + c.map_of_mul_left_rel_one Inv.inv (fun x => by simp only [inv_mul_cancel, c.refl 1]) h /-- Multiplicative congruence relations preserve division. -/ @[to_additive "Additive congruence relations preserve subtraction."] diff --git a/Mathlib/GroupTheory/Coprod/Basic.lean b/Mathlib/GroupTheory/Coprod/Basic.lean index 51cb84a45a464..2349d4a1dad9b 100644 --- a/Mathlib/GroupTheory/Coprod/Basic.lean +++ b/Mathlib/GroupTheory/Coprod/Basic.lean @@ -566,11 +566,11 @@ variable {G H : Type*} [Group G] [Group H] @[to_additive] theorem mk_of_inv_mul : ∀ x : G ⊕ H, mk (of (x.map Inv.inv Inv.inv)) * mk (of x) = 1 - | Sum.inl _ => map_mul_eq_one inl (mul_left_inv _) - | Sum.inr _ => map_mul_eq_one inr (mul_left_inv _) + | Sum.inl _ => map_mul_eq_one inl (inv_mul_cancel _) + | Sum.inr _ => map_mul_eq_one inr (inv_mul_cancel _) @[to_additive] -theorem con_mul_left_inv (x : FreeMonoid (G ⊕ H)) : +theorem con_inv_mul_cancel (x : FreeMonoid (G ⊕ H)) : coprodCon G H (ofList (x.toList.map (Sum.map Inv.inv Inv.inv)).reverse * x) 1 := by rw [← mk_eq_mk, map_mul, map_one] induction x with @@ -582,7 +582,7 @@ theorem con_mul_left_inv (x : FreeMonoid (G ⊕ H)) : @[to_additive] instance : Inv (G ∗ H) where inv := Quotient.map' (fun w => ofList (w.toList.map (Sum.map Inv.inv Inv.inv)).reverse) fun _ _ ↦ - (coprodCon G H).map_of_mul_left_rel_one _ con_mul_left_inv + (coprodCon G H).map_of_mul_left_rel_one _ con_inv_mul_cancel @[to_additive] theorem inv_def (w : FreeMonoid (G ⊕ H)) : @@ -591,7 +591,7 @@ theorem inv_def (w : FreeMonoid (G ⊕ H)) : @[to_additive] instance : Group (G ∗ H) where - mul_left_inv := mk_surjective.forall.2 fun x => mk_eq_mk.2 (con_mul_left_inv x) + inv_mul_cancel := mk_surjective.forall.2 fun x => mk_eq_mk.2 (con_inv_mul_cancel x) @[to_additive (attr := simp)] theorem closure_range_inl_union_inr : diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index 719e0b7b18fae..6d942b7c97ac8 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -236,14 +236,14 @@ theorem inv_def (x : CoprodI G) : rfl instance : Group (CoprodI G) := - { mul_left_inv := by + { inv_mul_cancel := by intro m rw [inv_def] induction m using CoprodI.induction_on with | h_one => rw [MonoidHom.map_one, MulOpposite.unop_one, one_mul] | h_of m ih => change of _⁻¹ * of _ = 1 - rw [← of.map_mul, mul_left_inv, of.map_one] + rw [← of.map_mul, inv_mul_cancel, of.map_one] | h_mul x y ihx ihy => rw [MonoidHom.map_mul, MulOpposite.unop_mul, mul_assoc, ← mul_assoc _ x y, ihx, one_mul, ihy] } diff --git a/Mathlib/GroupTheory/Coset.lean b/Mathlib/GroupTheory/Coset.lean index c21b21711d37b..a4c0c84b15acb 100644 --- a/Mathlib/GroupTheory/Coset.lean +++ b/Mathlib/GroupTheory/Coset.lean @@ -196,7 +196,7 @@ theorem leftCoset_eq_iff {x y : α} : x • (s : Set α) = y • s ↔ x⁻¹ * constructor · intro h apply (h y).mpr - rw [mul_left_inv] + rw [inv_mul_cancel] exact s.one_mem · intro h z rw [← mul_inv_cancel_right x⁻¹ y] @@ -210,7 +210,7 @@ theorem rightCoset_eq_iff {x y : α} : op x • (s : Set α) = op y • s ↔ y constructor · intro h apply (h y).mpr - rw [mul_right_inv] + rw [mul_inv_cancel] exact s.one_mem · intro h z rw [← inv_mul_cancel_left y x⁻¹] diff --git a/Mathlib/GroupTheory/CosetCover.lean b/Mathlib/GroupTheory/CosetCover.lean index b5c8e0f4f04ee..2accbbd944243 100644 --- a/Mathlib/GroupTheory/CosetCover.lean +++ b/Mathlib/GroupTheory/CosetCover.lean @@ -271,7 +271,7 @@ theorem leftCoset_cover_filter_FiniteIndex_aux f, K, hHD, ← (ht i hi _).2, hi, hfi, hkfi] · rw [hdensity] refine le_of_mul_le_mul_right ?_ (Nat.cast_pos.mpr (Nat.pos_of_ne_zero hD.finiteIndex)) - rw [one_mul, mul_assoc, inv_mul_cancel (Nat.cast_ne_zero.mpr hD.finiteIndex), mul_one, + rw [one_mul, mul_assoc, inv_mul_cancel₀ (Nat.cast_ne_zero.mpr hD.finiteIndex), mul_one, Nat.cast_le] exact index_le_of_leftCoset_cover_const hcovers' · rw [hdensity, mul_inv_eq_one₀ (Nat.cast_ne_zero.mpr hD.finiteIndex), @@ -343,7 +343,7 @@ theorem exists_index_le_card_of_leftCoset_cover : | inr hindex => exact inv_lt_inv_of_lt (by exact_mod_cast hs') (by exact_mod_cast h i hi ⟨hindex⟩) apply (Finset.sum_lt_sum_of_nonempty hs hlt).trans_eq - rw [Finset.sum_const, nsmul_eq_mul, mul_inv_cancel (Nat.cast_ne_zero.mpr hs'.ne')] + rw [Finset.sum_const, nsmul_eq_mul, mul_inv_cancel₀ (Nat.cast_ne_zero.mpr hs'.ne')] end diff --git a/Mathlib/GroupTheory/DoubleCoset.lean b/Mathlib/GroupTheory/DoubleCoset.lean index 8b283e296e543..4b564ca369672 100644 --- a/Mathlib/GroupTheory/DoubleCoset.lean +++ b/Mathlib/GroupTheory/DoubleCoset.lean @@ -144,7 +144,7 @@ theorem union_quotToDoset (H K : Subgroup G) : ⋃ q, quotToDoset H K q = Set.un use mk H K x obtain ⟨h, k, h3, h4, h5⟩ := mk_out'_eq_mul H K x refine ⟨h⁻¹, H.inv_mem h3, k⁻¹, K.inv_mem h4, ?_⟩ - simp only [h5, Subgroup.coe_mk, ← mul_assoc, one_mul, mul_left_inv, mul_inv_cancel_right] + simp only [h5, Subgroup.coe_mk, ← mul_assoc, one_mul, inv_mul_cancel, mul_inv_cancel_right] theorem doset_union_rightCoset (H K : Subgroup G) (a : G) : ⋃ k : K, op (a * k) • ↑H = doset a H K := by @@ -166,10 +166,10 @@ theorem doset_union_leftCoset (H K : Subgroup G) (a : G) : constructor · rintro ⟨y, h_h⟩ refine ⟨y, y.2, a⁻¹ * y⁻¹ * x, h_h, ?_⟩ - simp only [← mul_assoc, one_mul, mul_right_inv, mul_inv_cancel_right, InvMemClass.coe_inv] + simp only [← mul_assoc, one_mul, mul_inv_cancel, mul_inv_cancel_right, InvMemClass.coe_inv] · rintro ⟨x, hx, y, hy, hxy⟩ refine ⟨⟨x, hx⟩, ?_⟩ - simp only [hxy, ← mul_assoc, hy, one_mul, mul_left_inv, Subgroup.coe_mk, inv_mul_cancel_right] + simp only [hxy, ← mul_assoc, hy, one_mul, inv_mul_cancel, Subgroup.coe_mk, inv_mul_cancel_right] theorem left_bot_eq_left_quot (H : Subgroup G) : Quotient (⊥ : Subgroup G).1 (H : Set G) = (G ⧸ H) := by diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index eca08fb36875e..74da573c0472e 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -522,7 +522,7 @@ instance : Group (FreeGroup α) where mul_assoc := by rintro ⟨L₁⟩ ⟨L₂⟩ ⟨L₃⟩; simp one_mul := by rintro ⟨L⟩; rfl mul_one := by rintro ⟨L⟩; simp [one_eq_mk] - mul_left_inv := by + inv_mul_cancel := by rintro ⟨L⟩ exact List.recOn L rfl fun ⟨x, b⟩ tl ih => diff --git a/Mathlib/GroupTheory/GroupAction/Basic.lean b/Mathlib/GroupTheory/GroupAction/Basic.lean index 71bc41b29ddaf..70bf6218e1236 100644 --- a/Mathlib/GroupTheory/GroupAction/Basic.lean +++ b/Mathlib/GroupTheory/GroupAction/Basic.lean @@ -697,8 +697,8 @@ lemma stabilizer_mul_eq_right [Group α] [SMulCommClass G α α] (a b : α) : theorem stabilizer_smul_eq_stabilizer_map_conj (g : G) (a : α) : stabilizer G (g • a) = (stabilizer G a).map (MulAut.conj g).toMonoidHom := by ext h - rw [mem_stabilizer_iff, ← smul_left_cancel_iff g⁻¹, smul_smul, smul_smul, smul_smul, mul_left_inv, - one_smul, ← mem_stabilizer_iff, Subgroup.mem_map_equiv, MulAut.conj_symm_apply] + rw [mem_stabilizer_iff, ← smul_left_cancel_iff g⁻¹, smul_smul, smul_smul, smul_smul, + inv_mul_cancel, one_smul, ← mem_stabilizer_iff, Subgroup.mem_map_equiv, MulAut.conj_symm_apply] /-- A bijection between the stabilizers of two elements in the same orbit. -/ noncomputable def stabilizerEquivStabilizerOfOrbitRel {a b : α} (h : (orbitRel G α).Rel a b) : @@ -721,7 +721,7 @@ theorem stabilizer_vadd_eq_stabilizer_map_conj (g : G) (a : α) : stabilizer G (g +ᵥ a) = (stabilizer G a).map (AddAut.conj g).toAddMonoidHom := by ext h rw [mem_stabilizer_iff, ← vadd_left_cancel_iff (-g), vadd_vadd, vadd_vadd, vadd_vadd, - add_left_neg, zero_vadd, ← mem_stabilizer_iff, AddSubgroup.mem_map_equiv, + neg_add_cancel, zero_vadd, ← mem_stabilizer_iff, AddSubgroup.mem_map_equiv, AddAut.conj_symm_apply] /-- A bijection between the stabilizers of two elements in the same orbit. -/ diff --git a/Mathlib/GroupTheory/GroupAction/ConjAct.lean b/Mathlib/GroupTheory/GroupAction/ConjAct.lean index 460fd55fa25e6..b16c6ce0f2902 100644 --- a/Mathlib/GroupTheory/GroupAction/ConjAct.lean +++ b/Mathlib/GroupTheory/GroupAction/ConjAct.lean @@ -249,7 +249,7 @@ instance : MulDistribMulAction (ConjAct G) G where smul_mul := by simp only [smul_def] simp only [mul_assoc, inv_mul_cancel_left, forall_const, «forall»] - smul_one := by simp only [smul_def, mul_one, mul_right_inv, «forall», forall_const] + smul_one := by simp only [smul_def, mul_one, mul_inv_cancel, «forall», forall_const] one_smul := by simp only [smul_def, ofConjAct_one, one_mul, inv_one, mul_one, forall_const] mul_smul := by simp only [smul_def] diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 68a6f02b68d09..a03534bfbc10f 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -285,7 +285,7 @@ instance isPretransitive_quotient (G) [Group G] (H : Subgroup G) : IsPretransiti exists_smul_eq := by { rintro ⟨x⟩ ⟨y⟩ refine ⟨y * x⁻¹, QuotientGroup.eq.mpr ?_⟩ - simp only [smul_eq_mul, H.one_mem, mul_left_inv, inv_mul_cancel_right]} + simp only [smul_eq_mul, H.one_mem, inv_mul_cancel, inv_mul_cancel_right]} variable {α} diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index ff288a79a6b5f..a389d1051072c 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -150,7 +150,7 @@ theorem index_eq_two_iff : H.index = 2 ↔ ∃ a, ∀ b, Xor' (b * a ∈ H) (b ⟨fun ha b => ⟨fun hba hb => ?_, fun hb => ?_⟩, fun ha => ⟨?_, fun b hb => ?_⟩⟩ · exact ha.1 ((mul_mem_cancel_left hb).1 hba) · exact inv_inv b ▸ ha.2 _ (mt (inv_mem_iff (x := b)).1 hb) - · rw [← inv_mem_iff (x := a), ← ha, inv_mul_self] + · rw [← inv_mem_iff (x := a), ← ha, inv_mul_cancel] exact one_mem _ · rwa [ha, inv_mem_iff (x := b)] diff --git a/Mathlib/GroupTheory/Perm/Basic.lean b/Mathlib/GroupTheory/Perm/Basic.lean index 94a39ade46477..fbe49dd2f60d0 100644 --- a/Mathlib/GroupTheory/Perm/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Basic.lean @@ -34,7 +34,7 @@ instance permGroup : Group (Perm α) where mul_assoc f g h := (trans_assoc _ _ _).symm one_mul := trans_refl mul_one := refl_trans - mul_left_inv := self_trans_symm + inv_mul_cancel := self_trans_symm npow n f := f ^ n npow_succ n f := coe_fn_injective $ Function.iterate_succ _ _ zpow := zpowRec fun n f ↦ f ^ n diff --git a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean index 4cc974bdf7585..e57ccd1bc59bb 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean @@ -394,7 +394,7 @@ theorem isCycle_swap_mul_aux₂ {α : Type*} [DecidableEq α] : isCycle_swap_mul_aux₁ n hb (show (f⁻¹ ^ n) (f⁻¹ x) = f⁻¹ b by rw [← zpow_natCast, ← h, ← mul_apply, ← mul_apply, ← mul_apply, zpow_negSucc, - ← inv_pow, pow_succ, mul_assoc, mul_assoc, inv_mul_self, mul_one, zpow_natCast, + ← inv_pow, pow_succ, mul_assoc, mul_assoc, inv_mul_cancel, mul_one, zpow_natCast, ← pow_succ', ← pow_succ]) have h : (swap x (f⁻¹ x) * f⁻¹) (f x) = f⁻¹ x := by rw [mul_apply, inv_apply_self, swap_apply_left] diff --git a/Mathlib/GroupTheory/Perm/Cycle/Type.lean b/Mathlib/GroupTheory/Perm/Cycle/Type.lean index d577d2e703957..0b0d600b9a079 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Type.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Type.lean @@ -377,7 +377,7 @@ by appending the inverse of the product of `v`. -/ @[simps] def vectorEquiv : 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_self]⟩ + rw [mem_iff, Vector.toList_cons, List.prod_cons, inv_mul_cancel]⟩ invFun v := v.1.tail left_inv v := v.tail_cons v.toList.prod⁻¹ right_inv v := Subtype.ext <| diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean index b6db3f380021a..6f3baeb1be190 100644 --- a/Mathlib/GroupTheory/PushoutI.lean +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -513,7 +513,7 @@ noncomputable def consRecOn {motive : NormalWord d → Sort _} (w : NormalWord d (h3 _ _ (List.mem_cons_self _ _))] · apply d.injective i simp only [cons, equiv_fst_eq_mul_inv, MonoidHom.apply_ofInjective_symm, - map_one, mul_one, mul_right_inv, (equiv_snd_eq_self_iff_mem (d.compl i) (one_mem _)).2 + map_one, mul_one, mul_inv_cancel, (equiv_snd_eq_self_iff_mem (d.compl i) (one_mem _)).2 (h3 _ _ (List.mem_cons_self _ _))] · rwa [← SetLike.mem_coe, ← coe_equiv_snd_eq_one_iff_mem (d.compl i) (d.one_mem _), @@ -684,7 +684,7 @@ theorem inf_of_range_eq_base_range have := hw.eq_empty_of_mem_range hφ (by simp only [Word.prod, List.map_cons, List.prod_cons, List.prod_nil, List.map_nil, map_mul, ofCoprodI_of, hg₁, hg₂, map_inv, map_one, mul_one, - mul_inv_self, one_mem]) + mul_inv_cancel, one_mem]) simp [w, Word.empty] at this) (le_inf (by rw [← of_comp_eq_base i] diff --git a/Mathlib/GroupTheory/QuotientGroup.lean b/Mathlib/GroupTheory/QuotientGroup.lean index 649d28c78dc0d..b1570f17976f1 100644 --- a/Mathlib/GroupTheory/QuotientGroup.lean +++ b/Mathlib/GroupTheory/QuotientGroup.lean @@ -334,7 +334,7 @@ theorem kerLift_mk' (g : G) : (kerLift φ) (mk g) = φ g := @[to_additive] theorem kerLift_injective : Injective (kerLift φ) := fun a b => Quotient.inductionOn₂' a b fun a b (h : φ a = φ b) => - Quotient.sound' <| by rw [leftRel_apply, mem_ker, φ.map_mul, ← h, φ.map_inv, inv_mul_self] + Quotient.sound' <| by rw [leftRel_apply, mem_ker, φ.map_mul, ← h, φ.map_inv, inv_mul_cancel] -- Note that `ker φ` isn't definitionally `ker (φ.rangeRestrict)` -- so there is a bit of annoying code duplication here @@ -348,7 +348,7 @@ theorem rangeKerLift_injective : Injective (rangeKerLift φ) := fun a b => Quotient.inductionOn₂' a b fun a b (h : φ.rangeRestrict a = φ.rangeRestrict b) => Quotient.sound' <| by rw [leftRel_apply, ← ker_rangeRestrict, mem_ker, φ.rangeRestrict.map_mul, ← h, - φ.rangeRestrict.map_inv, inv_mul_self] + φ.rangeRestrict.map_inv, inv_mul_cancel] @[to_additive] theorem rangeKerLift_surjective : Surjective (rangeKerLift φ) := by @@ -523,7 +523,7 @@ noncomputable def quotientInfEquivProdNormalQuotient (H N : Subgroup G) [N.Norma change Setoid.r _ _ rw [leftRel_apply] change h⁻¹ * (h * n) ∈ N - rwa [← mul_assoc, inv_mul_self, one_mul] + rwa [← mul_assoc, inv_mul_cancel, one_mul] (quotientMulEquivOfEq (by simp [φ, ← comap_ker])).trans (quotientKerEquivOfSurjective φ φ_surjective) diff --git a/Mathlib/GroupTheory/Schreier.lean b/Mathlib/GroupTheory/Schreier.lean index 982802b84a20c..9ff2f73e3ac1d 100644 --- a/Mathlib/GroupTheory/Schreier.lean +++ b/Mathlib/GroupTheory/Schreier.lean @@ -99,7 +99,7 @@ theorem closure_mul_image_eq (hR : R ∈ rightTransversals (H : Set G)) (hR1 : ( suffices (⟨r, hr⟩ : R) = (⟨1, hR1⟩ : R) by simpa only [show r = 1 from Subtype.ext_iff.mp this, mul_one] apply (mem_rightTransversals_iff_existsUnique_mul_inv_mem.mp hR r).unique - · rw [Subtype.coe_mk, mul_inv_self] + · rw [Subtype.coe_mk, mul_inv_cancel] exact H.one_mem · rw [Subtype.coe_mk, inv_one, mul_one] exact (H.mul_mem_cancel_left (hU hg)).mp hh diff --git a/Mathlib/GroupTheory/SchurZassenhaus.lean b/Mathlib/GroupTheory/SchurZassenhaus.lean index e865c0472905e..dfa79ebca9a3b 100644 --- a/Mathlib/GroupTheory/SchurZassenhaus.lean +++ b/Mathlib/GroupTheory/SchurZassenhaus.lean @@ -50,7 +50,7 @@ theorem smul_diff_smul' [hH : Normal H] (g : Gᵐᵒᵖ) : { toFun := fun h => ⟨g.unop⁻¹ * h * g.unop, hH.mem_comm ((congr_arg (· ∈ H) (mul_inv_cancel_left _ _)).mpr (SetLike.coe_mem _))⟩ - map_one' := by rw [Subtype.ext_iff, coe_mk, coe_one, mul_one, inv_mul_self] + map_one' := by rw [Subtype.ext_iff, coe_mk, coe_one, mul_one, inv_mul_cancel] map_mul' := fun h₁ h₂ => by simp only [Subtype.ext_iff, coe_mk, coe_mul, mul_assoc, mul_inv_cancel_left] } refine (Fintype.prod_equiv (MulAction.toPerm g).symm _ _ fun x ↦ ?_).trans (map_prod ϕ _ _).symm @@ -103,7 +103,7 @@ theorem exists_smul_eq (hH : Nat.Coprime (Nat.card H) H.index) (α β : H.Quotie (diff_inv _ _ _).symm.trans (inv_eq_one.mpr ((smul_diff' β α ((powCoprime hH).symm (diff (MonoidHom.id H) β α))⁻¹).trans - (by rw [inv_pow, ← powCoprime_apply hH, Equiv.apply_symm_apply, mul_inv_self])))⟩) + (by rw [inv_pow, ← powCoprime_apply hH, Equiv.apply_symm_apply, mul_inv_cancel])))⟩) theorem isComplement'_stabilizer_of_coprime {α : H.QuotientDiff} (hH : Nat.Coprime (Nat.card H) H.index) : IsComplement' H (stabilizer G α) := diff --git a/Mathlib/GroupTheory/SemidirectProduct.lean b/Mathlib/GroupTheory/SemidirectProduct.lean index 9d702707e2f17..904d910b72cfa 100644 --- a/Mathlib/GroupTheory/SemidirectProduct.lean +++ b/Mathlib/GroupTheory/SemidirectProduct.lean @@ -90,7 +90,7 @@ instance : Group (N ⋊[φ] G) where mul_assoc a b c := SemidirectProduct.ext (by simp [mul_assoc]) (by simp [mul_assoc]) one_mul a := SemidirectProduct.ext (by simp) (one_mul a.2) mul_one a := SemidirectProduct.ext (by simp) (mul_one _) - mul_left_inv a := SemidirectProduct.ext (by simp) (by simp) + inv_mul_cancel a := SemidirectProduct.ext (by simp) (by simp) instance : Inhabited (N ⋊[φ] G) := ⟨1⟩ diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 71c6a80871b74..3be24db573211 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -522,9 +522,9 @@ theorem commutative_of_cyclic_center_quotient [IsCyclic H] (f : G →* H) (hf : have hm : x ^ m = f a := by simpa [Subtype.ext_iff] using hm have hn : x ^ n = f b := by simpa [Subtype.ext_iff] using hn have ha : y ^ (-m) * a ∈ center G := - hf (by rw [f.mem_ker, f.map_mul, f.map_zpow, hxy, zpow_neg x m, hm, inv_mul_self]) + hf (by rw [f.mem_ker, f.map_mul, f.map_zpow, hxy, zpow_neg x m, hm, inv_mul_cancel]) have hb : y ^ (-n) * b ∈ center G := - hf (by rw [f.mem_ker, f.map_mul, f.map_zpow, hxy, zpow_neg x n, hn, inv_mul_self]) + hf (by rw [f.mem_ker, f.map_mul, f.map_zpow, hxy, zpow_neg x n, hn, inv_mul_cancel]) calc a * b = y ^ m * (y ^ (-m) * a * y ^ n) * (y ^ (-n) * b) := by simp [mul_assoc] _ = y ^ m * (y ^ n * (y ^ (-m) * a)) * (y ^ (-n) * b) := by rw [mem_center_iff.1 ha] diff --git a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean index 357e55005fb13..26ff5fc37728e 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean @@ -67,9 +67,9 @@ instance : Group (DihedralGroup n) where · exact congr_arg r (add_zero a) · exact congr_arg sr (add_zero a) inv := inv - mul_left_inv := by + inv_mul_cancel := by rintro (a | a) - · exact congr_arg r (neg_add_self a) + · exact congr_arg r (neg_add_cancel a) · exact congr_arg r (sub_self a) @[simp] diff --git a/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean b/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean index 770017188fed3..be7d8cba6dabf 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean @@ -100,9 +100,9 @@ instance : Group (QuaternionGroup n) where · exact congr_arg a (add_zero i) · exact congr_arg xa (add_zero i) inv := inv - mul_left_inv := by + inv_mul_cancel := by rintro (i | i) - · exact congr_arg a (neg_add_self i) + · exact congr_arg a (neg_add_cancel i) · exact congr_arg a (sub_self (n + i)) @[simp] diff --git a/Mathlib/GroupTheory/Submonoid/Inverses.lean b/Mathlib/GroupTheory/Submonoid/Inverses.lean index 34db037aab4f1..469f9ece7f56c 100644 --- a/Mathlib/GroupTheory/Submonoid/Inverses.lean +++ b/Mathlib/GroupTheory/Submonoid/Inverses.lean @@ -37,7 +37,7 @@ namespace Submonoid noncomputable instance [Monoid M] : Group (IsUnit.submonoid M) := { inferInstanceAs (Monoid (IsUnit.submonoid M)) with inv := fun x ↦ ⟨x.prop.unit⁻¹.val, x.prop.unit⁻¹.isUnit⟩ - mul_left_inv := fun x ↦ + inv_mul_cancel := fun x ↦ Subtype.ext ((Units.val_mul x.prop.unit⁻¹ _).trans x.prop.unit.inv_val) } @[to_additive] @@ -191,11 +191,11 @@ theorem leftInv_eq_inv : S.leftInv = S⁻¹ := Submonoid.ext fun _ ↦ ⟨fun h ↦ Submonoid.mem_inv.mpr ((inv_eq_of_mul_eq_one_right h.choose_spec).symm ▸ h.choose.prop), - fun h ↦ ⟨⟨_, h⟩, mul_right_inv _⟩⟩ + fun h ↦ ⟨⟨_, h⟩, mul_inv_cancel _⟩⟩ @[to_additive (attr := simp)] theorem fromLeftInv_eq_inv (x : S.leftInv) : (S.fromLeftInv x : M) = (x : M)⁻¹ := by - rw [← mul_right_inj (x : M), mul_right_inv, mul_fromLeftInv] + rw [← mul_right_inj (x : M), mul_inv_cancel, mul_fromLeftInv] end Group @@ -205,7 +205,7 @@ variable [CommGroup M] (S : Submonoid M) (hS : S ≤ IsUnit.submonoid M) @[to_additive (attr := simp)] theorem leftInvEquiv_symm_eq_inv (x : S) : ((S.leftInvEquiv hS).symm x : M) = (x : M)⁻¹ := by - rw [← mul_right_inj (x : M), mul_right_inv, mul_leftInvEquiv_symm] + rw [← mul_right_inj (x : M), mul_inv_cancel, mul_leftInvEquiv_symm] end CommGroup diff --git a/Mathlib/GroupTheory/Torsion.lean b/Mathlib/GroupTheory/Torsion.lean index d3eef9e840874..ec5b8b580bb9e 100644 --- a/Mathlib/GroupTheory/Torsion.lean +++ b/Mathlib/GroupTheory/Torsion.lean @@ -64,7 +64,7 @@ open Monoid noncomputable def IsTorsion.group [Monoid G] (tG : IsTorsion G) : Group G := { ‹Monoid G› with inv := fun g => g ^ (orderOf g - 1) - mul_left_inv := fun g => by + inv_mul_cancel := fun g => by erw [← pow_succ, tsub_add_cancel_of_le, pow_orderOf_eq_one] exact (tG g).orderOf_pos } diff --git a/Mathlib/GroupTheory/Transfer.lean b/Mathlib/GroupTheory/Transfer.lean index fcfd842450b0b..f1d47f133261b 100644 --- a/Mathlib/GroupTheory/Transfer.lean +++ b/Mathlib/GroupTheory/Transfer.lean @@ -118,7 +118,7 @@ theorem transfer_eq_prod_quotient_orbitRel_zpowers_quot [FiniteIndex H] (g : G) rw [Fintype.prod_eq_single (0 : ZMod (Function.minimalPeriod (g • ·) q.out')) _] · simp only [if_pos, ZMod.cast_zero, zpow_zero, one_mul, mul_assoc] · intro k hk - simp only [if_neg hk, inv_mul_self] + simp only [if_neg hk, inv_mul_cancel] exact map_one ϕ /-- Auxiliary lemma in order to state `transfer_eq_pow`. -/ diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean index c8fd0d0b5ebb8..86b064cb61eac 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean @@ -326,7 +326,7 @@ instance group : Group (P₁ ≃ᵃ[k] P₁) where mul_assoc e₁ e₂ e₃ := trans_assoc _ _ _ one_mul := trans_refl mul_one := refl_trans - mul_left_inv := self_trans_symm + inv_mul_cancel := self_trans_symm theorem one_def : (1 : P₁ ≃ᵃ[k] P₁) = refl k P₁ := rfl diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean index e8ffe0b27a7f5..3dff1b6db9f72 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineSubspace.lean @@ -1620,7 +1620,7 @@ scoped[Affine] infixl:50 " ∥ " => AffineSubspace.Parallel theorem Parallel.symm {s₁ s₂ : AffineSubspace k P} (h : s₁ ∥ s₂) : s₂ ∥ s₁ := by rcases h with ⟨v, rfl⟩ refine ⟨-v, ?_⟩ - rw [map_map, ← coe_trans_to_affineMap, ← constVAdd_add, neg_add_self, constVAdd_zero, + rw [map_map, ← coe_trans_to_affineMap, ← constVAdd_add, neg_add_cancel, constVAdd_zero, coe_refl_to_affineMap, map_id] theorem parallel_comm {s₁ s₂ : AffineSubspace k P} : s₁ ∥ s₂ ↔ s₂ ∥ s₁ := diff --git a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean index 2121be8f2b0b7..621360ad2235f 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean @@ -727,7 +727,7 @@ theorem sum_centroidWeights_eq_one_of_card_ne_zero [CharZero k] (h : card s ≠ ∑ i ∈ s, s.centroidWeights k i = 1 := by -- Porting note: `simp` cannot find `mul_inv_cancel` and does not use `norm_cast` simp only [centroidWeights_apply, sum_const, nsmul_eq_mul, ne_eq, Nat.cast_eq_zero, card_eq_zero] - refine mul_inv_cancel ?_ + refine mul_inv_cancel₀ ?_ norm_cast /-- In the characteristic zero case, the weights in the centroid sum diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index f5ffb8af6f5a6..826df0fc1aed6 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -105,7 +105,7 @@ theorem affineIndependent_iff_linearIndependent_vsub (p : ι → P) (i1 : ι) : rw [hfdef] dsimp only rw [dif_pos rfl] - exact neg_add_self _ + exact neg_add_cancel _ have hs2 : s2.weightedVSub p f = (0 : V) := by set f2 : ι → V := fun x => f x • (p x -ᵥ p i1) with hf2def set g2 : { x // x ≠ i1 } → V := fun x => g x • (p x -ᵥ p i1) diff --git a/Mathlib/LinearAlgebra/AffineSpace/Midpoint.lean b/Mathlib/LinearAlgebra/AffineSpace/Midpoint.lean index 978ba6c76216f..d57b4f3a3c569 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Midpoint.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Midpoint.lean @@ -192,7 +192,7 @@ theorem midpoint_eq_smul_add (x y : V) : midpoint R x y = (⅟ 2 : R) • (x + y @[simp] theorem midpoint_self_neg (x : V) : midpoint R x (-x) = 0 := by - rw [midpoint_eq_smul_add, add_neg_self, smul_zero] + rw [midpoint_eq_smul_add, add_neg_cancel, smul_zero] @[simp] theorem midpoint_neg_self (x : V) : midpoint R (-x) x = 0 := by simpa using midpoint_self_neg R (-x) diff --git a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean index 6da58edbd959e..deb0d98dc3cb3 100644 --- a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean +++ b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean @@ -79,7 +79,7 @@ theorem domCoprod.summand_add_swap_smul_eq_zero (a : Mᵢ [⋀^ιa]→ₗ[R'] N simp only [one_mul, neg_mul, Function.comp_apply, Units.neg_smul, Perm.coe_mul, Units.val_neg, MultilinearMap.smul_apply, MultilinearMap.neg_apply, MultilinearMap.domDomCongr_apply, MultilinearMap.domCoprod_apply] - convert add_right_neg (G := N₁ ⊗[R'] N₂) _ using 6 <;> + convert add_neg_cancel (G := N₁ ⊗[R'] N₂) _ using 6 <;> · ext k rw [Equiv.apply_swap_eq_self hv] diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean index 175d5466da2fe..b1e419160e7e2 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean @@ -334,11 +334,11 @@ def changeFormEquiv : CliffordAlgebra Q ≃ₗ[R] CliffordAlgebra Q' := left_inv := fun x => by dsimp only exact (changeForm_changeForm _ _ x).trans <| - by simp_rw [(add_neg_self B), changeForm_self_apply] + by simp_rw [(add_neg_cancel B), changeForm_self_apply] right_inv := fun x => by dsimp only exact (changeForm_changeForm _ _ x).trans <| - by simp_rw [(add_left_neg B), changeForm_self_apply] } + by simp_rw [(neg_add_cancel B), changeForm_self_apply] } @[simp] theorem changeFormEquiv_symm : diff --git a/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean b/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean index b87f530235f73..6b0db8a0b52de 100644 --- a/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean +++ b/Mathlib/LinearAlgebra/Dimension/DivisionRing.lean @@ -155,7 +155,7 @@ theorem linearIndependent_of_top_le_span_of_card_eq_finrank {ι : Type*} [Fintyp calc (b i + (g i)⁻¹ • (s.erase i).sum fun j => g j • b j) = (g i)⁻¹ • (g i • b i + (s.erase i).sum fun j => g j • b j) := by - rw [smul_add, ← mul_smul, inv_mul_cancel gx_ne_zero, one_smul] + rw [smul_add, ← mul_smul, inv_mul_cancel₀ gx_ne_zero, one_smul] _ = (g i)⁻¹ • (0 : V) := congr_arg _ ?_ _ = 0 := smul_zero _ -- And then it's just a bit of manipulation with finite sums. diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index fd39101c95f7f..535f8379922d6 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -1316,7 +1316,7 @@ lemma range_eq_top_of_ne_zero : LinearMap.range f = ⊤ := by obtain ⟨v, hv⟩ : ∃ v, f v ≠ 0 := by contrapose! hf; ext v; simpa using hf v rw [eq_top_iff] - exact fun x _ ↦ ⟨x • (f v)⁻¹ • v, by simp [inv_mul_cancel hv]⟩ + exact fun x _ ↦ ⟨x • (f v)⁻¹ • v, by simp [inv_mul_cancel₀ hv]⟩ lemma finrank_ker_add_one_of_ne_zero : finrank K (LinearMap.ker f) + 1 = finrank K V₁ := by diff --git a/Mathlib/LinearAlgebra/Lagrange.lean b/Mathlib/LinearAlgebra/Lagrange.lean index 2fb50d66771f6..d3e6374765d4c 100644 --- a/Mathlib/LinearAlgebra/Lagrange.lean +++ b/Mathlib/LinearAlgebra/Lagrange.lean @@ -176,7 +176,7 @@ theorem eval_basisDivisor_right : eval y (basisDivisor x y) = 0 := by theorem eval_basisDivisor_left_of_ne (hxy : x ≠ y) : eval x (basisDivisor x y) = 1 := by simp only [basisDivisor, eval_mul, eval_C, eval_sub, eval_X] - exact inv_mul_cancel (sub_ne_zero_of_ne hxy) + exact inv_mul_cancel₀ (sub_ne_zero_of_ne hxy) end BasisDivisor @@ -594,7 +594,7 @@ theorem eval_basis_not_at_node (hi : i ∈ s) (hxi : x ≠ v i) : eval x (Lagrange.basis s v i) = eval x (nodal s v) * (nodalWeight s v i * (x - v i)⁻¹) := by rw [mul_comm, basis_eq_prod_sub_inv_mul_nodal_div hi, eval_mul, eval_C, ← nodal_erase_eq_nodal_div hi, eval_nodal, eval_nodal, mul_assoc, ← mul_prod_erase _ _ hi, ← - mul_assoc (x - v i)⁻¹, inv_mul_cancel (sub_ne_zero_of_ne hxi), one_mul] + mul_assoc (x - v i)⁻¹, inv_mul_cancel₀ (sub_ne_zero_of_ne hxi), one_mul] theorem interpolate_eq_nodalWeight_mul_nodal_div_X_sub_C : interpolate s v r = ∑ i ∈ s, C (nodalWeight s v i) * (nodal s v / (X - C (v i))) * C (r i) := diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 89caba69587a6..e6a3d680dc9fc 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -1276,7 +1276,7 @@ theorem LinearIndependent.pair_iff' {x y : V} (hx : x ≠ 0) : by_cases ht : t = 0 · exact ⟨by simpa [ht, hx] using hst, ht⟩ apply_fun (t⁻¹ • ·) at hst - simp only [smul_add, smul_smul, inv_mul_cancel ht, one_smul, smul_zero] at hst + simp only [smul_add, smul_smul, inv_mul_cancel₀ ht, one_smul, smul_zero] at hst cases H (-(t⁻¹ * s)) (by rwa [neg_smul, neg_eq_iff_eq_neg, eq_neg_iff_add_eq_zero]) theorem linearIndependent_fin_cons {n} {v : Fin n → V} : diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean index 0c727049fde61..3b0529e1973d4 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean @@ -315,7 +315,7 @@ lemma reverse_charpoly (M : Matrix n n R) : let t_inv : R[T;T⁻¹] := T (-1) let p : R[T;T⁻¹] := det (scalar n t - M.map LaurentPolynomial.C) let q : R[T;T⁻¹] := det (1 - scalar n t * M.map LaurentPolynomial.C) - have ht : t_inv * t = 1 := by rw [← T_add, add_left_neg, T_zero] + have ht : t_inv * t = 1 := by rw [← T_add, neg_add_cancel, T_zero] have hp : toLaurentAlg M.charpoly = p := by simp [p, charpoly, charmatrix, AlgHom.map_det, map_sub, map_smul'] have hq : toLaurentAlg M.charpolyRev = q := by diff --git a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean index 63d7bd3ecf2ea..301e6932dcf83 100644 --- a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean +++ b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean @@ -159,8 +159,8 @@ variable {n : Type u} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] @[coe] def coeToGL (A : SpecialLinearGroup n R) : GL n R := ⟨↑A, ↑A⁻¹, - congr_arg ((↑) : _ → Matrix n n R) (mul_right_inv A), - congr_arg ((↑) : _ → Matrix n n R) (mul_left_inv A)⟩ + congr_arg ((↑) : _ → Matrix n n R) (mul_inv_cancel A), + congr_arg ((↑) : _ → Matrix n n R) (inv_mul_cancel A)⟩ instance hasCoeToGeneralLinearGroup : Coe (SpecialLinearGroup n R) (GL n R) := ⟨coeToGL⟩ diff --git a/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean b/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean index e233de0d891e6..03ef3530c8616 100644 --- a/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean +++ b/Mathlib/LinearAlgebra/Matrix/SchurComplement.lean @@ -78,7 +78,7 @@ def fromBlocksZero₂₁Invertible (A : Matrix m m α) (B : Matrix m n α) (D : [Invertible A] [Invertible D] : Invertible (fromBlocks A B 0 D) := invertibleOfLeftInverse _ (fromBlocks (⅟ A) (-(⅟ A * B * ⅟ D)) 0 (⅟ D)) <| by simp_rw [fromBlocks_multiply, Matrix.mul_zero, Matrix.zero_mul, zero_add, add_zero, - Matrix.neg_mul, invOf_mul_self, Matrix.mul_invOf_mul_self_cancel, add_right_neg, + Matrix.neg_mul, invOf_mul_self, Matrix.mul_invOf_mul_self_cancel, add_neg_cancel, fromBlocks_one] /-- A lower-block-triangular matrix is invertible if its diagonal is. -/ @@ -88,7 +88,7 @@ def fromBlocksZero₁₂Invertible (A : Matrix m m α) (C : Matrix n m α) (D : (fromBlocks (⅟ A) 0 (-(⅟ D * C * ⅟ A)) (⅟ D)) <| by -- a symmetry argument is more work than just copying the proof simp_rw [fromBlocks_multiply, Matrix.mul_zero, Matrix.zero_mul, zero_add, add_zero, - Matrix.neg_mul, invOf_mul_self, Matrix.mul_invOf_mul_self_cancel, add_left_neg, + Matrix.neg_mul, invOf_mul_self, Matrix.mul_invOf_mul_self_cancel, neg_add_cancel, fromBlocks_one] theorem invOf_fromBlocks_zero₂₁_eq (A : Matrix m m α) (B : Matrix m n α) (D : Matrix n n α) @@ -500,7 +500,7 @@ theorem PosSemidef.fromBlocks₁₁ [Fintype m] [DecidableEq m] [Fintype n] {A : constructor · refine fun h => ⟨h.1, fun x => ?_⟩ have := h.2 (-((A⁻¹ * B) *ᵥ x) ⊕ᵥ x) - rw [dotProduct_mulVec, schur_complement_eq₁₁ B D _ _ hA.1, neg_add_self, dotProduct_zero, + rw [dotProduct_mulVec, schur_complement_eq₁₁ B D _ _ hA.1, neg_add_cancel, dotProduct_zero, zero_add] at this rw [dotProduct_mulVec]; exact this · refine fun h => ⟨h.1, fun x => ?_⟩ diff --git a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean index 11c0dfe89d9f0..bf7e0f5c040f9 100644 --- a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean +++ b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean @@ -177,7 +177,7 @@ instance monoid : Monoid (SpecialLinearGroup n R) := instance : Group (SpecialLinearGroup n R) := { SpecialLinearGroup.monoid, SpecialLinearGroup.hasInv with - mul_left_inv := fun A => by + inv_mul_cancel := fun A => by ext1 simp [adjugate_mul] } @@ -185,8 +185,8 @@ instance : Group (SpecialLinearGroup n R) := def toLin' : SpecialLinearGroup n R →* (n → R) ≃ₗ[R] n → R where toFun A := LinearEquiv.ofLinear (Matrix.toLin' ↑ₘA) (Matrix.toLin' ↑ₘA⁻¹) - (by rw [← toLin'_mul, ← coe_mul, mul_right_inv, coe_one, toLin'_one]) - (by rw [← toLin'_mul, ← coe_mul, mul_left_inv, coe_one, toLin'_one]) + (by rw [← toLin'_mul, ← coe_mul, mul_inv_cancel, coe_one, toLin'_one]) + (by rw [← toLin'_mul, ← coe_mul, inv_mul_cancel, coe_one, toLin'_one]) map_one' := LinearEquiv.toLinearMap_injective Matrix.toLin'_one map_mul' A B := LinearEquiv.toLinearMap_injective <| Matrix.toLin'_mul ↑ₘA ↑ₘB diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 0037b0403d11e..9e188e11220cf 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -1213,7 +1213,7 @@ theorem sub_apply (m : ∀ i, M₁ i) : (f - g) m = f m - g m := instance : AddCommGroup (MultilinearMap R M₁ M₂) := { MultilinearMap.addCommMonoid with - add_left_neg := fun a => MultilinearMap.ext fun v => add_left_neg _ + neg_add_cancel := fun a => MultilinearMap.ext fun v => neg_add_cancel _ sub_eq_add_neg := fun a b => MultilinearMap.ext fun v => sub_eq_add_neg _ _ zsmul := fun n f => { toFun := fun m => n • f m @@ -1235,7 +1235,7 @@ variable [Semiring R] [∀ i, AddCommGroup (M₁ i)] [AddCommGroup M₂] [∀ i, theorem map_neg [DecidableEq ι] (m : ∀ i, M₁ i) (i : ι) (x : M₁ i) : f (update m i (-x)) = -f (update m i x) := eq_neg_of_add_eq_zero_left <| by - rw [← MultilinearMap.map_add, add_left_neg, f.map_coord_zero i (update_same i 0 m)] + rw [← MultilinearMap.map_add, neg_add_cancel, f.map_coord_zero i (update_same i 0 m)] @[simp] theorem map_sub [DecidableEq ι] (m : ∀ i, M₁ i) (i : ι) (x y : M₁ i) : diff --git a/Mathlib/LinearAlgebra/Orientation.lean b/Mathlib/LinearAlgebra/Orientation.lean index 789d425dda083..68891b346f5df 100644 --- a/Mathlib/LinearAlgebra/Orientation.lean +++ b/Mathlib/LinearAlgebra/Orientation.lean @@ -174,7 +174,7 @@ theorem orientation_unitsSMul (e : Basis ι R M) (w : ι → Units R) : rw [Basis.orientation, Basis.orientation, smul_rayOfNeZero, ray_eq_iff, e.det.eq_smul_basis_det (e.unitsSMul w), det_unitsSMul_self, Units.smul_def, smul_smul] norm_cast - simp only [mul_left_inv, Units.val_one, one_smul] + simp only [inv_mul_cancel, Units.val_one, one_smul] exact SameRay.rfl @[simp] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index bddba63da1c4f..f3758444777a0 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -1222,7 +1222,7 @@ theorem exists_orthogonal_basis [hK : Invertible (2 : K)] {B : LinearMap.BilinFo refine ⟨-B x y / B x x, fun z hz => ?_⟩ obtain ⟨c, rfl⟩ := Submodule.mem_span_singleton.1 hz rw [IsOrtho, map_smul, smul_apply, map_add, map_smul, smul_eq_mul, smul_eq_mul, - div_mul_cancel₀ _ hx, add_neg_self, mul_zero]) + div_mul_cancel₀ _ hx, add_neg_cancel, mul_zero]) refine ⟨b, ?_⟩ rw [Basis.coe_mkFinCons] intro j i diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean b/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean index 980063ffab082..3285feacba1f0 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Complex.lean @@ -58,7 +58,7 @@ noncomputable def isometryEquivSumSquares (w' : ι → ℂ) : suffices v j * v j = w j ^ (-(1 / 2 : ℂ)) * w j ^ (-(1 / 2 : ℂ)) * w j * v j * v j by rw [this]; ring rw [← Complex.cpow_add _ _ (w j).ne_zero, show -(1 / 2 : ℂ) + -(1 / 2) = -1 by simp [← two_mul], - Complex.cpow_neg_one, inv_mul_cancel (w j).ne_zero, one_mul] + Complex.cpow_neg_one, inv_mul_cancel₀ (w j).ne_zero, one_mul] /-- The isometry between a weighted sum of squares on the complex numbers and the sum of squares, i.e. `weightedSumSquares` with weight `fun (i : ι) => 1`. -/ diff --git a/Mathlib/LinearAlgebra/SymplecticGroup.lean b/Mathlib/LinearAlgebra/SymplecticGroup.lean index 5639903ee2c37..c638f98ae8bd9 100644 --- a/Mathlib/LinearAlgebra/SymplecticGroup.lean +++ b/Mathlib/LinearAlgebra/SymplecticGroup.lean @@ -178,7 +178,7 @@ theorem inv_eq_symplectic_inv (A : Matrix (l ⊕ l) (l ⊕ l) R) (hA : A ∈ sym instance : Group (symplecticGroup l R) := { SymplecticGroup.hasInv, Submonoid.toMonoid _ with - mul_left_inv := fun A => by + inv_mul_cancel := fun A => by apply Subtype.ext simp only [Submonoid.coe_one, Submonoid.coe_mul, Matrix.neg_mul, coe_inv] exact inv_left_mul_aux A.2 } diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index 0dbf00d45ee83..8cc95132b62e1 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -1400,10 +1400,10 @@ variable {R} instance neg : Neg (M ⊗[R] N) where neg := Neg.aux R -protected theorem add_left_neg (x : M ⊗[R] N) : -x + x = 0 := +protected theorem neg_add_cancel (x : M ⊗[R] N) : -x + x = 0 := x.induction_on (by rw [add_zero]; apply (Neg.aux R).map_zero) - (fun x y => by convert (add_tmul (R := R) (-x) x y).symm; rw [add_left_neg, zero_tmul]) + (fun x y => by convert (add_tmul (R := R) (-x) x y).symm; rw [neg_add_cancel, zero_tmul]) fun x y hx hy => by suffices -x + x + (-y + y) = 0 by rw [← this] @@ -1418,13 +1418,13 @@ instance addCommGroup : AddCommGroup (M ⊗[R] N) := neg := Neg.neg sub := _ sub_eq_add_neg := fun _ _ => rfl - add_left_neg := fun x => TensorProduct.add_left_neg x + neg_add_cancel := fun x => TensorProduct.neg_add_cancel x zsmul := fun n v => n • v zsmul_zero' := by simp [TensorProduct.zero_smul] zsmul_succ' := by simp [add_comm, TensorProduct.one_smul, TensorProduct.add_smul] zsmul_neg' := fun n x => by change (-n.succ : ℤ) • x = -(((n : ℤ) + 1) • x) - rw [← zero_add (_ • x), ← TensorProduct.add_left_neg ((n.succ : ℤ) • x), add_assoc, + rw [← zero_add (_ • x), ← TensorProduct.neg_add_cancel ((n.succ : ℤ) • x), add_assoc, ← add_smul, ← sub_eq_add_neg, sub_self, zero_smul, add_zero] rfl } diff --git a/Mathlib/LinearAlgebra/UnitaryGroup.lean b/Mathlib/LinearAlgebra/UnitaryGroup.lean index d9a72a12126f8..3e19993891bae 100644 --- a/Mathlib/LinearAlgebra/UnitaryGroup.lean +++ b/Mathlib/LinearAlgebra/UnitaryGroup.lean @@ -142,11 +142,11 @@ def toLinearEquiv (A : unitaryGroup n α) : (n → α) ≃ₗ[α] n → α := left_inv := fun x => calc (toLin' A⁻¹).comp (toLin' A) x = (toLin' (A⁻¹ * A)) x := by rw [← toLin'_mul] - _ = x := by rw [mul_left_inv, toLin'_one, id_apply] + _ = x := by rw [inv_mul_cancel, toLin'_one, id_apply] right_inv := fun x => calc (toLin' A).comp (toLin' A⁻¹) x = toLin' (A * A⁻¹) x := by rw [← toLin'_mul] - _ = x := by rw [mul_right_inv, toLin'_one, id_apply] } + _ = x := by rw [mul_inv_cancel, toLin'_one, id_apply] } /-- `Matrix.unitaryGroup.toGL` is the map from the unitary group to the general linear group -/ def toGL (A : unitaryGroup n α) : GeneralLinearGroup α (n → α) := diff --git a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean index 61e01c28bc811..9311803c85be1 100644 --- a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean +++ b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean @@ -82,7 +82,7 @@ theorem centerAndRescale_center : a.centerAndRescale.c (last N) = 0 := by theorem centerAndRescale_radius {N : ℕ} {τ : ℝ} (a : SatelliteConfig E N τ) : a.centerAndRescale.r (last N) = 1 := by - simp [SatelliteConfig.centerAndRescale, inv_mul_cancel (a.rpos _).ne'] + simp [SatelliteConfig.centerAndRescale, inv_mul_cancel₀ (a.rpos _).ne'] end SatelliteConfig diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index 2170ef3b5f865..5cc54ff5324a5 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -604,7 +604,7 @@ theorem rnDeriv_smul_right (ν μ : Measure α) [IsFiniteMeasure ν] rwa [add_right_inj] at this rw [← ν.haveLebesgueDecomposition_add (r • μ), singularPart_smul_right _ _ _ hr, ENNReal.smul_def r, withDensity_smul_measure, ← ENNReal.smul_def, ← smul_assoc, - smul_eq_mul, inv_mul_cancel hr, one_smul] + smul_eq_mul, inv_mul_cancel₀ hr, one_smul] exact ν.haveLebesgueDecomposition_add μ /-- Radon-Nikodym derivative with respect to the scalar multiple of a measure. @@ -691,7 +691,7 @@ theorem exists_positive_of_not_mutuallySingular (μ ν : Measure α) [IsFiniteMe have hb₁ : (0 : ℝ) < (νA : ℝ)⁻¹ := by rw [_root_.inv_pos]; exact hb have h' : 1 / (↑n + 1) * νA < c := by rw [← NNReal.coe_lt_coe, ← mul_lt_mul_right hb₁, NNReal.coe_mul, mul_assoc, ← - NNReal.coe_inv, ← NNReal.coe_mul, _root_.mul_inv_cancel, ← NNReal.coe_mul, mul_one, + NNReal.coe_inv, ← NNReal.coe_mul, mul_inv_cancel₀, ← NNReal.coe_mul, mul_one, NNReal.coe_inv] · exact hn · exact hb.ne' @@ -990,7 +990,7 @@ theorem rnDeriv_smul_right' (ν μ : Measure α) [SigmaFinite ν] [SigmaFinite rwa [add_right_inj] at this rw [← ν.haveLebesgueDecomposition_add (r • μ), singularPart_smul_right _ _ _ hr, ENNReal.smul_def r, withDensity_smul_measure, ← ENNReal.smul_def, ← smul_assoc, - smul_eq_mul, inv_mul_cancel hr, one_smul] + smul_eq_mul, inv_mul_cancel₀ hr, one_smul] exact ν.haveLebesgueDecomposition_add μ · exact (measurable_rnDeriv _ _).aemeasurable · exact (measurable_rnDeriv _ _).aemeasurable.const_smul _ diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean index cdaa8675a1517..5feff3cfacb85 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean @@ -248,7 +248,7 @@ theorem condexp_bot' [hμ : NeZero μ] (f : α → F') : rw [h_eq] have h_integral : ∫ x, (μ[f|⊥]) x ∂μ = ∫ x, f x ∂μ := integral_condexp bot_le simp_rw [h_eq, integral_const] at h_integral - rw [← h_integral, ← smul_assoc, smul_eq_mul, inv_mul_cancel, one_smul] + rw [← h_integral, ← smul_assoc, smul_eq_mul, inv_mul_cancel₀, one_smul] rw [Ne, ENNReal.toReal_eq_zero_iff, not_or] exact ⟨NeZero.ne _, measure_ne_top μ Set.univ⟩ diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean index 4c7a244bca900..3b2163fa2524d 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean @@ -206,7 +206,7 @@ theorem Integrable.uniformIntegrable_condexp {ι : Type*} [IsFiniteMeasure μ] { rw [ENNReal.div_le_iff_le_mul (Or.inl (ENNReal.coe_ne_zero.2 hCpos.ne')) (Or.inl ENNReal.coe_lt_top.ne), hC, Nonneg.inv_mk, ENNReal.coe_mul, ENNReal.coe_toNNReal hg.eLpNorm_lt_top.ne, ← mul_assoc, ← - ENNReal.ofReal_eq_coe_nnreal, ← ENNReal.ofReal_mul hδ.le, mul_inv_cancel hδ.ne', + ENNReal.ofReal_eq_coe_nnreal, ← ENNReal.ofReal_mul hδ.le, mul_inv_cancel₀ hδ.ne', ENNReal.ofReal_one, one_mul] exact eLpNorm_one_condexp_le_eLpNorm _ refine ⟨C, fun n => le_trans ?_ (h {x : α | C ≤ ‖(μ[g|ℱ n]) x‖₊} (hmeas n C) (this n))⟩ diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index fb4fe7466b781..7cdd71d229171 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -129,7 +129,7 @@ theorem Memℒp.aestronglyMeasurable {f : α → E} {p : ℝ≥0∞} (h : Memℒ theorem lintegral_rpow_nnnorm_eq_rpow_eLpNorm' {f : α → F} (hq0_lt : 0 < q) : ∫⁻ a, (‖f a‖₊ : ℝ≥0∞) ^ q ∂μ = eLpNorm' f q μ ^ q := by - rw [eLpNorm', ← ENNReal.rpow_mul, one_div, inv_mul_cancel, ENNReal.rpow_one] + rw [eLpNorm', ← ENNReal.rpow_mul, one_div, inv_mul_cancel₀, ENNReal.rpow_one] exact (ne_of_lt hq0_lt).symm @[deprecated (since := "2024-07-27")] @@ -350,7 +350,7 @@ theorem eLpNorm'_const (c : F) (hq_pos : 0 < q) : congr rw [← ENNReal.rpow_mul] suffices hq_cancel : q * (1 / q) = 1 by rw [hq_cancel, ENNReal.rpow_one] - rw [one_div, mul_inv_cancel (ne_of_lt hq_pos).symm] + rw [one_div, mul_inv_cancel₀ (ne_of_lt hq_pos).symm] @[deprecated (since := "2024-07-27")] alias snorm'_const := eLpNorm'_const @@ -361,7 +361,7 @@ theorem eLpNorm'_const' [IsFiniteMeasure μ] (c : F) (hc_ne_zero : c ≠ 0) (hq_ · congr rw [← ENNReal.rpow_mul] suffices hp_cancel : q * (1 / q) = 1 by rw [hp_cancel, ENNReal.rpow_one] - rw [one_div, mul_inv_cancel hq_ne_zero] + rw [one_div, mul_inv_cancel₀ hq_ne_zero] · rw [Ne, ENNReal.rpow_eq_top_iff, not_or, not_and_or, not_and_or] constructor · left @@ -637,7 +637,7 @@ theorem eLpNorm'_norm_rpow (f : α → F) (p q : ℝ) (hq_pos : 0 < q) : simp_rw [eLpNorm'] rw [← ENNReal.rpow_mul, ← one_div_mul_one_div] simp_rw [one_div] - rw [mul_assoc, inv_mul_cancel hq_pos.ne.symm, mul_one] + rw [mul_assoc, inv_mul_cancel₀ hq_pos.ne.symm, mul_one] congr ext1 x simp_rw [← ofReal_norm_eq_coe_nnnorm] @@ -1082,7 +1082,7 @@ theorem eLpNorm'_le_nnreal_smul_eLpNorm'_of_ae_le_mul {f : α → F} {g : α → simp_rw [eLpNorm'] rw [← ENNReal.rpow_le_rpow_iff hp, ENNReal.smul_def, smul_eq_mul, ENNReal.mul_rpow_of_nonneg _ _ hp.le] - simp_rw [← ENNReal.rpow_mul, one_div, inv_mul_cancel hp.ne.symm, ENNReal.rpow_one, + simp_rw [← ENNReal.rpow_mul, one_div, inv_mul_cancel₀ hp.ne.symm, ENNReal.rpow_one, ENNReal.coe_rpow_of_nonneg _ hp.le, ← lintegral_const_mul' _ _ ENNReal.coe_ne_top, ← ENNReal.coe_mul] apply lintegral_mono_ae @@ -1255,7 +1255,7 @@ theorem le_eLpNorm_of_bddBelow (hp : p ≠ 0) (hp' : p ≠ ∞) {f : α → F} ( rw [ENNReal.smul_def, smul_eq_mul, eLpNorm_eq_lintegral_rpow_nnnorm hp hp', one_div, ENNReal.le_rpow_inv_iff (ENNReal.toReal_pos hp hp'), ENNReal.mul_rpow_of_nonneg _ _ ENNReal.toReal_nonneg, ← ENNReal.rpow_mul, - inv_mul_cancel (ENNReal.toReal_pos hp hp').ne.symm, ENNReal.rpow_one, ← setLIntegral_const, + inv_mul_cancel₀ (ENNReal.toReal_pos hp hp').ne.symm, ENNReal.rpow_one, ← setLIntegral_const, ← lintegral_indicator _ hs] refine lintegral_mono_ae ?_ filter_upwards [hf] with x hx @@ -1347,7 +1347,7 @@ theorem ae_bdd_liminf_atTop_of_eLpNorm_bdd {p : ℝ≥0∞} (hp : p ≠ 0) {f : refine (OrderIso.liminf_apply (ENNReal.orderIsoRpow p.toReal _) ?_ ?_ ?_ ?_).symm <;> isBoundedDefault rw [this] at hx - rw [← ENNReal.rpow_one (liminf (fun n => ‖f n x‖₊) atTop), ← mul_inv_cancel hppos.ne.symm, + rw [← ENNReal.rpow_one (liminf (fun n => ‖f n x‖₊) atTop), ← mul_inv_cancel₀ hppos.ne.symm, ENNReal.rpow_mul] exact ENNReal.rpow_lt_top_of_nonneg (inv_nonneg.2 hppos.le) hx.ne diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean index 340db78f39239..74680c8ebc928 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean @@ -53,7 +53,7 @@ theorem eLpNorm'_le_eLpNormEssSup_mul_rpow_measure_univ {q : ℝ} (hq_pos : 0 < have h_nnnorm_le_eLpNorm_ess_sup := coe_nnnorm_ae_le_eLpNormEssSup f μ exact h_nnnorm_le_eLpNorm_ess_sup.mono fun x hx => by gcongr rw [eLpNorm', ← ENNReal.rpow_one (eLpNormEssSup f μ)] - nth_rw 2 [← mul_inv_cancel (ne_of_lt hq_pos).symm] + nth_rw 2 [← mul_inv_cancel₀ (ne_of_lt hq_pos).symm] rw [ENNReal.rpow_mul, one_div, ← ENNReal.mul_rpow_of_nonneg _ _ (by simp [hq_pos.le] : 0 ≤ q⁻¹)] gcongr rwa [lintegral_const] at h_le @@ -204,7 +204,7 @@ theorem eLpNorm_le_eLpNorm_top_mul_eLpNorm (p : ℝ≥0∞) (f : α → E) {g : swap · rw [one_div_nonneg] exact ENNReal.toReal_nonneg - rw [← ENNReal.rpow_mul, one_div, mul_inv_cancel, ENNReal.rpow_one] + rw [← ENNReal.rpow_mul, one_div, mul_inv_cancel₀, ENNReal.rpow_one] rw [Ne, ENNReal.toReal_eq_zero_iff, not_or] exact ⟨hp_zero, hp_top⟩ diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 8ff7405363c51..77bda2bdf1659 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -953,7 +953,7 @@ theorem memℒp_norm_rpow_iff {q : ℝ≥0∞} {f : α → E} (hf : AEStronglyMe convert h.norm_rpow_div q⁻¹ using 1 · ext x rw [Real.norm_eq_abs, Real.abs_rpow_of_nonneg (norm_nonneg _), ← Real.rpow_mul (abs_nonneg _), - ENNReal.toReal_inv, mul_inv_cancel, abs_of_nonneg (norm_nonneg _), Real.rpow_one] + ENNReal.toReal_inv, mul_inv_cancel₀, abs_of_nonneg (norm_nonneg _), Real.rpow_one] simp [ENNReal.toReal_eq_zero_iff, not_or, q_zero, q_top] · rw [div_eq_mul_inv, inv_inv, div_eq_mul_inv, mul_assoc, ENNReal.inv_mul_cancel q_zero q_top, mul_one] @@ -1339,7 +1339,7 @@ theorem eLpNorm'_lim_le_liminf_eLpNorm' {E} [NormedAddCommGroup E] {f : ℕ → refine (h_rpow_mono.orderIsoOfSurjective _ h_rpow_surj).liminf_apply ?_ ?_ ?_ ?_ all_goals isBoundedDefault rw [h_pow_liminf] - simp_rw [eLpNorm', ← ENNReal.rpow_mul, one_div, inv_mul_cancel hp_pos.ne.symm, ENNReal.rpow_one] + simp_rw [eLpNorm', ← ENNReal.rpow_mul, one_div, inv_mul_cancel₀ hp_pos.ne.symm, ENNReal.rpow_one] @[deprecated (since := "2024-07-27")] alias snorm'_lim_le_liminf_snorm' := eLpNorm'_lim_le_liminf_eLpNorm' diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index d84856b7e6f41..c5f62b96d4297 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -237,7 +237,7 @@ theorem norm_approxBounded_le {β} {f : α → β} [SeminormedAddCommGroup β] [ · rwa [one_le_div (lt_of_le_of_ne (norm_nonneg _) (Ne.symm h0))] · rw [min_eq_right _] · rw [norm_div, norm_norm, mul_comm, mul_div, div_eq_mul_inv, mul_comm, ← mul_assoc, - inv_mul_cancel h0, one_mul, Real.norm_of_nonneg hc] + inv_mul_cancel₀ h0, one_mul, Real.norm_of_nonneg hc] · rwa [div_le_one (lt_of_le_of_ne (norm_nonneg _) (Ne.symm h0))] theorem _root_.stronglyMeasurable_bot_iff [Nonempty β] [T2Space β] : diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean index 85ecc189b07e7..4251891161db2 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Lemmas.lean @@ -96,6 +96,6 @@ theorem aestronglyMeasurable_withDensity_iff {E : Type*} [NormedAddCommGroup E] refine ⟨fun x => (f x : ℝ)⁻¹ • g' x, hf.coe_nnreal_real.inv.stronglyMeasurable.smul g'meas, ?_⟩ rw [EventuallyEq, ae_withDensity_iff hf.coe_nnreal_ennreal] filter_upwards [hg'] with x hx h'x - rw [← hx, smul_smul, _root_.inv_mul_cancel, one_smul] + rw [← hx, smul_smul, inv_mul_cancel₀, one_smul] simp only [Ne, ENNReal.coe_eq_zero] at h'x simpa only [NNReal.coe_eq_zero, Ne] using h'x diff --git a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean index 5c422b2b762b1..f337d65513ca3 100644 --- a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean @@ -878,7 +878,7 @@ theorem UniformIntegrable.spec' (hp : p ≠ 0) (hp' : p ≠ ∞) (hf : ∀ i, St specialize this (2 * max M 1 * δ⁻¹ ^ (1 / p.toReal)) rw [ENNReal.coe_rpow_of_nonneg _ (one_div_nonneg.2 ENNReal.toReal_nonneg), ← ENNReal.coe_smul, smul_eq_mul, mul_assoc, NNReal.inv_rpow, - inv_mul_cancel (NNReal.rpow_pos (NNReal.coe_pos.1 hδpos)).ne.symm, mul_one, ENNReal.coe_mul, + inv_mul_cancel₀ (NNReal.rpow_pos (NNReal.coe_pos.1 hδpos)).ne.symm, mul_one, ENNReal.coe_mul, ← NNReal.inv_rpow] at this refine (lt_of_le_of_lt (le_trans (hM <| ℐ <| 2 * max M 1 * δ⁻¹ ^ (1 / p.toReal)) (le_max_left (M : ℝ≥0∞) 1)) diff --git a/Mathlib/MeasureTheory/Integral/CircleIntegral.lean b/Mathlib/MeasureTheory/Integral/CircleIntegral.lean index d7f84de54f927..1a25d84a8bf81 100644 --- a/Mathlib/MeasureTheory/Integral/CircleIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/CircleIntegral.lean @@ -492,7 +492,7 @@ theorem norm_cauchyPowerSeries_le (f : ℂ → E) (c : ℂ) (R : ℝ) (n : ℕ) _ ≤ ((2 * π)⁻¹ * ∫ θ : ℝ in (0)..2 * π, ‖f (circleMap c R θ)‖) * |R|⁻¹ ^ n := by rcases eq_or_ne R 0 with (rfl | hR) · cases n <;> simp [-mul_inv_rev] - rw [← mul_assoc, inv_mul_cancel (Real.two_pi_pos.ne.symm), one_mul] + rw [← mul_assoc, inv_mul_cancel₀ (Real.two_pi_pos.ne.symm), one_mul] apply norm_nonneg · rw [mul_inv_cancel_left₀, mul_assoc, mul_comm (|R|⁻¹ ^ n)] rwa [Ne, _root_.abs_eq_zero] diff --git a/Mathlib/MeasureTheory/Integral/Gamma.lean b/Mathlib/MeasureTheory/Integral/Gamma.lean index 90fa55677e921..e45745db0df27 100644 --- a/Mathlib/MeasureTheory/Integral/Gamma.lean +++ b/Mathlib/MeasureTheory/Integral/Gamma.lean @@ -43,7 +43,7 @@ theorem integral_rpow_mul_exp_neg_mul_rpow {p q b : ℝ} (hp : 0 < p) (hq : - 1 _ = ∫ x in Ioi (0 : ℝ), b ^ (-p⁻¹ * q) * ((b ^ p⁻¹ * x) ^ q * rexp (-(b ^ p⁻¹ * x) ^ p)) := by refine setIntegral_congr measurableSet_Ioi (fun _ hx => ?_) rw [mul_rpow _ (le_of_lt hx), mul_rpow _ (le_of_lt hx), ← rpow_mul, ← rpow_mul, - inv_mul_cancel, rpow_one, mul_assoc, ← mul_assoc, ← rpow_add, neg_mul p⁻¹, add_left_neg, + inv_mul_cancel₀, rpow_one, mul_assoc, ← mul_assoc, ← rpow_add, neg_mul p⁻¹, neg_add_cancel, rpow_zero, one_mul, neg_mul] all_goals positivity _ = (b ^ p⁻¹)⁻¹ * ∫ x in Ioi (0 : ℝ), b ^ (-p⁻¹ * q) * (x ^ q * rexp (-x ^ p)) := by diff --git a/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean b/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean index f1c89ef6de1a1..874a14fc17cc0 100644 --- a/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean +++ b/Mathlib/MeasureTheory/Integral/LebesgueNormedSpace.lean @@ -40,6 +40,6 @@ theorem aemeasurable_withDensity_iff {E : Type*} [NormedAddCommGroup E] [NormedS rw [EventuallyEq, ae_withDensity_iff hf.coe_nnreal_ennreal] filter_upwards [hg'] intro x hx h'x - rw [← hx, smul_smul, _root_.inv_mul_cancel, one_smul] + rw [← hx, smul_smul, inv_mul_cancel₀, one_smul] simp only [Ne, ENNReal.coe_eq_zero] at h'x simpa only [NNReal.coe_eq_zero, Ne] using h'x diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index cf188e01e74b8..af565327ab251 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -288,7 +288,7 @@ theorem tendsto_setIntegral_pow_smul_of_unique_maximum_of_isCompact_of_measure_n apply (hμ u u_open x₀_u).trans_le exact measure_mono fun x hx => ⟨ne_of_gt (pow_pos (a := c x) (hu hx) _), hx.2⟩ have hiφ : ∀ n, ∫ x in s, φ n x ∂μ = 1 := fun n => by - rw [integral_mul_left, inv_mul_cancel (P n).ne'] + rw [integral_mul_left, inv_mul_cancel₀ (P n).ne'] have A : ∀ u : Set α, IsOpen u → x₀ ∈ u → TendstoUniformlyOn φ 0 atTop (s \ u) := by intro u u_open x₀u obtain ⟨t, t_pos, tx₀, ht⟩ : ∃ t, 0 ≤ t ∧ t < c x₀ ∧ ∀ x ∈ s \ u, c x ≤ t := by @@ -448,7 +448,7 @@ theorem tendsto_integral_comp_smul_smul_of_integrable apply this.congr' filter_upwards [Ioi_mem_atTop 0] with c (hc : 0 < c) rw [integral_mul_left, setIntegral_comp_smul_of_pos _ _ _ hc, smul_eq_mul, ← mul_assoc, - mul_inv_cancel (by positivity), _root_.smul_closedBall _ _ zero_le_one] + mul_inv_cancel₀ (by positivity), _root_.smul_closedBall _ _ zero_le_one] simp [abs_of_nonneg hc.le] · filter_upwards [Ioi_mem_atTop 0] with c (hc : 0 < c) exact (I.comp_smul hc.ne').aestronglyMeasurable.const_mul _ diff --git a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean index 006c73306931e..23979e07b5cf8 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean @@ -155,7 +155,7 @@ theorem integrable_comp_smul_iff {E : Type*} [NormedAddCommGroup E] [NormedSpace ∀ {g : E → F} (_ : Integrable g μ) {S : ℝ} (_ : S ≠ 0), Integrable (fun x => g (S • x)) μ by refine ⟨fun hf => ?_, fun hf => this hf hR⟩ convert this hf (inv_ne_zero hR) - rw [← mul_smul, mul_inv_cancel hR, one_smul] + rw [← mul_smul, mul_inv_cancel₀ hR, one_smul] -- now prove intro g hg S hS let t := ((Homeomorph.smul (isUnit_iff_ne_zero.2 hS).unit).toMeasurableEquiv : E ≃ᵐ E) diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index fdad18683daec..cd5d79900442d 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -167,7 +167,7 @@ lemma integral_isMulLeftInvariant_isMulRightInvariant_combo exact h'g.comp_homeomorph ((Homeomorph.inv G).trans (Homeomorph.mulRight x)) calc ∫ x, f x ∂μ = ∫ x, f x * (D x)⁻¹ * D x ∂μ := by - congr with x; rw [mul_assoc, inv_mul_cancel (D_pos x).ne', mul_one] + congr with x; rw [mul_assoc, inv_mul_cancel₀ (D_pos x).ne', mul_one] _ = ∫ x, (∫ y, f x * (D x)⁻¹ * g (y⁻¹ * x) ∂ν) ∂μ := by simp_rw [integral_mul_left] _ = ∫ y, (∫ x, f x * (D x)⁻¹ * g (y⁻¹ * x) ∂μ) ∂ν := by apply integral_integral_swap_of_hasCompactSupport diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index 8bb812ec959ad..28921d6562aa6 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -590,7 +590,7 @@ theorem hausdorffMeasure_zero_or_top {d₁ d₂ : ℝ} (h : d₁ < d₂) (s : Se · have : (r : ℝ≥0∞) ≠ 0 := by simpa only [ENNReal.coe_eq_zero, Ne] using hr₀ rw [← ENNReal.rpow_sub _ _ this ENNReal.coe_ne_top] refine (ENNReal.rpow_lt_rpow hrc (sub_pos.2 h)).le.trans ?_ - rw [← ENNReal.rpow_mul, inv_mul_cancel (sub_pos.2 h).ne', ENNReal.rpow_one] + rw [← ENNReal.rpow_mul, inv_mul_cancel₀ (sub_pos.2 h).ne', ENNReal.rpow_one] /-- Hausdorff measure `μH[d] s` is monotone in `d`. -/ theorem hausdorffMeasure_mono {d₁ d₂ : ℝ} (h : d₁ ≤ d₂) (s : Set X) : μH[d₂] s ≤ μH[d₁] s := by @@ -1017,7 +1017,7 @@ theorem hausdorffMeasure_smul_right_image [NormedAddCommGroup E] [NormedSpace · exact hausdorffMeasure_real.symm have iso_smul : Isometry (LinearMap.toSpanSingleton ℝ E (‖v‖⁻¹ • v)) := by refine AddMonoidHomClass.isometry_of_norm _ fun x => (norm_smul _ _).trans ?_ - rw [norm_smul, norm_inv, norm_norm, inv_mul_cancel hn, mul_one, LinearMap.id_apply] + rw [norm_smul, norm_inv, norm_norm, inv_mul_cancel₀ hn, mul_one, LinearMap.id_apply] rw [Set.image_smul, Measure.hausdorffMeasure_smul₀ zero_le_one hn, nnnorm_norm, NNReal.rpow_one, iso_smul.hausdorffMeasure_image (Or.inl <| zero_le_one' ℝ)] diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean index 504b07f6391be..afb91c570543c 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean @@ -268,7 +268,7 @@ theorem map_volume_mul_left {a : ℝ} (h : a ≠ 0) : Measure.map (a * ·) volume = ENNReal.ofReal |a⁻¹| • volume := by conv_rhs => rw [← Real.smul_map_volume_mul_left h, smul_smul, ← ENNReal.ofReal_mul (abs_nonneg _), ← - abs_mul, inv_mul_cancel h, abs_one, ENNReal.ofReal_one, one_smul] + abs_mul, inv_mul_cancel₀ h, abs_one, ENNReal.ofReal_one, one_smul] @[simp] theorem volume_preimage_mul_left {a : ℝ} (h : a ≠ 0) (s : Set ℝ) : @@ -323,7 +323,7 @@ theorem smul_map_diagonal_volume_pi [DecidableEq ι] {D : ι → ℝ} (h : det ( simp only [det_diagonal, Ne] at h exact Finset.prod_ne_zero_iff.1 h i (Finset.mem_univ i) rw [volume_preimage_mul_left A, ← mul_assoc, ← ENNReal.ofReal_mul (abs_nonneg _), ← abs_mul, - mul_inv_cancel A, abs_one, ENNReal.ofReal_one, one_mul] + mul_inv_cancel₀ A, abs_one, ENNReal.ofReal_one, one_mul] rw [this, volume_pi_pi, Finset.abs_prod, ENNReal.ofReal_prod_of_nonneg fun i _ => abs_nonneg (D i), ← Finset.prod_mul_distrib] simp only [B] @@ -358,7 +358,7 @@ theorem map_matrix_volume_pi_eq_smul_volume_pi [DecidableEq ι] {M : Matrix ι apply diagonal_transvection_induction_of_det_ne_zero _ M hM · intro D hD conv_rhs => rw [← smul_map_diagonal_volume_pi hD] - rw [smul_smul, ← ENNReal.ofReal_mul (abs_nonneg _), ← abs_mul, inv_mul_cancel hD, abs_one, + rw [smul_smul, ← ENNReal.ofReal_mul (abs_nonneg _), ← abs_mul, inv_mul_cancel₀ hD, abs_one, ENNReal.ofReal_one, one_smul] · intro t simp_rw [Matrix.TransvectionStruct.det, _root_.inv_one, abs_one, ENNReal.ofReal_one, one_smul, diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index e40e2e449318e..666c2b7b7b87f 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -189,7 +189,7 @@ theorem addHaar_submodule {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] simpa only [sub_eq_zero, Ne] using (pow_right_strictAnti cpos cone).injective.ne hmn.symm have : x ∈ s := by convert s.smul_mem (c ^ n - c ^ m)⁻¹ A - rw [smul_smul, inv_mul_cancel H, one_smul] + rw [smul_smul, inv_mul_cancel₀ H, one_smul] exact hx this /-- A strict affine subspace has measure zero. -/ @@ -653,7 +653,7 @@ theorem tendsto_addHaar_inter_smul_zero_of_density_zero_aux2 (s : Set E) (x : E) ENNReal.ofReal_eq_zero, inv_eq_zero, inv_pow, Ne, or_self_iff, mul_eq_zero] · refine (smul_set_mono t_bound).trans_eq ?_ rw [smul_closedBall _ _ Rpos.le, smul_zero, Real.norm_of_nonneg (inv_nonneg.2 Rpos.le), - inv_mul_cancel Rpos.ne'] + inv_mul_cancel₀ Rpos.ne'] have B : Tendsto (fun r : ℝ => R * r) (𝓝[>] 0) (𝓝[>] (R * 0)) := by apply tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within · exact (tendsto_const_nhds.mul tendsto_id).mono_left nhdsWithin_le_nhds @@ -666,9 +666,9 @@ theorem tendsto_addHaar_inter_smul_zero_of_density_zero_aux2 (s : Set E) (x : E) filter_upwards [self_mem_nhdsWithin] rintro r - have T : (R * r) • t' = r • t := by - rw [mul_comm, ht', smul_smul, mul_assoc, mul_inv_cancel Rpos.ne', mul_one] + rw [mul_comm, ht', smul_smul, mul_assoc, mul_inv_cancel₀ Rpos.ne', mul_one] have U : (R * r) • u' = r • u := by - rw [mul_comm, hu', smul_smul, mul_assoc, mul_inv_cancel Rpos.ne', mul_one] + rw [mul_comm, hu', smul_smul, mul_assoc, mul_inv_cancel₀ Rpos.ne', mul_one] dsimp rw [T, U] diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 921ddc85dc021..1e59d29ef22d1 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -1728,7 +1728,7 @@ theorem pairwise_aedisjoint_of_aedisjoint_forall_ne_one {G α : Type*} [Group G] exact hg.symm have : (g₂⁻¹ • ·) ⁻¹' (g • s ∩ s) = g₁ • s ∩ g₂ • s := by rw [preimage_eq_iff_eq_image (MulAction.bijective g₂⁻¹), image_smul, smul_set_inter, smul_smul, - smul_smul, inv_mul_self, one_smul] + smul_smul, inv_mul_cancel, one_smul] change μ (g₁ • s ∩ g₂ • s) = 0 exact this ▸ (h_qmp g₂⁻¹).preimage_null (h_ae_disjoint g hg) diff --git a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean index 8e9f5d37c6482..e990eebe4d4d6 100644 --- a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean @@ -368,7 +368,7 @@ theorem normalize_eq_inv_mass_smul_of_nonzero (nonzero : μ ≠ 0) : μ.normalize.toFiniteMeasure = μ.mass⁻¹ • μ := by nth_rw 3 [μ.self_eq_mass_smul_normalize] rw [← smul_assoc] - simp only [μ.mass_nonzero_iff.mpr nonzero, Algebra.id.smul_eq_mul, inv_mul_cancel, Ne, + simp only [μ.mass_nonzero_iff.mpr nonzero, Algebra.id.smul_eq_mul, inv_mul_cancel₀, Ne, not_false_iff, one_smul] theorem toMeasure_normalize_eq_of_nonzero (nonzero : μ ≠ 0) : diff --git a/Mathlib/MeasureTheory/Measure/Tilted.lean b/Mathlib/MeasureTheory/Measure/Tilted.lean index c5d01ef4df8f6..0dcb5897a5937 100644 --- a/Mathlib/MeasureTheory/Measure/Tilted.lean +++ b/Mathlib/MeasureTheory/Measure/Tilted.lean @@ -63,7 +63,7 @@ lemma tilted_const' (μ : Measure α) (c : ℝ) : · simp only [h_univ, ENNReal.top_toReal, zero_mul, log_zero, div_zero, ENNReal.ofReal_zero, zero_smul, ENNReal.inv_top] congr - rw [div_eq_mul_inv, mul_inv, mul_comm, mul_assoc, inv_mul_cancel (exp_pos _).ne', mul_one, + rw [div_eq_mul_inv, mul_inv, mul_comm, mul_assoc, inv_mul_cancel₀ (exp_pos _).ne', mul_one, ← ENNReal.toReal_inv, ENNReal.ofReal_toReal] simp [h0.out] @@ -269,7 +269,7 @@ lemma tilted_tilted (hf : Integrable (fun x ↦ exp (f x)) μ) (g : α → ℝ) field_simp ring_nf congr 1 - rw [mul_assoc, mul_inv_cancel, mul_one] + rw [mul_assoc, mul_inv_cancel₀, mul_one] exact (integral_exp_pos hf).ne' lemma tilted_comm (hf : Integrable (fun x ↦ exp (f x)) μ) {g : α → ℝ} diff --git a/Mathlib/ModelTheory/Algebra/Field/Basic.lean b/Mathlib/ModelTheory/Algebra/Field/Basic.lean index afa727e1810e7..dba9790f9c93a 100644 --- a/Mathlib/ModelTheory/Algebra/Field/Basic.lean +++ b/Mathlib/ModelTheory/Algebra/Field/Basic.lean @@ -39,7 +39,7 @@ Language.ring.Sentence` -/ inductive FieldAxiom : Type | addAssoc : FieldAxiom | zeroAdd : FieldAxiom - | addLeftNeg : FieldAxiom + | negAddCancel : FieldAxiom | mulAssoc : FieldAxiom | mulComm : FieldAxiom | oneMul : FieldAxiom @@ -52,7 +52,7 @@ inductive FieldAxiom : Type def FieldAxiom.toSentence : FieldAxiom → Language.ring.Sentence | .addAssoc => ∀' ∀' ∀' (((&0 + &1) + &2) =' (&0 + (&1 + &2))) | .zeroAdd => ∀' (((0 : Language.ring.Term _) + &0) =' &0) - | .addLeftNeg => ∀' ∀' ((-&0 + &0) =' 0) + | .negAddCancel => ∀' ∀' ((-&0 + &0) =' 0) | .mulAssoc => ∀' ∀' ∀' (((&0 * &1) * &2) =' (&0 * (&1 * &2))) | .mulComm => ∀' ∀' ((&0 * &1) =' (&1 * &0)) | .oneMul => ∀' (((1 : Language.ring.Term _) * &0) =' &0) @@ -66,7 +66,7 @@ def FieldAxiom.toProp (K : Type*) [Add K] [Mul K] [Neg K] [Zero K] [One K] : FieldAxiom → Prop | .addAssoc => ∀ x y z : K, (x + y) + z = x + (y + z) | .zeroAdd => ∀ x : K, 0 + x = x - | .addLeftNeg => ∀ x : K, -x + x = 0 + | .negAddCancel => ∀ x : K, -x + x = 0 | .mulAssoc => ∀ x y z : K, (x * y) * z = x * (y * z) | .mulComm => ∀ x y : K, x * y = y * x | .oneMul => ∀ x : K, 1 * x = x @@ -114,7 +114,7 @@ noncomputable abbrev fieldOfModelField (K : Type*) [Language.ring.Structure K] Field.ofMinimalAxioms K addAssoc.toProp_of_model zeroAdd.toProp_of_model - addLeftNeg.toProp_of_model + negAddCancel.toProp_of_model mulAssoc.toProp_of_model mulComm.toProp_of_model oneMul.toProp_of_model @@ -147,10 +147,10 @@ instance [Field K] [CompatibleRing K] : Theory.field.Model K := rw [a.realize_toSentence_iff_toProp (K := K)] cases a with | existsPairNE => exact exists_pair_ne K - | existsInv => exact fun x hx0 => ⟨x⁻¹, mul_inv_cancel hx0⟩ + | existsInv => exact fun x hx0 => ⟨x⁻¹, mul_inv_cancel₀ hx0⟩ | addAssoc => exact add_assoc | zeroAdd => exact zero_add - | addLeftNeg => exact add_left_neg + | negAddCancel => exact neg_add_cancel | mulAssoc => exact mul_assoc | mulComm => exact mul_comm | oneMul => exact one_mul diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index bb4265f91c63c..ffd2563801e40 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -231,7 +231,7 @@ instance [NegZeroClass R] : Neg (ArithmeticFunction R) where instance [AddGroup R] : AddGroup (ArithmeticFunction R) := { ArithmeticFunction.instAddMonoid with - add_left_neg := fun _ => ext fun _ => add_left_neg _ + neg_add_cancel := fun _ => ext fun _ => neg_add_cancel _ zsmul := zsmulRec } instance [AddCommGroup R] : AddCommGroup (ArithmeticFunction R) := @@ -367,7 +367,7 @@ instance [CommSemiring R] : CommSemiring (ArithmeticFunction R) := instance [CommRing R] : CommRing (ArithmeticFunction R) := { ArithmeticFunction.instSemiring with - add_left_neg := add_left_neg + neg_add_cancel := neg_add_cancel mul_comm := mul_comm zsmul := (· • ·) } @@ -1076,7 +1076,7 @@ theorem moebius_mul_coe_zeta : (μ * ζ : ArithmeticFunction ℤ) = 1 := by rw [coe_mul_zeta_apply, sum_divisors_prime_pow hp, sum_range_succ'] simp_rw [Nat.pow_zero, moebius_apply_one, moebius_apply_prime_pow hp (Nat.succ_ne_zero _), Nat.succ_inj', sum_ite_eq', mem_range, - if_pos hn, add_left_neg] + if_pos hn, neg_add_cancel] rw [one_apply_ne] rw [Ne, pow_eq_one_iff] · exact hp.ne_one diff --git a/Mathlib/NumberTheory/ClassNumber/Finite.lean b/Mathlib/NumberTheory/ClassNumber/Finite.lean index 20c12455860af..5bc0f41a96267 100644 --- a/Mathlib/NumberTheory/ClassNumber/Finite.lean +++ b/Mathlib/NumberTheory/ClassNumber/Finite.lean @@ -194,7 +194,7 @@ theorem exists_mem_finsetApprox (a : S) {b} (hb : b ≠ (0 : R)) : have := normBound_pos abv bS have := abv.nonneg b rw [ε_eq, Algebra.smul_def, eq_intCast, mul_rpow, ← rpow_mul, div_mul_cancel₀, rpow_neg_one, - mul_left_comm, mul_inv_cancel, mul_one, rpow_natCast] <;> + mul_left_comm, mul_inv_cancel₀, mul_one, rpow_natCast] <;> try norm_cast; omega · exact Iff.mpr Int.cast_nonneg this · linarith diff --git a/Mathlib/NumberTheory/Cyclotomic/Three.lean b/Mathlib/NumberTheory/Cyclotomic/Three.lean index 034ec4ce85a40..a022fd1146b1b 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Three.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Three.lean @@ -144,7 +144,7 @@ lemma lambda_dvd_or_dvd_sub_one_or_dvd_add_one [NumberField K] [IsCyclotomicExte rw [RingHom.map_sub, h, RingHom.map_one, sub_self] · right; right refine Ideal.mem_span_singleton.1 <| Ideal.Quotient.eq_zero_iff_mem.1 ?_ - rw [RingHom.map_add, h, RingHom.map_one, add_left_neg] + rw [RingHom.map_add, h, RingHom.map_one, neg_add_cancel] /-- We have that `η ^ 2 + η + 1 = 0`. -/ lemma eta_sq_add_eta_add_one : (η : 𝓞 K) ^ 2 + η + 1 = 0 := by diff --git a/Mathlib/NumberTheory/Dioph.lean b/Mathlib/NumberTheory/Dioph.lean index 7bdaca5ee4308..d0e7977dfd15d 100644 --- a/Mathlib/NumberTheory/Dioph.lean +++ b/Mathlib/NumberTheory/Dioph.lean @@ -175,7 +175,7 @@ instance : AddCommGroup (Poly α) where zero_add _ := by ext; simp_rw [add_apply, zero_apply, zero_add] add_comm _ _ := by ext; simp_rw [add_apply, add_comm] add_assoc _ _ _ := by ext; simp_rw [add_apply, ← add_assoc] - add_left_neg _ := by ext; simp_rw [add_apply, neg_apply, add_left_neg, zero_apply] + neg_add_cancel _ := by ext; simp_rw [add_apply, neg_apply, neg_add_cancel, zero_apply] instance : AddGroupWithOne (Poly α) := { (inferInstance : AddCommGroup (Poly α)) with diff --git a/Mathlib/NumberTheory/GaussSum.lean b/Mathlib/NumberTheory/GaussSum.lean index 8dcfe9dbeb4e8..5140078baa06e 100644 --- a/Mathlib/NumberTheory/GaussSum.lean +++ b/Mathlib/NumberTheory/GaussSum.lean @@ -111,7 +111,7 @@ private theorem gaussSum_mul_aux {χ : MulChar R R'} (hχ : χ ≠ 1) (ψ : AddC Finset.sum_const_zero, map_zero_eq_one, mul_one, χ.sum_eq_zero_of_ne_one hχ] · -- case `b ≠ 0` refine (Fintype.sum_bijective _ (mulLeft_bijective₀ b hb) _ _ fun x ↦ ?_).symm - rw [mul_assoc, mul_comm x, ← mul_assoc, mul_inv_cancel hb, one_mul, mul_sub, mul_one] + rw [mul_assoc, mul_comm x, ← mul_assoc, mul_inv_cancel₀ hb, one_mul, mul_sub, mul_one] /-- We have `gaussSum χ ψ * gaussSum χ⁻¹ ψ⁻¹ = Fintype.card R` when `χ` is nontrivial and `ψ` is primitive (and `R` is a field). -/ diff --git a/Mathlib/NumberTheory/Harmonic/GammaDeriv.lean b/Mathlib/NumberTheory/Harmonic/GammaDeriv.lean index 0434d0130737f..35d6bf5409607 100644 --- a/Mathlib/NumberTheory/Harmonic/GammaDeriv.lean +++ b/Mathlib/NumberTheory/Harmonic/GammaDeriv.lean @@ -106,7 +106,7 @@ lemma hasDerivAt_Gamma_one_half : HasDerivAt Gamma (-√π * (γ + 2 * log 2)) ( rw [deriv_mul, Gamma_one_half_eq, add_assoc, ← mul_add, deriv_comp_add_const, (by norm_num : 1/2 + 1/2 = (1 : ℝ)), Gamma_one, mul_one, - eulerMascheroniConstant_eq_neg_deriv, add_neg_self, mul_zero, add_zero] + eulerMascheroniConstant_eq_neg_deriv, add_neg_cancel, mul_zero, add_zero] · apply h_diff; norm_num -- s = 1 · apply h_diff; norm_num -- s = 1/2 · exact ((h_diff (by norm_num)).hasDerivAt.comp_add_const).differentiableAt -- s = 1 @@ -220,7 +220,7 @@ lemma hasDerivAt_Gammaℝ_one : HasDerivAt Gammaℝ (-(γ + log (4 * π)) / 2) 1 refine HasDerivAt.congr_deriv (hf.mul hg) ?_ simp only [f] rw [Gamma_one_half_eq, aux, div_mul_cancel₀ _ aux2, neg_div _ (1 : ℂ), cpow_neg, aux, - mul_div_assoc, ← mul_assoc, mul_neg, inv_mul_cancel aux2, neg_one_mul, ← neg_div, + mul_div_assoc, ← mul_assoc, mul_neg, inv_mul_cancel₀ aux2, neg_one_mul, ← neg_div, ← _root_.add_div, ← neg_add, add_comm, add_assoc, ← ofReal_log Real.pi_pos.le, ← ofReal_ofNat, ← ofReal_log zero_le_two, ← ofReal_mul, ← Nat.cast_ofNat (R := ℝ), ← Real.log_pow, ← ofReal_add, diff --git a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean index b7ad4b399bedb..14bfeef13fb1d 100644 --- a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean +++ b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean @@ -144,7 +144,7 @@ lemma hf_zero (P : WeakFEPair E) (r : ℝ) : have h_nv : P.ε⁻¹ * ↑(x ^ P.k) ≠ 0 := mul_ne_zero P.symm.hε h_nv2 specialize hC' hx simp_rw [Function.comp_apply, ← one_div, P.h_feq' _ hx] at hC' - rw [← ((mul_inv_cancel h_nv).symm ▸ one_smul ℂ P.g₀ :), mul_smul _ _ P.g₀, ← smul_sub, norm_smul, + rw [← ((mul_inv_cancel₀ h_nv).symm ▸ one_smul ℂ P.g₀ :), mul_smul _ _ P.g₀, ← smul_sub, norm_smul, ← le_div_iff' (lt_of_le_of_ne (norm_nonneg _) (norm_ne_zero_iff.mpr h_nv).symm)] at hC' convert hC' using 1 · congr 3 @@ -430,7 +430,7 @@ theorem functional_equation (s : ℂ) : have := P.functional_equation₀ s rw [P.Λ₀_eq, P.symm_Λ₀_eq, sub_sub_cancel] at this rwa [smul_add, smul_add, ← mul_smul, mul_one_div, ← mul_smul, ← mul_div_assoc, - mul_inv_cancel P.hε, add_assoc, add_comm (_ • _), add_assoc, add_left_inj] at this + mul_inv_cancel₀ P.hε, add_assoc, add_comm (_ • _), add_assoc, add_left_inj] at this /-- The residue of `Λ` at `s = k` is equal to `ε • g₀`. -/ theorem Λ_residue_k : diff --git a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean index 7a65bfbbc95a0..48fb9b87f7ce9 100644 --- a/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean +++ b/Mathlib/NumberTheory/LSeries/HurwitzZetaEven.lean @@ -154,7 +154,7 @@ lemma evenKernel_functional_equation (a : UnitAddCircle) (x : ℝ) : rw [mul_pow, mul_pow, I_sq, div_eq_iff hx'] ring rw [h1, h2, h3, h4, ← mul_assoc, mul_comm (cexp _), mul_assoc _ (cexp _) (cexp _), - ← Complex.exp_add, neg_add_self, Complex.exp_zero, mul_one, ofReal_div, ofReal_one] + ← Complex.exp_add, neg_add_cancel, Complex.exp_zero, mul_one, ofReal_div, ofReal_one] end kernel_defs @@ -198,7 +198,7 @@ lemma hasSum_int_evenKernel₀ (a : ℝ) {t : ℝ} (ht : 0 < t) : split_ifs with h · obtain ⟨k, rfl⟩ := h simp_rw [← Int.cast_add, Int.cast_eq_zero, add_eq_zero_iff_eq_neg] - simpa only [Int.cast_add, neg_mul, Int.cast_neg, add_left_neg, ne_eq, OfNat.ofNat_ne_zero, + simpa only [Int.cast_add, neg_mul, Int.cast_neg, neg_add_cancel, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, zero_pow, mul_zero, zero_mul, Real.exp_zero] using hasSum_ite_sub_hasSum (hasSum_int_evenKernel (k : ℝ) ht) (-k) · suffices ∀ (n : ℤ), n + a ≠ 0 by simpa [this] using hasSum_int_evenKernel a ht @@ -657,7 +657,7 @@ lemma differentiableAt_hurwitzZetaEven_sub_one_div (a : UnitAddCircle) : simp_rw [← sub_div, div_eq_mul_inv _ (Gammaℝ _)] refine DifferentiableAt.mul ?_ differentiable_Gammaℝ_inv.differentiableAt simp_rw [completedHurwitzZetaEven_eq, sub_sub, add_assoc] - conv => enter [2, s, 2]; rw [← neg_sub, div_neg, neg_add_self, add_zero] + conv => enter [2, s, 2]; rw [← neg_sub, div_neg, neg_add_cancel, add_zero] exact (differentiable_completedHurwitzZetaEven₀ a _).sub <| (differentiableAt_const _).div differentiableAt_id one_ne_zero diff --git a/Mathlib/NumberTheory/LSeries/MellinEqDirichlet.lean b/Mathlib/NumberTheory/LSeries/MellinEqDirichlet.lean index fe9a56fdf81bc..7f8d7a14ef207 100644 --- a/Mathlib/NumberTheory/LSeries/MellinEqDirichlet.lean +++ b/Mathlib/NumberTheory/LSeries/MellinEqDirichlet.lean @@ -42,7 +42,7 @@ lemma hasSum_mellin {a : ι → ℂ} {p : ι → ℝ} {F : ℝ → ℂ} {s : ℂ refine (IntegrableOn.congr_fun (this.const_mul (1 / p i ^ (s - 1))) (fun t (ht : 0 < t) ↦ ?_) measurableSet_Ioi).const_mul _ simp_rw [mul_comm (↑(rexp _) : ℂ), ← mul_assoc, neg_mul, ofReal_mul] - rw [mul_cpow_ofReal_nonneg hpi.le ht.le, ← mul_assoc, one_div, inv_mul_cancel, one_mul] + rw [mul_cpow_ofReal_nonneg hpi.le ht.le, ← mul_assoc, one_div, inv_mul_cancel₀, one_mul] rw [Ne, cpow_eq_zero_iff, not_and_or] exact Or.inl (ofReal_ne_zero.mpr hpi.ne') · -- summability of integrals of norms diff --git a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean index 642151f5382bc..51f004593f0b5 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean @@ -69,7 +69,7 @@ theorem to_mulShift_inj_of_isPrimitive {ψ : AddChar R R'} (hψ : IsPrimitive ψ Function.Injective ψ.mulShift := by intro a b h apply_fun fun x => x * mulShift ψ (-b) at h - simp only [mulShift_mul, mulShift_zero, add_right_neg, mulShift_apply] at h + simp only [mulShift_mul, mulShift_zero, add_neg_cancel, mulShift_apply] at h simpa [← sub_eq_add_neg, sub_eq_zero] using (hψ · h) -- `AddCommGroup.equiv_direct_sum_zmod_of_fintype` diff --git a/Mathlib/NumberTheory/LegendreSymbol/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/Basic.lean index 7a21e08d10bab..f0a2686941e1c 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/Basic.lean @@ -212,7 +212,7 @@ theorem eq_one_of_sq_sub_mul_sq_eq_zero {p : ℕ} [Fact p.Prime] {a : ℤ} (ha : apply_fun (· * y⁻¹ ^ 2) at hxy simp only [zero_mul] at hxy rw [(by ring : (x ^ 2 - ↑a * y ^ 2) * y⁻¹ ^ 2 = (x * y⁻¹) ^ 2 - a * (y * y⁻¹) ^ 2), - mul_inv_cancel hy, one_pow, mul_one, sub_eq_zero, pow_two] at hxy + mul_inv_cancel₀ hy, one_pow, mul_one, sub_eq_zero, pow_two] at hxy exact (eq_one_iff p ha).mpr ⟨x * y⁻¹, hxy.symm⟩ /-- The Legendre symbol `legendreSym p a = 1` if there is a solution in `ℤ/pℤ` diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index d856b7b460b30..b73957e3efc28 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -243,7 +243,7 @@ theorem quadraticChar_card_sqrts (hF : ringChar F ≠ 2) (a : F) : rw [h₁, List.toFinset_cons, List.toFinset_cons, List.toFinset_nil] exact card_pair (Ne.symm (mt (Ring.eq_self_iff_eq_zero_of_char_ne_two hF).mp h₀)) · rw [quadraticChar_neg_one_iff_not_isSquare.mpr h] - simp only [add_left_neg, Int.natCast_eq_zero, card_eq_zero] + simp only [neg_add_cancel, Int.natCast_eq_zero, card_eq_zero] ext1 -- Porting note(https://github.com/leanprover-community/mathlib4/issues/5026): -- added (Set.mem_toFinset), Set.mem_setOf diff --git a/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean b/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean index 77b9e83e7512d..ac2e9e02f07f7 100644 --- a/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean +++ b/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean @@ -124,7 +124,7 @@ theorem mul_rat (h : LiouvilleWith p x) (hr : r ≠ 0) : LiouvilleWith p (x * r) `x` satisfies the same condition. -/ theorem mul_rat_iff (hr : r ≠ 0) : LiouvilleWith p (x * r) ↔ LiouvilleWith p x := ⟨fun h => by - simpa only [mul_assoc, ← Rat.cast_mul, mul_inv_cancel hr, Rat.cast_one, mul_one] using + simpa only [mul_assoc, ← Rat.cast_mul, mul_inv_cancel₀ hr, Rat.cast_one, mul_one] using h.mul_rat (inv_ne_zero hr), fun h => h.mul_rat hr⟩ @@ -280,7 +280,7 @@ protected theorem irrational (h : LiouvilleWith p x) (hp : 1 < p) : Irrational x rcases eq_or_ne r 0 with (rfl | h0) · refine h.ne_cast_int hp 0 ?_; rw [Rat.cast_zero, Int.cast_zero] · refine (h.mul_rat (inv_ne_zero h0)).ne_cast_int hp 1 ?_ - rw [Rat.cast_inv, mul_inv_cancel] + rw [Rat.cast_inv, mul_inv_cancel₀] exacts [Int.cast_one.symm, Rat.cast_ne_zero.mpr h0] end LiouvilleWith diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 66e2180484330..e14a58682b9f1 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -143,7 +143,7 @@ theorem tendsto_normSq_coprime_pair : dsimp only [Pi.smul_apply, LinearMap.pi_apply, smul_eq_mul] fin_cases i · show (z : ℂ).im⁻¹ * (f c).im = c 0 - rw [f_def, add_im, im_ofReal_mul, ofReal_im, add_zero, mul_left_comm, inv_mul_cancel hz, + rw [f_def, add_im, im_ofReal_mul, ofReal_im, add_zero, mul_left_comm, inv_mul_cancel₀ hz, mul_one] · show (z : ℂ).im⁻¹ * ((z : ℂ) * conj (f c)).im = c 1 rw [f_def, RingHom.map_add, RingHom.map_mul, mul_add, mul_left_comm, mul_conj, conj_ofReal, diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean index 9895ab0022fde..91bae3555b311 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/Defs.lean @@ -75,10 +75,10 @@ def gammaSetEquiv (γ : SL(2, ℤ)) : gammaSet N a ≃ gammaSet N (a ᵥ* γ) wh have := vecMul_SL2_mem_gammaSet v.2 γ⁻¹ rw [vecMul_vecMul, ← SpecialLinearGroup.coe_mul] at this simpa only [SpecialLinearGroup.map_apply_coe, RingHom.mapMatrix_apply, Int.coe_castRingHom, - map_inv, mul_right_inv, SpecialLinearGroup.coe_one, vecMul_one]⟩ - left_inv v := by simp_rw [vecMul_vecMul, ← SpecialLinearGroup.coe_mul, mul_inv_self, + map_inv, mul_inv_cancel, SpecialLinearGroup.coe_one, vecMul_one]⟩ + left_inv v := by simp_rw [vecMul_vecMul, ← SpecialLinearGroup.coe_mul, mul_inv_cancel, SpecialLinearGroup.coe_one, vecMul_one] - right_inv v := by simp_rw [vecMul_vecMul, ← SpecialLinearGroup.coe_mul, inv_mul_self, + right_inv v := by simp_rw [vecMul_vecMul, ← SpecialLinearGroup.coe_mul, inv_mul_cancel, SpecialLinearGroup.coe_one, vecMul_one] end gamma_action diff --git a/Mathlib/NumberTheory/ModularForms/JacobiTheta/TwoVariable.lean b/Mathlib/NumberTheory/ModularForms/JacobiTheta/TwoVariable.lean index a727659658f0b..10883cf58e5b4 100644 --- a/Mathlib/NumberTheory/ModularForms/JacobiTheta/TwoVariable.lean +++ b/Mathlib/NumberTheory/ModularForms/JacobiTheta/TwoVariable.lean @@ -227,7 +227,7 @@ lemma summable_jacobiTheta₂'_term_iff (z τ : ℂ) : simp_rw [norm_mul, norm_I, norm_real, mul_one, norm_of_nonneg pi_pos.le, ← ofReal_ofNat, norm_real, norm_of_nonneg two_pos.le, ← ofReal_intCast, norm_real, Real.norm_eq_abs, ← Int.cast_abs, ← mul_assoc _ (2 * π), - inv_mul_cancel (mul_pos two_pos pi_pos).ne', one_mul] + inv_mul_cancel₀ (mul_pos two_pos pi_pos).ne', one_mul] rw [← Int.cast_one, Int.cast_le] exact Int.one_le_abs hn · refine fun hτ ↦ ((summable_pow_mul_jacobiTheta₂_term_bound diff --git a/Mathlib/NumberTheory/ModularForms/SlashActions.lean b/Mathlib/NumberTheory/ModularForms/SlashActions.lean index 4a597199dca6a..5f521ac77b3f0 100644 --- a/Mathlib/NumberTheory/ModularForms/SlashActions.lean +++ b/Mathlib/NumberTheory/ModularForms/SlashActions.lean @@ -57,7 +57,7 @@ open scoped ModularForm theorem SlashAction.neg_slash {β G α γ : Type*} [Group G] [AddGroup α] [SMul γ α] [SlashAction β G α γ] (k : β) (g : G) (a : α) : (-a) ∣[k;γ] g = -a ∣[k;γ] g := eq_neg_of_add_eq_zero_left <| by - rw [← SlashAction.add_slash, add_left_neg, SlashAction.zero_slash] + rw [← SlashAction.add_slash, neg_add_cancel, SlashAction.zero_slash] @[simp] theorem SlashAction.smul_slash_of_tower {R β G α : Type*} (γ : Type*) [Group G] [AddGroup α] diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index 500ac261a79da..f9a5760d1bcda 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -323,7 +323,7 @@ noncomputable instance commGroup : CommGroup (MulChar R R') := { one := 1 mul := (· * ·) inv := Inv.inv - mul_left_inv := inv_mul + inv_mul_cancel := inv_mul mul_assoc := by intro χ₁ χ₂ χ₃ ext a @@ -491,7 +491,7 @@ theorem IsQuadratic.inv {χ : MulChar R R'} (hχ : χ.IsQuadratic) : χ⁻¹ = /-- The square of a quadratic character is the trivial character. -/ theorem IsQuadratic.sq_eq_one {χ : MulChar R R'} (hχ : χ.IsQuadratic) : χ ^ 2 = 1 := by - rw [← mul_left_inv χ, pow_two, hχ.inv] + rw [← inv_mul_cancel χ, pow_two, hχ.inv] /-- The `p`th power of a quadratic character is itself, when `p` is the (prime) characteristic of the target ring. -/ diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index e452129c03f2a..f34495bae5580 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -140,8 +140,8 @@ theorem adjust_f {w₁ : InfinitePlace K} (B : ℝ≥0) (hf : ∀ w, w ≠ w₁ · exact fun w hw => Function.update_noteq hw _ f · rw [← Finset.mul_prod_erase Finset.univ _ (Finset.mem_univ w₁), Function.update_same, Finset.prod_congr rfl fun w hw => by rw [Function.update_noteq (Finset.ne_of_mem_erase hw)], - ← NNReal.rpow_natCast, ← NNReal.rpow_mul, inv_mul_cancel, NNReal.rpow_one, mul_assoc, - inv_mul_cancel, mul_one] + ← NNReal.rpow_natCast, ← NNReal.rpow_mul, inv_mul_cancel₀, NNReal.rpow_one, mul_assoc, + inv_mul_cancel₀, mul_one] · rw [Finset.prod_ne_zero_iff] exact fun w hw => pow_ne_zero _ (hf w (Finset.ne_of_mem_erase hw)) · rw [mult]; split_ifs <;> norm_num @@ -615,7 +615,7 @@ theorem exists_ne_zero_mem_ideal_of_norm_le {B : ℝ} obtain ⟨a, ha, rfl⟩ := hx refine ⟨a, ha, by simpa using h_nz, ?_⟩ rw [← rpow_natCast, ← rpow_le_rpow_iff (by simp only [Rat.cast_abs, abs_nonneg]) - (rpow_nonneg h2 _) h1, ← rpow_mul h2, mul_inv_cancel (Nat.cast_ne_zero.mpr + (rpow_nonneg h2 _) h1, ← rpow_mul h2, mul_inv_cancel₀ (Nat.cast_ne_zero.mpr (ne_of_gt finrank_pos)), rpow_one, le_div_iff' (Nat.cast_pos.mpr finrank_pos)] refine le_trans ?_ ((convexBodySum_mem K B).mp h_mem) rw [← le_div_iff' (Nat.cast_pos.mpr finrank_pos), ← sum_mult_eq, Nat.cast_sum] diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index 954e318cf0962..9bcddc91ed612 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -113,7 +113,7 @@ theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal refine le_of_eq ?_ rw [convexBodySum_volume, ← ENNReal.ofReal_pow (by positivity), ← Real.rpow_natCast, ← Real.rpow_mul toReal_nonneg, div_mul_cancel₀, Real.rpow_one, ofReal_toReal, mul_comm, - mul_assoc, ← coe_mul, inv_mul_cancel (convexBodySumFactor_ne_zero K), ENNReal.coe_one, + mul_assoc, ← coe_mul, inv_mul_cancel₀ (convexBodySumFactor_ne_zero K), ENNReal.coe_one, mul_one] · exact mul_ne_top (ne_of_lt (minkowskiBound_lt_top K I)) coe_ne_top · exact (Nat.cast_ne_zero.mpr (ne_of_gt finrank_pos)) @@ -288,7 +288,7 @@ theorem rank_le_rankOfDiscrBdd : refine lt_of_le_of_lt ?_ (mul_lt_mul_of_pos_left (Real.rpow_lt_rpow_of_exponent_lt h₂ h) (by positivity : (0 : ℝ) < 4 / 9)) rw [Real.rpow_logb (lt_trans zero_lt_one h₂) (ne_of_gt h₂) (by positivity), ← mul_assoc, - ← inv_div, inv_mul_cancel (by norm_num), one_mul, Int.cast_natCast] + ← inv_div, inv_mul_cancel₀ (by norm_num), one_mul, Int.cast_natCast] · refine div_nonneg (Real.log_nonneg ?_) (Real.log_nonneg (le_of_lt h₂)) rw [mul_comm, ← mul_div_assoc, _root_.le_div_iff (by positivity), one_mul, ← _root_.div_le_iff (by positivity)] diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 5a609f303cc0d..db8688cc1ec6d 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -297,7 +297,7 @@ theorem exists_unit (w₁ : InfinitePlace K) : · calc _ = w (algebraMap (𝓞 K) K (seq K w₁ hB m) * (algebraMap (𝓞 K) K (seq K w₁ hB n))⁻¹) := by rw [← congr_arg (algebraMap (𝓞 K) K) hu.choose_spec, mul_comm, map_mul (algebraMap _ _), - ← mul_assoc, inv_mul_cancel (seq_ne_zero K w₁ hB n), one_mul] + ← mul_assoc, inv_mul_cancel₀ (seq_ne_zero K w₁ hB n), one_mul] _ = w (algebraMap (𝓞 K) K (seq K w₁ hB m)) * w (algebraMap (𝓞 K) K (seq K w₁ hB n))⁻¹ := _root_.map_mul _ _ _ _ < 1 := by diff --git a/Mathlib/NumberTheory/Padics/Hensel.lean b/Mathlib/NumberTheory/Padics/Hensel.lean index 3a6b3479d5a4f..80a8898200c2a 100644 --- a/Mathlib/NumberTheory/Padics/Hensel.lean +++ b/Mathlib/NumberTheory/Padics/Hensel.lean @@ -188,7 +188,7 @@ private def calc_eval_z' {z z' z1 : ℤ_[p]} (hz' : z' = z - z1) {n} (hz : ih n (Subtype.ext <| by simp only [PadicInt.coe_neg, PadicInt.coe_mul, Subtype.coe_mk]) _ = -F.eval z := by simp only [mul_div_cancel₀ _ hdzne', Subtype.coe_eta] - exact ⟨q, by simpa only [sub_eq_add_neg, this, hz', add_right_neg, neg_sq, zero_add] using hq⟩ + exact ⟨q, by simpa only [sub_eq_add_neg, this, hz', add_neg_cancel, neg_sq, zero_add] using hq⟩ private def calc_eval_z'_norm {z z' z1 : ℤ_[p]} {n} (hz : ih n z) {q} (heq : F.eval z' = q * z1 ^ 2) diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 9bcf9525b8ada..2076ea8bd3206 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -392,7 +392,7 @@ theorem mul_inv : ∀ {z : ℤ_[p]}, ‖z‖ = 1 → z * z.inv = 1 dsimp only rw [dif_pos h] apply Subtype.ext_iff_val.2 - simp [mul_inv_cancel hk] + simp [mul_inv_cancel₀ hk] theorem inv_mul {z : ℤ_[p]} (hz : ‖z‖ = 1) : z.inv * z = 1 := by rw [mul_comm, mul_inv hz] @@ -509,7 +509,7 @@ theorem norm_lt_one_iff_dvd (x : ℤ_[p]) : ‖x‖ < 1 ↔ ↑p ∣ x := by have := norm_le_pow_iff_mem_span_pow x 1 rw [Ideal.mem_span_singleton, pow_one] at this rw [← this, norm_le_pow_iff_norm_lt_pow_add_one] - simp only [zpow_zero, Int.ofNat_zero, Int.ofNat_succ, add_left_neg, zero_add] + simp only [zpow_zero, Int.ofNat_zero, Int.ofNat_succ, neg_add_cancel, zero_add] @[simp] theorem pow_p_dvd_int_iff (n : ℕ) (a : ℤ) : (p : ℤ_[p]) ^ n ∣ a ↔ (p ^ n : ℤ) ∣ a := by @@ -612,7 +612,7 @@ instance isFractionRing : IsFractionRing ℤ_[p] ℚ_[p] where rw [h0, norm_zero] at hx exact hx zero_le_one rw [ha, padicNormE.mul, padicNormE.norm_p_pow, Padic.norm_eq_pow_val hx, ← zpow_add', - hn_coe, neg_neg, add_left_neg, zpow_zero] + hn_coe, neg_neg, neg_add_cancel, zpow_zero] exact Or.inl (Nat.cast_ne_zero.mpr (NeZero.ne p)) use (⟨a, le_of_eq ha_norm⟩, diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 15cddbfe318d9..368932b6dad66 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -352,7 +352,7 @@ protected theorem pow {q : ℚ} (hq : q ≠ 0) {k : ℕ} : protected theorem inv (q : ℚ) : padicValRat p q⁻¹ = -padicValRat p q := by by_cases hq : q = 0 · simp [hq] - · rw [eq_neg_iff_add_eq_zero, ← padicValRat.mul (inv_ne_zero hq) hq, inv_mul_cancel hq, + · rw [eq_neg_iff_add_eq_zero, ← padicValRat.mul (inv_ne_zero hq) hq, inv_mul_cancel₀ hq, padicValRat.one] /-- A rewrite lemma for `padicValRat p (q / r)` with conditions `q ≠ 0`, `r ≠ 0`. -/ diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index 80b91e674b707..3012b77db01b5 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -624,7 +624,7 @@ theorem eq_pow_of_nonneg {a₁ : Solution₁ d} (h : IsFundamental a₁) {a : So lift (a * a₁⁻¹).x to ℕ using hxx₁.le with x' hx' -- Porting note: `ih` has its arguments in a different order compared to lean 3. obtain ⟨n, hn⟩ := ih x' (mod_cast hxx₂.trans_eq hax'.symm) hyy hx' hxx₁ - exact ⟨n + 1, by rw [pow_succ', ← hn, mul_comm a, ← mul_assoc, mul_inv_self, one_mul]⟩ + exact ⟨n + 1, by rw [pow_succ', ← hn, mul_comm a, ← mul_assoc, mul_inv_cancel, one_mul]⟩ /-- Every solution is, up to a sign, a power of a given fundamental solution. -/ theorem eq_zpow_or_neg_zpow {a₁ : Solution₁ d} (h : IsFundamental a₁) (a : Solution₁ d) : diff --git a/Mathlib/NumberTheory/PellMatiyasevic.lean b/Mathlib/NumberTheory/PellMatiyasevic.lean index 10ce7abc5d7bd..806449a094861 100644 --- a/Mathlib/NumberTheory/PellMatiyasevic.lean +++ b/Mathlib/NumberTheory/PellMatiyasevic.lean @@ -280,7 +280,7 @@ theorem eq_pell_lem : ∀ (n) (b : ℤ√(d a1)), 1 ≤ b → IsPell b → (mul_le_mul_of_nonneg_left hn (le_trans zero_le_one h1)) a1p erw [bm, one_mul, mul_assoc, Eq.trans (mul_comm _ _) a1m, mul_one] at t exact ha t - simp only [sub_self, sub_neg_eq_add] at y0l; simp only [Zsqrtd.neg_re, add_right_neg, + simp only [sub_self, sub_neg_eq_add] at y0l; simp only [Zsqrtd.neg_re, add_neg_cancel, Zsqrtd.neg_im, neg_neg] at yl2 exact match y, y0l, (yl2 : (⟨_, _⟩ : ℤ√_) < ⟨_, _⟩) with diff --git a/Mathlib/NumberTheory/Wilson.lean b/Mathlib/NumberTheory/Wilson.lean index e4cefac632d88..fb2001f449b1a 100644 --- a/Mathlib/NumberTheory/Wilson.lean +++ b/Mathlib/NumberTheory/Wilson.lean @@ -90,7 +90,7 @@ theorem prime_of_fac_equiv_neg_one (h : ((n - 1)! : ZMod n) = -1) (h1 : n ≠ 1) obtain ⟨m, hm1, hm2 : 1 < m, hm3⟩ := exists_dvd_of_not_prime2 h1 h2 have hm : m ∣ (n - 1)! := Nat.dvd_factorial (pos_of_gt hm2) (le_pred_of_lt hm3) refine hm2.ne' (Nat.dvd_one.mp ((Nat.dvd_add_right hm).mp (hm1.trans ?_))) - rw [← ZMod.natCast_zmod_eq_zero_iff_dvd, cast_add, cast_one, h, add_left_neg] + rw [← ZMod.natCast_zmod_eq_zero_iff_dvd, cast_add, cast_one, h, neg_add_cancel] /-- **Wilson's Theorem**: For `n ≠ 1`, `(n-1)!` is congruent to `-1` modulo `n` iff n is prime. -/ theorem prime_iff_fac_equiv_neg_one (h : n ≠ 1) : Prime n ↔ ((n - 1)! : ZMod n) = -1 := by diff --git a/Mathlib/NumberTheory/Zsqrtd/Basic.lean b/Mathlib/NumberTheory/Zsqrtd/Basic.lean index 1a153a411f763..a697ad60f1bbe 100644 --- a/Mathlib/NumberTheory/Zsqrtd/Basic.lean +++ b/Mathlib/NumberTheory/Zsqrtd/Basic.lean @@ -139,7 +139,7 @@ instance addCommGroup : AddCommGroup (ℤ√d) := by add_assoc := ?_ zero_add := ?_ add_zero := ?_ - add_left_neg := ?_ + neg_add_cancel := ?_ add_comm := ?_ } <;> intros <;> ext <;> diff --git a/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean b/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean index 7d9de95aa9292..a22d6696a31a2 100644 --- a/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean +++ b/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean @@ -50,7 +50,7 @@ theorem mod_four_eq_three_of_nat_prime_of_prime (p : ℕ) [hp : Fact p.Prime] exact ⟨k.val, k.val_lt, ZMod.natCast_zmod_val k⟩ have hpk : p ∣ k ^ 2 + 1 := by rw [pow_two, ← CharP.cast_eq_zero_iff (ZMod p) p, Nat.cast_add, Nat.cast_mul, Nat.cast_one, - ← hk, add_left_neg] + ← hk, neg_add_cancel] have hkmul : (k ^ 2 + 1 : ℤ[i]) = ⟨k, 1⟩ * ⟨k, -1⟩ := by ext <;> simp [sq] have hkltp : 1 + k * k < p * p := calc diff --git a/Mathlib/Order/Filter/FilterProduct.lean b/Mathlib/Order/Filter/FilterProduct.lean index 5d0a1167920d6..827c876602408 100644 --- a/Mathlib/Order/Filter/FilterProduct.lean +++ b/Mathlib/Order/Filter/FilterProduct.lean @@ -37,7 +37,7 @@ instance instGroupWithZero [GroupWithZero β] : GroupWithZero β* where __ := instDivInvMonoid __ := instMonoidWithZero mul_inv_cancel f := inductionOn f fun f hf ↦ coe_eq.2 <| (φ.em fun y ↦ f y = 0).elim - (fun H ↦ (hf <| coe_eq.2 H).elim) fun H ↦ H.mono fun x ↦ mul_inv_cancel + (fun H ↦ (hf <| coe_eq.2 H).elim) fun H ↦ H.mono fun x ↦ mul_inv_cancel₀ inv_zero := coe_eq.2 <| by simp only [Function.comp, inv_zero, EventuallyEq.rfl] instance instDivisionSemiring [DivisionSemiring β] : DivisionSemiring β* where diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index ba315579aa20a..52f1af5d388e6 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -502,7 +502,7 @@ instance instDivisionMonoid [DivisionMonoid G] : DivisionMonoid (Germ l G) where @[to_additive] instance instGroup [Group G] : Group (Germ l G) := - { mul_left_inv := Quotient.ind' fun _ => congrArg ofFun <| mul_left_inv _ } + { inv_mul_cancel := Quotient.ind' fun _ => congrArg ofFun <| inv_mul_cancel _ } @[to_additive] instance instCommGroup [CommGroup G] : CommGroup (Germ l G) := diff --git a/Mathlib/Order/RelIso/Group.lean b/Mathlib/Order/RelIso/Group.lean index 56ebf504387ae..352ac47981f5a 100644 --- a/Mathlib/Order/RelIso/Group.lean +++ b/Mathlib/Order/RelIso/Group.lean @@ -22,7 +22,7 @@ instance : Group (r ≃r r) where mul_assoc f₁ f₂ f₃ := rfl one_mul f := ext fun _ => rfl mul_one f := ext fun _ => rfl - mul_left_inv f := ext f.symm_apply_apply + inv_mul_cancel f := ext f.symm_apply_apply @[simp] theorem coe_one : ((1 : r ≃r r) : α → α) = id := diff --git a/Mathlib/Probability/Distributions/Gaussian.lean b/Mathlib/Probability/Distributions/Gaussian.lean index 4c319723ef538..b31dba6a5ddf3 100644 --- a/Mathlib/Probability/Distributions/Gaussian.lean +++ b/Mathlib/Probability/Distributions/Gaussian.lean @@ -140,7 +140,7 @@ lemma gaussianPDFReal_inv_mul {μ : ℝ} {v : ℝ≥0} {c : ℝ} (hc : c ≠ 0) ring_nf calc (Real.sqrt ↑v)⁻¹ * (Real.sqrt 2)⁻¹ * (Real.sqrt π)⁻¹ = (Real.sqrt ↑v)⁻¹ * (Real.sqrt 2)⁻¹ * (Real.sqrt π)⁻¹ * (|c| * |c|⁻¹) := by - rw [mul_inv_cancel, mul_one] + rw [mul_inv_cancel₀, mul_one] simp only [ne_eq, abs_eq_zero, hc, not_false_eq_true] _ = (Real.sqrt ↑v)⁻¹ * (Real.sqrt 2)⁻¹ * (Real.sqrt π)⁻¹ * |c| * |c|⁻¹ := by ring · congr 1 @@ -302,7 +302,7 @@ lemma gaussianReal_map_const_mul (c : ℝ) : Equiv.coe_fn_symm_mk, gaussianPDFReal_inv_mul hc] congr with x suffices |c⁻¹| * |c| = 1 by rw [← mul_assoc, this, one_mul] - rw [abs_inv, inv_mul_cancel] + rw [abs_inv, inv_mul_cancel₀] rwa [ne_eq, abs_eq_zero] /-- The map of a Gaussian distribution by multiplication by a constant is a Gaussian. -/ diff --git a/Mathlib/Probability/Kernel/RadonNikodym.lean b/Mathlib/Probability/Kernel/RadonNikodym.lean index ab83c54d2f019..7f5426267ce77 100644 --- a/Mathlib/Probability/Kernel/RadonNikodym.lean +++ b/Mathlib/Probability/Kernel/RadonNikodym.lean @@ -302,7 +302,7 @@ lemma singularPart_compl_mutuallySingularSetSlice (κ η : Kernel α γ) [IsSFin simp only [mem_compl_iff, mutuallySingularSetSlice, mem_setOf, not_le] at hx simp_rw [rnDeriv] rw [← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, - mul_inv_cancel, one_mul, tsub_self, Pi.zero_apply] + mul_inv_cancel₀, one_mul, tsub_self, Pi.zero_apply] · simp only [ne_eq, sub_eq_zero, hx.ne', not_false_eq_true] · simp only [sub_nonneg, hx.le] · simp only [sub_pos, hx] @@ -375,7 +375,7 @@ lemma withDensity_rnDeriv_of_subset_compl_mutuallySingularSetSlice _ = ∫⁻ x in s, ENNReal.ofReal (rnDerivAux κ (κ + η) a x) ∂(κ + η) a := by refine setLIntegral_congr_fun hsm (ae_of_all _ fun x hx ↦ ?_) rw [h_coe, ← ENNReal.ofReal_div_of_pos, div_eq_inv_mul, ← ENNReal.ofReal_mul, ← mul_assoc, - mul_inv_cancel, one_mul] + mul_inv_cancel₀, one_mul] · rw [ne_eq, sub_eq_zero] exact (hs' x hx).ne' · simp [(hs' x hx).le] diff --git a/Mathlib/Probability/Variance.lean b/Mathlib/Probability/Variance.lean index de8f0b0c69fc3..b737727f95f56 100644 --- a/Mathlib/Probability/Variance.lean +++ b/Mathlib/Probability/Variance.lean @@ -60,7 +60,7 @@ theorem _root_.MeasureTheory.Memℒp.evariance_lt_top [IsFiniteMeasure μ] (hX : have := ENNReal.pow_lt_top (hX.sub <| memℒp_const <| μ[X]).2 2 rw [eLpNorm_eq_lintegral_rpow_nnnorm two_ne_zero ENNReal.two_ne_top, ← ENNReal.rpow_two] at this simp only [ENNReal.toReal_ofNat, Pi.sub_apply, ENNReal.one_toReal, one_div] at this - rw [← ENNReal.rpow_mul, inv_mul_cancel (two_ne_zero : (2 : ℝ) ≠ 0), ENNReal.rpow_one] at this + rw [← ENNReal.rpow_mul, inv_mul_cancel₀ (two_ne_zero : (2 : ℝ) ≠ 0), ENNReal.rpow_one] at this simp_rw [ENNReal.rpow_two] at this exact this @@ -246,7 +246,7 @@ theorem meas_ge_le_evariance_div_sq {X : Ω → ℝ} (hX : AEStronglyMeasurable ENNReal.toReal_ofNat, one_div, Pi.sub_apply] rw [div_eq_mul_inv, ENNReal.inv_pow, mul_comm, ENNReal.rpow_two] congr - simp_rw [← ENNReal.rpow_mul, inv_mul_cancel (two_ne_zero : (2 : ℝ) ≠ 0), ENNReal.rpow_two, + simp_rw [← ENNReal.rpow_mul, inv_mul_cancel₀ (two_ne_zero : (2 : ℝ) ≠ 0), ENNReal.rpow_two, ENNReal.rpow_one, evariance] /-- **Chebyshev's inequality**: one can control the deviation probability of a real random variable diff --git a/Mathlib/RepresentationTheory/Action/Basic.lean b/Mathlib/RepresentationTheory/Action/Basic.lean index 1d925a3a897c2..6002bcee80785 100644 --- a/Mathlib/RepresentationTheory/Action/Basic.lean +++ b/Mathlib/RepresentationTheory/Action/Basic.lean @@ -51,8 +51,8 @@ def ρAut {G : Grp.{u}} (A : Action V (MonCat.of G)) : G ⟶ Grp.of (Aut A.V) wh toFun g := { hom := A.ρ g inv := A.ρ (g⁻¹ : G) - hom_inv_id := (A.ρ.map_mul (g⁻¹ : G) g).symm.trans (by rw [inv_mul_self, ρ_one]) - inv_hom_id := (A.ρ.map_mul g (g⁻¹ : G)).symm.trans (by rw [mul_inv_self, ρ_one]) } + hom_inv_id := (A.ρ.map_mul (g⁻¹ : G) g).symm.trans (by rw [inv_mul_cancel, ρ_one]) + inv_hom_id := (A.ρ.map_mul g (g⁻¹ : G)).symm.trans (by rw [mul_inv_cancel, ρ_one]) } map_one' := Aut.ext A.ρ.map_one map_mul' x y := Aut.ext (A.ρ.map_mul x y) diff --git a/Mathlib/RepresentationTheory/Action/Limits.lean b/Mathlib/RepresentationTheory/Action/Limits.lean index 63eb1d784aeb4..60d9a97e9d74f 100644 --- a/Mathlib/RepresentationTheory/Action/Limits.lean +++ b/Mathlib/RepresentationTheory/Action/Limits.lean @@ -244,7 +244,7 @@ instance : Preadditive (Action V G) where zero_add := by intros; ext; exact zero_add _ add_zero := by intros; ext; exact add_zero _ add_assoc := by intros; ext; exact add_assoc _ _ _ - add_left_neg := by intros; ext; exact add_left_neg _ + neg_add_cancel := by intros; ext; exact neg_add_cancel _ add_comm := by intros; ext; exact add_comm _ _ } add_comp := by intros; ext; exact Preadditive.add_comp _ _ _ _ _ _ comp_add := by intros; ext; exact Preadditive.comp_add _ _ _ _ _ _ diff --git a/Mathlib/RepresentationTheory/Action/Monoidal.lean b/Mathlib/RepresentationTheory/Action/Monoidal.lean index dc555899d679c..831b1acf591af 100644 --- a/Mathlib/RepresentationTheory/Action/Monoidal.lean +++ b/Mathlib/RepresentationTheory/Action/Monoidal.lean @@ -295,13 +295,13 @@ noncomputable def leftRegularTensorIso (G : Type u) [Group G] (X : Action (Type funext x refine Prod.ext rfl ?_ change (X.ρ x.1 * X.ρ (x.1⁻¹ : G)) x.2 = x.2 - rw [← X.ρ.map_mul, mul_inv_self, X.ρ.map_one, MonCat.one_of, End.one_def, types_id_apply] + rw [← X.ρ.map_mul, mul_inv_cancel, X.ρ.map_one, MonCat.one_of, End.one_def, types_id_apply] inv_hom_id := by apply Hom.ext funext x refine Prod.ext rfl ?_ change (X.ρ (x.1⁻¹ : G) * X.ρ x.1) x.2 = x.2 - rw [← X.ρ.map_mul, inv_mul_self, X.ρ.map_one, MonCat.one_of, End.one_def, types_id_apply] + rw [← X.ρ.map_mul, inv_mul_cancel, X.ρ.map_one, MonCat.one_of, End.one_def, types_id_apply] /-- The natural isomorphism of `G`-sets `Gⁿ⁺¹ ≅ G × Gⁿ`, where `G` acts by left multiplication on each factor. -/ diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index bd825851e6089..c0689ba27a5a7 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -240,7 +240,7 @@ theorem mem_oneCocycles_iff (f : G → A) : @[simp] theorem oneCocycles_map_inv (f : oneCocycles A) (g : G) : A.ρ g (f.1 g⁻¹) = - f.1 g := by - rw [← add_eq_zero_iff_eq_neg, ← oneCocycles_map_one f, ← mul_inv_self g, + rw [← add_eq_zero_iff_eq_neg, ← oneCocycles_map_one f, ← mul_inv_cancel g, (mem_oneCocycles_iff f.1).1 f.2 g g⁻¹] theorem oneCocycles_map_mul_of_isTrivial [A.IsTrivial] (f : oneCocycles A) (g h : G) : @@ -302,7 +302,7 @@ lemma twoCocycles_ρ_map_inv_sub_map_inv (f : twoCocycles A) (g : G) : A.ρ g (f.1 (g⁻¹, g)) - f.1 (g, g⁻¹) = f.1 (1, 1) - f.1 (g, 1) := by have := (mem_twoCocycles_iff f.1).1 f.2 g g⁻¹ g - simp only [mul_right_inv, mul_left_inv, twoCocycles_map_one_fst _ g] + simp only [mul_inv_cancel, inv_mul_cancel, twoCocycles_map_one_fst _ g] at this exact sub_eq_sub_iff_add_eq_add.2 this.symm @@ -418,12 +418,12 @@ variable {G A : Type*} [Group G] [AddCommGroup A] [MulAction G A] @[scoped simp] theorem map_inv_of_isOneCocycle {f : G → A} (hf : IsOneCocycle f) (g : G) : g • f g⁻¹ = - f g := by - rw [← add_eq_zero_iff_eq_neg, ← map_one_of_isOneCocycle hf, ← mul_inv_self g, hf g g⁻¹] + rw [← add_eq_zero_iff_eq_neg, ← map_one_of_isOneCocycle hf, ← mul_inv_cancel g, hf g g⁻¹] theorem smul_map_inv_sub_map_inv_of_isTwoCocycle {f : G × G → A} (hf : IsTwoCocycle f) (g : G) : g • f (g⁻¹, g) - f (g, g⁻¹) = f (1, 1) - f (g, 1) := by have := hf g g⁻¹ g - simp only [mul_right_inv, mul_left_inv, map_one_fst_of_isTwoCocycle hf g] at this + simp only [mul_inv_cancel, inv_mul_cancel, map_one_fst_of_isTwoCocycle hf g] at this exact sub_eq_sub_iff_add_eq_add.2 this.symm end @@ -541,13 +541,13 @@ variable {G M : Type*} [Group G] [CommGroup M] [MulAction G M] @[scoped simp] theorem map_inv_of_isMulOneCocycle {f : G → M} (hf : IsMulOneCocycle f) (g : G) : g • f g⁻¹ = (f g)⁻¹ := by - rw [← mul_eq_one_iff_eq_inv, ← map_one_of_isMulOneCocycle hf, ← mul_inv_self g, hf g g⁻¹] + rw [← mul_eq_one_iff_eq_inv, ← map_one_of_isMulOneCocycle hf, ← mul_inv_cancel g, hf g g⁻¹] theorem smul_map_inv_div_map_inv_of_isMulTwoCocycle {f : G × G → M} (hf : IsMulTwoCocycle f) (g : G) : g • f (g⁻¹, g) / f (g, g⁻¹) = f (1, 1) / f (g, 1) := by have := hf g g⁻¹ g - simp only [mul_right_inv, mul_left_inv, map_one_fst_of_isMulTwoCocycle hf g] at this + simp only [mul_inv_cancel, inv_mul_cancel, map_one_fst_of_isMulTwoCocycle hf g] at this exact div_eq_div_iff_mul_eq_mul.2 this.symm end diff --git a/Mathlib/RepresentationTheory/Maschke.lean b/Mathlib/RepresentationTheory/Maschke.lean index 1c9739d487652..f35efe285d18d 100644 --- a/Mathlib/RepresentationTheory/Maschke.lean +++ b/Mathlib/RepresentationTheory/Maschke.lean @@ -77,7 +77,7 @@ section theorem conjugate_i (h : ∀ v : V, (π : W → V) (i v) = v) (g : G) (v : V) : (conjugate π g : W → V) (i v) = v := by - rw [conjugate_apply, ← i.map_smul, h, ← mul_smul, single_mul_single, mul_one, mul_left_inv, + rw [conjugate_apply, ← i.map_smul, h, ← mul_smul, single_mul_single, mul_one, inv_mul_cancel, ← one_def, one_smul] end diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index 3bc762a3a2580..8cf8b61925ff3 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -88,12 +88,12 @@ theorem of_ρ_apply {V : Type u} [AddCommGroup V] [Module k V] (ρ : Representat @[simp] theorem ρ_inv_self_apply {G : Type u} [Group G] (A : Rep k G) (g : G) (x : A) : A.ρ g⁻¹ (A.ρ g x) = x := - show (A.ρ g⁻¹ * A.ρ g) x = x by rw [← map_mul, inv_mul_self, map_one, LinearMap.one_apply] + show (A.ρ g⁻¹ * A.ρ g) x = x by rw [← map_mul, inv_mul_cancel, map_one, LinearMap.one_apply] @[simp] theorem ρ_self_inv_apply {G : Type u} [Group G] {A : Rep k G} (g : G) (x : A) : A.ρ g (A.ρ g⁻¹ x) = x := - show (A.ρ g * A.ρ g⁻¹) x = x by rw [← map_mul, mul_inv_self, map_one, LinearMap.one_apply] + show (A.ρ g * A.ρ g⁻¹) x = x by rw [← map_mul, mul_inv_cancel, map_one, LinearMap.one_apply] theorem hom_comm_apply {A B : Rep k G} (f : A ⟶ B) (g : G) (x : A) : f.hom (A.ρ g x) = B.ρ g (f.hom x) := @@ -354,7 +354,7 @@ def homEquiv (A B C : Rep k G) : (A ⊗ B ⟶ C) ≃ (B ⟶ (Rep.ihom A).obj C) change f.hom (_ ⊗ₜ[k] _) = C.ρ g (f.hom (_ ⊗ₜ[k] _)) rw [← hom_comm_apply] change _ = f.hom ((A.ρ g * A.ρ g⁻¹) y ⊗ₜ[k] _) - simp only [← map_mul, mul_inv_self, map_one] + simp only [← map_mul, mul_inv_cancel, map_one] rfl } invFun f := { hom := TensorProduct.uncurry k _ _ _ f.hom.flip diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index e72d6a31ef4c3..96954689661a2 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -454,7 +454,7 @@ theorem Subalgebra.isField_of_algebraic [Algebra.IsAlgebraic K L] : IsField A := { show Nontrivial A by infer_instance, Subalgebra.toCommRing A with mul_inv_cancel := fun {a} ha => ⟨⟨a⁻¹, A.inv_mem_of_algebraic (Algebra.IsAlgebraic.isAlgebraic (a : L))⟩, - Subtype.ext (mul_inv_cancel (mt (Subalgebra.coe_eq_zero _).mp ha))⟩ } + Subtype.ext (mul_inv_cancel₀ (mt (Subalgebra.coe_eq_zero _).mp ha))⟩ } end Field diff --git a/Mathlib/RingTheory/Binomial.lean b/Mathlib/RingTheory/Binomial.lean index 82a59a2f62044..481d3ed558e54 100644 --- a/Mathlib/RingTheory/Binomial.lean +++ b/Mathlib/RingTheory/Binomial.lean @@ -289,7 +289,7 @@ theorem smeval_ascPochhammer_self_neg : ∀ n : ℕ, theorem smeval_ascPochhammer_succ_neg (n : ℕ) : smeval (ascPochhammer ℕ (n + 1)) (-n : ℤ) = 0 := by rw [ascPochhammer_succ_right, smeval_mul, smeval_add, smeval_X, ← C_eq_natCast, smeval_C, - pow_zero, pow_one, Nat.cast_id, nsmul_eq_mul, mul_one, add_left_neg, mul_zero] + pow_zero, pow_one, Nat.cast_id, nsmul_eq_mul, mul_one, neg_add_cancel, mul_zero] theorem smeval_ascPochhammer_neg_add (n : ℕ) : ∀ k : ℕ, smeval (ascPochhammer ℕ (n + k + 1)) (-n : ℤ) = 0 diff --git a/Mathlib/RingTheory/Coprime/Basic.lean b/Mathlib/RingTheory/Coprime/Basic.lean index 9e261696f011a..ea3e3dce7eff2 100644 --- a/Mathlib/RingTheory/Coprime/Basic.lean +++ b/Mathlib/RingTheory/Coprime/Basic.lean @@ -232,7 +232,7 @@ variable {R G : Type*} [CommSemiring R] [Group G] [MulAction G R] [SMulCommClass theorem isCoprime_group_smul_left : IsCoprime (x • y) z ↔ IsCoprime y z := ⟨fun ⟨a, b, h⟩ => ⟨x • a, b, by rwa [smul_mul_assoc, ← mul_smul_comm]⟩, fun ⟨a, b, h⟩ => - ⟨x⁻¹ • a, b, by rwa [smul_mul_smul, inv_mul_self, one_smul]⟩⟩ + ⟨x⁻¹ • a, b, by rwa [smul_mul_smul, inv_mul_cancel, one_smul]⟩⟩ theorem isCoprime_group_smul_right : IsCoprime y (x • z) ↔ IsCoprime y z := isCoprime_comm.trans <| (isCoprime_group_smul_left x z y).trans isCoprime_comm diff --git a/Mathlib/RingTheory/DedekindDomain/Different.lean b/Mathlib/RingTheory/DedekindDomain/Different.lean index e4d4261b6eac5..915f686f016f1 100644 --- a/Mathlib/RingTheory/DedekindDomain/Different.lean +++ b/Mathlib/RingTheory/DedekindDomain/Different.lean @@ -319,24 +319,24 @@ variable (I) lemma inv_le_dual : I⁻¹ ≤ dual A K I := - if hI : I = 0 then by simp [hI] else le_dual_inv_aux A K hI (le_of_eq (mul_inv_cancel hI)) + if hI : I = 0 then by simp [hI] else le_dual_inv_aux A K hI (le_of_eq (mul_inv_cancel₀ hI)) lemma dual_inv_le : (dual A K I)⁻¹ ≤ I := by by_cases hI : I = 0; · simp [hI] convert mul_right_mono ((dual A K I)⁻¹) (mul_left_mono I (inv_le_dual A K I)) using 1 - · simp only [mul_inv_cancel hI, one_mul] - · simp only [mul_inv_cancel (dual_ne_zero A K (hI := hI)), mul_assoc, mul_one] + · simp only [mul_inv_cancel₀ hI, one_mul] + · simp only [mul_inv_cancel₀ (dual_ne_zero A K (hI := hI)), mul_assoc, mul_one] lemma dual_eq_mul_inv : dual A K I = dual A K 1 * I⁻¹ := by by_cases hI : I = 0; · simp [hI] apply le_antisymm · suffices dual A K I * I ≤ dual A K 1 by - convert mul_right_mono I⁻¹ this using 1; simp only [mul_inv_cancel hI, mul_one, mul_assoc] + convert mul_right_mono I⁻¹ this using 1; simp only [mul_inv_cancel₀ hI, mul_one, mul_assoc] rw [← le_dual_iff A K hI] - rw [le_dual_iff A K hI, mul_assoc, inv_mul_cancel hI, mul_one] + rw [le_dual_iff A K hI, mul_assoc, inv_mul_cancel₀ hI, mul_one] variable {I} @@ -348,7 +348,7 @@ lemma dual_div_dual : lemma dual_mul_self (hI : I ≠ 0) : dual A K I * I = dual A K 1 := by - rw [dual_eq_mul_inv, mul_assoc, inv_mul_cancel hI, mul_one] + rw [dual_eq_mul_inv, mul_assoc, inv_mul_cancel₀ hI, mul_one] lemma self_mul_dual (hI : I ≠ 0) : I * dual A K I = dual A K 1 := by @@ -362,7 +362,7 @@ variable (I) @[simp] lemma dual_dual : dual A K (dual A K I) = I := by - rw [dual_eq_mul_inv, dual_eq_mul_inv A K (I := I), mul_inv, inv_inv, ← mul_assoc, mul_inv_cancel, + rw [dual_eq_mul_inv, dual_eq_mul_inv A K (I := I), mul_inv, inv_inv, ← mul_assoc, mul_inv_cancel₀, one_mul] rw [dual_ne_zero_iff] exact one_ne_zero diff --git a/Mathlib/RingTheory/DedekindDomain/Factorization.lean b/Mathlib/RingTheory/DedekindDomain/Factorization.lean index c07bfc797cc24..21265293630ad 100644 --- a/Mathlib/RingTheory/DedekindDomain/Factorization.lean +++ b/Mathlib/RingTheory/DedekindDomain/Factorization.lean @@ -403,7 +403,7 @@ theorem count_neg_zpow (n : ℤ) (I : FractionalIdeal R⁰ K) : · rw [hn, neg_zero, zpow_zero, count_one, neg_zero] · rw [hI, zero_zpow n hn, zero_zpow (-n) (neg_ne_zero.mpr hn), count_zero, neg_zero] · rw [eq_neg_iff_add_eq_zero, ← count_mul K v (zpow_ne_zero _ hI) (zpow_ne_zero _ hI), - ← zpow_add₀ hI, neg_add_self, zpow_zero] + ← zpow_add₀ hI, neg_add_cancel, zpow_zero] exact count_one K v theorem count_inv (I : FractionalIdeal R⁰ K) : @@ -499,7 +499,7 @@ theorem count_mono {I J} (hI : I ≠ 0) (h : I ≤ J) : count K v J ≤ count K by_cases hJ : J = 0 · exact (hI (FractionalIdeal.le_zero_iff.mp (h.trans hJ.le))).elim have := FractionalIdeal.mul_le_mul_left h J⁻¹ - rw [inv_mul_cancel hJ, FractionalIdeal.le_one_iff_exists_coeIdeal] at this + rw [inv_mul_cancel₀ hJ, FractionalIdeal.le_one_iff_exists_coeIdeal] at this obtain ⟨J', hJ'⟩ := this rw [← mul_inv_cancel_left₀ hJ I, ← hJ', count_mul K v hJ, le_add_iff_nonneg_right] · exact count_coe_nonneg K v J' diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index d71fe641b714b..5cdfa4e768306 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -145,7 +145,7 @@ theorem coe_ideal_span_singleton_div_self {x : R₁} (hx : x ≠ 0) : theorem spanSingleton_mul_inv {x : K} (hx : x ≠ 0) : spanSingleton R₁⁰ x * (spanSingleton R₁⁰ x)⁻¹ = 1 := by - rw [spanSingleton_inv, spanSingleton_mul_spanSingleton, mul_inv_cancel hx, spanSingleton_one] + rw [spanSingleton_inv, spanSingleton_mul_spanSingleton, mul_inv_cancel₀ hx, spanSingleton_one] theorem coe_ideal_span_singleton_mul_inv {x : R₁} (hx : x ≠ 0) : (Ideal.span ({x} : Set R₁) : FractionalIdeal R₁⁰ K) * @@ -167,7 +167,7 @@ theorem mul_generator_self_inv {R₁ : Type*} [CommRing R₁] [Algebra R₁ K] [ I * spanSingleton _ (generator (I : Submodule R₁ K))⁻¹ = 1 := by -- Rewrite only the `I` that appears alone. conv_lhs => congr; rw [eq_spanSingleton_of_principal I] - rw [spanSingleton_mul_spanSingleton, mul_inv_cancel, spanSingleton_one] + rw [spanSingleton_mul_spanSingleton, mul_inv_cancel₀, spanSingleton_one] intro generator_I_eq_zero apply h rw [eq_spanSingleton_of_principal I, generator_I_eq_zero, spanSingleton_zero] @@ -416,7 +416,7 @@ lemma not_inv_le_one_of_ne_bot [IsDedekindDomain A] {I : Ideal A} exact Submodule.smul_mem_smul h_Iy hbZ rw [Ideal.mem_span_singleton'] at h_yb rcases h_yb with ⟨c, hc⟩ - rw [← hc, RingHom.map_mul, mul_assoc, mul_inv_cancel hnz_fa, mul_one] + rw [← hc, RingHom.map_mul, mul_assoc, mul_inv_cancel₀ hnz_fa, mul_one] apply coe_mem_one · refine mt (mem_one_iff _).mp ?_ rintro ⟨x', h₂_abs⟩ @@ -498,7 +498,7 @@ protected theorem mul_inv_cancel [IsDedekindDomain A] {I : FractionalIdeal A⁰ exact ⟨spanSingleton A⁰ (algebraMap _ _ a) * (J : FractionalIdeal A⁰ K)⁻¹, h₂⟩ subst hJ rw [mul_assoc, mul_left_comm (J : FractionalIdeal A⁰ K), coe_ideal_mul_inv, mul_one, - spanSingleton_mul_spanSingleton, inv_mul_cancel, spanSingleton_one] + spanSingleton_mul_spanSingleton, inv_mul_cancel₀, spanSingleton_one] · exact mt ((injective_iff_map_eq_zero (algebraMap A K)).mp (IsFractionRing.injective A K) _) ha · exact coeIdeal_ne_zero.mp (right_ne_zero_of_mul hne) @@ -588,11 +588,11 @@ theorem Ideal.dvd_iff_le {I J : Ideal A} : I ∣ J ↔ J ≤ I := have hI' : (I : FractionalIdeal A⁰ (FractionRing A)) ≠ 0 := coeIdeal_ne_zero.mpr hI have : (I : FractionalIdeal A⁰ (FractionRing A))⁻¹ * J ≤ 1 := le_trans (mul_left_mono (↑I)⁻¹ ((coeIdeal_le_coeIdeal _).mpr h)) - (le_of_eq (inv_mul_cancel hI')) + (le_of_eq (inv_mul_cancel₀ hI')) obtain ⟨H, hH⟩ := le_one_iff_exists_coeIdeal.mp this use H refine coeIdeal_injective (show (J : FractionalIdeal A⁰ (FractionRing A)) = ↑(I * H) from ?_) - rw [coeIdeal_mul, hH, ← mul_assoc, mul_inv_cancel hI', one_mul]⟩ + rw [coeIdeal_mul, hH, ← mul_assoc, mul_inv_cancel₀ hI', one_mul]⟩ theorem Ideal.dvdNotUnit_iff_lt {I J : Ideal A} : DvdNotUnit I J ↔ J < I := ⟨fun ⟨hI, H, hunit, hmul⟩ => diff --git a/Mathlib/RingTheory/Derivation/Basic.lean b/Mathlib/RingTheory/Derivation/Basic.lean index 8b4495bf94265..31605cb3d9c2b 100644 --- a/Mathlib/RingTheory/Derivation/Basic.lean +++ b/Mathlib/RingTheory/Derivation/Basic.lean @@ -398,7 +398,7 @@ variable {K : Type*} [Field K] [Module K M] [Algebra R K] (D : Derivation R K M) theorem leibniz_inv (a : K) : D a⁻¹ = -a⁻¹ ^ 2 • D a := by rcases eq_or_ne a 0 with (rfl | ha) · simp - · exact D.leibniz_of_mul_eq_one (inv_mul_cancel ha) + · exact D.leibniz_of_mul_eq_one (inv_mul_cancel₀ ha) theorem leibniz_div (a b : K) : D (a / b) = b⁻¹ ^ 2 • (b • D a - a • D b) := by simp only [div_eq_mul_inv, leibniz, leibniz_inv, inv_pow, neg_smul, smul_neg, smul_smul, add_comm, diff --git a/Mathlib/RingTheory/FractionalIdeal/Operations.lean b/Mathlib/RingTheory/FractionalIdeal/Operations.lean index 6836737387416..56982ee5c698d 100644 --- a/Mathlib/RingTheory/FractionalIdeal/Operations.lean +++ b/Mathlib/RingTheory/FractionalIdeal/Operations.lean @@ -736,7 +736,7 @@ theorem div_spanSingleton (J : FractionalIdeal R₁⁰ K) (d : K) : rw [coe_mul, one_div_spanSingleton, h_xd] exact Submodule.mul_mem_mul (mem_spanSingleton_self R₁⁰ _) hx · rw [le_div_iff_mul_le h_spand, mul_assoc, mul_left_comm, one_div_spanSingleton, - spanSingleton_mul_spanSingleton, inv_mul_cancel hd, spanSingleton_one, mul_one] + spanSingleton_mul_spanSingleton, inv_mul_cancel₀ hd, spanSingleton_one, mul_one] theorem exists_eq_spanSingleton_mul (I : FractionalIdeal R₁⁰ K) : ∃ (a : R₁) (aI : Ideal R₁), a ≠ 0 ∧ I = spanSingleton R₁⁰ (algebraMap R₁ K a)⁻¹ * aI := by @@ -753,12 +753,12 @@ theorem exists_eq_spanSingleton_mul (I : FractionalIdeal R₁⁰ K) : rw [Algebra.smul_def] at hx' refine ⟨algebraMap R₁ K x', (mem_coeIdeal _).mpr ⟨x', mem_singleton_mul.mpr ?_, rfl⟩, ?_⟩ · exact ⟨x, hx, hx'⟩ - · rw [hx', ← mul_assoc, inv_mul_cancel map_a_nonzero, one_mul] + · rw [hx', ← mul_assoc, inv_mul_cancel₀ map_a_nonzero, one_mul] · rintro ⟨y, hy, rfl⟩ obtain ⟨x', hx', rfl⟩ := (mem_coeIdeal _).mp hy obtain ⟨y', hy', hx'⟩ := mem_singleton_mul.mp hx' rw [Algebra.linearMap_apply] at hx' - rwa [hx', ← mul_assoc, inv_mul_cancel map_a_nonzero, one_mul] + rwa [hx', ← mul_assoc, inv_mul_cancel₀ map_a_nonzero, one_mul] /-- If `I` is a nonzero fractional ideal, `a ∈ R`, and `J` is an ideal of `R` such that @@ -857,7 +857,7 @@ theorem isNoetherian_spanSingleton_inv_to_map_mul (x : R₁) {I : FractionalIdea obtain ⟨s, hs⟩ := hI _ hJ use s * {(algebraMap R₁ K x)⁻¹} rw [Finset.coe_mul, Finset.coe_singleton, ← span_mul_span, hs, ← coe_spanSingleton R₁⁰, ← - coe_mul, mul_assoc, spanSingleton_mul_spanSingleton, mul_inv_cancel h_gx, spanSingleton_one, + coe_mul, mul_assoc, spanSingleton_mul_spanSingleton, mul_inv_cancel₀ h_gx, spanSingleton_one, mul_one] /-- Every fractional ideal of a noetherian integral domain is noetherian. -/ diff --git a/Mathlib/RingTheory/FreeCommRing.lean b/Mathlib/RingTheory/FreeCommRing.lean index f6594407ec2f7..eeaec370fa743 100644 --- a/Mathlib/RingTheory/FreeCommRing.lean +++ b/Mathlib/RingTheory/FreeCommRing.lean @@ -94,7 +94,7 @@ protected theorem induction_on {C : FreeCommRing α → Prop} (z : FreeCommRing C z := have hn : ∀ x, C x → C (-x) := fun x ih => neg_one_mul x ▸ hm _ _ hn1 ih have h1 : C 1 := neg_neg (1 : FreeCommRing α) ▸ hn _ hn1 - FreeAbelianGroup.induction_on z (add_left_neg (1 : FreeCommRing α) ▸ ha _ _ hn1 h1) + FreeAbelianGroup.induction_on z (neg_add_cancel (1 : FreeCommRing α) ▸ ha _ _ hn1 h1) (fun m => Multiset.induction_on m h1 fun a m ih => by convert hm (of a) _ (hb a) ih apply of_cons) diff --git a/Mathlib/RingTheory/FreeRing.lean b/Mathlib/RingTheory/FreeRing.lean index 570d7148f3688..1ec0166f4f531 100644 --- a/Mathlib/RingTheory/FreeRing.lean +++ b/Mathlib/RingTheory/FreeRing.lean @@ -57,7 +57,7 @@ protected theorem induction_on {C : FreeRing α → Prop} (z : FreeRing α) (hn1 C z := have hn : ∀ x, C x → C (-x) := fun x ih => neg_one_mul x ▸ hm _ _ hn1 ih have h1 : C 1 := neg_neg (1 : FreeRing α) ▸ hn _ hn1 - FreeAbelianGroup.induction_on z (add_left_neg (1 : FreeRing α) ▸ ha _ _ hn1 h1) + FreeAbelianGroup.induction_on z (neg_add_cancel (1 : FreeRing α) ▸ ha _ _ hn1 h1) (fun m => List.recOn m h1 fun a m ih => by -- Porting note: in mathlib, convert was not necessary, `exact hm _ _ (hb a) ih` worked fine convert hm _ _ (hb a) ih diff --git a/Mathlib/RingTheory/HahnSeries/Addition.lean b/Mathlib/RingTheory/HahnSeries/Addition.lean index e177e516056f3..61a63a5fceb4a 100644 --- a/Mathlib/RingTheory/HahnSeries/Addition.lean +++ b/Mathlib/RingTheory/HahnSeries/Addition.lean @@ -244,9 +244,9 @@ instance : Neg (HahnSeries Γ R) where instance : AddGroup (HahnSeries Γ R) := { inferInstanceAs (AddMonoid (HahnSeries Γ R)) with zsmul := zsmulRec - add_left_neg := fun x => by + neg_add_cancel := fun x => by ext - apply add_left_neg } + apply neg_add_cancel } @[simp] theorem neg_coeff' {x : HahnSeries Γ R} : (-x).coeff = -x.coeff := diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index e27893105186a..732438273dbdb 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -199,9 +199,9 @@ instance : Neg (SummableFamily Γ R α) := instance : AddCommGroup (SummableFamily Γ R α) := { inferInstanceAs (AddCommMonoid (SummableFamily Γ R α)) with zsmul := zsmulRec - add_left_neg := fun a => by + neg_add_cancel := fun a => by ext - apply add_left_neg } + apply neg_add_cancel } @[simp] theorem coe_neg : ⇑(-s) = -s := @@ -471,12 +471,12 @@ theorem unit_aux (x : HahnSeries Γ R) {r : R} (hr : r * x.leadingCoeff = 1) : by_cases h : x = 0; · simp [h] rw [← order_eq_orderTop_of_ne h, orderTop_single (fun _ => by simp_all only [zero_mul, zero_ne_one]), ← @WithTop.coe_add, - WithTop.coe_nonneg, add_left_neg] + WithTop.coe_nonneg, neg_add_cancel] · apply coeff_orderTop_ne h.symm simp only [C_apply, single_mul_single, zero_add, mul_one, sub_coeff', Pi.sub_apply, one_coeff, ↓reduceIte] have hrc := mul_coeff_order_add_order ((single (-x.order)) r) x - rw [order_single hrz, leadingCoeff_of_single, neg_add_self, hr] at hrc + rw [order_single hrz, leadingCoeff_of_single, neg_add_cancel, hr] at hrc rw [hrc, sub_self] theorem isUnit_iff {x : HahnSeries Γ R} : IsUnit x ↔ IsUnit (x.leadingCoeff) := by @@ -503,12 +503,12 @@ instance instField [Field R] : Field (HahnSeries Γ R) where if x0 : x = 0 then 0 else (single (-x.order)) (x.leadingCoeff)⁻¹ * - (SummableFamily.powers _ (unit_aux x (inv_mul_cancel (leadingCoeff_ne_iff.mpr x0)))).hsum + (SummableFamily.powers _ (unit_aux x (inv_mul_cancel₀ (leadingCoeff_ne_iff.mpr x0)))).hsum inv_zero := dif_pos rfl mul_inv_cancel x x0 := (congr rfl (dif_neg x0)).trans $ by have h := SummableFamily.one_sub_self_mul_hsum_powers - (unit_aux x (inv_mul_cancel (leadingCoeff_ne_iff.mpr x0))) + (unit_aux x (inv_mul_cancel₀ (leadingCoeff_ne_iff.mpr x0))) rw [sub_sub_cancel] at h rw [← mul_assoc, mul_comm x, h] nnqsmul := _ diff --git a/Mathlib/RingTheory/Henselian.lean b/Mathlib/RingTheory/Henselian.lean index 86cea0f1943d0..85b5017c6df8f 100644 --- a/Mathlib/RingTheory/Henselian.lean +++ b/Mathlib/RingTheory/Henselian.lean @@ -213,7 +213,7 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] · erw [Finset.sum_range_succ] rw [Finset.range_one, Finset.sum_singleton, taylor_coeff_zero, taylor_coeff_one, pow_zero, pow_one, mul_one, mul_neg, - mul_left_comm, Ring.mul_inverse_cancel _ (hf'c n), mul_one, add_neg_self] + mul_left_comm, Ring.mul_inverse_cancel _ (hf'c n), mul_one, add_neg_cancel] exact Ideal.zero_mem _ · refine Submodule.sum_mem _ ?_ simp only [Finset.mem_Ico] diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 1d74105198490..b0846e78aa5bb 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -610,7 +610,7 @@ theorem ker_isMaximal_of_surjective {R K F : Type*} [Ring R] [Field K] rw [H] refine J.sub_mem (J.mul_mem_left _ hxJ) (hJ ?_) rw [mem_ker] - simp only [hy, map_sub, map_one, map_mul, inv_mul_cancel (mt (mem_ker f).mpr hxf), sub_self] + simp only [hy, map_sub, map_one, map_mul, inv_mul_cancel₀ (mt (mem_ker f).mpr hxf), sub_self] end RingHom diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index f68f32b12f19e..0715d3ae16630 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -148,7 +148,7 @@ theorem single_order_mul_powerSeriesPart (x : LaurentSeries R) : theorem ofPowerSeries_powerSeriesPart (x : LaurentSeries R) : ofPowerSeries ℤ R x.powerSeriesPart = single (-x.order) 1 * x := by refine Eq.trans ?_ (congr rfl x.single_order_mul_powerSeriesPart) - rw [← mul_assoc, single_mul_single, neg_add_self, mul_one, ← C_apply, C_one, one_mul] + rw [← mul_assoc, single_mul_single, neg_add_cancel, mul_one, ← C_apply, C_one, one_mul] end Semiring @@ -478,7 +478,7 @@ theorem valuation_single_zpow (s : ℤ) : Valued.v (HahnSeries.single s (1 : K) : LaurentSeries K) = Multiplicative.ofAdd (-(s : ℤ)) := by have : Valued.v (1 : LaurentSeries K) = (1 : ℤₘ₀) := Valued.v.map_one - rw [← single_zero_one, ← add_right_neg s, ← mul_one 1, ← single_mul_single, map_mul, + rw [← single_zero_one, ← add_neg_cancel s, ← mul_one 1, ← single_mul_single, map_mul, mul_eq_one_iff_eq_inv₀] at this · rw [this] induction' s with s s diff --git a/Mathlib/RingTheory/Localization/NumDen.lean b/Mathlib/RingTheory/Localization/NumDen.lean index c9104241095e7..10a844a4ca9d4 100644 --- a/Mathlib/RingTheory/Localization/NumDen.lean +++ b/Mathlib/RingTheory/Localization/NumDen.lean @@ -105,7 +105,7 @@ theorem isInteger_of_isUnit_den {x : K} (h : IsUnit (den A x : A)) : IsInteger A refine _root_.trans ?_ (mk'_num_den A x) rw [map_mul, map_units_inv, hd] apply mul_left_cancel₀ d_ne_zero - rw [← mul_assoc, mul_inv_cancel d_ne_zero, one_mul, mk'_spec'] + rw [← mul_assoc, mul_inv_cancel₀ d_ne_zero, one_mul, mk'_spec'] theorem isUnit_den_iff (x : K) : IsUnit (den A x : A) ↔ IsLocalization.IsInteger A x where mp h := by diff --git a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean index 2c19ea9565c11..a7e076eb389ea 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Inverse.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Inverse.lean @@ -281,7 +281,7 @@ instance : InvOneClass (MvPowerSeries σ k) := theorem C_inv (r : k) : (C σ k r)⁻¹ = C σ k r⁻¹ := by rcases eq_or_ne r 0 with (rfl | hr) · simp - rw [MvPowerSeries.inv_eq_iff_mul_eq_one, ← map_mul, inv_mul_cancel hr, map_one] + rw [MvPowerSeries.inv_eq_iff_mul_eq_one, ← map_mul, inv_mul_cancel₀ hr, map_one] simpa using hr @[simp] diff --git a/Mathlib/RingTheory/OreLocalization/Basic.lean b/Mathlib/RingTheory/OreLocalization/Basic.lean index 24690095025f6..04732f2a57a29 100644 --- a/Mathlib/RingTheory/OreLocalization/Basic.lean +++ b/Mathlib/RingTheory/OreLocalization/Basic.lean @@ -472,7 +472,7 @@ def universalMulHom (hf : ∀ s : S, f s = fS s) : R[S⁻¹] →* T where have : (fS ⟨t * s, ht⟩ : T) = f t * fS s := by simp only [← hf, MonoidHom.map_mul] conv_rhs => - rw [MonoidHom.map_mul, ← one_mul (f r), ← Units.val_one, ← mul_right_inv (fS s)] + rw [MonoidHom.map_mul, ← one_mul (f r), ← Units.val_one, ← mul_inv_cancel (fS s)] rw [Units.val_mul, mul_assoc, ← mul_assoc _ (fS s : T), ← this, ← mul_assoc] simp only [one_mul, Units.inv_mul] map_one' := by beta_reduce; rw [OreLocalization.one_def, liftExpand_of]; simp @@ -505,7 +505,7 @@ theorem universalMulHom_unique (φ : R[S⁻¹] →* T) (huniv : ∀ r : R, φ (n φ = universalMulHom f fS hf := by ext x; induction' x with r s rw [universalMulHom_apply, ← huniv r, numeratorHom_apply, ← one_mul (φ (r /ₒ s)), ← - Units.val_one, ← mul_left_inv (fS s), Units.val_mul, mul_assoc, ← hf, ← huniv, ← φ.map_mul, + Units.val_one, ← inv_mul_cancel (fS s), Units.val_mul, mul_assoc, ← hf, ← huniv, ← φ.map_mul, numeratorHom_apply, OreLocalization.mul_cancel] end UMP @@ -868,7 +868,7 @@ instance instNegOreLocalization : Neg X[S⁻¹] := protected theorem neg_def (r : X) (s : S) : -(r /ₒ s) = -r /ₒ s := by with_unfolding_all rfl -protected theorem add_left_neg (x : X[S⁻¹]) : -x + x = 0 := by +protected theorem neg_add_cancel (x : X[S⁻¹]) : -x + x = 0 := by induction' x with r s; simp /-- `zsmul` of `OreLocalization` -/ @@ -877,7 +877,7 @@ protected def zsmul : ℤ → X[S⁻¹] → X[S⁻¹] := zsmulRec unseal OreLocalization.zsmul in instance instAddGroupOreLocalization : AddGroup X[S⁻¹] where - add_left_neg := OreLocalization.add_left_neg + neg_add_cancel := OreLocalization.neg_add_cancel zsmul := OreLocalization.zsmul end AddGroup diff --git a/Mathlib/RingTheory/OreLocalization/Ring.lean b/Mathlib/RingTheory/OreLocalization/Ring.lean index 3d54c272d5b5d..952d2bd534a8f 100644 --- a/Mathlib/RingTheory/OreLocalization/Ring.lean +++ b/Mathlib/RingTheory/OreLocalization/Ring.lean @@ -149,7 +149,7 @@ def universalHom : R[S⁻¹] →+* T := Submonoid.smul_def] simp only [mul_inv_rev, MonoidHom.map_mul, RingHom.map_add, RingHom.map_mul, Units.val_mul] rw [mul_add, mul_assoc, ← mul_assoc _ (f s₃), hf, ← Units.val_mul] - simp only [one_mul, mul_left_inv, Units.val_one] + simp only [one_mul, inv_mul_cancel, Units.val_one] congr 1 rw [← mul_assoc] congr 1 diff --git a/Mathlib/RingTheory/Polynomial/Dickson.lean b/Mathlib/RingTheory/Polynomial/Dickson.lean index 5f712f2e31027..6f3bfc0a478bf 100644 --- a/Mathlib/RingTheory/Polynomial/Dickson.lean +++ b/Mathlib/RingTheory/Polynomial/Dickson.lean @@ -201,7 +201,7 @@ theorem dickson_one_one_zmod_p (p : ℕ) [Fact p.Prime] : dickson 1 (1 : ZMod p) apply @Set.Infinite.mono _ { x : K | ∃ y, x = y + y⁻¹ ∧ y ≠ 0 } · rintro _ ⟨x, rfl, hx⟩ simp only [eval_X, eval_pow, Set.mem_setOf_eq, @add_pow_char K _ p, - dickson_one_one_eval_add_inv _ _ (mul_inv_cancel hx), inv_pow, ZMod.castHom_apply, + dickson_one_one_eval_add_inv _ _ (mul_inv_cancel₀ hx), inv_pow, ZMod.castHom_apply, ZMod.cast_one'] -- Now we need to show that the set of such `x` is infinite. -- If the set is finite, then we will show that `K` is also finite. @@ -233,7 +233,7 @@ theorem dickson_one_one_zmod_p (p : ℕ) [Fact p.Prime] : dickson 1 (1 : ZMod p) by_cases hy : y = 0 · simp only [hy, eq_self_iff_true, or_true_iff] apply or_congr _ Iff.rfl - rw [← mul_left_inj' hy, eq_comm, ← sub_eq_zero, add_mul, inv_mul_cancel hy] + rw [← mul_left_inj' hy, eq_comm, ← sub_eq_zero, add_mul, inv_mul_cancel₀ hy] apply eq_iff_eq_cancel_right.mpr ring -- Finally, we prove the claim that our finite union of finite sets covers all of `K`. diff --git a/Mathlib/RingTheory/Polynomial/GaussLemma.lean b/Mathlib/RingTheory/Polynomial/GaussLemma.lean index 181633aa094d2..ea918e9caf17c 100644 --- a/Mathlib/RingTheory/Polynomial/GaussLemma.lean +++ b/Mathlib/RingTheory/Polynomial/GaussLemma.lean @@ -78,7 +78,7 @@ theorem IsIntegrallyClosed.eq_map_mul_C_of_dvd [IsIntegrallyClosed R] {f : R[X]} suffices lem : ∃ g' : R[X], g'.map (algebraMap R K) = g * C g.leadingCoeff⁻¹ by obtain ⟨g', hg'⟩ := lem use g' - rw [hg', mul_assoc, ← C_mul, inv_mul_cancel (leadingCoeff_ne_zero.mpr g_ne_0), C_1, mul_one] + rw [hg', mul_assoc, ← C_mul, inv_mul_cancel₀ (leadingCoeff_ne_zero.mpr g_ne_0), C_1, mul_one] have g_mul_dvd : g * C g.leadingCoeff⁻¹ ∣ f.map (algebraMap R K) := by rwa [Associated.dvd_iff_dvd_left (show Associated (g * C g.leadingCoeff⁻¹) g from _)] rw [associated_mul_isUnit_left_iff] diff --git a/Mathlib/RingTheory/Valuation/Integers.lean b/Mathlib/RingTheory/Valuation/Integers.lean index d14885dd76c84..ae1c553270829 100644 --- a/Mathlib/RingTheory/Valuation/Integers.lean +++ b/Mathlib/RingTheory/Valuation/Integers.lean @@ -116,7 +116,7 @@ theorem dvd_of_le (hv : Integers v O) {x y : O} hx.symm ▸ dvd_zero y) fun hy : algebraMap O F y ≠ 0 => have : v ((algebraMap O F y)⁻¹ * algebraMap O F x) ≤ 1 := by - rw [← v.map_one, ← inv_mul_cancel hy, v.map_mul, v.map_mul] + rw [← v.map_one, ← inv_mul_cancel₀ hy, v.map_mul, v.map_mul] exact mul_le_mul_left' h _ let ⟨z, hz⟩ := hv.3 this ⟨z, hv.1 <| ((algebraMap O F).map_mul y z).symm ▸ hz.symm ▸ (mul_inv_cancel_left₀ hy _).symm⟩ diff --git a/Mathlib/RingTheory/Valuation/ValuationRing.lean b/Mathlib/RingTheory/Valuation/ValuationRing.lean index 976b2a387a2a7..ed5c3bf093421 100644 --- a/Mathlib/RingTheory/Valuation/ValuationRing.lean +++ b/Mathlib/RingTheory/Valuation/ValuationRing.lean @@ -162,7 +162,7 @@ noncomputable instance linearOrderedCommGroupWithZero : apply Quotient.sound' use 1 simp only [one_smul, ne_eq] - apply (mul_inv_cancel _).symm + apply (mul_inv_cancel₀ _).symm contrapose ha simp only [Classical.not_not] at ha ⊢ rw [ha] diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index 0dad8f77f0f4f..fb8a0c376dfbb 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -136,7 +136,7 @@ instance : IsFractionRing A K where by_cases h : z = 0; · use (0, 1); simp [h] cases' A.mem_or_inv_mem z with hh hh · use (⟨z, hh⟩, 1); simp - · refine ⟨⟨1, ⟨⟨_, hh⟩, ?_⟩⟩, mul_inv_cancel h⟩ + · refine ⟨⟨1, ⟨⟨_, hh⟩, ?_⟩⟩, mul_inv_cancel₀ h⟩ exact mem_nonZeroDivisors_iff_ne_zero.2 fun c => h (inv_eq_zero.mp (congr_arg Subtype.val c)) exists_of_eq {a b} h := ⟨1, by ext; simpa using h⟩ diff --git a/Mathlib/SetTheory/Game/Basic.lean b/Mathlib/SetTheory/Game/Basic.lean index 175b69c1a18db..13796f6637a38 100644 --- a/Mathlib/SetTheory/Game/Basic.lean +++ b/Mathlib/SetTheory/Game/Basic.lean @@ -68,7 +68,7 @@ instance instAddCommGroupWithOneGame : AddCommGroupWithOne Game where add_assoc := by rintro ⟨x⟩ ⟨y⟩ ⟨z⟩ exact Quot.sound add_assoc_equiv - add_left_neg := Quotient.ind <| fun x => Quot.sound (add_left_neg_equiv x) + neg_add_cancel := Quotient.ind <| fun x => Quot.sound (neg_add_cancel_equiv x) add_comm := by rintro ⟨x⟩ ⟨y⟩ exact Quot.sound add_comm_equiv diff --git a/Mathlib/SetTheory/Game/Impartial.lean b/Mathlib/SetTheory/Game/Impartial.lean index 815f1fae0b646..24c984856e4bb 100644 --- a/Mathlib/SetTheory/Game/Impartial.lean +++ b/Mathlib/SetTheory/Game/Impartial.lean @@ -132,7 +132,7 @@ theorem not_fuzzy_zero_iff : ¬G ‖ 0 ↔ (G ≈ 0) := ⟨(equiv_or_fuzzy_zero G).resolve_right, Equiv.not_fuzzy⟩ theorem add_self : G + G ≈ 0 := - Equiv.trans (add_congr_left (neg_equiv_self G)) (add_left_neg_equiv G) + Equiv.trans (add_congr_left (neg_equiv_self G)) (neg_add_cancel_equiv G) -- Porting note: Changed `⟦G⟧` to `(⟦G⟧ : Quotient setoid)` @[simp] diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index 6238d450593da..b7f067d54e9eb 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -1479,38 +1479,38 @@ termination_by x y z => (x, y, z) theorem add_assoc_equiv {x y z : PGame} : x + y + z ≈ x + (y + z) := (addAssocRelabelling x y z).equiv -theorem add_left_neg_le_zero : ∀ x : PGame, -x + x ≤ 0 +theorem neg_add_cancel_le_zero : ∀ x : PGame, -x + x ≤ 0 | ⟨xl, xr, xL, xR⟩ => le_zero.2 fun i => by cases' i with i i · -- If Left played in -x, Right responds with the same move in x. refine ⟨@toRightMovesAdd _ ⟨_, _, _, _⟩ (Sum.inr i), ?_⟩ - convert @add_left_neg_le_zero (xR i) + convert @neg_add_cancel_le_zero (xR i) apply add_moveRight_inr · -- If Left in x, Right responds with the same move in -x. dsimp refine ⟨@toRightMovesAdd ⟨_, _, _, _⟩ _ (Sum.inl i), ?_⟩ - convert @add_left_neg_le_zero (xL i) + convert @neg_add_cancel_le_zero (xL i) apply add_moveRight_inl -theorem zero_le_add_left_neg (x : PGame) : 0 ≤ -x + x := by +theorem zero_le_neg_add_cancel (x : PGame) : 0 ≤ -x + x := by rw [← neg_le_neg_iff, neg_zero] - exact neg_add_le.trans (add_left_neg_le_zero _) + exact neg_add_le.trans (neg_add_cancel_le_zero _) -theorem add_left_neg_equiv (x : PGame) : -x + x ≈ 0 := - ⟨add_left_neg_le_zero x, zero_le_add_left_neg x⟩ +theorem neg_add_cancel_equiv (x : PGame) : -x + x ≈ 0 := + ⟨neg_add_cancel_le_zero x, zero_le_neg_add_cancel x⟩ -theorem add_right_neg_le_zero (x : PGame) : x + -x ≤ 0 := - add_comm_le.trans (add_left_neg_le_zero x) +theorem add_neg_cancel_le_zero (x : PGame) : x + -x ≤ 0 := + add_comm_le.trans (neg_add_cancel_le_zero x) -theorem zero_le_add_right_neg (x : PGame) : 0 ≤ x + -x := - (zero_le_add_left_neg x).trans add_comm_le +theorem zero_le_add_neg_cancel (x : PGame) : 0 ≤ x + -x := + (zero_le_neg_add_cancel x).trans add_comm_le -theorem add_right_neg_equiv (x : PGame) : x + -x ≈ 0 := - ⟨add_right_neg_le_zero x, zero_le_add_right_neg x⟩ +theorem add_neg_cancel_equiv (x : PGame) : x + -x ≈ 0 := + ⟨add_neg_cancel_le_zero x, zero_le_add_neg_cancel x⟩ theorem sub_self_equiv : ∀ (x : PGame), x - x ≈ 0 := - add_right_neg_equiv + add_neg_cancel_equiv private theorem add_le_add_right' : ∀ {x y z : PGame}, x ≤ y → x + z ≤ y + z | mk xl xr xL xR, mk yl yr yL yR, mk zl zr zL zR => fun h => by @@ -1546,11 +1546,11 @@ theorem add_lf_add_right {y z : PGame} (h : y ⧏ z) (x) : y + x ⧏ z + x := fun w => calc z ≤ z + 0 := (addZeroRelabelling _).symm.le - _ ≤ z + (x + -x) := add_le_add_left (zero_le_add_right_neg x) _ + _ ≤ z + (x + -x) := add_le_add_left (zero_le_add_neg_cancel x) _ _ ≤ z + x + -x := (addAssocRelabelling _ _ _).symm.le _ ≤ y + x + -x := add_le_add_right w _ _ ≤ y + (x + -x) := (addAssocRelabelling _ _ _).le - _ ≤ y + 0 := add_le_add_left (add_right_neg_le_zero x) _ + _ ≤ y + 0 := add_le_add_left (add_neg_cancel_le_zero x) _ _ ≤ y := (addZeroRelabelling _).le theorem add_lf_add_left {y z : PGame} (h : y ⧏ z) (x) : x + y ⧏ x + z := by @@ -1589,32 +1589,32 @@ theorem sub_congr_right {x y z : PGame} : (y ≈ z) → (x - y ≈ x - z) := sub_congr equiv_rfl theorem le_iff_sub_nonneg {x y : PGame} : x ≤ y ↔ 0 ≤ y - x := - ⟨fun h => (zero_le_add_right_neg x).trans (add_le_add_right h _), fun h => + ⟨fun h => (zero_le_add_neg_cancel x).trans (add_le_add_right h _), fun h => calc x ≤ 0 + x := (zeroAddRelabelling x).symm.le _ ≤ y - x + x := add_le_add_right h _ _ ≤ y + (-x + x) := (addAssocRelabelling _ _ _).le - _ ≤ y + 0 := add_le_add_left (add_left_neg_le_zero x) _ + _ ≤ y + 0 := add_le_add_left (neg_add_cancel_le_zero x) _ _ ≤ y := (addZeroRelabelling y).le ⟩ theorem lf_iff_sub_zero_lf {x y : PGame} : x ⧏ y ↔ 0 ⧏ y - x := - ⟨fun h => (zero_le_add_right_neg x).trans_lf (add_lf_add_right h _), fun h => + ⟨fun h => (zero_le_add_neg_cancel x).trans_lf (add_lf_add_right h _), fun h => calc x ≤ 0 + x := (zeroAddRelabelling x).symm.le _ ⧏ y - x + x := add_lf_add_right h _ _ ≤ y + (-x + x) := (addAssocRelabelling _ _ _).le - _ ≤ y + 0 := add_le_add_left (add_left_neg_le_zero x) _ + _ ≤ y + 0 := add_le_add_left (neg_add_cancel_le_zero x) _ _ ≤ y := (addZeroRelabelling y).le ⟩ theorem lt_iff_sub_pos {x y : PGame} : x < y ↔ 0 < y - x := - ⟨fun h => lt_of_le_of_lt (zero_le_add_right_neg x) (add_lt_add_right h _), fun h => + ⟨fun h => lt_of_le_of_lt (zero_le_add_neg_cancel x) (add_lt_add_right h _), fun h => calc x ≤ 0 + x := (zeroAddRelabelling x).symm.le _ < y - x + x := add_lt_add_right h _ _ ≤ y + (-x + x) := (addAssocRelabelling _ _ _).le - _ ≤ y + 0 := add_le_add_left (add_left_neg_le_zero x) _ + _ ≤ y + 0 := add_le_add_left (neg_add_cancel_le_zero x) _ _ ≤ y := (addZeroRelabelling y).le ⟩ diff --git a/Mathlib/SetTheory/Surreal/Basic.lean b/Mathlib/SetTheory/Surreal/Basic.lean index 89c9fbb1fb781..bbb97f8232f9c 100644 --- a/Mathlib/SetTheory/Surreal/Basic.lean +++ b/Mathlib/SetTheory/Surreal/Basic.lean @@ -341,7 +341,7 @@ instance orderedAddCommGroup : OrderedAddCommGroup Surreal where zero_add := by rintro ⟨a⟩; exact Quotient.sound (zero_add_equiv a) add_zero := by rintro ⟨a⟩; exact Quotient.sound (add_zero_equiv a) neg := Neg.neg - add_left_neg := by rintro ⟨a⟩; exact Quotient.sound (add_left_neg_equiv a) + neg_add_cancel := by rintro ⟨a⟩; exact Quotient.sound (neg_add_cancel_equiv a) add_comm := by rintro ⟨_⟩ ⟨_⟩; exact Quotient.sound add_comm_equiv le := (· ≤ ·) lt := (· < ·) diff --git a/Mathlib/Tactic/Group.lean b/Mathlib/Tactic/Group.lean index ddcd1d9a35841..1f0943d59fc51 100644 --- a/Mathlib/Tactic/Group.lean +++ b/Mathlib/Tactic/Group.lean @@ -56,7 +56,7 @@ macro_rules one_zpow, zpow_zero, zpow_one, mul_zpow_neg_one, ← mul_assoc, ← zpow_add, ← zpow_add_one, ← zpow_one_add, zpow_trick, zpow_trick_one, zpow_trick_one', - tsub_self, sub_self, add_neg_self, neg_add_self] + tsub_self, sub_self, add_neg_cancel, neg_add_cancel] $[at $location]?) /-- Auxiliary tactic for the `group` tactic. Calls `ring_nf` to normalize exponents. -/ diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index 1cec60f0872db..de9f0789c24cb 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -219,7 +219,7 @@ instance : CommGroup (ContinuousMonoidHom A E) where one_mul f := ext fun x => one_mul (f x) mul_one f := ext fun x => mul_one (f x) inv f := (inv E).comp f - mul_left_inv f := ext fun x => mul_left_inv (f x) + inv_mul_cancel f := ext fun x => inv_mul_cancel (f x) @[to_additive] instance : TopologicalSpace (ContinuousMonoidHom A B) := diff --git a/Mathlib/Topology/Algebra/Field.lean b/Mathlib/Topology/Algebra/Field.lean index feeac9976e3d4..6a6d37b9907fc 100644 --- a/Mathlib/Topology/Algebra/Field.lean +++ b/Mathlib/Topology/Algebra/Field.lean @@ -23,13 +23,13 @@ variable {K : Type*} [DivisionRing K] [TopologicalSpace K] inverse images of compact sets are compact. -/ theorem Filter.tendsto_cocompact_mul_left₀ [ContinuousMul K] {a : K} (ha : a ≠ 0) : Filter.Tendsto (fun x : K => a * x) (Filter.cocompact K) (Filter.cocompact K) := - Filter.tendsto_cocompact_mul_left (inv_mul_cancel ha) + Filter.tendsto_cocompact_mul_left (inv_mul_cancel₀ ha) /-- Right-multiplication by a nonzero element of a topological division ring is proper, i.e., inverse images of compact sets are compact. -/ theorem Filter.tendsto_cocompact_mul_right₀ [ContinuousMul K] {a : K} (ha : a ≠ 0) : Filter.Tendsto (fun x : K => x * a) (Filter.cocompact K) (Filter.cocompact K) := - Filter.tendsto_cocompact_mul_right (mul_inv_cancel ha) + Filter.tendsto_cocompact_mul_right (mul_inv_cancel₀ ha) variable (K) diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 99692392d933b..8b8f83b53c3bc 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -950,7 +950,7 @@ lemma Filter.tendsto_div_const_iff {G : Type*} {b : G} (hb : b ≠ 0) {c : G} {f : α → G} {l : Filter α} : Tendsto (f · / b) l (𝓝 (c / b)) ↔ Tendsto f l (𝓝 c) := by refine ⟨fun h ↦ ?_, fun h ↦ Filter.Tendsto.div_const' h b⟩ - convert h.div_const' b⁻¹ with k <;> rw [div_div, mul_inv_cancel hb, div_one] + convert h.div_const' b⁻¹ with k <;> rw [div_div, mul_inv_cancel₀ hb, div_one] lemma Filter.tendsto_sub_const_iff {G : Type*} [AddCommGroup G] [TopologicalSpace G] [ContinuousSub G] @@ -1332,7 +1332,7 @@ theorem TopologicalGroup.t2Space_of_one_sep (H : ∀ x : G, x ≠ 1 → ∃ U suffices T1Space G from inferInstance refine t1Space_iff_specializes_imp_eq.2 fun x y hspec ↦ by_contra fun hne ↦ ?_ rcases H (x * y⁻¹) (by rwa [Ne, mul_inv_eq_one]) with ⟨U, hU₁, hU⟩ - exact hU <| mem_of_mem_nhds <| hspec.map (continuous_mul_right y⁻¹) (by rwa [mul_inv_self]) + exact hU <| mem_of_mem_nhds <| hspec.map (continuous_mul_right y⁻¹) (by rwa [mul_inv_cancel]) /-- Given a neighborhood `U` of the identity, one may find a neighborhood `V` of the identity which is closed, symmetric, and satisfies `V * V ⊆ U`. -/ diff --git a/Mathlib/Topology/Algebra/GroupCompletion.lean b/Mathlib/Topology/Algebra/GroupCompletion.lean index 72ddaa26dcdc5..234486fbb1753 100644 --- a/Mathlib/Topology/Algebra/GroupCompletion.lean +++ b/Mathlib/Topology/Algebra/GroupCompletion.lean @@ -148,12 +148,12 @@ instance : SubNegMonoid (Completion α) := instance addGroup : AddGroup (Completion α) := { (inferInstance : SubNegMonoid <| Completion α) with - add_left_neg := fun a ↦ + neg_add_cancel := fun a ↦ Completion.induction_on a (isClosed_eq (continuous_map₂ Completion.continuous_map continuous_id) continuous_const) fun a ↦ show -(a : Completion α) + a = 0 by - rw_mod_cast [add_left_neg] + rw_mod_cast [neg_add_cancel] rfl } instance uniformAddGroup : UniformAddGroup (Completion α) := diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index f4ca295850c70..ca811c376a395 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -638,19 +638,19 @@ instance addCommMonoid : AddCommMonoid (M₁ →SL[σ₁₂] M₂) where zero_add := by intros ext - apply_rules [zero_add, add_assoc, add_zero, add_left_neg, add_comm] + apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] add_zero := by intros ext - apply_rules [zero_add, add_assoc, add_zero, add_left_neg, add_comm] + apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] add_comm := by intros ext - apply_rules [zero_add, add_assoc, add_zero, add_left_neg, add_comm] + apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] add_assoc := by intros ext - apply_rules [zero_add, add_assoc, add_zero, add_left_neg, add_comm] + apply_rules [zero_add, add_assoc, add_zero, neg_add_cancel, add_comm] nsmul := (· • ·) nsmul_zero f := by ext @@ -1263,7 +1263,7 @@ instance addCommGroup : AddCommGroup (M →SL[σ₁₂] M₂) where zsmul_zero' f := by ext; simp zsmul_succ' n f := by ext; simp [add_smul, add_comm] zsmul_neg' n f := by ext; simp [add_smul] - add_left_neg _ := by ext; apply add_left_neg + neg_add_cancel _ := by ext; apply neg_add_cancel theorem sub_apply (f g : M →SL[σ₁₂] M₂) (x : M) : (f - g) x = f x - g x := rfl @@ -1989,7 +1989,7 @@ instance automorphismGroup : Group (M₁ ≃L[R₁] M₁) where one_mul f := by ext rfl - mul_left_inv f := by + inv_mul_cancel f := by ext x exact f.left_inv x diff --git a/Mathlib/Topology/Algebra/Module/Cardinality.lean b/Mathlib/Topology/Algebra/Module/Cardinality.lean index c63a9c32cc65d..c45fb8205558f 100644 --- a/Mathlib/Topology/Algebra/Module/Cardinality.lean +++ b/Mathlib/Topology/Algebra/Module/Cardinality.lean @@ -87,8 +87,8 @@ lemma cardinal_eq_of_mem_nhds_zero have : (c^n • s :) ≃ s := { toFun := fun x ↦ ⟨(c^n)⁻¹ • x.1, (mem_smul_set_iff_inv_smul_mem₀ (cn_ne n) _ _).1 x.2⟩ invFun := fun x ↦ ⟨(c^n) • x.1, smul_mem_smul_set x.2⟩ - left_inv := fun x ↦ by simp [smul_smul, mul_inv_cancel (cn_ne n)] - right_inv := fun x ↦ by simp [smul_smul, inv_mul_cancel (cn_ne n)] } + left_inv := fun x ↦ by simp [smul_smul, mul_inv_cancel₀ (cn_ne n)] + right_inv := fun x ↦ by simp [smul_smul, inv_mul_cancel₀ (cn_ne n)] } exact Cardinal.mk_congr this apply (Cardinal.mk_of_countable_eventually_mem A B).symm diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index 1cfd1e13e0032..e7d18534a96a9 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -101,7 +101,7 @@ theorem unique_topology_of_t2 {t : TopologicalSpace 𝕜} (h₁ : @TopologicalAd -- `ξ₀ ∈ 𝓑 ⊆ {ξ₀}ᶜ`, which is a contradiction. by_contra! h suffices (ξ₀ * ξ⁻¹) • ξ ∈ balancedCore 𝕜 {ξ₀}ᶜ by - rw [smul_eq_mul 𝕜, mul_assoc, inv_mul_cancel hξ0, mul_one] at this + rw [smul_eq_mul 𝕜, mul_assoc, inv_mul_cancel₀ hξ0, mul_one] at this exact not_mem_compl_iff.mpr (mem_singleton ξ₀) ((balancedCore_subset _) this) -- For that, we use that `𝓑` is balanced : since `‖ξ₀‖ < ε < ‖ξ‖`, we have `‖ξ₀ / ξ‖ ≤ 1`, -- hence `ξ₀ = (ξ₀ / ξ) • ξ ∈ 𝓑` because `ξ ∈ 𝓑`. diff --git a/Mathlib/Topology/Algebra/UniformField.lean b/Mathlib/Topology/Algebra/UniformField.lean index 96baeef3010b7..c626826669020 100644 --- a/Mathlib/Topology/Algebra/UniformField.lean +++ b/Mathlib/Topology/Algebra/UniformField.lean @@ -140,7 +140,7 @@ theorem mul_hatInv_cancel {x : hat K} (x_ne : x ≠ 0) : x * hatInv x = 1 := by rw [mem_compl_singleton_iff] at z_ne dsimp [f] rw [hatInv_extends z_ne, ← coe_mul] - rw [mul_inv_cancel z_ne, coe_one] + rw [mul_inv_cancel₀ z_ne, coe_one] replace fxclo := closure_mono this fxclo rwa [closure_singleton, mem_singleton_iff] at fxclo diff --git a/Mathlib/Topology/Algebra/Valued/ValuedField.lean b/Mathlib/Topology/Algebra/Valued/ValuedField.lean index 01fd1a0decbca..615d15555a9be 100644 --- a/Mathlib/Topology/Algebra/Valued/ValuedField.lean +++ b/Mathlib/Topology/Algebra/Valued/ValuedField.lean @@ -58,8 +58,8 @@ theorem Valuation.inversion_estimate {x y : K} {γ : Γ₀ˣ} (y_ne : y ≠ 0) rw [h, v.map_zero] at key exact v.zero_iff.1 key.symm have decomp : x⁻¹ - y⁻¹ = x⁻¹ * (y - x) * y⁻¹ := by - rw [mul_sub_left_distrib, sub_mul, mul_assoc, show y * y⁻¹ = 1 from mul_inv_cancel y_ne, - show x⁻¹ * x = 1 from inv_mul_cancel x_ne, mul_one, one_mul] + rw [mul_sub_left_distrib, sub_mul, mul_assoc, show y * y⁻¹ = 1 from mul_inv_cancel₀ y_ne, + show x⁻¹ * x = 1 from inv_mul_cancel₀ x_ne, mul_one, one_mul] calc v (x⁻¹ - y⁻¹) = v (x⁻¹ * (y - x) * y⁻¹) := by rw [decomp] _ = v x⁻¹ * (v <| y - x) * v y⁻¹ := by repeat' rw [Valuation.map_mul] @@ -233,13 +233,13 @@ theorem continuous_extension : Continuous (Valued.extension : hat K → Γ₀) : have nhds_right : (fun x => x * x₀) '' V' ∈ 𝓝 x₀ := by have l : Function.LeftInverse (fun x : hat K => x * x₀⁻¹) fun x : hat K => x * x₀ := by intro x - simp only [mul_assoc, mul_inv_cancel h, mul_one] + simp only [mul_assoc, mul_inv_cancel₀ h, mul_one] have r : Function.RightInverse (fun x : hat K => x * x₀⁻¹) fun x : hat K => x * x₀ := by intro x - simp only [mul_assoc, inv_mul_cancel h, mul_one] + simp only [mul_assoc, inv_mul_cancel₀ h, mul_one] have c : Continuous fun x : hat K => x * x₀⁻¹ := continuous_id.mul continuous_const rw [image_eq_preimage_of_inverse l r] - rw [← mul_inv_cancel h] at V'_in + rw [← mul_inv_cancel₀ h] at V'_in exact c.continuousAt V'_in have : ∃ z₀ : K, ∃ y₀ ∈ V', ↑z₀ = y₀ * x₀ ∧ z₀ ≠ 0 := by rcases Completion.denseRange_coe.mem_nhds nhds_right with ⟨z₀, y₀, y₀_in, H : y₀ * x₀ = z₀⟩ @@ -256,10 +256,10 @@ theorem continuous_extension : Continuous (Valued.extension : hat K → Γ₀) : apply hV have : (z₀⁻¹ : K) = (z₀ : hat K)⁻¹ := map_inv₀ (Completion.coeRingHom : K →+* hat K) z₀ rw [Completion.coe_mul, this, ha, hz₀, mul_inv, mul_comm y₀⁻¹, ← mul_assoc, mul_assoc y, - mul_inv_cancel h, mul_one] + mul_inv_cancel₀ h, mul_one] solve_by_elim calc - v a = v (a * z₀⁻¹ * z₀) := by rw [mul_assoc, inv_mul_cancel z₀_ne, mul_one] + v a = v (a * z₀⁻¹ * z₀) := by rw [mul_assoc, inv_mul_cancel₀ z₀_ne, mul_one] _ = v (a * z₀⁻¹) * v z₀ := Valuation.map_mul _ _ _ _ = v z₀ := by rw [this, one_mul] diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index d938d43269972..dba770f9f562c 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -458,7 +458,7 @@ theorem trans_continuous_family {ι : Type*} [TopologicalSpace ι] (continuous_id.prod_map <| (continuous_const.mul continuous_subtype_val).sub continuous_const) · rintro st hst - simp [hst, mul_inv_cancel (two_ne_zero' ℝ)] + simp [hst, mul_inv_cancel₀ (two_ne_zero' ℝ)] @[continuity] theorem _root_.Continuous.path_trans {f : Y → Path x y} {g : Y → Path y z} : diff --git a/Mathlib/Topology/ContinuousFunction/Ideals.lean b/Mathlib/Topology/ContinuousFunction/Ideals.lean index 2bda7f44edbd2..e40d94ca79034 100644 --- a/Mathlib/Topology/ContinuousFunction/Ideals.lean +++ b/Mathlib/Topology/ContinuousFunction/Ideals.lean @@ -173,7 +173,7 @@ theorem exists_mul_le_one_eqOn_ge (f : C(X, ℝ≥0)) {c : ℝ≥0} (hc : 0 < c) fun x hx => by simpa only [coe_const, mul_apply, coe_mk, Pi.inv_apply, Pi.sup_apply, Function.const_apply, sup_eq_left.mpr (Set.mem_setOf.mp hx), ne_eq, Pi.one_apply] - using inv_mul_cancel (hc.trans_le hx).ne' ⟩ + using inv_mul_cancel₀ (hc.trans_le hx).ne' ⟩ variable [CompactSpace X] [T2Space X] diff --git a/Mathlib/Topology/Instances/AddCircle.lean b/Mathlib/Topology/Instances/AddCircle.lean index 21303dc60bf02..988b0aced090f 100644 --- a/Mathlib/Topology/Instances/AddCircle.lean +++ b/Mathlib/Topology/Instances/AddCircle.lean @@ -363,7 +363,7 @@ instance : DivisibleBy (AddCircle p) ℤ where div_cancel {n} x hn := by replace hn : (n : 𝕜) ≠ 0 := by norm_cast change n • QuotientAddGroup.mk' _ ((n : 𝕜)⁻¹ * ↑(equivIco p 0 x)) = x - rw [← map_zsmul, ← smul_mul_assoc, zsmul_eq_mul, mul_inv_cancel hn, one_mul] + rw [← map_zsmul, ← smul_mul_assoc, zsmul_eq_mul, mul_inv_cancel₀ hn, one_mul] exact (equivIco p 0).symm_apply_apply x end FloorRing diff --git a/Mathlib/Topology/MetricSpace/DilationEquiv.lean b/Mathlib/Topology/MetricSpace/DilationEquiv.lean index 9cbb3a37ad723..d510cd58fb9e7 100644 --- a/Mathlib/Topology/MetricSpace/DilationEquiv.lean +++ b/Mathlib/Topology/MetricSpace/DilationEquiv.lean @@ -75,7 +75,7 @@ def symm (e : X ≃ᵈ Y) : Y ≃ᵈ X where edist_eq' := by refine ⟨(ratio e)⁻¹, inv_ne_zero <| ratio_ne_zero e, e.surjective.forall₂.2 fun x y ↦ ?_⟩ simp_rw [Equiv.toFun_as_coe, Equiv.symm_apply_apply, coe_toEquiv, edist_eq] - rw [← mul_assoc, ← ENNReal.coe_mul, inv_mul_cancel (ratio_ne_zero e), + rw [← mul_assoc, ← ENNReal.coe_mul, inv_mul_cancel₀ (ratio_ne_zero e), ENNReal.coe_one, one_mul] @[simp] theorem symm_symm (e : X ≃ᵈ Y) : e.symm.symm = e := rfl @@ -142,7 +142,7 @@ instance : Group (X ≃ᵈ X) where one_mul _ := rfl mul_one _ := rfl inv := symm - mul_left_inv := self_trans_symm + inv_mul_cancel := self_trans_symm theorem mul_def (e e' : X ≃ᵈ X) : e * e' = e'.trans e := rfl theorem one_def : (1 : X ≃ᵈ X) = refl X := rfl diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean index c81ccacd3fc2e..fbeb7156e42d0 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean @@ -743,7 +743,7 @@ instance : SecondCountableTopology GHSpace := by _ ≤ 1 := le_of_lt (abs_sub_lt_one_of_floor_eq_floor this) calc |dist x y - dist (Ψ x) (Ψ y)| = ε * ε⁻¹ * |dist x y - dist (Ψ x) (Ψ y)| := by - rw [mul_inv_cancel (ne_of_gt εpos), one_mul] + rw [mul_inv_cancel₀ (ne_of_gt εpos), one_mul] _ = ε * (|ε⁻¹| * |dist x y - dist (Ψ x) (Ψ y)|) := by rw [abs_of_nonneg (le_of_lt (inv_pos.2 εpos)), mul_assoc] _ ≤ ε * 1 := mul_le_mul_of_nonneg_left I (le_of_lt εpos) @@ -914,7 +914,7 @@ theorem totallyBounded {t : Set GHSpace} {C : ℝ} {u : ℕ → ℝ} {K : ℕ _ ≤ 1 := le_of_lt (abs_sub_lt_one_of_floor_eq_floor this) calc |dist x y - dist (Ψ x) (Ψ y)| = ε * ε⁻¹ * |dist x y - dist (Ψ x) (Ψ y)| := by - rw [mul_inv_cancel (ne_of_gt εpos), one_mul] + rw [mul_inv_cancel₀ (ne_of_gt εpos), one_mul] _ = ε * (|ε⁻¹| * |dist x y - dist (Ψ x) (Ψ y)|) := by rw [abs_of_nonneg (le_of_lt (inv_pos.2 εpos)), mul_assoc] _ ≤ ε * 1 := mul_le_mul_of_nonneg_left I (le_of_lt εpos) diff --git a/Mathlib/Topology/MetricSpace/IsometricSMul.lean b/Mathlib/Topology/MetricSpace/IsometricSMul.lean index 010c943fca0f3..1d8b4a236200b 100644 --- a/Mathlib/Topology/MetricSpace/IsometricSMul.lean +++ b/Mathlib/Topology/MetricSpace/IsometricSMul.lean @@ -105,7 +105,7 @@ theorem edist_div_right [DivInvMonoid M] [PseudoEMetricSpace M] [IsometricSMul M @[to_additive (attr := simp)] theorem edist_inv_inv [PseudoEMetricSpace G] [IsometricSMul G G] [IsometricSMul Gᵐᵒᵖ G] (a b : G) : edist a⁻¹ b⁻¹ = edist a b := by - rw [← edist_mul_left a, ← edist_mul_right _ _ b, mul_right_inv, one_mul, inv_mul_cancel_right, + rw [← edist_mul_left a, ← edist_mul_right _ _ b, mul_inv_cancel, one_mul, inv_mul_cancel_right, edist_comm] @[to_additive] diff --git a/Mathlib/Topology/MetricSpace/Isometry.lean b/Mathlib/Topology/MetricSpace/Isometry.lean index 39d21d78785e2..9e0c828a91991 100644 --- a/Mathlib/Topology/MetricSpace/Isometry.lean +++ b/Mathlib/Topology/MetricSpace/Isometry.lean @@ -451,7 +451,7 @@ instance : Group (α ≃ᵢ α) where mul_assoc e₁ e₂ e₃ := rfl one_mul e := ext fun _ => rfl mul_one e := ext fun _ => rfl - mul_left_inv e := ext e.symm_apply_apply + inv_mul_cancel e := ext e.symm_apply_apply @[simp] theorem coe_one : ⇑(1 : α ≃ᵢ α) = id := rfl From 2215b9e43b7cd64e43162ea8125a7f88f93ab85a Mon Sep 17 00:00:00 2001 From: Ralf Stephan Date: Mon, 12 Aug 2024 15:04:35 +0000 Subject: [PATCH 214/359] doc(Tactics): fix typo (#15739) Fix minor typo in `compute_degree` docstring. --- Mathlib/Tactic/ComputeDegree.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Tactic/ComputeDegree.lean b/Mathlib/Tactic/ComputeDegree.lean index 3dc94baf457c7..792092c33497a 100644 --- a/Mathlib/Tactic/ComputeDegree.lean +++ b/Mathlib/Tactic/ComputeDegree.lean @@ -444,7 +444,7 @@ The tactic may leave goals of the form `d' = d` `d' ≤ d`, or `r ≠ 0`, where `WithBot ℕ` is the tactic's guess of the degree, and `r` is the coefficient's guess of the leading coefficient of `f`. -`compute_degree` applies `norm_num` to the left-hand side of all side goals, trying to clos them. +`compute_degree` applies `norm_num` to the left-hand side of all side goals, trying to close them. The variant `compute_degree!` first applies `compute_degree`. Then it uses `norm_num` on all the whole remaining goals and tries `assumption`. From eb190041ae3e3b951a84c597df7c311752cae9f8 Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 12 Aug 2024 15:28:14 +0000 Subject: [PATCH 215/359] perf(Geometry/Euclidean): squeeze slow field_simp calls (#13497) Replace them by the equivalent `simp` call and squeeze that. We leave `field_simp` calls slower than roughly 150ms in place. A few calls cannot easily be replaced; we leave these alone as well. Speeds up all affected files, including `Euclidean/Circumcenter` by 4%, `Inversion/Basic` by 27% and `Triangle` by 52%. --- .../Euclidean/Angle/Oriented/Basic.lean | 5 +++- .../Euclidean/Angle/Unoriented/Basic.lean | 3 ++- Mathlib/Geometry/Euclidean/Circumcenter.lean | 5 +++- .../Geometry/Euclidean/Inversion/Basic.lean | 13 ++++++++-- .../Euclidean/Inversion/ImageHyperplane.lean | 1 + Mathlib/Geometry/Euclidean/MongePoint.lean | 11 ++++++-- .../Geometry/Euclidean/Sphere/Ptolemy.lean | 1 + Mathlib/Geometry/Euclidean/Triangle.lean | 25 ++++++++++++++++--- 8 files changed, 53 insertions(+), 11 deletions(-) diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean index 0e765f0f53f1a..eb15d01c7e2d7 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean @@ -547,7 +547,10 @@ theorem inner_eq_norm_mul_norm_mul_cos_oangle (x y : V) : have : ‖y‖ ≠ 0 := by simpa using hy rw [oangle, Real.Angle.cos_coe, Complex.cos_arg, o.abs_kahler] · simp only [kahler_apply_apply, real_smul, add_re, ofReal_re, mul_re, I_re, ofReal_im] - field_simp + -- TODO(#15486): used to be `field_simp`; replaced by `simp only ...` to speed up + -- Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [mul_zero, I_im, mul_one, sub_self, add_zero, + mul_div_assoc', mul_div_cancel_left₀] · exact o.kahler_ne_zero hx hy /-- The cosine of the oriented angle between two nonzero vectors is the inner product divided by diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Basic.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Basic.lean index 082c66d8c5662..c55f777c49a2f 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Basic.lean @@ -178,7 +178,8 @@ theorem sin_angle_mul_norm_mul_norm (x y : V) : rw [hx, inner_zero_left, zero_mul, neg_zero] · rw [norm_eq_zero] at hy rw [hy, inner_zero_right, zero_mul, neg_zero] - · field_simp [h] + · -- takes 600ms; squeezing the "equivalent" simp call yields an invalid result + field_simp [h] ring_nf /-- The angle between two vectors is zero if and only if they are diff --git a/Mathlib/Geometry/Euclidean/Circumcenter.lean b/Mathlib/Geometry/Euclidean/Circumcenter.lean index 991e11164415f..9961e0140f7fa 100644 --- a/Mathlib/Geometry/Euclidean/Circumcenter.lean +++ b/Mathlib/Geometry/Euclidean/Circumcenter.lean @@ -118,7 +118,10 @@ theorem existsUnique_dist_eq_of_insert {s : AffineSubspace ℝ P} dist_sq_smul_orthogonal_vadd_smul_orthogonal_vadd (orthogonalProjection_mem p) hcc _ _ (vsub_orthogonalProjection_mem_direction_orthogonal s p), ← dist_eq_norm_vsub V p, dist_comm _ cc] - field_simp [ycc₂, hy0] + -- TODO(#15486): used to be `field_simp`, but was really slow + -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [div_div, sub_div', one_mul, mul_div_assoc', + div_mul_eq_mul_div, add_div', eq_div_iff, div_eq_iff, ycc₂] ring · rw [dist_sq_eq_dist_orthogonalProjection_sq_add_dist_orthogonalProjection_sq _ (hps hp1), orthogonalProjection_vadd_smul_vsub_orthogonalProjection _ _ hcc, Subtype.coe_mk, diff --git a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean index 5d868b2a052ee..ea3592eaa2d83 100644 --- a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean @@ -92,7 +92,12 @@ theorem dist_inversion_center (c x : P) (R : ℝ) : dist (inversion c R x) c = R rcases eq_or_ne x c with (rfl | hx) · simp have : dist x c ≠ 0 := dist_ne_zero.2 hx - field_simp [inversion, norm_smul, abs_div, ← dist_eq_norm_vsub, sq, mul_assoc] + -- was `field_simp [inversion, norm_smul, abs_div, ← dist_eq_norm_vsub, sq, mul_assoc]`, + -- but really slow. Replaced by `simp only ...` to speed up. + -- TODO(#15486): reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [inversion, sq, mul_div_assoc', div_mul_eq_mul_div, + div_div, dist_vadd_left, norm_smul, norm_div, norm_mul, Real.norm_eq_abs, abs_mul_abs_self, + abs_dist, ← dist_eq_norm_vsub, mul_assoc, eq_div_iff, div_eq_iff] /-- Distance from the center of an inversion to the image of a point under the inversion. This formula accidentally works for `x = c`. -/ @@ -159,7 +164,11 @@ theorem dist_inversion_mul_dist_center_eq (hx : x ≠ c) (hy : y ≠ c) : conv in dist _ y => rw [← inversion_inversion c hR y] rw [dist_inversion_inversion hx hy', dist_inversion_center] have : dist x c ≠ 0 := dist_ne_zero.2 hx - field_simp; ring + -- used to be `field_simp`, but was really slow; replaced by `simp only ...` to speed up + -- TODO(#15486): reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [mul_div_assoc', div_div_eq_mul_div, div_mul_eq_mul_div, + div_eq_iff] + ring /-! ### Ptolemy's inequality diff --git a/Mathlib/Geometry/Euclidean/Inversion/ImageHyperplane.lean b/Mathlib/Geometry/Euclidean/Inversion/ImageHyperplane.lean index e441919cf0ed2..fd018a7a039bf 100644 --- a/Mathlib/Geometry/Euclidean/Inversion/ImageHyperplane.lean +++ b/Mathlib/Geometry/Euclidean/Inversion/ImageHyperplane.lean @@ -39,6 +39,7 @@ theorem inversion_mem_perpBisector_inversion_iff (hR : R ≠ 0) (hx : x ≠ c) ( rw [mem_perpBisector_iff_dist_eq, dist_inversion_inversion hx hy, dist_inversion_center] have hx' := dist_ne_zero.2 hx have hy' := dist_ne_zero.2 hy + -- takes 300ms, but the "equivalent" simp call fails -> hard to speed up field_simp [mul_assoc, mul_comm, hx, hx.symm, eq_comm] /-- The inversion with center `c` and radius `R` maps a sphere passing through the center to a diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index 77b59b018017f..95434096a7f94 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -114,7 +114,10 @@ theorem sum_mongePointWeightsWithCircumcenter (n : ℕ) : nsmul_eq_mul] -- Porting note: replaced -- have hn1 : (n + 1 : ℝ) ≠ 0 := mod_cast Nat.succ_ne_zero _ - field_simp [n.cast_add_one_ne_zero] + -- TODO(#15486): used to be `field_simp [n.cast_add_one_ne_zero]`, but was really slow + -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [Nat.cast_add, Nat.cast_ofNat, Nat.cast_one, + inv_eq_one_div, mul_div_assoc', mul_one, add_div', div_mul_cancel₀, div_eq_iff, one_mul] ring /-- The Monge point of an (n+2)-simplex, in terms of @@ -142,7 +145,11 @@ theorem mongePoint_eq_affineCombination_of_pointsWithCircumcenter {n : ℕ} -- have hn3 : (n + 2 + 1 : ℝ) ≠ 0 := mod_cast Nat.succ_ne_zero _ have hn3 : (n + 2 + 1 : ℝ) ≠ 0 := by norm_cast field_simp [hn1, hn3, mul_comm] - · field_simp [hn1] + · -- TODO(#15486): used to be `field_simp [hn1]`, but was really slow + -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only + [Nat.cast_add, Nat.cast_ofNat, Nat.cast_one, zero_sub, mul_neg, mul_one, neg_div', + neg_add_rev, div_add', one_mul, eq_div_iff, div_mul_cancel₀] ring /-- The weights for the Monge point of an (n+2)-simplex, minus the diff --git a/Mathlib/Geometry/Euclidean/Sphere/Ptolemy.lean b/Mathlib/Geometry/Euclidean/Sphere/Ptolemy.lean index aa5f3e7cb0a63..3a06ee64c0d95 100644 --- a/Mathlib/Geometry/Euclidean/Sphere/Ptolemy.lean +++ b/Mathlib/Geometry/Euclidean/Sphere/Ptolemy.lean @@ -64,6 +64,7 @@ theorem mul_dist_add_mul_dist_eq_mul_dist_of_cospherical {a b c d p : P} all_goals field_simp [mul_comm, hmul] have h₃ : dist d p = dist a p * dist c p / dist b p := by field_simp [mul_comm, hmul] have h₄ : ∀ x y : ℝ, x * (y * x) = x * x * y := fun x y => by rw [mul_left_comm, mul_comm] + -- takes 450ms, but the "equivalent" simp call leaves some remaining goals field_simp [h₁, h₂, dist_eq_add_dist_of_angle_eq_pi hbpd, h₃, hbp, dist_comm a b, h₄, ← sq, dist_sq_mul_dist_add_dist_sq_mul_dist b, hapc] diff --git a/Mathlib/Geometry/Euclidean/Triangle.lean b/Mathlib/Geometry/Euclidean/Triangle.lean index 768058a8cbe5d..b91f8c483a927 100644 --- a/Mathlib/Geometry/Euclidean/Triangle.lean +++ b/Mathlib/Geometry/Euclidean/Triangle.lean @@ -134,7 +134,10 @@ theorem cos_angle_sub_add_angle_sub_rev_eq_neg_cos_angle {x y : V} (hx : x ≠ 0 Real.mul_self_sqrt (sub_nonneg_of_le (real_inner_mul_inner_self_le x y)), real_inner_self_eq_norm_mul_norm, real_inner_self_eq_norm_mul_norm, real_inner_eq_norm_mul_self_add_norm_mul_self_sub_norm_sub_mul_self_div_two] - field_simp [hxn, hyn, hxyn] + -- TODO(#15486): used to be `field_simp [hxn, hyn, hxyn]`, but was really slow + -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [sub_div', div_div, mul_div_assoc', + div_mul_eq_mul_div, div_sub', neg_div', neg_sub, eq_div_iff, div_eq_iff] ring /-- The sine of the sum of two angles in a possibly degenerate @@ -174,7 +177,10 @@ theorem sin_angle_sub_add_angle_sub_rev_eq_sin_angle {x y : V} (hx : x ≠ 0) (h inner_sub_right, inner_sub_right, inner_sub_right, inner_sub_right, real_inner_comm x y, H3, H4, real_inner_self_eq_norm_mul_norm, real_inner_self_eq_norm_mul_norm, real_inner_eq_norm_mul_self_add_norm_mul_self_sub_norm_sub_mul_self_div_two] - field_simp [hxn, hyn, hxyn] + -- TODO(#15486): used to be `field_simp [hxn, hyn, hxyn]`, but was really slow + -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [mul_div_assoc', div_mul_eq_mul_div, div_div, + sub_div', Real.sqrt_div', Real.sqrt_mul_self, add_div', div_add', eq_div_iff, div_eq_iff] ring /-- The cosine of the sum of the angles of a possibly degenerate @@ -322,8 +328,19 @@ theorem dist_sq_add_dist_sq_eq_two_mul_dist_midpoint_sq_add_half_dist_sq (a b c simp only [m, dist_left_midpoint, dist_right_midpoint, Real.norm_two] at hm calc dist a b ^ 2 + dist a c ^ 2 = 2 / dist b c * (dist a b ^ 2 * - ((2 : ℝ)⁻¹ * dist b c) + dist a c ^ 2 * (2⁻¹ * dist b c)) := by field_simp; ring - _ = 2 * (dist a (midpoint ℝ b c) ^ 2 + (dist b c / 2) ^ 2) := by rw [hm]; field_simp; ring + ((2 : ℝ)⁻¹ * dist b c) + dist a c ^ 2 * (2⁻¹ * dist b c)) := by + -- TODO(#15486): used to be `field_simp`, but was really slow + -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [inv_eq_one_div, div_mul_eq_mul_div, one_mul, + mul_div_assoc', add_div', div_mul_cancel₀, div_div, eq_div_iff] + ring + _ = 2 * (dist a (midpoint ℝ b c) ^ 2 + (dist b c / 2) ^ 2) := by + rw [hm] + -- TODO(#15486): used to be `field_simp`, but was really slow + -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [inv_eq_one_div, div_mul_eq_mul_div, one_mul, + mul_div_assoc', div_div, add_div', div_pow, eq_div_iff, div_eq_iff] + ring theorem dist_mul_of_eq_angle_of_dist_mul (a b c a' b' c' : P) (r : ℝ) (h : ∠ a' b' c' = ∠ a b c) (hab : dist a' b' = r * dist a b) (hcb : dist c' b' = r * dist c b) : From 75da87b914b453245c47afd74134ad25b62953f1 Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 12 Aug 2024 15:28:15 +0000 Subject: [PATCH 216/359] perf(Geometry/Manifold/Instances/Sphere): speed up (#15488) Speed up the file by 9% by doing the following - squeeze simps where the full form is short and this is a clear win - "squeeze" three really slow field_simp calls: replace them by the equivalent simp invocation and squeeze that - avoid one slow `convert` - squeeze two really slow nlinarith calls (they're still slow, but much better) - replace one congr! by congr This is a reduced version of #13456, which is hopefully less controversial. --- .../Geometry/Manifold/Instances/Sphere.lean | 44 ++++++++++++------- 1 file changed, 28 insertions(+), 16 deletions(-) diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index cd4fb9f824108..83c790d52ad36 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -127,7 +127,7 @@ theorem stereoInvFunAux_mem (hv : ‖v‖ = 1) {w : E} (hw : w ∈ (ℝ ∙ v) simp only [mem_sphere_zero_iff_norm, norm_smul, Real.norm_eq_abs, abs_inv, this, abs_of_pos h₁, stereoInvFunAux_apply, inv_mul_cancel₀ h₁.ne'] suffices ‖(4 : ℝ) • w + (‖w‖ ^ 2 - 4) • v‖ ^ 2 = (‖w‖ ^ 2 + 4) ^ 2 by - simpa [sq_eq_sq_iff_abs_eq_abs, abs_of_pos h₁] using this + simpa only [sq_eq_sq_iff_abs_eq_abs, abs_norm, abs_of_pos h₁] using this rw [Submodule.mem_orthogonal_singleton_iff_inner_left] at hw simp [norm_add_sq_real, norm_smul, inner_smul_left, inner_smul_right, hw, mul_pow, Real.norm_eq_abs, hv] @@ -137,7 +137,7 @@ theorem hasFDerivAt_stereoInvFunAux (v : E) : HasFDerivAt (stereoInvFunAux v) (ContinuousLinearMap.id ℝ E) 0 := by have h₀ : HasFDerivAt (fun w : E => ‖w‖ ^ 2) (0 : E →L[ℝ] ℝ) 0 := by convert (hasStrictFDerivAt_norm_sq (0 : E)).hasFDerivAt - simp + simp only [map_zero, smul_zero] have h₁ : HasFDerivAt (fun w : E => (‖w‖ ^ 2 + 4)⁻¹) (0 : E →L[ℝ] ℝ) 0 := by convert (hasFDerivAt_inv _).comp _ (h₀.add (hasFDerivAt_const 4 0)) <;> simp have h₂ : HasFDerivAt (fun w => (4 : ℝ) • w + (‖w‖ ^ 2 - 4) • v) @@ -154,7 +154,7 @@ theorem hasFDerivAt_stereoInvFunAux_comp_coe (v : E) : HasFDerivAt (stereoInvFunAux v ∘ ((↑) : (ℝ ∙ v)ᗮ → E)) (ℝ ∙ v)ᗮ.subtypeL 0 := by have : HasFDerivAt (stereoInvFunAux v) (ContinuousLinearMap.id ℝ E) ((ℝ ∙ v)ᗮ.subtypeL 0) := hasFDerivAt_stereoInvFunAux v - convert this.comp (0 : (ℝ ∙ v)ᗮ) (by apply ContinuousLinearMap.hasFDerivAt) + 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 ℝ @@ -216,13 +216,20 @@ theorem stereo_left_inv (hv : ‖v‖ = 1) {x : sphere (0 : E) 1} (hx : (x : E) linarith -- the core of the problem is these two algebraic identities: have h₁ : (2 ^ 2 / (1 - a) ^ 2 * ‖y‖ ^ 2 + 4)⁻¹ * 4 * (2 / (1 - a)) = 1 := by - field_simp; simp only [Submodule.coe_norm] at *; nlinarith + -- TODO(#15486): used to be `field_simp`, but was really slow + -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [AddSubgroupClass.coe_norm, div_mul_eq_mul_div, + div_add', inv_div, mul_div_assoc', div_div, div_eq_iff, one_mul] + simp only [Submodule.coe_norm] at *; nlinarith only [pythag] have h₂ : (2 ^ 2 / (1 - a) ^ 2 * ‖y‖ ^ 2 + 4)⁻¹ * (2 ^ 2 / (1 - a) ^ 2 * ‖y‖ ^ 2 - 4) = a := by - field_simp + -- TODO(#15486): used to be `field_simp`, but was really slow + -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [AddSubgroupClass.coe_norm, div_mul_eq_mul_div, + div_add', inv_div, div_sub', mul_div_assoc', div_div, div_eq_iff] transitivity (1 - a) ^ 2 * (a * (2 ^ 2 * ‖y‖ ^ 2 + 4 * (1 - a) ^ 2)) · congr simp only [Submodule.coe_norm] at * - nlinarith + nlinarith only [pythag] ring! convert congr_arg₂ Add.add (congr_arg (fun t => t • (y : E)) h₁) (congr_arg (fun t => t • v) h₂) using 1 @@ -233,13 +240,18 @@ theorem stereo_left_inv (hv : ‖v‖ = 1) {x : sphere (0 : E) 1} (hx : (x : E) ac_rfl -- Porting note: this branch did not exit in ml3 · rw [split, add_comm] - congr! + congr dsimp rw [one_smul] theorem stereo_right_inv (hv : ‖v‖ = 1) (w : (ℝ ∙ v)ᗮ) : stereoToFun v (stereoInvFun hv w) = w := by have : 2 / (1 - (‖(w : E)‖ ^ 2 + 4)⁻¹ * (‖(w : E)‖ ^ 2 - 4)) * (‖(w : E)‖ ^ 2 + 4)⁻¹ * 4 = 1 := by - field_simp; ring + -- TODO(#15486): used to be `field_simp`, but was really slow + -- replaced by `simp only ...` to speed up. Reinstate `field_simp` once it is faster. + simp (disch := field_simp_discharge) only [inv_eq_one_div, div_mul_eq_mul_div, one_mul, + sub_div', add_sub_sub_cancel, div_div_eq_mul_div, mul_div_assoc', mul_one, div_div, + div_eq_iff] + ring convert congr_arg (· • w) this · have h₁ : orthogonalProjection (ℝ ∙ v)ᗮ v = 0 := orthogonalProjection_orthogonalComplement_singleton_eq_zero v @@ -463,6 +475,12 @@ theorem contMDiff_neg_sphere {n : ℕ} [Fact (finrank ℝ E = n + 1)] : apply contDiff_neg.contMDiff.comp _ exact contMDiff_coe_sphere +private lemma stereographic'_neg {n : ℕ} [Fact (finrank ℝ E = n + 1)] (v : sphere (0 : E) 1) : + stereographic' n (-v) v = 0 := by + dsimp [stereographic'] + simp only [AddEquivClass.map_eq_zero_iff] + apply stereographic_neg_apply + /-- Consider the differential of the inclusion of the sphere in `E` at the point `v` as a continuous linear map from `TangentSpace (𝓡 n) v` to `E`. The range of this map is the orthogonal complement of `v` in `E`. @@ -486,10 +504,7 @@ theorem range_mfderiv_coe_sphere {n : ℕ} [Fact (finrank ℝ E = n + 1)] (v : s suffices LinearMap.range (fderiv ℝ ((stereoInvFunAux (-v : E) ∘ (↑)) ∘ U.symm) 0) = (ℝ ∙ (v : E))ᗮ by convert this using 3 - show stereographic' n (-v) v = 0 - dsimp [stereographic'] - simp only [AddEquivClass.map_eq_zero_iff] - apply stereographic_neg_apply + apply stereographic'_neg have : HasFDerivAt (stereoInvFunAux (-v : E) ∘ (Subtype.val : (ℝ ∙ (↑(-v) : E))ᗮ → E)) (ℝ ∙ (↑(-v) : E))ᗮ.subtypeL (U.symm 0) := by @@ -523,10 +538,7 @@ theorem mfderiv_coe_sphere_injective {n : ℕ} [Fact (finrank ℝ E = n + 1)] (v (𝕜 := ℝ) n (ne_zero_of_mem_unit_sphere (-v))).repr suffices Injective (fderiv ℝ ((stereoInvFunAux (-v : E) ∘ (↑)) ∘ U.symm) 0) by convert this using 3 - show stereographic' n (-v) v = 0 - dsimp [stereographic'] - simp only [AddEquivClass.map_eq_zero_iff] - apply stereographic_neg_apply + apply stereographic'_neg have : HasFDerivAt (stereoInvFunAux (-v : E) ∘ (Subtype.val : (ℝ ∙ (↑(-v) : E))ᗮ → E)) (ℝ ∙ (↑(-v) : E))ᗮ.subtypeL (U.symm 0) := by convert hasFDerivAt_stereoInvFunAux_comp_coe (-v : E) From d6d9c22637923ee6c76efb6a56e0c56b83bb7461 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 12 Aug 2024 15:39:15 +0000 Subject: [PATCH 217/359] chore: remove unnecessary _root_ qualifications (#15704) --- Mathlib/Algebra/Group/Equiv/Basic.lean | 8 ++++---- Mathlib/Data/NNReal/Basic.lean | 6 +++--- Mathlib/NumberTheory/KummerDedekind.lean | 2 +- Mathlib/RingTheory/Algebraic.lean | 6 +++--- Mathlib/RingTheory/ClassGroup.lean | 2 +- Mathlib/RingTheory/Ideal/Norm.lean | 4 ++-- Mathlib/RingTheory/Kaehler/Basic.lean | 4 ++-- Mathlib/RingTheory/Localization/Submodule.lean | 8 ++++---- Mathlib/RingTheory/MvPolynomial/Localization.lean | 6 +++--- .../RingTheory/Polynomial/Eisenstein/IsIntegral.lean | 6 +++--- Mathlib/RingTheory/RootsOfUnity/Basic.lean | 6 +++--- Mathlib/RingTheory/TensorProduct/Basic.lean | 10 +++++----- 12 files changed, 34 insertions(+), 34 deletions(-) diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index f39dd56b339f8..6e39d0838d7ec 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -213,7 +213,7 @@ theorem coe_toMulHom {f : M ≃* N} : (f.toMulHom : M → N) = f := rfl /-- A multiplicative isomorphism preserves multiplication. -/ @[to_additive "An additive isomorphism preserves addition."] protected theorem map_mul (f : M ≃* N) : ∀ x y, f (x * y) = f x * f y := - _root_.map_mul f + map_mul f attribute [deprecated map_mul (since := "2024-08-08")] MulEquiv.map_mul attribute [deprecated map_add (since := "2024-08-08")] AddEquiv.map_add @@ -458,7 +458,7 @@ lemma comp_right_injective (e : M ≃* N) : Injective fun f : P →* M ↦ (e : @[to_additive "An additive isomorphism of additive monoids sends `0` to `0` (and is hence an additive monoid isomorphism)."] -protected theorem map_one (h : M ≃* N) : h 1 = 1 := _root_.map_one h +protected theorem map_one (h : M ≃* N) : h 1 = 1 := map_one h @[to_additive] protected theorem map_eq_one_iff (h : M ≃* N) {x : M} : h x = 1 ↔ x = 1 := @@ -578,13 +578,13 @@ def piUnique {ι : Type*} (M : ι → Type*) [∀ j, Mul (M j)] [Unique ι] : @[to_additive "An additive equivalence of additive groups preserves negation."] protected theorem map_inv [Group G] [DivisionMonoid H] (h : G ≃* H) (x : G) : h x⁻¹ = (h x)⁻¹ := - _root_.map_inv h x + map_inv h x /-- A multiplicative equivalence of groups preserves division. -/ @[to_additive "An additive equivalence of additive groups preserves subtractions."] protected theorem map_div [Group G] [DivisionMonoid H] (h : G ≃* H) (x y : G) : h (x / y) = h x / h y := - _root_.map_div h x y + map_div h x y end MulEquiv diff --git a/Mathlib/Data/NNReal/Basic.lean b/Mathlib/Data/NNReal/Basic.lean index 18de3f526e468..86a021f070b0d 100644 --- a/Mathlib/Data/NNReal/Basic.lean +++ b/Mathlib/Data/NNReal/Basic.lean @@ -1063,15 +1063,15 @@ theorem NNReal.exists_lt_of_strictMono [h : Nontrivial Γ₀ˣ] {f : Γ₀ →* have hfu : f u < 1 := by rw [hu] split_ifs with hu1 - · rw [← _root_.map_one f]; exact hf hu1 + · rw [← map_one f]; exact hf hu1 · have hfg0 : f g ≠ 0 := fun h0 ↦ (Units.ne_zero g) ((map_eq_zero f).mp h0) have hg1' : 1 < g := lt_of_le_of_ne (not_lt.mp hu1) hg1.symm - rw [Units.val_inv_eq_inv_val, map_inv₀, inv_lt_one_iff hfg0, ← _root_.map_one f] + rw [Units.val_inv_eq_inv_val, map_inv₀, inv_lt_one_iff hfg0, ← map_one f] exact hf hg1' obtain ⟨n, hn⟩ := exists_pow_lt_of_lt_one hr hfu use u ^ n - rwa [Units.val_pow_eq_pow_val, _root_.map_pow] + rwa [Units.val_pow_eq_pow_val, map_pow] /-- If `Γ₀ˣ` is nontrivial and `f : Γ₀ →*₀ ℝ≥0` is strictly monotone, then for any positive real `r`, there exists `d : Γ₀ˣ` with `f d < r`. -/ diff --git a/Mathlib/NumberTheory/KummerDedekind.lean b/Mathlib/NumberTheory/KummerDedekind.lean index 151ea19a4a0dc..747f0b3cbec26 100644 --- a/Mathlib/NumberTheory/KummerDedekind.lean +++ b/Mathlib/NumberTheory/KummerDedekind.lean @@ -97,7 +97,7 @@ lemma mem_coeSubmodule_conductor {L} [CommRing L] [Algebra S L] [Algebra R L] exact ⟨_, hy z, map_mul _ _ _⟩ · intro H obtain ⟨y, _, e⟩ := H 1 - rw [_root_.map_one, mul_one] at e + rw [map_one, mul_one] at e subst e simp only [← _root_.map_mul, (NoZeroSMulDivisors.algebraMap_injective S L).eq_iff, exists_eq_right] at H diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index 96954689661a2..0326b17fa9001 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -109,10 +109,10 @@ theorem isAlgebraic_zero [Nontrivial R] : IsAlgebraic R (0 : A) := /-- An element of `R` is algebraic, when viewed as an element of the `R`-algebra `A`. -/ theorem isAlgebraic_algebraMap [Nontrivial R] (x : R) : IsAlgebraic R (algebraMap R A x) := - ⟨_, X_sub_C_ne_zero x, by rw [_root_.map_sub, aeval_X, aeval_C, sub_self]⟩ + ⟨_, X_sub_C_ne_zero x, by rw [map_sub, aeval_X, aeval_C, sub_self]⟩ theorem isAlgebraic_one [Nontrivial R] : IsAlgebraic R (1 : A) := by - rw [← _root_.map_one (algebraMap R A)] + rw [← map_one (algebraMap R A)] exact isAlgebraic_algebraMap 1 theorem isAlgebraic_nat [Nontrivial R] (n : ℕ) : IsAlgebraic R (n : A) := by @@ -120,7 +120,7 @@ theorem isAlgebraic_nat [Nontrivial R] (n : ℕ) : IsAlgebraic R (n : A) := by exact isAlgebraic_algebraMap (Nat.cast n) theorem isAlgebraic_int [Nontrivial R] (n : ℤ) : IsAlgebraic R (n : A) := by - rw [← _root_.map_intCast (algebraMap R A)] + rw [← map_intCast (algebraMap R A)] exact isAlgebraic_algebraMap (Int.cast n) theorem isAlgebraic_rat (R : Type u) {A : Type v} [DivisionRing A] [Field R] [Algebra R A] (n : ℚ) : diff --git a/Mathlib/RingTheory/ClassGroup.lean b/Mathlib/RingTheory/ClassGroup.lean index 036bd50a67491..9037779c7cd2a 100644 --- a/Mathlib/RingTheory/ClassGroup.lean +++ b/Mathlib/RingTheory/ClassGroup.lean @@ -136,7 +136,7 @@ theorem ClassGroup.mk_eq_mk_of_coe_ideal {I J : (FractionalIdeal R⁰ <| Fractio theorem ClassGroup.mk_eq_one_of_coe_ideal {I : (FractionalIdeal R⁰ <| FractionRing R)ˣ} {I' : Ideal R} (hI : (I : FractionalIdeal R⁰ <| FractionRing R) = I') : ClassGroup.mk I = 1 ↔ ∃ x : R, x ≠ 0 ∧ I' = Ideal.span {x} := by - rw [← _root_.map_one (ClassGroup.mk (R := R) (K := FractionRing R)), + rw [← map_one (ClassGroup.mk (R := R) (K := FractionRing R)), ClassGroup.mk_eq_mk_of_coe_ideal hI (?_ : _ = ↑(⊤ : Ideal R))] any_goals rfl constructor diff --git a/Mathlib/RingTheory/Ideal/Norm.lean b/Mathlib/RingTheory/Ideal/Norm.lean index 68a098c40b8ed..b30d0820ffb94 100644 --- a/Mathlib/RingTheory/Ideal/Norm.lean +++ b/Mathlib/RingTheory/Ideal/Norm.lean @@ -222,7 +222,7 @@ theorem absNorm_apply (I : Ideal S) : absNorm I = cardQuot I := rfl theorem absNorm_bot : absNorm (⊥ : Ideal S) = 0 := by rw [← Ideal.zero_eq_bot, _root_.map_zero] @[simp] -theorem absNorm_top : absNorm (⊤ : Ideal S) = 1 := by rw [← Ideal.one_eq_top, _root_.map_one] +theorem absNorm_top : absNorm (⊤ : Ideal S) = 1 := by rw [← Ideal.one_eq_top, map_one] @[simp] theorem absNorm_eq_one_iff {I : Ideal S} : absNorm I = 1 ↔ I = ⊤ := by @@ -593,7 +593,7 @@ theorem spanNorm_eq (I : Ideal S) : spanNorm R I = relNorm R I := rfl @[simp] theorem relNorm_bot : relNorm R (⊥ : Ideal S) = ⊥ := by - simpa only [zero_eq_bot] using _root_.map_zero (relNorm R : Ideal S →*₀ _) + simpa only [zero_eq_bot] using map_zero (relNorm R : Ideal S →*₀ _) @[simp] theorem relNorm_top : relNorm R (⊤ : Ideal S) = ⊤ := by diff --git a/Mathlib/RingTheory/Kaehler/Basic.lean b/Mathlib/RingTheory/Kaehler/Basic.lean index c57b401f8224a..661532a73b713 100644 --- a/Mathlib/RingTheory/Kaehler/Basic.lean +++ b/Mathlib/RingTheory/Kaehler/Basic.lean @@ -835,7 +835,7 @@ theorem KaehlerDifferential.range_kerCotangentToTensor rw [KaehlerDifferential.ker_map_of_surjective R A B h] at hx obtain ⟨x, hx, rfl⟩ := hx simp only [TensorProduct.lid_symm_apply, LinearMap.rTensor_tmul, - Algebra.linearMap_apply, _root_.map_one] + Algebra.linearMap_apply, map_one] rw [← Finsupp.sum_single x, Finsupp.sum, ← Finset.sum_fiberwise_of_maps_to (fun _ ↦ Finset.mem_image_of_mem (algebraMap A B))] simp only [Function.comp_apply, map_sum (s := x.support.image (algebraMap A B)), @@ -843,7 +843,7 @@ theorem KaehlerDifferential.range_kerCotangentToTensor apply sum_mem intro c _ simp only [Finset.filter_congr_decidable, TensorProduct.lid_symm_apply, LinearMap.rTensor_tmul, - AlgHom.toLinearMap_apply, _root_.map_one, LinearMap.mem_range] + AlgHom.toLinearMap_apply, map_one, LinearMap.mem_range] simp only [map_sum, Finsupp.total_single] have : (x.support.filter (algebraMap A B · = c)).sum x ∈ RingHom.ker (algebraMap A B) := by simpa [Finsupp.mapDomain, Finsupp.sum, Finsupp.finset_sum_apply, RingHom.mem_ker, diff --git a/Mathlib/RingTheory/Localization/Submodule.lean b/Mathlib/RingTheory/Localization/Submodule.lean index eae7c15f91fdf..eecf67203492a 100644 --- a/Mathlib/RingTheory/Localization/Submodule.lean +++ b/Mathlib/RingTheory/Localization/Submodule.lean @@ -125,8 +125,8 @@ theorem mem_span_iff {N : Type*} [AddCommGroup N] [Module R N] [Module S N] [IsS · intro h refine Submodule.span_induction h ?_ ?_ ?_ ?_ · rintro x hx - exact ⟨x, Submodule.subset_span hx, 1, by rw [mk'_one, _root_.map_one, one_smul]⟩ - · exact ⟨0, Submodule.zero_mem _, 1, by rw [mk'_one, _root_.map_one, one_smul]⟩ + exact ⟨x, Submodule.subset_span hx, 1, by rw [mk'_one, map_one, one_smul]⟩ + · exact ⟨0, Submodule.zero_mem _, 1, by rw [mk'_one, map_one, one_smul]⟩ · rintro _ _ ⟨y, hy, z, rfl⟩ ⟨y', hy', z', rfl⟩ refine ⟨(z' : R) • y + (z : R) • y', @@ -134,8 +134,8 @@ theorem mem_span_iff {N : Type*} [AddCommGroup N] [Module R N] [Module S N] [IsS rw [smul_add, ← IsScalarTower.algebraMap_smul S (z : R), ← IsScalarTower.algebraMap_smul S (z' : R), smul_smul, smul_smul] congr 1 - · rw [← mul_one (1 : R), mk'_mul, mul_assoc, mk'_spec, _root_.map_one, mul_one, mul_one] - · rw [← mul_one (1 : R), mk'_mul, mul_right_comm, mk'_spec, _root_.map_one, mul_one, one_mul] + · rw [← mul_one (1 : R), mk'_mul, mul_assoc, mk'_spec, map_one, mul_one, mul_one] + · rw [← mul_one (1 : R), mk'_mul, mul_right_comm, mk'_spec, map_one, mul_one, one_mul] · rintro a _ ⟨y, hy, z, rfl⟩ obtain ⟨y', z', rfl⟩ := mk'_surjective M a refine ⟨y' • y, Submodule.smul_mem _ _ hy, z' * z, ?_⟩ diff --git a/Mathlib/RingTheory/MvPolynomial/Localization.lean b/Mathlib/RingTheory/MvPolynomial/Localization.lean index 6bb865ab1f60a..552085840486d 100644 --- a/Mathlib/RingTheory/MvPolynomial/Localization.lean +++ b/Mathlib/RingTheory/MvPolynomial/Localization.lean @@ -113,7 +113,7 @@ def auxInv : S →+* (MvPolynomial Unit R) ⧸ Ideal.span { C r * X () - 1 } := simp only [RingHom.coe_comp, Function.comp_apply, g] rw [isUnit_iff_exists_inv] use (Ideal.Quotient.mk _ <| X ()) - rw [← _root_.map_mul, ← map_one (Ideal.Quotient.mk _), Ideal.Quotient.mk_eq_mk_iff_sub_mem] + rw [← map_mul, ← map_one (Ideal.Quotient.mk _), Ideal.Quotient.mk_eq_mk_iff_sub_mem] exact Ideal.mem_span_singleton_self (C r * X () - 1) private lemma auxHom_auxInv : (auxHom S r).toRingHom.comp (auxInv S r) = RingHom.id S := by @@ -129,8 +129,8 @@ private lemma auxInv_auxHom : (auxInv S r).comp (auxHom (S := S) r).toRingHom = Function.comp_apply, auxHom_mk, aeval_X, RingHomCompTriple.comp_eq] erw [IsLocalization.lift_mk'_spec] simp only [map_one, RingHom.coe_comp, Function.comp_apply] - rw [← _root_.map_one (Ideal.Quotient.mk _)] - rw [← _root_.map_mul, Ideal.Quotient.mk_eq_mk_iff_sub_mem, ← Ideal.neg_mem_iff, neg_sub] + rw [← map_one (Ideal.Quotient.mk _)] + rw [← map_mul, Ideal.Quotient.mk_eq_mk_iff_sub_mem, ← Ideal.neg_mem_iff, neg_sub] exact Ideal.mem_span_singleton_self (C r * X x - 1) /-- The canonical algebra isomorphism from `MvPolynomial Unit R` quotiented by diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean index b85746a6d864f..50e6698cdb932 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean @@ -173,7 +173,7 @@ theorem dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt {B : Pow -- Do the computation in `K` so we can work in terms of `z` instead of `r`. apply IsFractionRing.injective R K - simp only [_root_.map_mul, _root_.map_pow, _root_.map_neg, _root_.map_one] + simp only [_root_.map_mul, map_pow, map_neg, map_one] -- Both sides are actually norms: calc _ = norm K (Q.coeff 0 • B.gen ^ n) := ?_ @@ -182,7 +182,7 @@ theorem dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt {B : Pow (congr_arg (norm K) (eq_sub_of_add_eq ?_)) _ = _ := ?_ · simp only [Algebra.smul_def, algebraMap_apply R K L, Algebra.norm_algebraMap, _root_.map_mul, - _root_.map_pow, finrank_K_L, PowerBasis.norm_gen_eq_coeff_zero_minpoly, + map_pow, finrank_K_L, PowerBasis.norm_gen_eq_coeff_zero_minpoly, minpoly.isIntegrallyClosed_eq_field_fractions' K hBint, coeff_map, ← hn] ring swap @@ -193,7 +193,7 @@ theorem dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt {B : Pow B.gen ^ n := ?_ _ = (Q.coeff 0 • B.gen ^ 0 + ∑ x ∈ (range (Q.natDegree + 1)).erase 0, Q.coeff x • B.gen ^ x) * B.gen ^ n := by - rw [_root_.pow_zero] + rw [pow_zero] _ = aeval B.gen Q * B.gen ^ n := ?_ _ = _ := by rw [hQ, Algebra.smul_mul_assoc] · have : ∀ i ∈ (range (Q.natDegree + 1)).erase 0, diff --git a/Mathlib/RingTheory/RootsOfUnity/Basic.lean b/Mathlib/RingTheory/RootsOfUnity/Basic.lean index fdfc31cb9f3fc..04f0e45847dbc 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Basic.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Basic.lean @@ -467,7 +467,7 @@ variable [FunLike F M N] theorem map_of_injective [MonoidHomClass F M N] (h : IsPrimitiveRoot ζ k) (hf : Injective f) : IsPrimitiveRoot (f ζ) k where - pow_eq_one := by rw [← map_pow, h.pow_eq_one, _root_.map_one] + pow_eq_one := by rw [← map_pow, h.pow_eq_one, map_one] dvd_of_pow_eq_one := by rw [h.eq_orderOf] intro l hl @@ -476,12 +476,12 @@ theorem map_of_injective [MonoidHomClass F M N] (h : IsPrimitiveRoot ζ k) (hf : theorem of_map_of_injective [MonoidHomClass F M N] (h : IsPrimitiveRoot (f ζ) k) (hf : Injective f) : IsPrimitiveRoot ζ k where - pow_eq_one := by apply_fun f; rw [map_pow, _root_.map_one, h.pow_eq_one] + pow_eq_one := by apply_fun f; rw [map_pow, map_one, h.pow_eq_one] dvd_of_pow_eq_one := by rw [h.eq_orderOf] intro l hl apply_fun f at hl - rw [map_pow, _root_.map_one] at hl + rw [map_pow, map_one] at hl exact orderOf_dvd_of_pow_eq_one hl theorem map_iff_of_injective [MonoidHomClass F M N] (hf : Injective f) : diff --git a/Mathlib/RingTheory/TensorProduct/Basic.lean b/Mathlib/RingTheory/TensorProduct/Basic.lean index e215d9a539db5..78c5b9dc2e6f9 100644 --- a/Mathlib/RingTheory/TensorProduct/Basic.lean +++ b/Mathlib/RingTheory/TensorProduct/Basic.lean @@ -462,7 +462,7 @@ theorem ext ⦃f g : (A ⊗[R] B) →ₐ[S] C⦄ ext a b have := congr_arg₂ HMul.hMul (AlgHom.congr_fun ha a) (AlgHom.congr_fun hb b) dsimp at * - rwa [← _root_.map_mul, ← _root_.map_mul, tmul_mul_tmul, _root_.one_mul, _root_.mul_one] at this + rwa [← _root_.map_mul, ← _root_.map_mul, tmul_mul_tmul, one_mul, mul_one] at this theorem ext' {g h : A ⊗[R] B →ₐ[S] C} (H : ∀ a b, g (a ⊗ₜ b) = h (a ⊗ₜ b)) : g = h := ext (AlgHom.ext fun _ => H _ _) (AlgHom.ext fun _ => H _ _) @@ -908,7 +908,7 @@ theorem map_range (f : A →ₐ[R] B) (g : C →ₐ[R] D) : apply le_antisymm · rw [← map_top, ← adjoin_tmul_eq_top, ← adjoin_image, adjoin_le_iff] rintro _ ⟨_, ⟨a, b, rfl⟩, rfl⟩ - rw [map_tmul, ← _root_.mul_one (f a), ← _root_.one_mul (g b), ← tmul_mul_tmul] + rw [map_tmul, ← mul_one (f a), ← one_mul (g b), ← tmul_mul_tmul] exact mul_mem_sup (AlgHom.mem_range_self _ a) (AlgHom.mem_range_self _ b) · rw [← map_comp_includeLeft f g, ← map_comp_includeRight f g] exact sup_le (AlgHom.range_comp_le_range _ _) (AlgHom.range_comp_le_range _ _) @@ -977,7 +977,7 @@ variable (R) def lmul' : S ⊗[R] S →ₐ[R] S := algHomOfLinearMapTensorProduct (LinearMap.mul' R S) (fun a₁ a₂ b₁ b₂ => by simp only [LinearMap.mul'_apply, mul_mul_mul_comm]) <| by - simp only [LinearMap.mul'_apply, _root_.mul_one] + simp only [LinearMap.mul'_apply, mul_one] variable {R} @@ -990,11 +990,11 @@ theorem lmul'_apply_tmul (a b : S) : lmul' (S := S) R (a ⊗ₜ[R] b) = a * b := @[simp] theorem lmul'_comp_includeLeft : (lmul' R : _ →ₐ[R] S).comp includeLeft = AlgHom.id R S := - AlgHom.ext <| _root_.mul_one + AlgHom.ext <| mul_one @[simp] theorem lmul'_comp_includeRight : (lmul' R : _ →ₐ[R] S).comp includeRight = AlgHom.id R S := - AlgHom.ext <| _root_.one_mul + AlgHom.ext <| one_mul /-- If `S` is commutative, for a pair of morphisms `f : A →ₐ[R] S`, `g : B →ₐ[R] S`, We obtain a map `A ⊗[R] B →ₐ[R] S` that commutes with `f`, `g` via `a ⊗ b ↦ f(a) * g(b)`. From f2cd6eff0575997cb37e755f7249be3d1c5e2f9c Mon Sep 17 00:00:00 2001 From: Etienne Date: Mon, 12 Aug 2024 16:36:26 +0000 Subject: [PATCH 218/359] doc: Add a note on how to use `CompactlyGeneratedSpace` (#15132) Following this discussion : https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Compactly.20generated.20spaces/near/453958229. Indicate that `CompactlyGeneratedSpace` should be used with an explicit universe parameter. --- Mathlib/Topology/Category/CompactlyGenerated.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Topology/Category/CompactlyGenerated.lean b/Mathlib/Topology/Category/CompactlyGenerated.lean index 43a1129864eb3..200b869bfa079 100644 --- a/Mathlib/Topology/Category/CompactlyGenerated.lean +++ b/Mathlib/Topology/Category/CompactlyGenerated.lean @@ -26,7 +26,8 @@ universe u w open CategoryTheory Topology TopologicalSpace -/-- The type of `u`-compactly generated `w`-small topological spaces. -/ +/-- `CompactlyGenerated.{u, w}` is the type of `u`-compactly generated `w`-small topological spaces. +This should always be used with explicit universe parameters. -/ structure CompactlyGenerated where /-- The underlying topological space of an object of `CompactlyGenerated`. -/ toTop : TopCat.{w} From 795a856503c1a55d06ab4f1cd61caa6131cbec45 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 12 Aug 2024 16:50:11 +0000 Subject: [PATCH 219/359] chore: uncomment some now viable assert_not_exists (#15715) Per [zulip](https://leanprover.zulipchat.com/#narrow/stream/345428-mathlib-reviewers/topic/.60assert_not_exists.60.20at.20top.20or.20bottom.20of.20file.3F/near/459960523) discussion. Co-authored-by: adomani Co-authored-by: Johan Commelin --- Mathlib/Algebra/Algebra/Defs.lean | 3 +-- Mathlib/Algebra/Associated/Basic.lean | 5 ++--- Mathlib/Algebra/Group/Aut.lean | 3 +-- Mathlib/Algebra/Group/Conj.lean | 3 +-- Mathlib/Algebra/Group/ConjFinite.lean | 2 +- Mathlib/Algebra/Group/Defs.lean | 5 ++--- Mathlib/Algebra/Group/Hom/Instances.lean | 3 +-- Mathlib/Algebra/Group/Subgroup/Basic.lean | 5 ++--- Mathlib/Algebra/Module/Basic.lean | 3 +-- Mathlib/Algebra/Order/CauSeq/Basic.lean | 3 +-- Mathlib/Algebra/Order/Group/Defs.lean | 9 ++++---- .../Algebra/Order/Group/Unbundled/Int.lean | 1 - Mathlib/Algebra/Order/Ring/Defs.lean | 4 ++-- .../Algebra/Order/Ring/Unbundled/Basic.lean | 1 - Mathlib/Algebra/Ring/Center.lean | 8 +++---- Mathlib/Algebra/Ring/Defs.lean | 22 ++++++++++--------- Mathlib/Algebra/Ring/Equiv.lean | 7 +++--- Mathlib/Algebra/Ring/Hom/Defs.lean | 9 ++++---- Mathlib/Analysis/Normed/Field/Basic.lean | 6 ++--- Mathlib/Analysis/Seminorm.lean | 4 ++-- Mathlib/CategoryTheory/Limits/Yoneda.lean | 4 ++-- Mathlib/Data/Complex/Basic.lean | 6 ++--- Mathlib/Data/Fin/Basic.lean | 4 +--- Mathlib/Data/Finset/Basic.lean | 1 - Mathlib/Data/List/Basic.lean | 3 +-- Mathlib/Data/Nat/Factors.lean | 4 ++-- .../FieldTheory/PolynomialGaloisGroup.lean | 3 +-- Mathlib/GroupTheory/Perm/Fin.lean | 3 +-- Mathlib/GroupTheory/Subgroup/Center.lean | 5 ++--- Mathlib/GroupTheory/Submonoid/Center.lean | 5 ++--- .../GroupTheory/Submonoid/Centralizer.lean | 5 ++--- Mathlib/GroupTheory/Subsemigroup/Center.lean | 5 ++--- .../GroupTheory/Subsemigroup/Centralizer.lean | 5 ++--- Mathlib/LinearAlgebra/Basis/Basic.lean | 3 +-- Mathlib/LinearAlgebra/Basis/Defs.lean | 5 ++--- Mathlib/MeasureTheory/Function/AEEqFun.lean | 6 ++--- .../Function/SpecialFunctions/Basic.lean | 9 ++++---- .../Function/StronglyMeasurable/Basic.lean | 5 ++--- Mathlib/NumberTheory/Wilson.lean | 3 +-- Mathlib/Order/CompleteLatticeIntervals.lean | 4 ++-- .../ConditionallyCompleteLattice/Basic.lean | 5 ++--- Mathlib/Order/Hom/Basic.lean | 5 ++--- Mathlib/Order/Interval/Finset/Nat.lean | 3 +-- Mathlib/Order/SemiconjSup.lean | 5 ++--- Mathlib/Util/AssertExists.lean | 2 ++ 45 files changed, 92 insertions(+), 122 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Defs.lean b/Mathlib/Algebra/Algebra/Defs.lean index 00f6fc494fb64..10be082a45ea6 100644 --- a/Mathlib/Algebra/Algebra/Defs.lean +++ b/Mathlib/Algebra/Algebra/Defs.lean @@ -82,6 +82,7 @@ the second approach only when you need to weaken a condition on either `R` or `A -/ assert_not_exists Field +assert_not_exists Module.End universe u v w u₁ v₁ @@ -380,5 +381,3 @@ theorem algebraMap.coe_smul (A B C : Type*) [SMul A B] [CommSemiring B] [Semirin ((a • b : B) : C) = (a • b) • 1 := Algebra.algebraMap_eq_smul_one _ _ = a • (b • 1) := smul_assoc .. _ = a • (b : C) := congrArg _ (Algebra.algebraMap_eq_smul_one b).symm - -assert_not_exists Module.End diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index cfa5e0393c92a..a37f091f6c2ac 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -28,6 +28,8 @@ Then we show that the quotient type `Associates` is a monoid and prove basic properties of this quotient. -/ +assert_not_exists OrderedCommMonoid +assert_not_exists Multiset variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} @@ -1144,6 +1146,3 @@ theorem dvd_prime_pow [CancelCommMonoidWithZero α] {p q : α} (hp : Prime p) (n exact ⟨i, hi.trans n.le_succ, hq⟩ end CancelCommMonoidWithZero - -assert_not_exists OrderedCommMonoid -assert_not_exists Multiset diff --git a/Mathlib/Algebra/Group/Aut.lean b/Mathlib/Algebra/Group/Aut.lean index 5f7cba0d26ab5..9a89372872dd6 100644 --- a/Mathlib/Algebra/Group/Aut.lean +++ b/Mathlib/Algebra/Group/Aut.lean @@ -26,8 +26,7 @@ equivalences (and other files that use them) before the group structure is defin MulAut, AddAut -/ --- TODO after #13161 --- assert_not_exists MonoidWithZero +assert_not_exists MonoidWithZero assert_not_exists Ring variable {A : Type*} {M : Type*} {G : Type*} diff --git a/Mathlib/Algebra/Group/Conj.lean b/Mathlib/Algebra/Group/Conj.lean index 0be980f48df22..77426efb55ff3 100644 --- a/Mathlib/Algebra/Group/Conj.lean +++ b/Mathlib/Algebra/Group/Conj.lean @@ -12,8 +12,7 @@ import Mathlib.Algebra.Group.Semiconj.Units See also `MulAut.conj` and `Quandle.conj`. -/ --- TODO: After #13027, --- assert_not_exists MonoidWithZero +assert_not_exists MonoidWithZero assert_not_exists Multiset universe u v diff --git a/Mathlib/Algebra/Group/ConjFinite.lean b/Mathlib/Algebra/Group/ConjFinite.lean index 5130245903e72..a432f381360cd 100644 --- a/Mathlib/Algebra/Group/ConjFinite.lean +++ b/Mathlib/Algebra/Group/ConjFinite.lean @@ -11,7 +11,7 @@ import Mathlib.Data.Fintype.Units # Conjugacy of elements of finite groups -/ --- TODO: After #13027, +-- TODO: the following `assert_not_exists` should work, but does not -- assert_not_exists MonoidWithZero variable {α : Type*} [Monoid α] diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index f005d732edc45..078d8fe493dd1 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -48,6 +48,8 @@ See the module `Algebra.AddTorsor` for a motivating example for the name `VAdd` assert_not_exists MonoidWithZero assert_not_exists DenselyOrdered +assert_not_exists Function.Injective.eq_iff +assert_not_exists IsCommutative universe u v w @@ -1179,6 +1181,3 @@ initialize_simps_projections Group initialize_simps_projections AddGroup initialize_simps_projections CommGroup initialize_simps_projections AddCommGroup - -assert_not_exists Function.Injective.eq_iff -assert_not_exists IsCommutative diff --git a/Mathlib/Algebra/Group/Hom/Instances.lean b/Mathlib/Algebra/Group/Hom/Instances.lean index 51101f1a4791d..3b1f5e56a0684 100644 --- a/Mathlib/Algebra/Group/Hom/Instances.lean +++ b/Mathlib/Algebra/Group/Hom/Instances.lean @@ -20,6 +20,7 @@ Finally, we provide the `Ring` structure on `AddMonoid.End`. -/ assert_not_exists AddMonoidWithOne +assert_not_exists Ring universe uM uN uP uQ @@ -233,5 +234,3 @@ theorem compr₂_apply [MulOneClass M] [MulOneClass N] [CommMonoid P] [CommMonoi rfl end MonoidHom - -assert_not_exists Ring diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 187ab2f00201a..c307e84d8c382 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -82,6 +82,8 @@ subgroup, subgroups -/ assert_not_exists OrderedAddCommMonoid +assert_not_exists Multiset +assert_not_exists Ring open Function open Int @@ -2904,6 +2906,3 @@ def noncenter (G : Type*) [Monoid G] : Set (ConjClasses G) := g ∈ noncenter G ↔ g.carrier.Nontrivial := Iff.rfl end ConjClasses - -assert_not_exists Multiset -assert_not_exists Ring diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index b5e777b79a1e3..91e832b7f7568 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -15,6 +15,7 @@ import Mathlib.GroupTheory.GroupAction.Group -/ assert_not_exists Nonneg.inv +assert_not_exists Multiset open Function Set @@ -202,5 +203,3 @@ lemma smul_indicator_one_apply (s : Set α) (r : R) (a : α) : end MulZeroOneClass end Set - -assert_not_exists Multiset diff --git a/Mathlib/Algebra/Order/CauSeq/Basic.lean b/Mathlib/Algebra/Order/CauSeq/Basic.lean index 9523087b69005..abb991002f5a5 100644 --- a/Mathlib/Algebra/Order/CauSeq/Basic.lean +++ b/Mathlib/Algebra/Order/CauSeq/Basic.lean @@ -35,6 +35,7 @@ assert_not_exists Finset assert_not_exists Module assert_not_exists Submonoid assert_not_exists FloorRing +assert_not_exists Module variable {α β : Type*} @@ -877,5 +878,3 @@ protected theorem sup_inf_distrib_right (a b c : CauSeq α abs) : a ⊓ b ⊔ c end Abs end CauSeq - -assert_not_exists Module diff --git a/Mathlib/Algebra/Order/Group/Defs.lean b/Mathlib/Algebra/Order/Group/Defs.lean index ee0b0e870b100..c486a4984cd8a 100644 --- a/Mathlib/Algebra/Order/Group/Defs.lean +++ b/Mathlib/Algebra/Order/Group/Defs.lean @@ -20,6 +20,10 @@ may differ between the multiplicative and the additive version of a lemma. The reason is that we did not want to change existing names in the library. -/ +/- +`NeZero` should not be needed at this point in the ordered algebraic hierarchy. +-/ +assert_not_exists NeZero open Function @@ -194,8 +198,3 @@ theorem one_le_inv_of_le_one : a ≤ 1 → 1 ≤ a⁻¹ := one_le_inv'.mpr end NormNumLemmas - -/- -`NeZero` should not be needed at this point in the ordered algebraic hierarchy. --/ -assert_not_exists NeZero diff --git a/Mathlib/Algebra/Order/Group/Unbundled/Int.lean b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean index 6082988933594..773e4ce5f46eb 100644 --- a/Mathlib/Algebra/Order/Group/Unbundled/Int.lean +++ b/Mathlib/Algebra/Order/Group/Unbundled/Int.lean @@ -23,7 +23,6 @@ See note [foundational algebra order theory]. -- We should need only a minimal development of sets in order to get here. assert_not_exists Set.Subsingleton - assert_not_exists Ring open Function Nat diff --git a/Mathlib/Algebra/Order/Ring/Defs.lean b/Mathlib/Algebra/Order/Ring/Defs.lean index 949a3f6e5a05c..61fcecb798dbf 100644 --- a/Mathlib/Algebra/Order/Ring/Defs.lean +++ b/Mathlib/Algebra/Order/Ring/Defs.lean @@ -103,6 +103,8 @@ immediate predecessors and what conditions are added to each of them. - `CommRing` & `IsDomain` & linear order structure -/ +assert_not_exists MonoidHom + open Function universe u @@ -422,5 +424,3 @@ instance (priority := 100) LinearOrderedCommRing.toStrictOrderedCommRing instance (priority := 100) LinearOrderedCommRing.toLinearOrderedCommSemiring [d : LinearOrderedCommRing α] : LinearOrderedCommSemiring α := { d, LinearOrderedRing.toLinearOrderedSemiring with } - -assert_not_exists MonoidHom diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean index 4ec7c3d1885a2..e9a10a1f07b78 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean @@ -113,7 +113,6 @@ TODO: the mixin assumptiosn can be relaxed in most cases -/ - assert_not_exists OrderedCommMonoid assert_not_exists MonoidHom diff --git a/Mathlib/Algebra/Ring/Center.lean b/Mathlib/Algebra/Ring/Center.lean index bf93ca846a37d..d28091e849958 100644 --- a/Mathlib/Algebra/Ring/Center.lean +++ b/Mathlib/Algebra/Ring/Center.lean @@ -11,6 +11,10 @@ import Mathlib.Data.Int.Cast.Lemmas -/ +-- Guard against import creep +assert_not_exists Finset +assert_not_exists Subsemigroup + variable {M : Type*} @@ -83,7 +87,3 @@ theorem neg_mem_center [NonUnitalNonAssocRing M] {a : M} (ha : a ∈ Set.center right_assoc _ _ := by rw [mul_neg, ha.right_assoc, mul_neg, mul_neg] end Set - --- Guard against import creep -assert_not_exists Finset -assert_not_exists Subsemigroup diff --git a/Mathlib/Algebra/Ring/Defs.lean b/Mathlib/Algebra/Ring/Defs.lean index 3cc194109119e..18b144825141a 100644 --- a/Mathlib/Algebra/Ring/Defs.lean +++ b/Mathlib/Algebra/Ring/Defs.lean @@ -30,6 +30,18 @@ the present file is about their interaction. `Semiring`, `CommSemiring`, `Ring`, `CommRing`, domain, `IsDomain`, nonzero, units -/ + +/-! +Previously an import dependency on `Mathlib.Algebra.Group.Basic` had crept in. +In general, the `.Defs` files in the basic algebraic hierarchy should only depend on earlier `.Defs` +files, without importing `.Basic` theory development. + +These `assert_not_exists` statements guard against this returning. +-/ +assert_not_exists DivisionMonoid.toDivInvOneMonoid +assert_not_exists mul_rotate + + universe u v w x variable {α : Type u} {β : Type v} {γ : Type w} {R : Type x} @@ -411,13 +423,3 @@ is cancellative on both sides. In other words, a nontrivial semiring `R` satisfy This is implemented as a mixin for `Semiring α`. To obtain an integral domain use `[CommRing α] [IsDomain α]`. -/ class IsDomain (α : Type u) [Semiring α] extends IsCancelMulZero α, Nontrivial α : Prop - -/-! -Previously an import dependency on `Mathlib.Algebra.Group.Basic` had crept in. -In general, the `.Defs` files in the basic algebraic hierarchy should only depend on earlier `.Defs` -files, without importing `.Basic` theory development. - -These `assert_not_exists` statements guard against this returning. --/ -assert_not_exists DivisionMonoid.toDivInvOneMonoid -assert_not_exists mul_rotate diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index 33303fe81c670..f08987f2fa38f 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -38,6 +38,9 @@ multiplication in `Equiv.Perm`, and multiplication in `CategoryTheory.End`, not Equiv, MulEquiv, AddEquiv, RingEquiv, MulAut, AddAut, RingAut -/ +-- guard against import creep +assert_not_exists Field +assert_not_exists Fintype variable {F α β R S S' : Type*} @@ -828,7 +831,3 @@ protected theorem isDomain {A : Type*} (B : Type*) [Semiring A] [Semiring B] [Is exists_pair_ne := ⟨e.symm 0, e.symm 1, e.symm.injective.ne zero_ne_one⟩ } end MulEquiv - --- guard against import creep -assert_not_exists Field -assert_not_exists Fintype diff --git a/Mathlib/Algebra/Ring/Hom/Defs.lean b/Mathlib/Algebra/Ring/Hom/Defs.lean index a440563a5296e..5b45c997cbab6 100644 --- a/Mathlib/Algebra/Ring/Hom/Defs.lean +++ b/Mathlib/Algebra/Ring/Hom/Defs.lean @@ -42,6 +42,10 @@ groups, we use the same structure `RingHom a β`, a.k.a. `α →+* β`, for both `RingHom`, `SemiringHom` -/ +assert_not_exists Function.Injective.mulZeroClass +assert_not_exists semigroupDvd +assert_not_exists Units.map +assert_not_exists Set.range open Function @@ -654,8 +658,3 @@ theorem coe_addMonoidHom_mkRingHomOfMulSelfOfTwoNeZero (h h_two h_one) : rfl end AddMonoidHom - -assert_not_exists Function.Injective.mulZeroClass -assert_not_exists semigroupDvd -assert_not_exists Units.map -assert_not_exists Set.range diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index a8160837b5391..fbcfe54ff4ad9 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -22,6 +22,9 @@ In this file we define (semi)normed rings and fields. We also prove some theorem definitions. -/ +-- Guard against import creep. +assert_not_exists RestrictScalars + variable {α : Type*} {β : Type*} {γ : Type*} {ι : Type*} open Filter Metric Bornology @@ -1206,6 +1209,3 @@ instance toNormedCommRing [NormedCommRing R] [SubringClass S R] (s : S) : Normed { SubringClass.toNormedRing s with mul_comm := mul_comm } end SubringClass - --- Guard against import creep. -assert_not_exists RestrictScalars diff --git a/Mathlib/Analysis/Seminorm.lean b/Mathlib/Analysis/Seminorm.lean index ed51cb1a81624..99aa49d8ceb0b 100644 --- a/Mathlib/Analysis/Seminorm.lean +++ b/Mathlib/Analysis/Seminorm.lean @@ -32,6 +32,8 @@ For a module over a normed ring: seminorm, locally convex, LCTVS -/ +assert_not_exists balancedCore + open NormedField Set Filter open scoped NNReal Pointwise Topology Uniformity @@ -1327,5 +1329,3 @@ lemma rescale_to_shell [NormedAddCommGroup F] [NormedSpace 𝕜 F] {c : 𝕜} (h rescale_to_shell_semi_normed hc εpos (norm_ne_zero_iff.mpr hx) end normSeminorm - -assert_not_exists balancedCore diff --git a/Mathlib/CategoryTheory/Limits/Yoneda.lean b/Mathlib/CategoryTheory/Limits/Yoneda.lean index 07b2b79a18c53..1e0a5bb19da0e 100644 --- a/Mathlib/CategoryTheory/Limits/Yoneda.lean +++ b/Mathlib/CategoryTheory/Limits/Yoneda.lean @@ -16,6 +16,8 @@ We calculate the colimit of `Y ↦ (X ⟶ Y)`, which is just `PUnit`. We also show the (co)yoneda embeddings preserve limits and jointly reflect them. -/ +assert_not_exists AddCommMonoid + open Opposite CategoryTheory Limits universe t w v u @@ -226,5 +228,3 @@ end Corepresentable end Functor end CategoryTheory - -assert_not_exists AddCommMonoid diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 2b1a91b9ab87f..6483469eea3bc 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -18,6 +18,9 @@ of characteristic zero. The result that the complex numbers are algebraically cl `FieldTheory.AlgebraicClosure`. -/ +assert_not_exists Multiset +assert_not_exists Algebra + open Set Function /-! ### Definition and basic arithmetic -/ @@ -809,6 +812,3 @@ unsafe instance instRepr : Repr ℂ where reprPrec f.re 65 ++ " + " ++ reprPrec f.im 70 ++ "*I" end Complex - -assert_not_exists Multiset -assert_not_exists Algebra diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 4ae9d41455378..0e9ba529e710b 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -68,7 +68,7 @@ This file expands on the development in the core library. -/ assert_not_exists Monoid - +assert_not_exists Fintype universe u v open Fin Nat Function @@ -544,8 +544,6 @@ def castLEEmb (h : n ≤ m) : Fin n ↪ Fin m where /- The next proof can be golfed a lot using `Fintype.card`. It is written this way to define `ENat.card` and `Nat.card` without a `Fintype` dependency (not done yet). -/ -assert_not_exists Fintype - lemma nonempty_embedding_iff : Nonempty (Fin n ↪ Fin m) ↔ n ≤ m := by refine ⟨fun h ↦ ?_, fun h ↦ ⟨castLEEmb h⟩⟩ induction n generalizing m with diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index b0dfbf24bd6a0..2eff3772f98da 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -119,7 +119,6 @@ finite sets, finset -- Note that we cannot use `List.sublists` itself as that is defined very early. assert_not_exists List.sublistsLen assert_not_exists Multiset.Powerset - assert_not_exists CompleteLattice open Multiset Subtype Nat Function diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index f0e9fc75a0292..ff79ec56c9722 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -21,6 +21,7 @@ import Batteries.Data.List.Perm assert_not_exists Set.range assert_not_exists GroupWithZero assert_not_exists Ring +assert_not_exists Lattice open Function @@ -2600,5 +2601,3 @@ lemma lookup_graph (f : α → β) {a : α} {as : List α} (h : a ∈ as) : end lookup end List - -assert_not_exists Lattice diff --git a/Mathlib/Data/Nat/Factors.lean b/Mathlib/Data/Nat/Factors.lean index 10a0826c4731c..3f99b646ebaa3 100644 --- a/Mathlib/Data/Nat/Factors.lean +++ b/Mathlib/Data/Nat/Factors.lean @@ -20,6 +20,8 @@ This file deals with the factors of natural numbers. -/ +assert_not_exists Multiset + open Bool Subtype open Nat @@ -312,5 +314,3 @@ theorem four_dvd_or_exists_odd_prime_and_dvd_of_two_lt {n : ℕ} (n2 : 2 < n) : · exact Or.inr ⟨p, hp, hdvd, hodd⟩ end Nat - -assert_not_exists Multiset diff --git a/Mathlib/FieldTheory/PolynomialGaloisGroup.lean b/Mathlib/FieldTheory/PolynomialGaloisGroup.lean index b445a15253fc7..3c936d1dc75f4 100644 --- a/Mathlib/FieldTheory/PolynomialGaloisGroup.lean +++ b/Mathlib/FieldTheory/PolynomialGaloisGroup.lean @@ -35,6 +35,7 @@ equals the number of real roots plus the number of roots not fixed by complex co -/ +assert_not_exists Real noncomputable section @@ -399,5 +400,3 @@ theorem prime_degree_dvd_card [CharZero F] (p_irr : Irreducible p) (p_deg : p.na end Gal end Polynomial - -assert_not_exists Real diff --git a/Mathlib/GroupTheory/Perm/Fin.lean b/Mathlib/GroupTheory/Perm/Fin.lean index bcf82c2c47d7d..b5ba110fa6dae 100644 --- a/Mathlib/GroupTheory/Perm/Fin.lean +++ b/Mathlib/GroupTheory/Perm/Fin.lean @@ -12,6 +12,7 @@ import Mathlib.Logic.Equiv.Fintype # Permutations of `Fin n` -/ +assert_not_exists LinearMap open Equiv @@ -275,5 +276,3 @@ theorem isThreeCycle_cycleRange_two {n : ℕ} : IsThreeCycle (cycleRange 2 : Per end Fin end CycleRange - -assert_not_exists LinearMap diff --git a/Mathlib/GroupTheory/Subgroup/Center.lean b/Mathlib/GroupTheory/Subgroup/Center.lean index e0be786c9f89a..a3b3323d2c5ff 100644 --- a/Mathlib/GroupTheory/Subgroup/Center.lean +++ b/Mathlib/GroupTheory/Subgroup/Center.lean @@ -12,6 +12,8 @@ import Mathlib.GroupTheory.Submonoid.Center -/ +assert_not_exists Multiset +assert_not_exists Ring open Function open Int @@ -142,6 +144,3 @@ theorem mk_bijOn (G : Type*) [Group G] : exact ⟨h, rfl⟩ end ConjClasses - -assert_not_exists Multiset -assert_not_exists Ring diff --git a/Mathlib/GroupTheory/Submonoid/Center.lean b/Mathlib/GroupTheory/Submonoid/Center.lean index ac76048902ea0..44028df37b868 100644 --- a/Mathlib/GroupTheory/Submonoid/Center.lean +++ b/Mathlib/GroupTheory/Submonoid/Center.lean @@ -18,6 +18,8 @@ We provide `Subgroup.center`, `AddSubgroup.center`, `Subsemiring.center`, and `S other files. -/ +-- Guard against import creep +assert_not_exists Finset namespace Submonoid @@ -124,6 +126,3 @@ def unitsCenterToCenterUnits [Monoid M] : (Submonoid.center M)ˣ →* Submonoid. theorem unitsCenterToCenterUnits_injective [Monoid M] : Function.Injective (unitsCenterToCenterUnits M) := fun _a _b h => Units.ext <| Subtype.ext <| congr_arg (Units.val ∘ Subtype.val) h - --- Guard against import creep -assert_not_exists Finset diff --git a/Mathlib/GroupTheory/Submonoid/Centralizer.lean b/Mathlib/GroupTheory/Submonoid/Centralizer.lean index c057fe0f0b26e..3ad0ece4f0e49 100644 --- a/Mathlib/GroupTheory/Submonoid/Centralizer.lean +++ b/Mathlib/GroupTheory/Submonoid/Centralizer.lean @@ -17,6 +17,8 @@ import Mathlib.GroupTheory.Submonoid.Center We provide `Subgroup.centralizer`, `AddSubgroup.centralizer` in other files. -/ +-- Guard against import creep +assert_not_exists Finset variable {M : Type*} {S T : Set M} @@ -89,6 +91,3 @@ lemma centralizer_centralizer_centralizer {s : Set M} : end end Submonoid - --- Guard against import creep -assert_not_exists Finset diff --git a/Mathlib/GroupTheory/Subsemigroup/Center.lean b/Mathlib/GroupTheory/Subsemigroup/Center.lean index 86b175a0e7560..d0970e7e4ca85 100644 --- a/Mathlib/GroupTheory/Subsemigroup/Center.lean +++ b/Mathlib/GroupTheory/Subsemigroup/Center.lean @@ -23,6 +23,8 @@ We provide `Submonoid.center`, `AddSubmonoid.center`, `Subgroup.center`, `AddSub [cabreragarciarodriguezpalacios2014] -/ +-- Guard against import creep +assert_not_exists Finset /-! ### `Set.center` as a `Subsemigroup`. -/ @@ -76,6 +78,3 @@ theorem center_eq_top : center M = ⊤ := end CommSemigroup end Subsemigroup - --- Guard against import creep -assert_not_exists Finset diff --git a/Mathlib/GroupTheory/Subsemigroup/Centralizer.lean b/Mathlib/GroupTheory/Subsemigroup/Centralizer.lean index 2c6163cce16b0..40211ea65474b 100644 --- a/Mathlib/GroupTheory/Subsemigroup/Centralizer.lean +++ b/Mathlib/GroupTheory/Subsemigroup/Centralizer.lean @@ -18,6 +18,8 @@ We provide `Monoid.centralizer`, `AddMonoid.centralizer`, `Subgroup.centralizer` `AddSubgroup.centralizer` in other files. -/ +-- Guard against import creep +assert_not_exists Finset variable {M : Type*} {S T : Set M} namespace Subsemigroup @@ -67,6 +69,3 @@ theorem centralizer_univ : centralizer Set.univ = center M := end end Subsemigroup - --- Guard against import creep -assert_not_exists Finset diff --git a/Mathlib/LinearAlgebra/Basis/Basic.lean b/Mathlib/LinearAlgebra/Basis/Basic.lean index 3ec32691a84b3..bbf77e0742586 100644 --- a/Mathlib/LinearAlgebra/Basis/Basic.lean +++ b/Mathlib/LinearAlgebra/Basis/Basic.lean @@ -16,6 +16,7 @@ Further results on bases in modules and vector spaces. -/ +assert_not_exists Ordinal noncomputable section @@ -559,5 +560,3 @@ theorem Basis.mem_span_iff_repr_mem (m : M) : exact smul_mem _ _ (subset_span (Set.mem_range_self i)) end RestrictScalars - -assert_not_exists Ordinal diff --git a/Mathlib/LinearAlgebra/Basis/Defs.lean b/Mathlib/LinearAlgebra/Basis/Defs.lean index b77f16e72455a..23deb6b640ceb 100644 --- a/Mathlib/LinearAlgebra/Basis/Defs.lean +++ b/Mathlib/LinearAlgebra/Basis/Defs.lean @@ -52,6 +52,8 @@ basis, bases -/ +assert_not_exists LinearIndependent +assert_not_exists Cardinal noncomputable section @@ -832,6 +834,3 @@ theorem sum_repr_mul_repr {ι'} [Fintype ι'] (b' : Basis ι' R M) (x : M) (i : end Basis end CommSemiring - -assert_not_exists LinearIndependent -assert_not_exists Cardinal diff --git a/Mathlib/MeasureTheory/Function/AEEqFun.lean b/Mathlib/MeasureTheory/Function/AEEqFun.lean index 4a391c3dc6ad6..cb6030efa880d 100644 --- a/Mathlib/MeasureTheory/Function/AEEqFun.lean +++ b/Mathlib/MeasureTheory/Function/AEEqFun.lean @@ -67,6 +67,9 @@ function space, almost everywhere equal, `L⁰`, ae_eq_fun -/ +-- Guard against import creep +assert_not_exists InnerProductSpace + noncomputable section open Topology Set Filter TopologicalSpace ENNReal EMetric MeasureTheory Function @@ -902,6 +905,3 @@ def toAEEqFunLinearMap : C(α, γ) →ₗ[𝕜] α →ₘ[μ] γ := map_smul' := fun c f => AEEqFun.smul_mk c f f.continuous.aestronglyMeasurable } end ContinuousMap - --- Guard against import creep -assert_not_exists InnerProductSpace diff --git a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean index 5b68867a440a9..43e55d7c27774 100644 --- a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean +++ b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean @@ -17,6 +17,10 @@ See also `MeasureTheory.Function.SpecialFunctions.Arctan` and `MeasureTheory.Function.SpecialFunctions.Inner`, which have been split off to minimize imports. -/ +-- Guard against import creep: +assert_not_exists InnerProductSpace +assert_not_exists Real.arctan +assert_not_exists FiniteDimensional.proper noncomputable section @@ -223,8 +227,3 @@ instance ENNReal.hasMeasurablePow : MeasurablePow ℝ≥0∞ ℝ := by exact Measurable.ite (measurableSet_singleton 0) measurable_const measurable_const end PowInstances - --- Guard against import creep: -assert_not_exists InnerProductSpace -assert_not_exists Real.arctan -assert_not_exists FiniteDimensional.proper diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index c5f62b96d4297..e5e60b719ee57 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -54,6 +54,8 @@ measurable functions, as a basis for the Bochner integral. -/ +-- Guard against import creep +assert_not_exists InnerProductSpace open MeasureTheory Filter TopologicalSpace Function Set MeasureTheory.Measure @@ -1875,6 +1877,3 @@ theorem stronglyMeasurable_uncurry_of_continuous_of_stronglyMeasurable {α β ι exact ((t_sf n).measurable.comp measurable_fst).subtype_mk end MeasureTheory - --- Guard against import creep -assert_not_exists InnerProductSpace diff --git a/Mathlib/NumberTheory/Wilson.lean b/Mathlib/NumberTheory/Wilson.lean index fb2001f449b1a..a5d360e59a890 100644 --- a/Mathlib/NumberTheory/Wilson.lean +++ b/Mathlib/NumberTheory/Wilson.lean @@ -24,6 +24,7 @@ This could be generalized to similar results about finite abelian groups. * Give `wilsons_lemma` a descriptive name. -/ +assert_not_exists legendre_sym.quadratic_reciprocity open Finset Nat FiniteField ZMod @@ -99,5 +100,3 @@ theorem prime_iff_fac_equiv_neg_one (h : n ≠ 1) : Prime n ↔ ((n - 1)! : ZMod exact ZMod.wilsons_lemma n end Nat - -assert_not_exists legendre_sym.quadratic_reciprocity diff --git a/Mathlib/Order/CompleteLatticeIntervals.lean b/Mathlib/Order/CompleteLatticeIntervals.lean index ddfb764e0a271..44fa32da464fc 100644 --- a/Mathlib/Order/CompleteLatticeIntervals.lean +++ b/Mathlib/Order/CompleteLatticeIntervals.lean @@ -20,6 +20,8 @@ Add appropriate instances for all `Set.Ixx`. This requires a refactor that will default values for `sSup` and `sInf`. -/ +assert_not_exists Multiset + open Set variable {ι : Sort*} {α : Type*} (s : Set α) @@ -267,5 +269,3 @@ theorem coe_biInf : (↑(⨅ i, ⨅ (_ : p i), f i) : α) = a ⊓ ⨅ i, ⨅ (_ end Set.Iic - -assert_not_exists Multiset diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index 0bd74776170ed..9ab2e4ae32457 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -29,6 +29,8 @@ while `csInf_le` is the same statement in conditionally complete lattices with an additional assumption that `s` is bounded below. -/ +-- Guard against import creep +assert_not_exists Multiset open Function OrderDual Set @@ -1636,6 +1638,3 @@ lemma iInf_coe_lt_top : ⨅ i, (f i : WithTop α) < ⊤ ↔ Nonempty ι := by end WithTop end WithTopBot - --- Guard against import creep -assert_not_exists Multiset diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index 951efec678d5d..5b9b773385b2c 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -69,6 +69,8 @@ We also define two functions to convert other bundled maps to `α →o β`: monotone map, bundled morphism -/ +-- Developments relating order homs and sets belong in `Order.Hom.Set` or later. +assert_not_exists Set.range open OrderDual @@ -1239,6 +1241,3 @@ theorem OrderIso.complementedLattice_iff (f : α ≃o β) : end BoundedOrder end LatticeIsos - --- Developments relating order homs and sets belong in `Order.Hom.Set` or later. -assert_not_exists Set.range diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index 1d1a980d68df0..94c2c0add091b 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -17,8 +17,7 @@ Some lemmas can be generalized using `OrderedGroup`, `CanonicallyOrderedCommMono and subsequently be moved upstream to `Order.Interval.Finset`. -/ --- TODO --- assert_not_exists Ring +assert_not_exists Ring open Finset Nat diff --git a/Mathlib/Order/SemiconjSup.lean b/Mathlib/Order/SemiconjSup.lean index 1a7a4802f052c..835ea7b923fb9 100644 --- a/Mathlib/Order/SemiconjSup.lean +++ b/Mathlib/Order/SemiconjSup.lean @@ -31,6 +31,8 @@ homeomorphisms of the circle, so in order to apply results from this file one ha homeomorphisms to the real line first. -/ +-- Guard against import creep +assert_not_exists Finset variable {α β γ : Type*} @@ -116,7 +118,4 @@ theorem csSup_div_semiconj [ConditionallyCompleteLattice α] [Group G] (f₁ f Function.Semiconj (fun x => ⨆ g' : G, (f₁ g')⁻¹ (f₂ g' x)) (f₂ g) (f₁ g) := semiconj_of_isLUB f₁ f₂ (fun x => isLUB_csSup (range_nonempty _) (hbdd x)) _ --- Guard against import creep -assert_not_exists Finset - end Function diff --git a/Mathlib/Util/AssertExists.lean b/Mathlib/Util/AssertExists.lean index da180c2eab390..747726c2b139b 100644 --- a/Mathlib/Util/AssertExists.lean +++ b/Mathlib/Util/AssertExists.lean @@ -47,6 +47,8 @@ it is probably because you have introduced new import dependencies to a file. In this case, you should refactor your work (for example by creating new files rather than adding imports to existing files). You should *not* delete the `assert_not_exists` statement without careful discussion ahead of time. + +`assert_not_exists` statements should generally live at the top of the file, after the module doc. -/ elab "assert_not_exists " n:ident : command => do let decl ← try liftCoreM <| realizeGlobalConstNoOverloadWithInfo n catch _ => return From c3a48abadbb2bb75b975e8eb6e6d59dc4178abcd Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Mon, 12 Aug 2024 17:47:56 +0000 Subject: [PATCH 220/359] feat(CategoryTheory): Composing `Comma.map` with `Comma.fst` and `Comma.snd` (#15730) --- Mathlib/CategoryTheory/Comma/Basic.lean | 30 +++++++++++++++++++++++-- 1 file changed, 28 insertions(+), 2 deletions(-) diff --git a/Mathlib/CategoryTheory/Comma/Basic.lean b/Mathlib/CategoryTheory/Comma/Basic.lean index f8f1cea64b759..23c4e18177717 100644 --- a/Mathlib/CategoryTheory/Comma/Basic.lean +++ b/Mathlib/CategoryTheory/Comma/Basic.lean @@ -50,12 +50,14 @@ namespace CategoryTheory open Category -- declare the `v`'s first; see `CategoryTheory.Category` for an explanation -universe v₁ v₂ v₃ v₄ v₅ u₁ u₂ u₃ u₄ u₅ +universe v₁ v₂ v₃ v₄ v₅ v₆ u₁ u₂ u₃ u₄ u₅ u₆ variable {A : Type u₁} [Category.{v₁} A] variable {B : Type u₂} [Category.{v₂} B] variable {T : Type u₃} [Category.{v₃} T] -variable {A' B' T' : Type*} [Category A'] [Category B'] [Category T'] +variable {A' : Type u₄} [Category.{v₄} A'] +variable {B' : Type u₅} [Category.{v₅} B'] +variable {T' : Type u₆} [Category.{v₆} T'] /-- The objects of the comma category are triples of an object `left : A`, an object `right : B` and a morphism `hom : L.obj left ⟶ R.obj right`. -/ @@ -251,6 +253,30 @@ noncomputable instance isEquivalenceMap [F₁.IsEquivalence] [F₂.IsEquivalence] [F.Faithful] [F.Full] [IsIso α] [IsIso β] : (map α β).IsEquivalence where +/-- The equality between `map α β ⋙ fst L' R'` and `fst L R ⋙ F₁`, +where `α : F₁ ⋙ L' ⟶ L ⋙ F`. -/ +@[simp] +theorem map_fst : map α β ⋙ fst L' R' = fst L R ⋙ F₁ := + rfl + +/-- The isomorphism between `map α β ⋙ fst L' R'` and `fst L R ⋙ F₁`, +where `α : F₁ ⋙ L' ⟶ L ⋙ F`. -/ +@[simps!] +def mapFst : map α β ⋙ fst L' R' ≅ fst L R ⋙ F₁ := + NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat) + +/-- The equality between `map α β ⋙ snd L' R'` and `snd L R ⋙ F₂`, +where `β : R ⋙ F ⟶ F₂ ⋙ R'`. -/ +@[simp] +theorem map_snd : map α β ⋙ snd L' R' = snd L R ⋙ F₂ := + rfl + +/-- The isomorphism between `map α β ⋙ snd L' R'` and `snd L R ⋙ F₂`, +where `β : R ⋙ F ⟶ F₂ ⋙ R'`. -/ +@[simps!] +def mapSnd : map α β ⋙ snd L' R' ≅ snd L R ⋙ F₂ := + NatIso.ofComponents (fun _ => Iso.refl _) (by aesop_cat) + end /-- A natural transformation `L₁ ⟶ L₂` induces a functor `Comma L₂ R ⥤ Comma L₁ R`. -/ From ef99038fb2ef6a4d5eac060da2b52078286d950b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 12 Aug 2024 18:29:50 +0000 Subject: [PATCH 221/359] chore: move to v4.11.0-rc2 (new `variable` command) (#15726) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Co-authored-by: Blizzard_inc Co-authored-by: Jeremy Tan Jie Rui Co-authored-by: Parcly Taxel Co-authored-by: Tomáš Skřivan Co-authored-by: mathlib-bors <150093616+mathlib-bors@users.noreply.github.com> --- Archive/Imo/Imo1981Q3.lean | 2 + Archive/Imo/Imo1986Q5.lean | 1 + Archive/Imo/Imo2024Q1.lean | 1 + Archive/Imo/Imo2024Q2.lean | 3 + Archive/Imo/Imo2024Q6.lean | 1 + Archive/Wiedijk100Theorems/BuffonsNeedle.lean | 8 + Archive/Wiedijk100Theorems/CubingACube.lean | 3 + .../Wiedijk100Theorems/FriendshipGraphs.lean | 15 + Mathlib/Algebra/Associated/Basic.lean | 7 +- Mathlib/Algebra/Category/ModuleCat/Free.lean | 12 + .../Category/ModuleCat/Presheaf/Sheafify.lean | 7 +- Mathlib/Algebra/Divisibility/Units.lean | 2 + .../Algebra/Group/Submonoid/Operations.lean | 8 +- .../Group/Subsemigroup/Operations.lean | 7 +- Mathlib/Algebra/GroupWithZero/InjSurj.lean | 2 + Mathlib/Algebra/Homology/BifunctorFlip.lean | 4 +- .../DerivedCategory/HomologySequence.lean | 3 + .../Homology/Embedding/IsSupported.lean | 5 +- .../Algebra/Homology/HomologySequence.lean | 2 + .../Homology/HomologySequenceLemmas.lean | 5 +- .../Homology/HomotopyCategory/ShortExact.lean | 1 + Mathlib/Algebra/Homology/HomotopyCofiber.lean | 12 +- Mathlib/Algebra/Homology/Localization.lean | 32 +- .../ShortComplex/HomologicalComplex.lean | 23 +- Mathlib/Algebra/Homology/TotalComplex.lean | 6 +- Mathlib/Algebra/Lie/Abelian.lean | 1 + Mathlib/Algebra/Lie/Engel.lean | 1 + Mathlib/Algebra/Lie/InvariantForm.lean | 9 +- Mathlib/Algebra/Lie/Nilpotent.lean | 2 + Mathlib/Algebra/Lie/Rank.lean | 2 + Mathlib/Algebra/Lie/Sl2.lean | 22 +- Mathlib/Algebra/Lie/TraceForm.lean | 31 +- Mathlib/Algebra/Lie/Weights/Basic.lean | 6 +- Mathlib/Algebra/Lie/Weights/Chain.lean | 25 +- Mathlib/Algebra/Module/LocalizedModule.lean | 2 + Mathlib/Algebra/Module/PID.lean | 6 +- Mathlib/Algebra/Module/Submodule/EqLocus.lean | 2 + .../Module/Submodule/Localization.lean | 1 + Mathlib/Algebra/Module/Submodule/Map.lean | 2 + Mathlib/Algebra/MonoidAlgebra/Degree.lean | 2 +- Mathlib/Algebra/Order/Pointwise.lean | 1 + .../Algebra/Order/Ring/Unbundled/Basic.lean | 2 +- Mathlib/Algebra/Order/ToIntervalMod.lean | 7 + Mathlib/Algebra/Polynomial/EraseLead.lean | 4 +- Mathlib/Algebra/Polynomial/Expand.lean | 2 +- Mathlib/Algebra/Polynomial/RingDivision.lean | 3 +- Mathlib/Algebra/Ring/Divisibility/Basic.lean | 6 +- Mathlib/Algebra/Ring/Hom/Defs.lean | 1 + Mathlib/Algebra/Ring/InjSurj.lean | 2 + Mathlib/Algebra/Star/NonUnitalSubalgebra.lean | 77 ++-- Mathlib/AlgebraicGeometry/AffineScheme.lean | 11 + .../Morphisms/RingHomProperties.lean | 2 + Mathlib/AlgebraicGeometry/Noetherian.lean | 1 + .../PrimeSpectrum/TensorProduct.lean | 3 +- .../ProjectiveSpectrum/Scheme.lean | 2 + Mathlib/AlgebraicGeometry/Pullbacks.lean | 1 + .../FundamentalGroupoid/InducedMaps.lean | 1 + Mathlib/Analysis/Analytic/Meromorphic.lean | 11 + Mathlib/Analysis/BoundedVariation.lean | 2 +- .../BoxIntegral/Partition/Filter.lean | 5 +- .../Instances.lean | 1 + .../NonUnital.lean | 9 + .../ContinuousFunctionalCalculus/Unital.lean | 11 + Mathlib/Analysis/CStarAlgebra/Module.lean | 1 + .../Analysis/Calculus/FDeriv/Analytic.lean | 2 + .../Analysis/Calculus/FDeriv/Symmetric.lean | 5 + .../InverseFunctionTheorem/Deriv.lean | 1 + .../Calculus/IteratedDeriv/Lemmas.lean | 17 +- Mathlib/Analysis/Calculus/MeanValue.lean | 7 + Mathlib/Analysis/Calculus/Rademacher.lean | 27 +- Mathlib/Analysis/Complex/Angle.lean | 1 - Mathlib/Analysis/Complex/Hadamard.lean | 10 +- Mathlib/Analysis/Complex/TaylorSeries.lean | 8 + Mathlib/Analysis/Complex/Tietze.lean | 3 +- Mathlib/Analysis/Convex/Function.lean | 6 - .../Analysis/Distribution/SchwartzSpace.lean | 32 +- Mathlib/Analysis/Fourier/AddCircle.lean | 2 +- Mathlib/Analysis/InnerProductSpace/Basic.lean | 13 +- .../Analysis/InnerProductSpace/Calculus.lean | 10 + .../InnerProductSpace/LinearPMap.lean | 3 + .../Analysis/InnerProductSpace/l2Space.lean | 16 +- Mathlib/Analysis/LocallyConvex/Barrelled.lean | 5 +- Mathlib/Analysis/Normed/Affine/Isometry.lean | 1 + .../Analysis/Normed/Algebra/Exponential.lean | 1 + Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 3 +- Mathlib/Analysis/Normed/Module/Basic.lean | 4 +- Mathlib/Analysis/Normed/Operator/Banach.lean | 39 +- .../Normed/Ring/SeminormFromConst.lean | 7 +- Mathlib/Analysis/NormedSpace/BallAction.lean | 2 + Mathlib/Analysis/NormedSpace/RCLike.lean | 2 +- Mathlib/Analysis/ODE/Gronwall.lean | 7 + Mathlib/Analysis/ODE/PicardLindelof.lean | 38 +- .../Analysis/SpecialFunctions/Log/Base.lean | 5 + .../SpecialFunctions/Pow/Integral.lean | 3 + Mathlib/Analysis/SpecificLimits/Basic.lean | 17 +- Mathlib/Analysis/Subadditive.lean | 2 + .../Abelian/DiagramLemmas/Four.lean | 1 + Mathlib/CategoryTheory/Abelian/Exact.lean | 1 + .../Abelian/ProjectiveResolution.lean | 14 +- Mathlib/CategoryTheory/CofilteredSystem.lean | 1 + Mathlib/CategoryTheory/ComposableArrows.lean | 1 - .../Functor/Derived/RightDerived.lean | 5 +- .../CategoryTheory/Functor/FullyFaithful.lean | 4 +- .../Functor/KanExtension/Basic.lean | 12 +- .../Functor/KanExtension/Pointwise.lean | 6 +- .../CategoryTheory/Galois/Decomposition.lean | 7 +- .../Galois/Prorepresentability.lean | 6 +- Mathlib/CategoryTheory/GradedObject.lean | 1 + .../GradedObject/Trifunctor.lean | 10 + .../CategoryTheory/GradedObject/Unitor.lean | 12 +- Mathlib/CategoryTheory/Limits/Final.lean | 2 + Mathlib/CategoryTheory/Limits/MonoCoprod.lean | 3 + .../Limits/Preserves/Shapes/Square.lean | 1 + .../CategoryTheory/Limits/Shapes/Types.lean | 1 + Mathlib/CategoryTheory/Limits/Types.lean | 1 + .../Localization/Adjunction.lean | 1 + .../Localization/Bousfield.lean | 1 + .../Localization/CalculusOfFractions.lean | 2 + .../CalculusOfFractions/Preadditive.lean | 8 +- .../Localization/Equivalence.lean | 1 + .../Localization/FiniteProducts.lean | 10 +- .../Localization/LocalizerMorphism.lean | 1 + .../CategoryTheory/Localization/SmallHom.lean | 1 + .../Localization/SmallShiftedHom.lean | 14 +- .../Localization/Triangulated.lean | 3 + .../Preadditive/AdditiveFunctor.lean | 1 + .../Sites/Coherent/ReflectsPrecoherent.lean | 1 + .../Sites/Coherent/ReflectsPreregular.lean | 1 + .../CategoryTheory/Sites/CoverLifting.lean | 1 + .../CategoryTheory/Sites/CoverPreserving.lean | 7 +- Mathlib/CategoryTheory/Sites/CoversTop.lean | 1 + .../CategoryTheory/Sites/DenseSubsite.lean | 53 +-- Mathlib/CategoryTheory/Sites/Equivalence.lean | 30 +- .../Sites/IsSheafOneHypercover.lean | 13 +- .../Sites/LocallyInjective.lean | 10 +- .../Sites/MayerVietorisSquare.lean | 1 + Mathlib/CategoryTheory/Sites/Preserves.lean | 80 +++-- Mathlib/CategoryTheory/Sites/Sheaf.lean | 1 + Mathlib/CategoryTheory/Sites/SheafHom.lean | 12 +- .../Triangulated/HomologicalFunctor.lean | 4 + .../Triangulated/Pretriangulated.lean | 1 + .../Triangulated/Subcategory.lean | 2 + .../Additive/AP/Three/Behrend.lean | 1 - Mathlib/Combinatorics/SimpleGraph/Maps.lean | 1 + .../Combinatorics/SimpleGraph/Operations.lean | 1 + .../SimpleGraph/Triangle/Tripartite.lean | 1 - Mathlib/Combinatorics/SimpleGraph/Turan.lean | 13 +- Mathlib/Combinatorics/SimpleGraph/Walk.lean | 4 +- .../Computability/AkraBazzi/AkraBazzi.lean | 106 +++--- Mathlib/Computability/TuringMachine.lean | 8 +- Mathlib/Data/FunLike/Equiv.lean | 2 +- Mathlib/Data/Int/CardIntervalMod.lean | 2 + Mathlib/Data/List/Basic.lean | 2 + Mathlib/Data/Real/ConjExponents.lean | 20 +- Mathlib/Data/Set/Lattice.lean | 1 + Mathlib/Data/Set/Sigma.lean | 1 - Mathlib/Data/Set/UnionLift.lean | 2 - Mathlib/Data/Setoid/Partition.lean | 7 +- Mathlib/Data/TwoPointing.lean | 1 + Mathlib/Data/Vector/MapLemmas.lean | 2 - Mathlib/Data/Vector/Snoc.lean | 1 - Mathlib/Deprecated/Group.lean | 29 +- Mathlib/Deprecated/Subgroup.lean | 1 + .../RotationNumber/TranslationNumber.lean | 4 +- Mathlib/FieldTheory/Extension.lean | 29 +- Mathlib/FieldTheory/KummerExtension.lean | 16 +- Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean | 70 ++-- Mathlib/FieldTheory/NormalClosure.lean | 2 + Mathlib/FieldTheory/PrimitiveElement.lean | 1 + Mathlib/FieldTheory/Separable.lean | 3 +- .../Geometry/Euclidean/Inversion/Basic.lean | 1 + Mathlib/Geometry/Euclidean/Sphere/Basic.lean | 1 + .../Geometry/Manifold/Algebra/LieGroup.lean | 2 + Mathlib/Geometry/Manifold/Algebra/Monoid.lean | 1 + Mathlib/Geometry/Manifold/Complex.lean | 4 +- .../Manifold/ContMDiff/NormedSpace.lean | 3 - Mathlib/Geometry/Manifold/IntegralCurve.lean | 7 +- .../Manifold/LocalInvariantProperties.lean | 2 + Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean | 1 + Mathlib/Geometry/Manifold/MFDeriv/Basic.lean | 337 +++++++++--------- .../Geometry/Manifold/PartitionOfUnity.lean | 12 +- .../Manifold/VectorBundle/SmoothSection.lean | 2 - .../Geometry/Manifold/WhitneyEmbedding.lean | 4 +- Mathlib/GroupTheory/CoprodI.lean | 47 ++- Mathlib/GroupTheory/CosetCover.lean | 5 +- Mathlib/GroupTheory/Coxeter/Inversion.lean | 2 + Mathlib/GroupTheory/EckmannHilton.lean | 2 + Mathlib/GroupTheory/PGroup.lean | 1 + Mathlib/GroupTheory/SchurZassenhaus.lean | 10 +- .../GroupTheory/SpecificGroups/Cyclic.lean | 29 +- Mathlib/GroupTheory/Transfer.lean | 1 + Mathlib/LinearAlgebra/AffineSpace/Basis.lean | 1 + .../BilinearForm/Properties.lean | 3 + .../CliffordAlgebra/Contraction.lean | 5 +- .../LinearAlgebra/CliffordAlgebra/Prod.lean | 1 + .../LinearAlgebra/Dimension/Localization.lean | 10 + Mathlib/LinearAlgebra/Dual.lean | 14 +- Mathlib/LinearAlgebra/LinearDisjoint.lean | 8 + Mathlib/LinearAlgebra/LinearPMap.lean | 1 + Mathlib/LinearAlgebra/Matrix/Block.lean | 2 - .../Matrix/Charpoly/LinearMap.lean | 8 +- Mathlib/LinearAlgebra/Matrix/PosDef.lean | 30 +- Mathlib/LinearAlgebra/PerfectPairing.lean | 3 + Mathlib/LinearAlgebra/Pi.lean | 8 +- Mathlib/LinearAlgebra/RootSystem/Basic.lean | 2 + .../RootSystem/RootPositive.lean | 1 + Mathlib/LinearAlgebra/Semisimple.lean | 2 + Mathlib/LinearAlgebra/SesquilinearForm.lean | 7 +- Mathlib/LinearAlgebra/Span.lean | 3 + .../TensorProduct/RightExactness.lean | 4 + Mathlib/Logic/Basic.lean | 6 +- Mathlib/Logic/Equiv/TransferInstance.lean | 1 + Mathlib/Logic/Function/Basic.lean | 2 + .../Constructions/HaarToSphere.lean | 10 +- .../Covering/Differentiation.lean | 1 + .../MeasureTheory/Covering/VitaliFamily.lean | 1 + .../ConditionalExpectation/Unique.lean | 2 + .../Function/SpecialFunctions/Basic.lean | 2 + Mathlib/MeasureTheory/Group/Action.lean | 13 +- .../Group/FundamentalDomain.lean | 56 +-- .../Integral/CircleIntegral.lean | 9 +- .../Integral/IntegralEqImproper.lean | 26 +- .../Integral/IntervalIntegral.lean | 46 +-- Mathlib/MeasureTheory/Integral/Periodic.lean | 7 +- .../MeasureTheory/Integral/SetIntegral.lean | 19 +- .../MeasureTheory/Integral/TorusIntegral.lean | 11 +- .../MeasureTheory/Measure/Haar/Quotient.lean | 41 ++- .../Measure/Lebesgue/VolumeOfBalls.lean | 5 +- Mathlib/MeasureTheory/Measure/Restrict.lean | 8 +- .../MeasureTheory/OuterMeasure/Induced.lean | 16 +- Mathlib/ModelTheory/Substructures.lean | 2 + Mathlib/NumberTheory/Cyclotomic/Basic.lean | 1 + .../Cyclotomic/PrimitiveRoots.lean | 19 +- Mathlib/NumberTheory/Dioph.lean | 7 +- .../DiophantineApproximation.lean | 12 +- Mathlib/NumberTheory/EulerProduct/Basic.lean | 15 +- Mathlib/NumberTheory/FLT/Three.lean | 1 + .../LegendreSymbol/AddCharacter.lean | 2 +- .../LegendreSymbol/QuadraticReciprocity.lean | 12 +- Mathlib/NumberTheory/Multiplicity.lean | 3 + .../NumberField/Discriminant.lean | 3 + Mathlib/NumberTheory/NumberField/House.lean | 6 + .../NumberField/Units/DirichletTheorem.lean | 1 + Mathlib/NumberTheory/Ostrowski.lean | 5 + Mathlib/NumberTheory/Padics/Hensel.lean | 131 +++---- Mathlib/NumberTheory/Padics/RingHoms.lean | 16 +- Mathlib/NumberTheory/PellMatiyasevic.lean | 2 + Mathlib/NumberTheory/PythagoreanTriples.lean | 14 +- Mathlib/Order/Atoms.lean | 2 + Mathlib/Order/Bounds/Basic.lean | 18 +- .../ConditionallyCompleteLattice/Basic.lean | 1 + Mathlib/Order/Extension/Well.lean | 1 + Mathlib/Order/GaloisConnection.lean | 28 +- Mathlib/Order/Hom/Lattice.lean | 3 +- Mathlib/Order/Interval/Finset/Fin.lean | 8 - Mathlib/Order/Interval/Set/Infinite.lean | 1 + .../Interval/Set/OrdConnectedComponent.lean | 1 - Mathlib/Order/SupIndep.lean | 4 + Mathlib/Order/WellFounded.lean | 8 +- .../Disintegration/MeasurableStieltjes.lean | 22 +- Mathlib/Probability/StrongLaw.lean | 11 +- .../AdicCompletion/AsTensorProduct.lean | 36 +- .../RingTheory/AdicCompletion/Exactness.lean | 3 + Mathlib/RingTheory/Adjoin/PowerBasis.lean | 11 +- Mathlib/RingTheory/AlgebraicIndependent.lean | 13 +- Mathlib/RingTheory/Artinian.lean | 1 + Mathlib/RingTheory/DedekindDomain/Ideal.lean | 12 +- .../DedekindDomain/IntegralClosure.lean | 1 + Mathlib/RingTheory/DedekindDomain/PID.lean | 1 + Mathlib/RingTheory/Etale/Basic.lean | 1 + Mathlib/RingTheory/Ideal/Maps.lean | 23 +- Mathlib/RingTheory/Ideal/MinimalPrime.lean | 3 +- Mathlib/RingTheory/Ideal/Over.lean | 11 +- .../IntegralClosure/IntegralRestrict.lean | 2 + Mathlib/RingTheory/IsAdjoinRoot.lean | 8 +- Mathlib/RingTheory/IsTensorProduct.lean | 5 + Mathlib/RingTheory/Jacobson.lean | 8 +- Mathlib/RingTheory/Localization/Basic.lean | 46 ++- .../RingTheory/Localization/Finiteness.lean | 1 + .../RingTheory/Localization/FractionRing.lean | 1 + Mathlib/RingTheory/Localization/Ideal.lean | 4 +- .../LocalizationLocalization.lean | 2 + Mathlib/RingTheory/Localization/Module.lean | 14 +- .../RingTheory/Localization/NormTrace.lean | 2 +- .../RingTheory/Localization/Submodule.lean | 5 +- Mathlib/RingTheory/Perfection.lean | 16 +- Mathlib/RingTheory/Polynomial/GaussLemma.lean | 9 +- .../Polynomial/SeparableDegree.lean | 4 +- Mathlib/RingTheory/QuotientNilpotent.lean | 3 +- Mathlib/RingTheory/Regular/IsSMulRegular.lean | 1 + Mathlib/RingTheory/RootsOfUnity/Minpoly.lean | 1 + Mathlib/RingTheory/Smooth/Basic.lean | 8 +- Mathlib/RingTheory/Smooth/Kaehler.lean | 48 ++- .../RingTheory/UniqueFactorizationDomain.lean | 10 +- Mathlib/RingTheory/Unramified/Basic.lean | 7 +- Mathlib/RingTheory/Valuation/Integral.lean | 4 +- Mathlib/RingTheory/WittVector/Truncated.lean | 30 +- .../Algebra/Module/FiniteDimension.lean | 17 +- Mathlib/Topology/Algebra/Monoid.lean | 1 + Mathlib/Topology/Algebra/MulAction.lean | 1 + .../Algebra/Nonarchimedean/Bases.lean | 1 + Mathlib/Topology/Algebra/UniformGroup.lean | 17 +- .../Topology/Category/Profinite/Nobeling.lean | 65 +++- Mathlib/Topology/Category/Stonean/Limits.lean | 1 + Mathlib/Topology/Connected/Clopen.lean | 1 + Mathlib/Topology/Covering.lean | 1 + Mathlib/Topology/LocalAtTarget.lean | 5 +- Mathlib/Topology/MetricSpace/Contracting.lean | 17 +- Mathlib/Topology/MetricSpace/Isometry.lean | 1 + Mathlib/Topology/Order/LawsonTopology.lean | 3 +- Mathlib/Topology/Order/LeftRightLim.lean | 2 + Mathlib/Topology/SeparatedMap.lean | 27 +- Mathlib/Topology/Separation.lean | 14 +- Mathlib/Topology/StoneCech.lean | 1 + Mathlib/Topology/Support.lean | 14 +- .../UniformSpace/UniformEmbedding.lean | 4 + lean-toolchain | 2 +- test/casesm.lean | 1 + test/fun_prop_dev.lean | 56 ++- 319 files changed, 2124 insertions(+), 1144 deletions(-) diff --git a/Archive/Imo/Imo1981Q3.lean b/Archive/Imo/Imo1981Q3.lean index 04bb0b4298d90..dfb8d03caa121 100644 --- a/Archive/Imo/Imo1981Q3.lean +++ b/Archive/Imo/Imo1981Q3.lean @@ -130,6 +130,7 @@ satisfying `NatPredicate m n N` are `fib K` and `fib (K+1)`, respectively. -/ variable {K : ℕ} (HK : N < fib K + fib (K + 1)) {N} +include HK in theorem m_n_bounds {m n : ℕ} (h1 : NatPredicate N m n) : m ≤ fib K ∧ n ≤ fib (K + 1) := by obtain ⟨k : ℕ, hm : m = fib k, hn : n = fib (k + 1)⟩ := h1.imp_fib m by_cases h2 : k < K + 1 @@ -156,6 +157,7 @@ theorem m_n_bounds {m n : ℕ} (h1 : NatPredicate N m n) : m ≤ fib K ∧ n ≤ We spell out the consequences of this result for `specifiedSet N` here. -/ variable {M : ℕ} (HM : M = fib K ^ 2 + fib (K + 1) ^ 2) +include HK HM theorem k_bound {m n : ℤ} (h1 : ProblemPredicate N m n) : m ^ 2 + n ^ 2 ≤ M := by have h2 : 0 ≤ m := h1.m_range.left.le diff --git a/Archive/Imo/Imo1986Q5.lean b/Archive/Imo/Imo1986Q5.lean index 411ec4c9e0395..9c265680eb14f 100644 --- a/Archive/Imo/Imo1986Q5.lean +++ b/Archive/Imo/Imo1986Q5.lean @@ -38,6 +38,7 @@ structure IsGood (f : ℝ≥0 → ℝ≥0) : Prop where namespace IsGood variable {f : ℝ≥0 → ℝ≥0} (hf : IsGood f) {x y : ℝ≥0} +include hf theorem map_add (x y : ℝ≥0) : f (x + y) = f (x * f y) * f y := (hf.map_add_rev x y).symm diff --git a/Archive/Imo/Imo2024Q1.lean b/Archive/Imo/Imo2024Q1.lean index 74975bfd430bd..cf6f08b6fa937 100644 --- a/Archive/Imo/Imo2024Q1.lean +++ b/Archive/Imo/Imo2024Q1.lean @@ -60,6 +60,7 @@ lemma condition_toIcoMod_iff {α : ℝ} : namespace Condition variable {α : ℝ} (hc : Condition α) +include hc lemma mem_Ico_one_of_mem_Ioo (h : α ∈ Set.Ioo 0 2) : α ∈ Set.Ico 1 2 := by rcases h with ⟨h0, h2⟩ diff --git a/Archive/Imo/Imo2024Q2.lean b/Archive/Imo/Imo2024Q2.lean index cf1a6f8be4e86..9bfc8bc8bb485 100644 --- a/Archive/Imo/Imo2024Q2.lean +++ b/Archive/Imo/Imo2024Q2.lean @@ -48,6 +48,7 @@ namespace Condition variable {a b : ℕ} (h : Condition a b) section +include h lemma a_pos : 0 < a := h.1 @@ -155,6 +156,8 @@ lemma ab_add_one_dvd_a_pow_large_n_0_add_b : a * b + 1 ∣ a ^ h.large_n_0 + b : rw [← h.gcd_eq_g h.N_le_large_n_0] exact Nat.gcd_dvd_left _ _ +include h + lemma ab_add_one_dvd_b_add_one : a * b + 1 ∣ b + 1 := by rw [add_comm b] suffices a * b + 1 ∣ a ^ h.large_n_0 + b by diff --git a/Archive/Imo/Imo2024Q6.lean b/Archive/Imo/Imo2024Q6.lean index 12e05f56b7108..fe5ae6338ae6c 100644 --- a/Archive/Imo/Imo2024Q6.lean +++ b/Archive/Imo/Imo2024Q6.lean @@ -39,6 +39,7 @@ symmetric form). -/ def Aquaesulian (f : G → G) : Prop := ∀ x y, f (f y + x) = f x + y ∨ f (f x + y) = f y + x variable {f : G → G} (h : Aquaesulian f) +include h lemma Aquaesulian.apply_apply_add (x : G) : f (f x + x) = f x + x := by rcases h x x with hx | hx <;> exact hx diff --git a/Archive/Wiedijk100Theorems/BuffonsNeedle.lean b/Archive/Wiedijk100Theorems/BuffonsNeedle.lean index 062587554806b..fafebe0649917 100644 --- a/Archive/Wiedijk100Theorems/BuffonsNeedle.lean +++ b/Archive/Wiedijk100Theorems/BuffonsNeedle.lean @@ -126,6 +126,7 @@ noncomputable def N : Ω → ℝ := needleCrossesIndicator l ∘ B -/ abbrev needleSpace : Set (ℝ × ℝ) := Set.Icc (-d / 2) (d / 2) ×ˢ Set.Icc 0 π +include hd in lemma volume_needleSpace : ℙ (needleSpace d) = ENNReal.ofReal (d * π) := by simp_rw [MeasureTheory.Measure.volume_eq_prod, MeasureTheory.Measure.prod_prod, Real.volume_Icc, ENNReal.ofReal_mul hd.le] @@ -157,6 +158,7 @@ lemma stronglyMeasurable_needleCrossesIndicator : · simp_rw [needleCrossesIndicator, Set.indicator_of_not_mem hp] at hxp apply Or.inl hxp.symm +include hd in lemma integrable_needleCrossesIndicator : MeasureTheory.Integrable (needleCrossesIndicator l) (Measure.prod @@ -185,6 +187,7 @@ lemma integrable_needleCrossesIndicator : neg_div, sub_neg_eq_add, add_halves, sub_zero, ← ENNReal.ofReal_mul hd.le, ENNReal.ofReal_lt_top] +include hd hB hBₘ in /-- This is a common step in both the short and the long case to simplify the expectation of the needle crossing a line to a double integral. @@ -230,6 +233,7 @@ lemma buffon_integral : · rw [if_neg h, if_neg (this.not.mp h)] simp_rw [indicator_eq, MeasureTheory.setIntegral_indicator measurableSet_Icc, Pi.one_apply] +include hl in /-- From `buffon_integral`, in both the short and the long case, we have ```lean @@ -247,6 +251,7 @@ lemma short_needle_inter_eq (h : l ≤ d) (θ : ℝ) : min_div_div_right zero_le_two, neg_mul, max_neg_neg, mul_comm, min_eq_right (mul_le_of_le_of_le_one_of_nonneg h θ.sin_le_one hl.le)] +include hd hBₘ hB hl in /-- Buffon's Needle, the short case (`l ≤ d`). The probability of the needle crossing a line equals `(2 * l) / (d * π)`. @@ -289,6 +294,7 @@ lemma integral_min_eq_two_mul : (by ring : -(π / 2) + π = π / 2), two_mul] all_goals exact intervalIntegrable_min_const_sin_mul d l _ _ +include hd hl in /-- The first of two adjacent integrals in the long case. In the range `(0)..(d / l).arcsin`, we have that `θ.sin * l ≤ d`, and thus the integral is `∫ θ in (0)..(d / l).arcsin, θ.sin * l`. @@ -306,6 +312,7 @@ lemma integral_zero_to_arcsin_min : rw [intervalIntegral.integral_congr this, intervalIntegral.integral_mul_const, integral_sin, Real.cos_zero, Real.cos_arcsin] +include hl in /-- The second of two adjacent integrals in the long case. In the range `(d / l).arcsin..(π / 2)`, we have that `d ≤ θ.sin * l`, and thus the integral is `∫ θ in (d / l).arcsin..(π / 2), d`. @@ -324,6 +331,7 @@ lemma integral_arcsin_to_pi_div_two_min (h : d ≤ l) : simp_rw [min_eq_left ((div_le_iff hl).mp ((Real.arcsin_le_iff_le_sin' hθ_mem).mp hθ₁))] rw [intervalIntegral.integral_congr this, intervalIntegral.integral_const, smul_eq_mul] +include hd hBₘ hB hl in /-- Buffon's Needle, the long case (`d ≤ l`). -/ diff --git a/Archive/Wiedijk100Theorems/CubingACube.lean b/Archive/Wiedijk100Theorems/CubingACube.lean index 9d7164a59e35c..f2710670e0f47 100644 --- a/Archive/Wiedijk100Theorems/CubingACube.lean +++ b/Archive/Wiedijk100Theorems/CubingACube.lean @@ -132,6 +132,7 @@ structure Correct (cs : ι → Cube n) : Prop where namespace Correct variable (h : Correct cs) +include h theorem toSet_subset_unitCube {i} : (cs i).toSet ⊆ unitCube.toSet := by convert h.iUnion_eq ▸ subset_iUnion _ i @@ -261,6 +262,7 @@ theorem t_le_t (hi : i ∈ bcubes cs c) (j : Fin n) : · simp [hw] section +include h v /-- Every cube in the valley must be smaller than it -/ theorem w_lt_w (hi : i ∈ bcubes cs c) : (cs i).w < c.w := by @@ -514,6 +516,7 @@ end variable [Finite ι] [Nontrivial ι] +include h in theorem strictAnti_sequenceOfCubes : StrictAnti <| decreasingSequence h := strictAnti_nat_of_succ_lt fun k => by let v := (sequenceOfCubes h k).2; dsimp only [decreasingSequence, sequenceOfCubes] diff --git a/Archive/Wiedijk100Theorems/FriendshipGraphs.lean b/Archive/Wiedijk100Theorems/FriendshipGraphs.lean index af46f3d333e29..a25f296a01c50 100644 --- a/Archive/Wiedijk100Theorems/FriendshipGraphs.lean +++ b/Archive/Wiedijk100Theorems/FriendshipGraphs.lean @@ -75,6 +75,7 @@ namespace Friendship variable (R) +include hG in /-- One characterization of a friendship graph is that there is exactly one walk of length 2 between distinct vertices. These walks are counted in off-diagonal entries of the square of the adjacency matrix, so for a friendship graph, those entries are all 1. -/ @@ -84,6 +85,7 @@ theorem adjMatrix_sq_of_ne {v w : V} (hvw : v ≠ w) : simp [commonNeighbors, neighborFinset_eq_filter, Finset.filter_filter, and_comm, ← neighborFinset_def] +include hG in /-- This calculation amounts to counting the number of length 3 walks between nonadjacent vertices. We use it to show that nonadjacent vertices have equal degrees. -/ theorem adjMatrix_pow_three_of_not_adj {v w : V} (non_adj : ¬G.Adj v w) : @@ -98,6 +100,7 @@ theorem adjMatrix_pow_three_of_not_adj {v w : V} (non_adj : ¬G.Adj v w) : variable {R} +include hG in /-- As `v` and `w` not being adjacent implies `degree G v = ((G.adjMatrix R) ^ 3) v w` and `degree G w = ((G.adjMatrix R) ^ 3) v w`, the degrees are equal if `((G.adjMatrix R) ^ 3) v w = ((G.adjMatrix R) ^ 3) w v` @@ -111,6 +114,7 @@ theorem degree_eq_of_not_adj {v w : V} (hvw : ¬G.Adj v w) : degree G v = degree simp only [pow_succ _ 2, sq, ← transpose_mul, transpose_apply] simp only [mul_assoc] +include hG in /-- Let `A` be the adjacency matrix of a graph `G`. If `G` is a friendship graph, then all of the off-diagonal entries of `A^2` are 1. If `G` is `d`-regular, then all of the diagonal entries of `A^2` are `d`. @@ -121,6 +125,7 @@ theorem adjMatrix_sq_of_regular (hd : G.IsRegularOfDegree d) : · rw [h, sq, adjMatrix_mul_self_apply_self, hd]; simp · rw [adjMatrix_sq_of_ne R hG h, of_apply, if_neg h] +include hG in theorem adjMatrix_sq_mod_p_of_regular {p : ℕ} (dmod : (d : ZMod p) = 1) (hd : G.IsRegularOfDegree d) : G.adjMatrix (ZMod p) ^ 2 = of fun _ _ => 1 := by simp [adjMatrix_sq_of_regular hG hd, dmod] @@ -129,6 +134,7 @@ section Nonempty variable [Nonempty V] +include hG in /-- If `G` is a friendship graph without a politician (a vertex adjacent to all others), then it is regular. We have shown that nonadjacent vertices of a friendship graph have the same degree, and if there isn't a politician, we can show this for adjacent vertices by finding a vertex @@ -161,6 +167,7 @@ theorem isRegularOf_not_existsPolitician (hG' : ¬ExistsPolitician G) : rw [key ((mem_commonNeighbors G).mpr ⟨hvx, G.symm hxw⟩), key ((mem_commonNeighbors G).mpr ⟨hvy, G.symm hcontra⟩)] +include hG in /-- Let `A` be the adjacency matrix of a `d`-regular friendship graph, and let `v` be a vector all of whose components are `1`. Then `v` is an eigenvector of `A ^ 2`, and we can compute the eigenvalue to be `d * d`, or as `d + (Fintype.card V - 1)`, so those quantities must be equal. @@ -179,6 +186,7 @@ theorem card_of_regular (hd : G.IsRegularOfDegree d) : d + (Fintype.card V - 1) card_neighborSet_eq_degree, hd v, Function.const_def, adjMatrix_mulVec_apply _ _ (mulVec _ _), mul_one, sum_const, Set.toFinset_card, Algebra.id.smul_eq_mul, Nat.cast_id] +include hG in /-- The size of a `d`-regular friendship graph is `1 mod (d-1)`, and thus `1 mod p` for a factor `p ∣ d-1`. -/ theorem card_mod_p_of_regular {p : ℕ} (dmod : (d : ZMod p) = 1) (hd : G.IsRegularOfDegree d) : @@ -202,6 +210,7 @@ theorem adjMatrix_mul_const_one_mod_p_of_regular {p : ℕ} (dmod : (d : ZMod p) G.adjMatrix (ZMod p) * of (fun _ _ => 1) = of (fun _ _ => 1) := by rw [adjMatrix_sq_mul_const_one_of_regular hd, dmod] +include hG in /-- Modulo a factor of `d-1`, the square and all higher powers of the adjacency matrix of a `d`-regular friendship graph reduce to the matrix whose entries are all 1. -/ theorem adjMatrix_pow_mod_p_of_regular {p : ℕ} (dmod : (d : ZMod p) = 1) @@ -217,6 +226,7 @@ theorem adjMatrix_pow_mod_p_of_regular {p : ℕ} (dmod : (d : ZMod p) = 1) variable [Nonempty V] +include hG in /-- This is the main proof. Assuming that `3 ≤ d`, we take `p` to be a prime factor of `d-1`. Then the `p`th power of the adjacency matrix of a `d`-regular friendship graph must have trace 1 mod `p`, but we can also show that the trace must be the `p`th power of the trace of the original @@ -249,6 +259,7 @@ theorem false_of_three_le_degree (hd : G.IsRegularOfDegree d) (h : 3 ≤ d) : Fa Nat.dvd_one, Nat.minFac_eq_one_iff] linarith +include hG in /-- If `d ≤ 1`, a `d`-regular friendship graph has at most one vertex, which is trivially a politician. -/ theorem existsPolitician_of_degree_le_one (hd : G.IsRegularOfDegree d) (hd1 : d ≤ 1) : @@ -270,6 +281,7 @@ theorem existsPolitician_of_degree_le_one (hd : G.IsRegularOfDegree d) (hd1 : d apply h apply Fintype.card_le_one_iff.mp this +include hG in /-- If `d = 2`, a `d`-regular friendship graph has 3 vertices, so it must be complete graph, and all the vertices are politicians. -/ theorem neighborFinset_eq_of_degree_eq_two (hd : G.IsRegularOfDegree 2) (v : V) : @@ -287,6 +299,7 @@ theorem neighborFinset_eq_of_degree_eq_two (hd : G.IsRegularOfDegree 2) (v : V) · dsimp only [IsRegularOfDegree, degree] at hd rw [hd] +include hG in theorem existsPolitician_of_degree_eq_two (hd : G.IsRegularOfDegree 2) : ExistsPolitician G := by have v := Classical.arbitrary V use v @@ -294,6 +307,7 @@ theorem existsPolitician_of_degree_eq_two (hd : G.IsRegularOfDegree 2) : ExistsP rw [← mem_neighborFinset, neighborFinset_eq_of_degree_eq_two hG hd v, Finset.mem_erase] exact ⟨hvw.symm, Finset.mem_univ _⟩ +include hG in theorem existsPolitician_of_degree_le_two (hd : G.IsRegularOfDegree d) (h : d ≤ 2) : ExistsPolitician G := by interval_cases d @@ -302,6 +316,7 @@ theorem existsPolitician_of_degree_le_two (hd : G.IsRegularOfDegree d) (h : d end Friendship +include hG in /-- **Friendship theorem**: We wish to show that a friendship graph has a politician (a vertex adjacent to all others). We proceed by contradiction, and assume the graph has no politician. We have already proven that a friendship graph with no politician is `d`-regular for some `d`, diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index a37f091f6c2ac..8ea3c00ea7004 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -45,6 +45,7 @@ def Prime (p : α) : Prop := namespace Prime variable {p : α} (hp : Prime p) +include hp theorem ne_zero : p ≠ 0 := hp.1 @@ -57,19 +58,19 @@ theorem not_dvd_one : ¬p ∣ 1 := theorem ne_one : p ≠ 1 := fun h => hp.2.1 (h.symm ▸ isUnit_one) -theorem dvd_or_dvd (hp : Prime p) {a b : α} (h : p ∣ a * b) : p ∣ a ∨ p ∣ b := +theorem dvd_or_dvd {a b : α} (h : p ∣ a * b) : p ∣ a ∨ p ∣ b := hp.2.2 a b h theorem dvd_mul {a b : α} : p ∣ a * b ↔ p ∣ a ∨ p ∣ b := ⟨hp.dvd_or_dvd, (Or.elim · (dvd_mul_of_dvd_left · _) (dvd_mul_of_dvd_right · _))⟩ -theorem isPrimal (hp : Prime p) : IsPrimal p := fun _a _b dvd ↦ (hp.dvd_or_dvd dvd).elim +theorem isPrimal : IsPrimal p := fun _a _b dvd ↦ (hp.dvd_or_dvd dvd).elim (fun h ↦ ⟨p, 1, h, one_dvd _, (mul_one p).symm⟩) fun h ↦ ⟨1, p, one_dvd _, h, (one_mul p).symm⟩ theorem not_dvd_mul {a b : α} (ha : ¬ p ∣ a) (hb : ¬ p ∣ b) : ¬ p ∣ a * b := hp.dvd_mul.not.mpr <| not_or.mpr ⟨ha, hb⟩ -theorem dvd_of_dvd_pow (hp : Prime p) {a : α} {n : ℕ} (h : p ∣ a ^ n) : p ∣ a := by +theorem dvd_of_dvd_pow {a : α} {n : ℕ} (h : p ∣ a ^ n) : p ∣ a := by induction n with | zero => rw [pow_zero] at h diff --git a/Mathlib/Algebra/Category/ModuleCat/Free.lean b/Mathlib/Algebra/Category/ModuleCat/Free.lean index 1b3e7db37673f..7c67bb3e86b4f 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Free.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Free.lean @@ -41,6 +41,9 @@ variable (hv : LinearIndependent R v) {u : ι ⊕ ι' → S.X₂} (hw : LinearIndependent R (S.g ∘ u ∘ Sum.inr)) (hm : Mono S.f) (huv : u ∘ Sum.inl = S.f ∘ v) +section +include hS hw huv + theorem disjoint_span_sum : Disjoint (span R (range (u ∘ Sum.inl))) (span R (range (u ∘ Sum.inr))) := by rw [huv, disjoint_comm] @@ -48,6 +51,8 @@ theorem disjoint_span_sum : Disjoint (span R (range (u ∘ Sum.inl))) rw [← LinearMap.range_coe, span_eq (LinearMap.range S.f), hS.moduleCat_range_eq_ker] exact range_ker_disjoint hw +include hv hm in + /-- In the commutative diagram ``` f g @@ -67,6 +72,9 @@ theorem linearIndependent_leftExact : LinearIndependent R u := by infer_instance exact hv +end + +include hS' hv in /-- Given a short exact sequence `0 ⟶ X₁ ⟶ X₂ ⟶ X₃ ⟶ 0` of `R`-modules and linearly independent families `v : ι → N` and `w : ι' → P`, we get a linearly independent family `ι ⊕ ι' → M` -/ theorem linearIndependent_shortExact {w : ι' → S.X₃} (hw : LinearIndependent R w) : @@ -81,6 +89,7 @@ end LinearIndependent section Span +include hS in /-- In the commutative diagram ``` f g @@ -124,6 +133,7 @@ theorem span_exact {β : Type*} {u : ι ⊕ β → S.X₂} (huv : u ∘ Sum.inl use cm.mapDomain (Sum.inr) rw [Finsupp.sum_mapDomain_index_inj Sum.inr_injective] +include hS in /-- Given an exact sequence `X₁ ⟶ X₂ ⟶ X₃ ⟶ 0` of `R`-modules and spanning families `v : ι → X₁` and `w : ι' → X₃`, we get a spanning family `ι ⊕ ι' → X₂` -/ theorem span_rightExact {w : ι' → S.X₃} (hv : ⊤ ≤ span R (range v)) @@ -147,6 +157,8 @@ def Basis.ofShortExact Basis.mk (linearIndependent_shortExact hS' bN.linearIndependent bP.linearIndependent) (span_rightExact hS'.exact (le_of_eq (bN.span_eq.symm)) (le_of_eq (bP.span_eq.symm)) hS'.epi_g) +include hS' + /-- In a short exact sequence `0 ⟶ X₁ ⟶ X₂ ⟶ X₃ ⟶ 0`, if `X₁` and `X₃` are free, then `X₂` is free. -/ theorem free_shortExact [Module.Free R S.X₁] [Module.Free R S.X₃] : diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean index 07f18f8c8fc8e..69db169f3d58c 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean @@ -53,8 +53,7 @@ variable {R₀ R : Cᵒᵖ ⥤ RingCat.{u}} (α : R₀ ⟶ R) [Presheaf.IsLocall [Presheaf.IsLocallyInjective J φ] (hA : Presheaf.IsSeparated J A) {X : C} (r : R.obj (Opposite.op X)) (m : A.obj (Opposite.op X)) {P : Presieve X} (r₀ : FamilyOfElements (R₀ ⋙ forget _) P) (m₀ : FamilyOfElements (M₀.presheaf ⋙ forget _) P) - (hr₀ : (r₀.map (whiskerRight α (forget _))).IsAmalgamation r) - (hm₀ : (m₀.map (whiskerRight φ (forget _))).IsAmalgamation m) +include hA lemma _root_.PresheafOfModules.Sheafify.app_eq_of_isLocallyInjective {Y : C} (r₀ r₀' : R₀.obj (Opposite.op Y)) @@ -84,6 +83,10 @@ lemma isCompatible_map_smul_aux {Y Z : C} (f : Y ⟶ X) (g : Z ⟶ Y) erw [NatTrans.naturality_apply] rfl +variable (hr₀ : (r₀.map (whiskerRight α (forget _))).IsAmalgamation r) + (hm₀ : (m₀.map (whiskerRight φ (forget _))).IsAmalgamation m) + +include hr₀ hm₀ in lemma isCompatible_map_smul : ((r₀.smul m₀).map (whiskerRight φ (forget _))).Compatible := by intro Y₁ Y₂ Z g₁ g₂ f₁ f₂ h₁ h₂ fac let a₁ := r₀ f₁ h₁ diff --git a/Mathlib/Algebra/Divisibility/Units.lean b/Mathlib/Algebra/Divisibility/Units.lean index 9c26ce76082e1..aee46556f879c 100644 --- a/Mathlib/Algebra/Divisibility/Units.lean +++ b/Mathlib/Algebra/Divisibility/Units.lean @@ -180,6 +180,8 @@ section IsUnit variable (hu : IsUnit x) +include hu + theorem isRelPrime_mul_unit_left_left : IsRelPrime (x * y) z ↔ IsRelPrime y z := ⟨IsRelPrime.of_mul_left_right, fun H _ h ↦ H (hu.dvd_mul_left.mp h)⟩ diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 50979132619e5..213f451a27df5 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -305,6 +305,7 @@ def gciMapComap (hf : Function.Injective f) : GaloisCoinsertion (map f) (comap f (gc_map_comap f).toGaloisCoinsertion fun S x => by simp [mem_comap, mem_map, hf.eq_iff] variable (hf : Function.Injective f) +include hf @[to_additive] theorem comap_map_eq_of_injective (S : Submonoid M) : (S.map f).comap f = S := @@ -346,15 +347,18 @@ end GaloisCoinsertion section GaloisInsertion -variable {ι : Type*} {f : F} (hf : Function.Surjective f) +variable {ι : Type*} {f : F} /-- `map f` and `comap f` form a `GaloisInsertion` when `f` is surjective. -/ @[to_additive " `map f` and `comap f` form a `GaloisInsertion` when `f` is surjective. "] -def giMapComap : GaloisInsertion (map f) (comap f) := +def giMapComap (hf : Function.Surjective f) : GaloisInsertion (map f) (comap f) := (gc_map_comap f).toGaloisInsertion fun S x h => let ⟨y, hy⟩ := hf x mem_map.2 ⟨y, by simp [hy, h]⟩ +variable (hf : Function.Surjective f) +include hf + @[to_additive] theorem map_comap_eq_of_surjective (S : Submonoid N) : (S.comap f).map f = S := (giMapComap hf).l_u_eq _ diff --git a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean index 5299435a31d16..0c804659e191b 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean @@ -301,14 +301,14 @@ theorem map_id (S : Subsemigroup M) : S.map (MulHom.id M) = S := section GaloisCoinsertion -variable {ι : Type*} {f : M →ₙ* N} +variable {ι : Type*} {f : M →ₙ* N} (hf : Function.Injective f) +include hf /-- `map f` and `comap f` form a `GaloisCoinsertion` when `f` is injective. -/ @[to_additive " `map f` and `comap f` form a `GaloisCoinsertion` when `f` is injective. "] -def gciMapComap (hf : Function.Injective f) : GaloisCoinsertion (map f) (comap f) := +def gciMapComap : GaloisCoinsertion (map f) (comap f) := (gc_map_comap f).toGaloisCoinsertion fun S x => by simp [mem_comap, mem_map, hf.eq_iff] -variable (hf : Function.Injective f) @[to_additive] theorem comap_map_eq_of_injective (S : Subsemigroup M) : (S.map f).comap f = S := @@ -353,6 +353,7 @@ end GaloisCoinsertion section GaloisInsertion variable {ι : Type*} {f : M →ₙ* N} (hf : Function.Surjective f) +include hf /-- `map f` and `comap f` form a `GaloisInsertion` when `f` is surjective. -/ @[to_additive " `map f` and `comap f` form a `GaloisInsertion` when `f` is surjective. "] diff --git a/Mathlib/Algebra/GroupWithZero/InjSurj.lean b/Mathlib/Algebra/GroupWithZero/InjSurj.lean index e8de1ecfd7972..d7a7ff367118a 100644 --- a/Mathlib/Algebra/GroupWithZero/InjSurj.lean +++ b/Mathlib/Algebra/GroupWithZero/InjSurj.lean @@ -48,6 +48,8 @@ section NoZeroDivisors variable [Mul M₀] [Zero M₀] [Mul M₀'] [Zero M₀'] (f : M₀ → M₀') (hf : Injective f) (zero : f 0 = 0) (mul : ∀ x y, f (x * y) = f x * f y) +include hf zero mul + /-- Pull back a `NoZeroDivisors` instance along an injective function. -/ protected theorem Function.Injective.noZeroDivisors [NoZeroDivisors M₀'] : NoZeroDivisors M₀ := { eq_zero_or_eq_zero_of_mul_eq_zero := fun {a b} H ↦ diff --git a/Mathlib/Algebra/Homology/BifunctorFlip.lean b/Mathlib/Algebra/Homology/BifunctorFlip.lean index c6149ca1d5248..1f1ea01d48205 100644 --- a/Mathlib/Algebra/Homology/BifunctorFlip.lean +++ b/Mathlib/Algebra/Homology/BifunctorFlip.lean @@ -28,13 +28,13 @@ variable {I₁ I₂ J : Type*} {c₁ : ComplexShape I₁} {c₂ : ComplexShape I (K₁ : HomologicalComplex C₁ c₁) (K₂ : HomologicalComplex C₂ c₂) (F : C₁ ⥤ C₂ ⥤ D) [F.PreservesZeroMorphisms] [∀ X₁, (F.obj X₁).PreservesZeroMorphisms] (c : ComplexShape J) [TotalComplexShape c₁ c₂ c] [TotalComplexShape c₂ c₁ c] - [TotalComplexShapeSymmetry c₁ c₂ c] [DecidableEq J] + [TotalComplexShapeSymmetry c₁ c₂ c] lemma hasMapBifunctor_flip_iff : HasMapBifunctor K₂ K₁ F.flip c ↔ HasMapBifunctor K₁ K₂ F c := (((F.mapBifunctorHomologicalComplex c₁ c₂).obj K₁).obj K₂).flip_hasTotal_iff c -variable [HasMapBifunctor K₁ K₂ F c] +variable [DecidableEq J] [HasMapBifunctor K₁ K₂ F c] instance : HasMapBifunctor K₂ K₁ F.flip c := by rw [hasMapBifunctor_flip_iff] diff --git a/Mathlib/Algebra/Homology/DerivedCategory/HomologySequence.lean b/Mathlib/Algebra/Homology/DerivedCategory/HomologySequence.lean index 01875bb37339d..454e664b1fd81 100644 --- a/Mathlib/Algebra/Homology/DerivedCategory/HomologySequence.lean +++ b/Mathlib/Algebra/Homology/DerivedCategory/HomologySequence.lean @@ -56,12 +56,15 @@ namespace HomologySequence variable (T : Triangle (DerivedCategory C)) (hT : T ∈ distTriang _) (n₀ n₁ : ℤ) (h : n₀ + 1 = n₁) +include hT /-- The connecting homomorphism on the homology sequence attached to a distinguished triangle in the derived category. -/ noncomputable def δ : (homologyFunctor C n₀).obj T.obj₃ ⟶ (homologyFunctor C n₁).obj T.obj₁ := (homologyFunctor C 0).shiftMap T.mor₃ n₀ n₁ (by rw [add_comm 1, h]) +include hT + @[reassoc (attr := simp)] lemma comp_δ : (homologyFunctor C n₀).map T.mor₂ ≫ δ T n₀ n₁ h = 0 := (homologyFunctor C 0).comp_homologySequenceδ _ hT _ _ h diff --git a/Mathlib/Algebra/Homology/Embedding/IsSupported.lean b/Mathlib/Algebra/Homology/Embedding/IsSupported.lean index 02f0b1a5dbd9a..1504e486bb998 100644 --- a/Mathlib/Algebra/Homology/Embedding/IsSupported.lean +++ b/Mathlib/Algebra/Homology/Embedding/IsSupported.lean @@ -45,6 +45,7 @@ lemma isZero_X_of_isStrictlySupported [K.IsStrictlySupported e] IsZero (K.X i') := IsStrictlySupported.isZero i' hi' +include e' in variable {K L} in lemma isStrictlySupported_of_iso [K.IsStrictlySupported e] : L.IsStrictlySupported e where isZero i' hi' := (K.isZero_X_of_isStrictlySupported e i' hi').of_iso @@ -60,9 +61,11 @@ lemma exactAt_of_isSupported [K.IsSupported e] (i' : ι') (hi' : ∀ i, e.f i K.ExactAt i' := IsSupported.exactAt i' hi' +include e' in variable {K L} in lemma isSupported_of_iso [K.IsSupported e] : L.IsSupported e where - exactAt i' hi' := (K.exactAt_of_isSupported e i' hi').of_iso e' + exactAt i' hi' := + (K.exactAt_of_isSupported e i' hi').of_iso e' instance [K.IsStrictlySupported e] : K.IsSupported e where exactAt i' hi' := by diff --git a/Mathlib/Algebra/Homology/HomologySequence.lean b/Mathlib/Algebra/Homology/HomologySequence.lean index 56dc6c76a78e9..7a54cb2191c0a 100644 --- a/Mathlib/Algebra/Homology/HomologySequence.lean +++ b/Mathlib/Algebra/Homology/HomologySequence.lean @@ -121,6 +121,7 @@ instance [K.HasHomology i] [K.HasHomology j] : dsimp infer_instance +include hij in /-- The diagram `K.homology i ⟶ K.opcycles i ⟶ K.cycles j ⟶ K.homology j` is exact when `c.Rel i j`. -/ lemma composableArrows₃_exact [CategoryWithHomology C] : @@ -296,6 +297,7 @@ lemma comp_δ : HomologicalComplex.homologyMap S.g i ≫ hS.δ i j hij = 0 := lemma homology_exact₁ : (ShortComplex.mk _ _ (δ_comp hS i j hij)).Exact := (snakeInput hS i j hij).L₂'_exact +include hS in /-- Exactness of `S.X₁.homology i ⟶ S.X₂.homology i ⟶ S.X₃.homology i`. -/ lemma homology_exact₂ : (ShortComplex.mk (HomologicalComplex.homologyMap S.f i) (HomologicalComplex.homologyMap S.g i) (by rw [← HomologicalComplex.homologyMap_comp, diff --git a/Mathlib/Algebra/Homology/HomologySequenceLemmas.lean b/Mathlib/Algebra/Homology/HomologySequenceLemmas.lean index 3964e891d0235..acf08cce2512f 100644 --- a/Mathlib/Algebra/Homology/HomologySequenceLemmas.lean +++ b/Mathlib/Algebra/Homology/HomologySequenceLemmas.lean @@ -60,7 +60,7 @@ variable (S) noncomputable def composableArrows₂ (i : ι) : ComposableArrows C 2 := mk₂ (homologyMap S.f i) (homologyMap S.g i) -lemma composableArrows₂_exact (i : ι) : +lemma composableArrows₂_exact (hS₁ : S₁.ShortExact) (i : ι) : (composableArrows₂ S₁ i).Exact := (hS₁.homology_exact₂ i).exact_toComposableArrows @@ -105,6 +105,8 @@ noncomputable def mapComposableArrows₅ (i j : ι) (hij : c.Rel i j) : attribute [local instance] epi_comp +include hS₁ hS₂ + lemma mono_homologyMap_τ₃ (i : ι) (h₁ : Epi (homologyMap φ.τ₁ i)) (h₂ : Mono (homologyMap φ.τ₂ i)) @@ -146,7 +148,6 @@ lemma epi_homologyMap_τ₃ (i : ι) have := epi_homologyMap_of_epi_of_not_rel S₂.g i (by simpa using hi) exact epi_of_epi_fac eq.symm - lemma isIso_homologyMap_τ₃ (i : ι) (h₁ : Epi (homologyMap φ.τ₁ i)) (h₂ : IsIso (homologyMap φ.τ₂ i)) diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/ShortExact.lean b/Mathlib/Algebra/Homology/HomotopyCategory/ShortExact.lean index 2dc4694c7f66e..e2dc209b90363 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/ShortExact.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/ShortExact.lean @@ -110,6 +110,7 @@ open ComposableArrows set_option simprocs false +include hS in lemma quasiIso_descShortComplex : QuasiIso (descShortComplex S) where quasiIsoAt n := by rw [quasiIsoAt_iff_isIso_homologyMap] diff --git a/Mathlib/Algebra/Homology/HomotopyCofiber.lean b/Mathlib/Algebra/Homology/HomotopyCofiber.lean index 6b4f8c2872560..88c9ce69aaee4 100644 --- a/Mathlib/Algebra/Homology/HomotopyCofiber.lean +++ b/Mathlib/Algebra/Homology/HomotopyCofiber.lean @@ -474,11 +474,12 @@ lemma inlX_nullHomotopy_f (i j : ι) (hij : c.Rel j i) : ← HomologicalComplex.comp_f_assoc, biprod.lift_snd, neg_f_apply, id_f, neg_comp, id_comp, inlX_π_assoc, zero_sub] -lemma biprod_lift_id_sub_id : biprod.lift (𝟙 K) (-𝟙 K) = biprod.inl - biprod.inr := - biprod.hom_ext _ _ (by simp) (by simp) +include hc lemma inrX_nullHomotopy_f (j : ι) : inrX K j ≫ (nullHomotopicMap K).f j = inrX K j ≫ (π K ≫ ι₀ K - 𝟙 _).f j := by + have : biprod.lift (𝟙 K) (-𝟙 K) = biprod.inl - biprod.inr := + biprod.hom_ext _ _ (by simp) (by simp) obtain ⟨i, hij⟩ := hc j dsimp [nullHomotopicMap] by_cases hj : ∃ (k : ι), c.Rel j k @@ -493,7 +494,7 @@ lemma inrX_nullHomotopy_f (j : ι) : · simp [ι₀] · dsimp simp only [inr_biprodXIso_inv_assoc, biprod_inr_snd_f_assoc, comp_sub, - biprod_inr_desc_f_assoc, id_f, id_comp, ι₀, comp_f, biprod_lift_id_sub_id, + biprod_inr_desc_f_assoc, id_f, id_comp, ι₀, comp_f, this, sub_f_apply, sub_comp, homotopyCofiber_X, homotopyCofiber.inr_f] · simp only [not_exists] at hj simp only [assoc, Homotopy.nullHomotopicMap'_f_of_not_rel_left hij hj, homotopyCofiber_X, @@ -502,7 +503,7 @@ lemma inrX_nullHomotopy_f (j : ι) : rw [← cancel_epi (biprodXIso K K j).inv] ext · simp - · simp [biprod_lift_id_sub_id] + · simp [this] lemma nullHomotopicMap_eq : nullHomotopicMap K = π K ≫ ι₀ K - 𝟙 _ := by ext i @@ -531,6 +532,7 @@ noncomputable def homotopy₀₁ : Homotopy (ι₀ K) (ι₁ K) := (Homotopy.ofEq (by simp)).trans (((πCompι₀Homotopy K hc).compLeft (ι₁ K)).trans (Homotopy.ofEq (by simp))) +include hc in lemma map_ι₀_eq_map_ι₁ {D : Type*} [Category D] (H : HomologicalComplex C c ⥤ D) (hH : (homotopyEquivalences C c).IsInvertedBy H) : H.map (ι₀ K) = H.map (ι₁ K) := by @@ -543,7 +545,7 @@ end cylinder /-- If a functor inverts homotopy equivalences, it sends homotopic maps to the same map. -/ lemma _root_.Homotopy.map_eq_of_inverts_homotopyEquivalences - {φ₀ φ₁ : F ⟶ G} (h : Homotopy φ₀ φ₁) (hc : ∀ j, ∃ i, c.Rel i j) + {φ₀ φ₁ : F ⟶ G} (h : Homotopy φ₀ φ₁)(hc : ∀ j, ∃ i, c.Rel i j) [∀ i, HasBinaryBiproduct (F.X i) (F.X i)] [HasHomotopyCofiber (biprod.lift (𝟙 F) (-𝟙 F))] {D : Type*} [Category D] (H : HomologicalComplex C c ⥤ D) diff --git a/Mathlib/Algebra/Homology/Localization.lean b/Mathlib/Algebra/Homology/Localization.lean index fe1e8360e7b00..197ad242ab355 100644 --- a/Mathlib/Algebra/Homology/Localization.lean +++ b/Mathlib/Algebra/Homology/Localization.lean @@ -30,26 +30,25 @@ open CategoryTheory Limits section variable (C : Type*) [Category C] {ι : Type*} (c : ComplexShape ι) [HasZeroMorphisms C] - [CategoryWithHomology C] [(HomologicalComplex.quasiIso C c).HasLocalization] + [CategoryWithHomology C] + +lemma HomologicalComplex.homologyFunctor_inverts_quasiIso (i : ι) : + (quasiIso C c).IsInvertedBy (homologyFunctor C c i) := fun _ _ _ hf => by + rw [mem_quasiIso_iff] at hf + dsimp + infer_instance + +variable [(HomologicalComplex.quasiIso C c).HasLocalization] /-- The category of homological complexes up to quasi-isomorphisms. -/ abbrev HomologicalComplexUpToQuasiIso := (HomologicalComplex.quasiIso C c).Localization' -variable {C c} - +variable {C c} in /-- The localization functor `HomologicalComplex C c ⥤ HomologicalComplexUpToQuasiIso C c`. -/ abbrev HomologicalComplexUpToQuasiIso.Q : HomologicalComplex C c ⥤ HomologicalComplexUpToQuasiIso C c := (HomologicalComplex.quasiIso C c).Q' -variable (C c) - -lemma HomologicalComplex.homologyFunctor_inverts_quasiIso (i : ι) : - (quasiIso C c).IsInvertedBy (homologyFunctor C c i) := fun _ _ _ hf => by - rw [mem_quasiIso_iff] at hf - dsimp - infer_instance - namespace HomologicalComplexUpToQuasiIso /-- The homology functor `HomologicalComplexUpToQuasiIso C c ⥤ C` for each `i : ι`. -/ @@ -84,9 +83,10 @@ end section variable (C : Type*) [Category C] {ι : Type*} (c : ComplexShape ι) [Preadditive C] - [CategoryWithHomology C] [(HomologicalComplex.quasiIso C c).HasLocalization] + [CategoryWithHomology C] -lemma HomologicalComplexUpToQuasiIso.Q_inverts_homotopyEquivalences : +lemma HomologicalComplexUpToQuasiIso.Q_inverts_homotopyEquivalences + [(HomologicalComplex.quasiIso C c).HasLocalization] : (HomologicalComplex.homotopyEquivalences C c).IsInvertedBy HomologicalComplexUpToQuasiIso.Q := MorphismProperty.IsInvertedBy.of_le _ _ _ @@ -148,7 +148,7 @@ class ComplexShape.QFactorsThroughHomotopy {ι : Type*} (c : ComplexShape ι) namespace HomologicalComplexUpToQuasiIso variable {C c} -variable [c.QFactorsThroughHomotopy C] +variable [(HomologicalComplex.quasiIso C c).HasLocalization] [c.QFactorsThroughHomotopy C] lemma Q_map_eq_of_homotopy {K L : HomologicalComplex C c} {f g : K ⟶ L} (h : Homotopy f g) : Q.map f = Q.map g := @@ -211,6 +211,7 @@ section Cylinder variable {ι : Type*} (c : ComplexShape ι) (hc : ∀ j, ∃ i, c.Rel i j) (C : Type*) [Category C] [Preadditive C] [HasBinaryBiproducts C] +include hc /-- The homotopy category satisfies the universal property of the localized category with respect to homotopy equivalences. -/ @@ -301,7 +302,6 @@ section variable [Preadditive C] [Preadditive D] [CategoryWithHomology C] [CategoryWithHomology D] - [(HomologicalComplex.quasiIso C c).HasLocalization] [(HomologicalComplex.quasiIso D c).HasLocalization] [F.Additive] [F.PreservesHomology] @@ -318,6 +318,8 @@ lemma mapHomologicalComplex_upToQuasiIso_Q_inverts_quasiIso : (F.mapHomologicalComplex c ⋙ HomologicalComplexUpToQuasiIso.Q) := by apply (F.mapHomologicalComplexUpToQuasiIsoLocalizerMorphism c).inverts +variable [(HomologicalComplex.quasiIso C c).HasLocalization] + /-- The functor `HomologicalComplexUpToQuasiIso C c ⥤ HomologicalComplexUpToQuasiIso D c` induced by a functor `F : C ⥤ D` which preserves homology. -/ noncomputable def mapHomologicalComplexUpToQuasiIso : diff --git a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean index 30b30ca0b6c18..c14da91511cae 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/HomologicalComplex.lean @@ -54,8 +54,6 @@ noncomputable def natIsoSc' (i j k : ι) (hi : c.prev j = i) (hk : c.next j = k) variable {C c} -section - variable (K L M : HomologicalComplex C c) (φ : K ⟶ L) (ψ : L ⟶ M) (i j k : ι) /-- The short complex `K.X i ⟶ K.X j ⟶ K.X k` for arbitrary indices `i`, `j` and `k`. -/ @@ -72,7 +70,7 @@ noncomputable abbrev isoSc' (hi : c.prev j = i) (hk : c.next j = k) : short complex `K.sc i` has. -/ abbrev HasHomology := (K.sc i).HasHomology -variable [K.HasHomology i] +section variable [K.HasHomology i] /-- The homology in degree `i` of a homological complex. -/ noncomputable def homology := (K.sc i).homology @@ -126,11 +124,16 @@ noncomputable def cyclesIsKernel (hj : c.next i = j) : obtain rfl := hj exact (K.sc i).cyclesIsKernel +end + @[reassoc (attr := simp)] lemma toCycles_i [K.HasHomology j] : K.toCycles i j ≫ K.iCycles j = K.d i j := liftCycles_i _ _ _ _ _ +section +variable [K.HasHomology i] + instance : Mono (K.iCycles i) := by dsimp only [iCycles] infer_instance @@ -139,6 +142,8 @@ instance : Epi (K.homologyπ i) := by dsimp only [homologyπ] infer_instance +end + @[reassoc (attr := simp)] lemma d_toCycles [K.HasHomology k] : K.d i j ≫ K.toCycles j k = 0 := by @@ -146,6 +151,9 @@ lemma d_toCycles [K.HasHomology k] : variable {i} +section +variable [K.HasHomology i] + @[reassoc] lemma comp_liftCycles {A' A : C} (k : A ⟶ K.X i) (j : ι) (hj : c.next i = j) (hk : k ≫ K.d i j = 0) (α : A' ⟶ A) : @@ -164,6 +172,8 @@ lemma liftCycles_homologyπ_eq_zero_of_boundary {A : C} (k : A ⟶ K.X i) (j : rw [← cancel_mono (K.iCycles i), zero_comp, liftCycles_i, hx] rw [this, zero_comp] +end + variable (i) @[reassoc (attr := simp)] @@ -178,6 +188,9 @@ noncomputable def homologyIsCokernel (hi : c.prev j = i) [K.HasHomology j] : subst hi exact (K.sc j).homologyIsCokernel +section +variable [K.HasHomology i] + /-- The opcycles in degree `i` of a homological complex. -/ noncomputable def opcycles := (K.sc i).opcycles @@ -456,11 +469,10 @@ end end -variable (K L : HomologicalComplex C c) (i j k : ι) - section variable (hj : c.next i = j) (h : K.d i j = 0) [K.HasHomology i] +include hj h lemma isIso_iCycles : IsIso (K.iCycles i) := by subst hj @@ -507,6 +519,7 @@ end section variable (hi : c.prev j = i) (h : K.d i j = 0) [K.HasHomology j] +include hi h lemma isIso_pOpcycles : IsIso (K.pOpcycles j) := by obtain rfl := hi diff --git a/Mathlib/Algebra/Homology/TotalComplex.lean b/Mathlib/Algebra/Homology/TotalComplex.lean index 3a338452f6582..73fb3959181c0 100644 --- a/Mathlib/Algebra/Homology/TotalComplex.lean +++ b/Mathlib/Algebra/Homology/TotalComplex.lean @@ -31,20 +31,20 @@ namespace HomologicalComplex₂ variable {C : Type*} [Category C] [Preadditive C] {I₁ I₂ I₁₂ : Type*} {c₁ : ComplexShape I₁} {c₂ : ComplexShape I₂} (K L M : HomologicalComplex₂ C c₁ c₂) (φ : K ⟶ L) (e : K ≅ L) (ψ : L ⟶ M) - (c₁₂ : ComplexShape I₁₂) [DecidableEq I₁₂] - [TotalComplexShape c₁ c₂ c₁₂] + (c₁₂ : ComplexShape I₁₂) [TotalComplexShape c₁ c₂ c₁₂] /-- A bicomplex has a total bicomplex if for any `i₁₂ : I₁₂`, the coproduct of the objects `(K.X i₁).X i₂` such that `ComplexShape.π c₁ c₂ c₁₂ ⟨i₁, i₂⟩ = i₁₂` exists. -/ abbrev HasTotal := K.toGradedObject.HasMap (ComplexShape.π c₁ c₂ c₁₂) +include e in variable {K L} in lemma hasTotal_of_iso [K.HasTotal c₁₂] : L.HasTotal c₁₂ := GradedObject.hasMap_of_iso (GradedObject.isoMk K.toGradedObject L.toGradedObject (fun ⟨i₁, i₂⟩ => (HomologicalComplex.eval _ _ i₁ ⋙ HomologicalComplex.eval _ _ i₂).mapIso e)) _ -variable [K.HasTotal c₁₂] +variable [DecidableEq I₁₂] [K.HasTotal c₁₂] section diff --git a/Mathlib/Algebra/Lie/Abelian.lean b/Mathlib/Algebra/Lie/Abelian.lean index af06765422898..9d26a486b09c5 100644 --- a/Mathlib/Algebra/Lie/Abelian.lean +++ b/Mathlib/Algebra/Lie/Abelian.lean @@ -252,6 +252,7 @@ namespace LieModule variable {R L} variable {x : L} (hx : x ∈ LieAlgebra.center R L) (y : L) +include hx lemma commute_toEnd_of_mem_center_left : Commute (toEnd R L M x) (toEnd R L M y) := by diff --git a/Mathlib/Algebra/Lie/Engel.lean b/Mathlib/Algebra/Lie/Engel.lean index b8a88f7e442bc..75a7f7b2c26dc 100644 --- a/Mathlib/Algebra/Lie/Engel.lean +++ b/Mathlib/Algebra/Lie/Engel.lean @@ -76,6 +76,7 @@ namespace LieSubmodule open LieModule variable {I : LieIdeal R L} {x : L} (hxI : (R ∙ x) ⊔ I = ⊤) +include hxI theorem exists_smul_add_of_span_sup_eq_top (y : L) : ∃ t : R, ∃ z ∈ I, y = t • x + z := by have hy : y ∈ (⊤ : Submodule R L) := Submodule.mem_top diff --git a/Mathlib/Algebra/Lie/InvariantForm.lean b/Mathlib/Algebra/Lie/InvariantForm.lean index fa9f9e8d7866d..ba8b199d8e77c 100644 --- a/Mathlib/Algebra/Lie/InvariantForm.lean +++ b/Mathlib/Algebra/Lie/InvariantForm.lean @@ -37,7 +37,7 @@ namespace InvariantForm section ring variable {R L M : Type*} -variable [CommRing R] [LieRing L] [LieAlgebra R L] +variable [CommRing R] [LieRing L] variable [AddCommGroup M] [Module R M] [LieRingModule L M] variable (Φ : LinearMap.BilinForm R M) (hΦ_nondeg : Φ.Nondegenerate) @@ -50,7 +50,7 @@ for all `x : L` and `y z : M` the condition `Φ ⁅x, y⁆ z = -Φ y ⁅x, z⁆` def _root_.LinearMap.BilinForm.lieInvariant : Prop := ∀ (x : L) (y z : M), Φ ⁅x, y⁆ z = -Φ y ⁅x, z⁆ -lemma _root_.LinearMap.BilinForm.lieInvariant_iff [LieModule R L M] : +lemma _root_.LinearMap.BilinForm.lieInvariant_iff [LieAlgebra R L] [LieModule R L M] : Φ.lieInvariant L ↔ Φ ∈ LieModule.maxTrivSubmodule R L (LinearMap.BilinForm R M) := by refine ⟨fun h x ↦ ?_, fun h x y z ↦ ?_⟩ · ext y z @@ -87,6 +87,8 @@ lemma mem_orthogonal (N : LieSubmodule R L M) (y : M) : y ∈ orthogonal Φ hΦ_inv N ↔ ∀ x ∈ N, Φ x y = 0 := by simp [orthogonal, LinearMap.BilinForm.isOrtho_def, LinearMap.BilinForm.mem_orthogonal_iff] +variable [LieAlgebra R L] + lemma orthogonal_disjoint (Φ : LinearMap.BilinForm R L) (hΦ_nondeg : Φ.Nondegenerate) (hΦ_inv : Φ.lieInvariant L) -- TODO: replace the following assumption by a typeclass assumption `[HasNonAbelianAtoms]` @@ -120,6 +122,7 @@ variable (Φ : LinearMap.BilinForm K L) (hΦ_nondeg : Φ.Nondegenerate) variable (hΦ_inv : Φ.lieInvariant L) (hΦ_refl : Φ.IsRefl) -- TODO: replace the following assumption by a typeclass assumption `[HasNonAbelianAtoms]` variable (hL : ∀ I : LieIdeal K L, IsAtom I → ¬IsLieAbelian I) +include hΦ_nondeg hΦ_refl hL open FiniteDimensional Submodule in lemma orthogonal_isCompl_coe_submodule (I : LieIdeal K L) (hI : IsAtom I) : @@ -134,6 +137,8 @@ lemma orthogonal_isCompl (I : LieIdeal K L) (hI : IsAtom I) : rw [LieSubmodule.isCompl_iff_coe_toSubmodule] exact orthogonal_isCompl_coe_submodule Φ hΦ_nondeg hΦ_inv hΦ_refl hL I hI +include hΦ_inv + lemma restrict_nondegenerate (I : LieIdeal K L) (hI : IsAtom I) : (Φ.restrict I).Nondegenerate := by rw [LinearMap.BilinForm.restrict_nondegenerate_iff_isCompl_orthogonal hΦ_refl] diff --git a/Mathlib/Algebra/Lie/Nilpotent.lean b/Mathlib/Algebra/Lie/Nilpotent.lean index 24358b6aad617..ad16dcc5bdd9a 100644 --- a/Mathlib/Algebra/Lie/Nilpotent.lean +++ b/Mathlib/Algebra/Lie/Nilpotent.lean @@ -527,6 +527,7 @@ variable [AddCommGroup M₂] [Module R M₂] [LieRingModule L₂ M₂] [LieModul variable {f : L →ₗ⁅R⁆ L₂} {g : M →ₗ[R] M₂} variable (hf : Surjective f) (hg : Surjective g) (hfg : ∀ x m, ⁅f x, g m⁆ = g ⁅x, m⁆) +include hf hg hfg in theorem Function.Surjective.lieModule_lcs_map_eq (k : ℕ) : (lowerCentralSeries R L M k : Submodule R M).map g = lowerCentralSeries R L₂ M₂ k := by induction' k with k ih @@ -550,6 +551,7 @@ theorem Function.Surjective.lieModule_lcs_map_eq (k : ℕ) : obtain ⟨y, rfl⟩ := hf x exact ⟨⁅y, n⁆, ⟨y, n, hn, rfl⟩, (hfg y n).symm⟩ +include hf hg hfg in theorem Function.Surjective.lieModuleIsNilpotent [IsNilpotent R L M] : IsNilpotent R L₂ M₂ := by obtain ⟨k, hk⟩ := id (by infer_instance : IsNilpotent R L M) use k diff --git a/Mathlib/Algebra/Lie/Rank.lean b/Mathlib/Algebra/Lie/Rank.lean index e4e3f5c3f6738..3fabb5568a84c 100644 --- a/Mathlib/Algebra/Lie/Rank.lean +++ b/Mathlib/Algebra/Lie/Rank.lean @@ -67,6 +67,7 @@ lemma rank_eq_natTrailingDegree [Nontrivial R] [DecidableEq ι] : open FiniteDimensional +include bₘ in lemma rank_le_card [Nontrivial R] : rank R L M ≤ Fintype.card ιₘ := nilRank_le_card _ bₘ @@ -139,6 +140,7 @@ lemma rank_eq_natTrailingDegree [Nontrivial R] [DecidableEq ι] : open FiniteDimensional +include b in lemma rank_le_card [Nontrivial R] : rank R L ≤ Fintype.card ι := nilRank_le_card _ b diff --git a/Mathlib/Algebra/Lie/Sl2.lean b/Mathlib/Algebra/Lie/Sl2.lean index e0da8e26ab1ff..04d470605c27a 100644 --- a/Mathlib/Algebra/Lie/Sl2.lean +++ b/Mathlib/Algebra/Lie/Sl2.lean @@ -93,10 +93,20 @@ lemma HasPrimitiveVectorWith.mk' [NoZeroSMulDivisors ℤ M] (t : IsSl2Triple h e namespace HasPrimitiveVectorWith -variable {m : M} {μ : R} {t : IsSl2Triple h e f} (P : HasPrimitiveVectorWith t m μ) - +variable {m : M} {μ : R} {t : IsSl2Triple h e f} local notation "ψ" n => ((toEnd R L M f) ^ n) m +-- Although this is true by definition, we include this lemma (and the assumption) to mirror the API +-- for `lie_h_pow_toEnd_f` and `lie_e_pow_succ_toEnd_f`. +set_option linter.unusedVariables false in +@[nolint unusedArguments] +lemma lie_f_pow_toEnd_f (P : HasPrimitiveVectorWith t m μ) (n : ℕ) : + ⁅f, ψ n⁆ = ψ (n + 1) := by + simp [pow_succ'] + +variable (P : HasPrimitiveVectorWith t m μ) +include P + lemma lie_h_pow_toEnd_f (n : ℕ) : ⁅h, ψ n⁆ = (μ - 2 * n) • ψ n := by induction n with @@ -107,14 +117,6 @@ lemma lie_h_pow_toEnd_f (n : ℕ) : congr ring --- Although this is true by definition, we include this lemma (and the assumption) to mirror the API --- for `lie_h_pow_toEnd_f` and `lie_e_pow_succ_toEnd_f`. -set_option linter.unusedVariables false in -@[nolint unusedArguments] -lemma lie_f_pow_toEnd_f (P : HasPrimitiveVectorWith t m μ) (n : ℕ) : - ⁅f, ψ n⁆ = ψ (n + 1) := by - simp [pow_succ'] - lemma lie_e_pow_succ_toEnd_f (n : ℕ) : ⁅e, ψ (n + 1)⁆ = ((n + 1) * (μ - n)) • ψ n := by induction n with diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index 6392851189478..735f6856fa8fe 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -34,7 +34,6 @@ We define the trace / Killing form in this file and prove some basic properties. variable (R K L M : Type*) [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] - [Module.Free R M] [Module.Finite R M] local notation "φ" => LieModule.toEnd R L M @@ -95,7 +94,8 @@ lemma traceForm_lieInvariant : (traceForm R L M).lieInvariant L := by rw [LieHom.lie_apply, LinearMap.sub_apply, Module.Dual.lie_apply, LinearMap.zero_apply, LinearMap.zero_apply, traceForm_apply_lie_apply', sub_self] -@[simp] lemma traceForm_eq_zero_of_isNilpotent [IsReduced R] [IsNilpotent R L M] : +@[simp] lemma traceForm_eq_zero_of_isNilpotent + [Module.Free R M] [Module.Finite R M] [IsReduced R] [IsNilpotent R L M] : traceForm R L M = 0 := by ext x y simp only [traceForm_apply_apply, LinearMap.zero_apply, ← isNilpotent_iff_eq_zero] @@ -103,7 +103,8 @@ lemma traceForm_lieInvariant : (traceForm R L M).lieInvariant L := by exact isNilpotent_toEnd_of_isNilpotent₂ R L M x y @[simp] -lemma traceForm_weightSpace_eq [IsDomain R] [IsPrincipalIdealRing R] +lemma traceForm_weightSpace_eq [Module.Free R M] + [IsDomain R] [IsPrincipalIdealRing R] [LieAlgebra.IsNilpotent R L] [IsNoetherian R M] [LinearWeights R L M] (χ : L → R) (x y : L) : traceForm R L (weightSpace M χ) x y = finrank R (weightSpace M χ) • (χ x * χ y) := by set d := finrank R (weightSpace M χ) @@ -210,7 +211,8 @@ open TensorProduct variable [LieAlgebra.IsNilpotent R L] [IsDomain R] [IsPrincipalIdealRing R] -lemma traceForm_eq_sum_weightSpaceOf [IsTriangularizable R L M] (z : L) : +lemma traceForm_eq_sum_weightSpaceOf + [NoZeroSMulDivisors R M] [IsNoetherian R M] [IsTriangularizable R L M] (z : L) : traceForm R L M = ∑ χ ∈ (finite_weightSpaceOf_ne_bot R L M z).toFinset, traceForm R L (weightSpaceOf M χ z) := by ext x y @@ -230,7 +232,7 @@ lemma traceForm_eq_sum_weightSpaceOf [IsTriangularizable R L M] (z : L) : -- In characteristic zero (or even just `LinearWeights R L M`) a stronger result holds (no -- `⊓ LieAlgebra.center R L`) TODO prove this using `LieModule.traceForm_eq_sum_finrank_nsmul_mul`. -lemma lowerCentralSeries_one_inf_center_le_ker_traceForm : +lemma lowerCentralSeries_one_inf_center_le_ker_traceForm [Module.Free R M] [Module.Finite R M] : lowerCentralSeries R L L 1 ⊓ LieAlgebra.center R L ≤ LinearMap.ker (traceForm R L M) := by /- Sketch of proof (due to Zassenhaus): @@ -277,8 +279,8 @@ lemma lowerCentralSeries_one_inf_center_le_ker_traceForm : · exact commute_toEnd_of_mem_center_right (A ⊗[R] M) hzc (1 ⊗ₜ x) /-- A nilpotent Lie algebra with a representation whose trace form is non-singular is Abelian. -/ -lemma isLieAbelian_of_ker_traceForm_eq_bot (h : LinearMap.ker (traceForm R L M) = ⊥) : - IsLieAbelian L := by +lemma isLieAbelian_of_ker_traceForm_eq_bot [Module.Free R M] [Module.Finite R M] + (h : LinearMap.ker (traceForm R L M) = ⊥) : IsLieAbelian L := by simpa only [← disjoint_lowerCentralSeries_maxTrivSubmodule_iff R L L, disjoint_iff_inf_le, LieIdeal.coe_to_lieSubalgebra_to_submodule, LieSubmodule.coeSubmodule_eq_bot_iff, h] using lowerCentralSeries_one_inf_center_le_ker_traceForm R L M @@ -290,6 +292,7 @@ namespace LieSubmodule open LieModule (traceForm) variable {R L M} +variable [Module.Free R M] [Module.Finite R M] variable [IsDomain R] [IsPrincipalIdealRing R] (N : LieSubmodule R L M) (I : LieIdeal R L) (h : I ≤ N.idealizer) (x : L) {y : L} (hy : y ∈ I) @@ -299,6 +302,7 @@ lemma trace_eq_trace_restrict_of_le_idealizer suffices ∀ m, ⁅x, ⁅y, m⁆⁆ ∈ N by simp [(φ x ∘ₗ φ y).trace_restrict_eq_of_forall_mem _ this] exact fun m ↦ N.lie_mem (h hy m) +include h in lemma traceForm_eq_of_le_idealizer : traceForm R I N = (traceForm R L M).restrict I := by ext ⟨x, hx⟩ ⟨y, hy⟩ @@ -306,6 +310,7 @@ lemma traceForm_eq_of_le_idealizer : rw [N.trace_eq_trace_restrict_of_le_idealizer I h x hy] rfl +include h hy in /-- Note that this result is slightly stronger than it might look at first glance: we only assume that `N` is trivial over `I` rather than all of `L`. This means that it applies in the important case of an Abelian ideal (which has `M = L` and `N = I`). -/ @@ -322,8 +327,6 @@ end LieSubmodule section LieAlgebra -variable [Module.Free R L] [Module.Finite R L] - /-- A finite, free (as an `R`-module) Lie algebra `L` carries a bilinear form on `L`. This is a specialisation of `LieModule.traceForm` to the adjoint representation of `L`. -/ @@ -363,16 +366,16 @@ lemma coe_killingCompl_top : ext x simp [LinearMap.ext_iff, LinearMap.BilinForm.IsOrtho, LieModule.traceForm_comm R L L x] -variable [IsDomain R] [IsPrincipalIdealRing R] +lemma restrict_killingForm : + (killingForm R L).restrict I = LieModule.traceForm R I L := + rfl + +variable [Module.Free R L] [Module.Finite R L] [IsDomain R] [IsPrincipalIdealRing R] lemma killingForm_eq : killingForm R I = (killingForm R L).restrict I := LieSubmodule.traceForm_eq_of_le_idealizer I I <| by simp -lemma restrict_killingForm : - (killingForm R L).restrict I = LieModule.traceForm R I L := - rfl - @[simp] lemma le_killingCompl_top_of_isLieAbelian [IsLieAbelian I] : I ≤ LieIdeal.killingCompl R L ⊤ := by intro x (hx : x ∈ I) diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index 404329ebc837e..9b309c2715513 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -45,7 +45,7 @@ Basic definitions and properties of the above ideas are provided in this file. lie character, eigenvalue, eigenspace, weight, weight vector, root, root vector -/ -variable {K R L M : Type*} [CommRing R] [LieRing L] [LieAlgebra R L] [LieAlgebra.IsNilpotent R L] +variable {K R L M : Type*} [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] namespace LieModule @@ -146,7 +146,7 @@ variable (M) `weightSpaceOf M χ x` is the maximal generalized `χ`-eigenspace of the action of `x` on `M`. It is a Lie submodule because `L` is nilpotent. -/ -def weightSpaceOf (χ : R) (x : L) : LieSubmodule R L M := +def weightSpaceOf [LieAlgebra.IsNilpotent R L] (χ : R) (x : L) : LieSubmodule R L M := { 𝕎(M, χ, x) with lie_mem := by intro y m hm @@ -158,6 +158,7 @@ def weightSpaceOf (χ : R) (x : L) : LieSubmodule R L M := end notation_weightSpaceOf variable (M) +variable [LieAlgebra.IsNilpotent R L] theorem mem_weightSpaceOf (χ : R) (x : L) (m : M) : m ∈ weightSpaceOf M χ x ↔ ∃ k : ℕ, ((toEnd R L M x - χ • ↑1) ^ k) m = 0 := by @@ -763,7 +764,6 @@ instance (N : LieSubmodule K L M) [IsTriangularizable K L M] : IsTriangularizabl See also `LieModule.iSup_weightSpace_eq_top'`. -/ lemma iSup_weightSpace_eq_top [IsTriangularizable K L M] : ⨆ χ : L → K, weightSpace M χ = ⊤ := by - clear! R -- cf https://github.com/leanprover/lean4/issues/2452 induction' h_dim : finrank K M using Nat.strong_induction_on with n ih generalizing M obtain h' | ⟨y : L, hy : ¬ ∃ φ, weightSpaceOf M φ y = ⊤⟩ := forall_or_exists_not (fun (x : L) ↦ ∃ (φ : K), weightSpaceOf M φ x = ⊤) diff --git a/Mathlib/Algebra/Lie/Weights/Chain.lean b/Mathlib/Algebra/Lie/Weights/Chain.lean index e17fe80cd32a6..beb8dcaeaa05e 100644 --- a/Mathlib/Algebra/Lie/Weights/Chain.lean +++ b/Mathlib/Algebra/Lie/Weights/Chain.lean @@ -55,6 +55,7 @@ variable [LieAlgebra.IsNilpotent R L] (χ₁ χ₂ : L → R) (p q : ℤ) section variable [NoZeroSMulDivisors ℤ R] [NoZeroSMulDivisors R M] [IsNoetherian R M] (hχ₁ : χ₁ ≠ 0) +include hχ₁ lemma eventually_weightSpace_smul_add_eq_bot : ∀ᶠ (k : ℕ) in Filter.atTop, weightSpace M (k • χ₁ + χ₂) = ⊥ := by @@ -222,7 +223,7 @@ section variable {M} variable [LieAlgebra.IsNilpotent R L] variable [NoZeroSMulDivisors ℤ R] [NoZeroSMulDivisors R M] [IsNoetherian R M] -variable (α : L → R) (β : Weight R L M) (hα : α ≠ 0) +variable (α : L → R) (β : Weight R L M) /-- This is the largest `n : ℕ` such that `i • α + β` is a weight for all `0 ≤ i ≤ n`. -/ noncomputable @@ -243,6 +244,10 @@ def chainBotCoeff : ℕ := chainTopCoeff (-α) β @[simp] lemma chainTopCoeff_zero : chainTopCoeff 0 β = 0 := dif_pos rfl @[simp] lemma chainBotCoeff_zero : chainBotCoeff 0 β = 0 := dif_pos neg_zero +section +variable (hα : α ≠ 0) +include hα + lemma chainTopCoeff_add_one : letI := Classical.propDecidable chainTopCoeff α β + 1 = @@ -268,6 +273,13 @@ lemma weightSpace_chainTopCoeff_add_one_zsmul_add : rw [← weightSpace_chainTopCoeff_add_one_nsmul_add α β hα, ← Nat.cast_smul_eq_nsmul ℤ, Nat.cast_add, Nat.cast_one] +lemma weightSpace_chainBotCoeff_sub_one_zsmul_sub : + weightSpace M ((-chainBotCoeff α β - 1 : ℤ) • α + β : L → R) = ⊥ := by + rw [sub_eq_add_neg, ← neg_add, neg_smul, ← smul_neg, chainBotCoeff, + weightSpace_chainTopCoeff_add_one_zsmul_add _ _ (by simpa using hα)] + +end + lemma weightSpace_nsmul_add_ne_bot_of_le {n} (hn : n ≤ chainTopCoeff α β) : weightSpace M (n • α + β : L → R) ≠ ⊥ := by by_cases hα : α = 0 @@ -276,11 +288,6 @@ lemma weightSpace_nsmul_add_ne_bot_of_le {n} (hn : n ≤ chainTopCoeff α β) : rw [← Nat.lt_succ, Nat.succ_eq_add_one, chainTopCoeff_add_one _ _ hα] at hn exact Nat.find_min (eventually_weightSpace_smul_add_eq_bot M α β hα).exists hn -lemma weightSpace_chainBotCoeff_sub_one_zsmul_sub : - weightSpace M ((-chainBotCoeff α β - 1 : ℤ) • α + β : L → R) = ⊥ := by - rw [sub_eq_add_neg, ← neg_add, neg_smul, ← smul_neg, chainBotCoeff, - weightSpace_chainTopCoeff_add_one_zsmul_add _ _ (by simpa using hα)] - lemma weightSpace_zsmul_add_ne_bot {n : ℤ} (hn : -chainBotCoeff α β ≤ n) (hn' : n ≤ chainTopCoeff α β) : weightSpace M (n • α + β : L → R) ≠ ⊥ := by @@ -317,6 +324,10 @@ lemma coe_chainTop' : (chainTop α β : L → R) = chainTopCoeff α β • α + @[simp] lemma chainTop_zero : chainTop 0 β = β := by ext; simp @[simp] lemma chainBot_zero : chainBot 0 β = β := by ext; simp +section +variable (hα : α ≠ 0) +include hα + lemma weightSpace_add_chainTop : weightSpace M (α + chainTop α β : L → R) = ⊥ := by rw [coe_chainTop', ← add_assoc, ← succ_nsmul', weightSpace_chainTopCoeff_add_one_nsmul_add _ _ hα] @@ -331,6 +342,8 @@ lemma chainTop_isNonZero' (hα' : weightSpace M α ≠ ⊥) : apply hα' rw [← add_zero (α : L → R), ← e, weightSpace_add_chainTop _ _ hα] +end + lemma chainTop_isNonZero (α β : Weight R L M) (hα : α.IsNonZero) : (chainTop α β).IsNonZero := chainTop_isNonZero' α β hα α.2 diff --git a/Mathlib/Algebra/Module/LocalizedModule.lean b/Mathlib/Algebra/Module/LocalizedModule.lean index ff78e7064cfa4..45ebbceedbb8a 100644 --- a/Mathlib/Algebra/Module/LocalizedModule.lean +++ b/Mathlib/Algebra/Module/LocalizedModule.lean @@ -884,9 +884,11 @@ noncomputable def linearEquiv [IsLocalizedModule S g] : M' ≃ₗ[R] M'' := variable {S} +include f in theorem smul_injective (s : S) : Function.Injective fun m : M' => s • m := ((Module.End_isUnit_iff _).mp (IsLocalizedModule.map_units f s)).injective +include f in theorem smul_inj (s : S) (m₁ m₂ : M') : s • m₁ = s • m₂ ↔ m₁ = m₂ := (smul_injective f s).eq_iff diff --git a/Mathlib/Algebra/Module/PID.lean b/Mathlib/Algebra/Module/PID.lean index 50be6d05c621a..65924b9cdde6d 100644 --- a/Mathlib/Algebra/Module/PID.lean +++ b/Mathlib/Algebra/Module/PID.lean @@ -53,7 +53,7 @@ assert_not_exists TopologicalSpace universe u v -variable {R : Type u} [CommRing R] [IsDomain R] [IsPrincipalIdealRing R] +variable {R : Type u} [CommRing R] [IsPrincipalIdealRing R] variable {M : Type v} [AddCommGroup M] [Module R M] variable {N : Type max u v} [AddCommGroup N] [Module R N] @@ -69,6 +69,8 @@ theorem Submodule.isSemisimple_torsionBy_of_irreducible {a : R} (h : Irreducible letI := Ideal.Quotient.field (R ∙ a) (submodule_torsionBy_orderIso a).complementedLattice +variable [IsDomain R] + /-- A finitely generated torsion module over a PID is an internal direct sum of its `p i ^ e i`-torsion submodules for some primes `p i` and numbers `e i`. -/ theorem Submodule.isInternal_prime_power_torsion_of_pid [DecidableEq (Ideal R)] [Module.Finite R M] @@ -105,6 +107,8 @@ variable [dec : ∀ x : M, Decidable (x = 0)] open Ideal Submodule.IsPrincipal +include hp + theorem _root_.Ideal.torsionOf_eq_span_pow_pOrder (x : M) : torsionOf R M x = span {p ^ pOrder hM x} := by classical diff --git a/Mathlib/Algebra/Module/Submodule/EqLocus.lean b/Mathlib/Algebra/Module/Submodule/EqLocus.lean index 2e49aec2e699e..397b15a79a6ac 100644 --- a/Mathlib/Algebra/Module/Submodule/EqLocus.lean +++ b/Mathlib/Algebra/Module/Submodule/EqLocus.lean @@ -64,11 +64,13 @@ theorem eqLocus_same (f : F) : eqLocus f f = ⊤ := eqLocus_eq_top.2 rfl theorem le_eqLocus {f g : F} {S : Submodule R M} : S ≤ eqLocus f g ↔ Set.EqOn f g S := Iff.rfl +include τ₁₂ in theorem eqOn_sup {f g : F} {S T : Submodule R M} (hS : Set.EqOn f g S) (hT : Set.EqOn f g T) : Set.EqOn f g ↑(S ⊔ T) := by rw [← le_eqLocus] at hS hT ⊢ exact sup_le hS hT +include τ₁₂ in theorem ext_on_codisjoint {f g : F} {S T : Submodule R M} (hST : Codisjoint S T) (hS : Set.EqOn f g S) (hT : Set.EqOn f g T) : f = g := DFunLike.ext _ _ fun _ ↦ eqOn_sup hS hT <| hST.eq_top.symm ▸ trivial diff --git a/Mathlib/Algebra/Module/Submodule/Localization.lean b/Mathlib/Algebra/Module/Submodule/Localization.lean index b26e173ed73bf..9530de7dfed44 100644 --- a/Mathlib/Algebra/Module/Submodule/Localization.lean +++ b/Mathlib/Algebra/Module/Submodule/Localization.lean @@ -157,6 +157,7 @@ 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]) +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) : diff --git a/Mathlib/Algebra/Module/Submodule/Map.lean b/Mathlib/Algebra/Module/Submodule/Map.lean index d701d5f5cc447..7e0f78656d3e7 100644 --- a/Mathlib/Algebra/Module/Submodule/Map.lean +++ b/Mathlib/Algebra/Module/Submodule/Map.lean @@ -257,6 +257,7 @@ def giMapComap (hf : Surjective f) : GaloisInsertion (map f) (comap f) := exact ⟨y, hx, rfl⟩ variable (hf : Surjective f) +include hf theorem map_comap_eq_of_surjective (p : Submodule R₂ M₂) : (p.comap f).map f = p := (giMapComap hf).l_u_eq _ @@ -304,6 +305,7 @@ def gciMapComap (hf : Injective f) : GaloisCoinsertion (map f) (comap f) := rwa [← hxy] variable (hf : Injective f) +include hf theorem comap_map_eq_of_injective (p : Submodule R M) : (p.map f).comap f = p := (gciMapComap hf).u_l_eq _ diff --git a/Mathlib/Algebra/MonoidAlgebra/Degree.lean b/Mathlib/Algebra/MonoidAlgebra/Degree.lean index e0bf576fc7e96..b5fb1511da9ed 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Degree.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Degree.lean @@ -256,7 +256,7 @@ theorem supDegree_withBot_some_comp {s : AddMonoidAlgebra R A} (hs : s.support.N unfold AddMonoidAlgebra.supDegree rw [← Finset.coe_sup' hs, Finset.sup'_eq_sup] -variable [AddZeroClass A] {p q : R[A]} +variable [AddZeroClass A] {p q : R[A]} @[simp] theorem supDegree_zero : (0 : R[A]).supDegree D = ⊥ := by simp [supDegree] diff --git a/Mathlib/Algebra/Order/Pointwise.lean b/Mathlib/Algebra/Order/Pointwise.lean index 0a1ea8acdeb08..2bda973451012 100644 --- a/Mathlib/Algebra/Order/Pointwise.lean +++ b/Mathlib/Algebra/Order/Pointwise.lean @@ -143,6 +143,7 @@ end ConditionallyCompleteLattice namespace LinearOrderedField variable {K : Type*} [LinearOrderedField K] {a b r : K} (hr : 0 < r) +include hr open Set diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean index e9a10a1f07b78..738f12558d80d 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean @@ -586,7 +586,7 @@ lemma mul_add_mul_lt_mul_add_mul [ExistsAddOfLE α] [MulPosStrictMono α] /-- Binary **rearrangement inequality**. -/ lemma mul_add_mul_lt_mul_add_mul' [ExistsAddOfLE α] [MulPosStrictMono α] - [CovariantClass α α (· + ·) (· < ·)] [ContravariantClass α α (· + ·) (· < ·)] + [CovariantClass α α (· + ·) (· < ·)] (hba : b < a) (hdc : d < c) : a * d + b * c < a * c + b * d := by rw [add_comm (a * d), add_comm (a * c)] exact mul_add_mul_lt_mul_add_mul hba hdc diff --git a/Mathlib/Algebra/Order/ToIntervalMod.lean b/Mathlib/Algebra/Order/ToIntervalMod.lean index 749e1b6de5802..1f303379050ac 100644 --- a/Mathlib/Algebra/Order/ToIntervalMod.lean +++ b/Mathlib/Algebra/Order/ToIntervalMod.lean @@ -38,6 +38,9 @@ section LinearOrderedAddCommGroup variable {α : Type*} [LinearOrderedAddCommGroup α] [hα : Archimedean α] {p : α} (hp : 0 < p) {a b c : α} {n : ℤ} +section +include hp + /-- The unique integer such that this multiple of `p`, subtracted from `b`, is in `Ico a (a + p)`. -/ def toIcoDiv (a b : α) : ℤ := @@ -706,6 +709,7 @@ theorem QuotientAddGroup.equivIocMod_coe (a b : α) : theorem QuotientAddGroup.equivIocMod_zero (a : α) : QuotientAddGroup.equivIocMod hp a 0 = ⟨toIocMod hp a 0, toIocMod_mem_Ioc hp a _⟩ := rfl +end /-! ### The circular order structure on `α ⧸ AddSubgroup.zmultiples p` @@ -714,6 +718,8 @@ theorem QuotientAddGroup.equivIocMod_zero (a : α) : section Circular +open AddCommGroup + private theorem toIxxMod_iff (x₁ x₂ x₃ : α) : toIcoMod hp x₁ x₂ ≤ toIocMod hp x₁ x₃ ↔ toIcoMod hp 0 (x₂ - x₁) + toIcoMod hp 0 (x₁ - x₃) ≤ p := by rw [toIcoMod_eq_sub, toIocMod_eq_sub _ x₁, add_le_add_iff_right, ← neg_sub x₁ x₃, toIocMod_neg, @@ -901,6 +907,7 @@ open Set Int section LinearOrderedAddCommGroup variable {α : Type*} [LinearOrderedAddCommGroup α] [Archimedean α] {p : α} (hp : 0 < p) (a : α) +include hp theorem iUnion_Ioc_add_zsmul : ⋃ n : ℤ, Ioc (a + n • p) (a + (n + 1) • p) = univ := by refine eq_univ_iff_forall.mpr fun b => mem_iUnion.mpr ?_ diff --git a/Mathlib/Algebra/Polynomial/EraseLead.lean b/Mathlib/Algebra/Polynomial/EraseLead.lean index 71851a0b1af06..56246cc180cc6 100644 --- a/Mathlib/Algebra/Polynomial/EraseLead.lean +++ b/Mathlib/Algebra/Polynomial/EraseLead.lean @@ -220,8 +220,8 @@ theorem natDegree_eraseLead_le_of_nextCoeff_eq_zero (h : f.nextCoeff = 0) : rw [eraseLead_coeff_of_ne _ (tsub_lt_self hf zero_lt_one).ne, ← nextCoeff_of_natDegree_pos hf] simp [nextCoeff_eq_zero, h, eq_zero_or_pos] -lemma two_le_natDegree_of_nextCoeff_eraseLead (hlead : f.eraseLead ≠ 0) (hnext : f.nextCoeff = 0) : - 2 ≤ f.natDegree := by +lemma two_le_natDegree_of_nextCoeff_eraseLead (hlead : f.eraseLead ≠ 0) + (hnext : f.nextCoeff = 0) : 2 ≤ f.natDegree := by contrapose! hlead rw [Nat.lt_succ_iff, Nat.le_one_iff_eq_zero_or_eq_one, natDegree_eq_zero, natDegree_eq_one] at hlead diff --git a/Mathlib/Algebra/Polynomial/Expand.lean b/Mathlib/Algebra/Polynomial/Expand.lean index d16d75a257d9b..d31b700d20a34 100644 --- a/Mathlib/Algebra/Polynomial/Expand.lean +++ b/Mathlib/Algebra/Polynomial/Expand.lean @@ -214,7 +214,7 @@ theorem expand_contract [CharP R p] [NoZeroDivisors R] {f : R[X]} (hf : Polynomi rw [hf, coeff_zero, zero_eq_mul] at this cases' this with h' · rw [h'] - rename_i _ _ _ _ h' + rename_i _ _ _ h' rw [← Nat.cast_succ, CharP.cast_eq_zero_iff R p] at h' exact absurd h' h diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index 32336ed620bbc..5732a7c7bd649 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -311,6 +311,7 @@ variable [CommSemiring R] {a p : R[X]} section Monic variable (hp : p.Monic) +include hp theorem Monic.C_dvd_iff_isUnit {a : R} : C a ∣ p ↔ IsUnit a := ⟨fun h => isUnit_iff_dvd_one.mpr <| @@ -681,7 +682,7 @@ theorem comp_eq_zero_iff : p.comp q = 0 ↔ p = 0 ∨ p.eval (q.coeff 0) = 0 ∧ · exact fun h => Or.rec (fun h => by rw [h, zero_comp]) (fun h => by rw [h.2, comp_C, h.1, C_0]) h -lemma aeval_ne_zero_of_isCoprime [CommSemiring R] [Nontrivial S] [Semiring S] [Algebra R S] +lemma aeval_ne_zero_of_isCoprime {R} [CommSemiring R] [Nontrivial S] [Semiring S] [Algebra R S] {p q : R[X]} (h : IsCoprime p q) (s : S) : aeval s p ≠ 0 ∨ aeval s q ≠ 0 := by by_contra! hpq rcases h with ⟨_, _, h⟩ diff --git a/Mathlib/Algebra/Ring/Divisibility/Basic.lean b/Mathlib/Algebra/Ring/Divisibility/Basic.lean index 90286ddcf3813..5f1dd2b07ea5a 100644 --- a/Mathlib/Algebra/Ring/Divisibility/Basic.lean +++ b/Mathlib/Algebra/Ring/Divisibility/Basic.lean @@ -20,13 +20,13 @@ variable {α β : Type*} section Semigroup -variable [Semigroup α] [Semigroup β] {F : Type*} [EquivLike F α β] [MulEquivClass F α β] (f : F) +variable [Semigroup α] [Semigroup β] {F : Type*} [EquivLike F α β] [MulEquivClass F α β] -theorem map_dvd_iff {a b} : f a ∣ f b ↔ a ∣ b := +theorem map_dvd_iff (f : F) {a b} : f a ∣ f b ↔ a ∣ b := let f := MulEquivClass.toMulEquiv f ⟨fun h ↦ by rw [← f.left_inv a, ← f.left_inv b]; exact map_dvd f.symm h, map_dvd f⟩ -theorem MulEquiv.decompositionMonoid [DecompositionMonoid β] : DecompositionMonoid α where +theorem MulEquiv.decompositionMonoid (f : F) [DecompositionMonoid β] : DecompositionMonoid α where primal a b c h := by rw [← map_dvd_iff f, map_mul] at h obtain ⟨a₁, a₂, h⟩ := DecompositionMonoid.primal _ h diff --git a/Mathlib/Algebra/Ring/Hom/Defs.lean b/Mathlib/Algebra/Ring/Hom/Defs.lean index 5b45c997cbab6..8b6d0d07f6d18 100644 --- a/Mathlib/Algebra/Ring/Hom/Defs.lean +++ b/Mathlib/Algebra/Ring/Hom/Defs.lean @@ -506,6 +506,7 @@ theorem codomain_trivial_iff_range_trivial : (0 : β) = 1 ↔ ∀ x, f x = 0 := theorem map_one_ne_zero [Nontrivial β] : f 1 ≠ 0 := mt f.codomain_trivial_iff_map_one_eq_zero.mpr zero_ne_one +include f in /-- If there is a homomorphism `f : α →+* β` and `β` is nontrivial, then `α` is nontrivial. -/ theorem domain_nontrivial [Nontrivial β] : Nontrivial α := ⟨⟨1, 0, mt (fun h => show f 1 = 0 by rw [h, map_zero]) f.map_one_ne_zero⟩⟩ diff --git a/Mathlib/Algebra/Ring/InjSurj.lean b/Mathlib/Algebra/Ring/InjSurj.lean index 659f7585b8136..a16aaee0516ab 100644 --- a/Mathlib/Algebra/Ring/InjSurj.lean +++ b/Mathlib/Algebra/Ring/InjSurj.lean @@ -15,6 +15,7 @@ variable {α β : Type*} namespace Function.Injective variable (f : β → α) (hf : Injective f) +include hf variable [Add β] [Mul β] @@ -195,6 +196,7 @@ end Function.Injective namespace Function.Surjective variable (f : α → β) (hf : Surjective f) +include hf variable [Add β] [Mul β] diff --git a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean index ab0f694338b18..b0c9fd1f0330b 100644 --- a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean @@ -598,12 +598,14 @@ end NonUnitalSubalgebra namespace NonUnitalStarAlgebra variable [CommSemiring R] [StarRing R] -variable [NonUnitalSemiring A] [StarRing A] -variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] -variable [NonUnitalSemiring B] [StarRing B] -variable [Module R B] [IsScalarTower R B B] [SMulCommClass R B B] [StarModule R B] +variable [NonUnitalSemiring A] [StarRing A] [Module R A] +variable [NonUnitalSemiring B] [StarRing B] [Module R B] variable [FunLike F A B] [NonUnitalAlgHomClass F R A B] [NonUnitalStarAlgHomClass F R A B] +section StarSubAlgebraA + +variable [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] + open scoped Pointwise open NonUnitalStarSubalgebra @@ -726,7 +728,8 @@ theorem mul_mem_sup {S T : NonUnitalStarSubalgebra R A} {x y : A} (hx : x ∈ S) x * y ∈ S ⊔ T := mul_mem (mem_sup_left hx) (mem_sup_right hy) -theorem map_sup (f : F) (S T : NonUnitalStarSubalgebra R A) : +theorem map_sup [IsScalarTower R B B] [SMulCommClass R B B] [StarModule R B] (f : F) + (S T : NonUnitalStarSubalgebra R A) : ((S ⊔ T).map f : NonUnitalStarSubalgebra R B) = S.map f ⊔ T.map f := (NonUnitalStarSubalgebra.gc_map_comap f).l_sup @@ -789,44 +792,54 @@ theorem eq_top_iff {S : NonUnitalStarSubalgebra R A} : S = ⊤ ↔ ∀ x : A, x ⟨fun h x => by rw [h]; exact mem_top, fun h => by ext x; exact ⟨fun _ => mem_top, fun _ => h x⟩⟩ -theorem range_top_iff_surjective (f : F) : - NonUnitalStarAlgHom.range f = (⊤ : NonUnitalStarSubalgebra R B) ↔ Function.Surjective f := - NonUnitalStarAlgebra.eq_top_iff - @[simp] theorem range_id : NonUnitalStarAlgHom.range (NonUnitalStarAlgHom.id R A) = ⊤ := SetLike.coe_injective Set.range_id @[simp] -theorem map_top (f : F) : (⊤ : NonUnitalStarSubalgebra R A).map f = NonUnitalStarAlgHom.range f := - SetLike.coe_injective Set.image_univ - -@[simp] -theorem map_bot (f : F) : (⊥ : NonUnitalStarSubalgebra R A).map f = ⊥ := +theorem map_bot [IsScalarTower R B B] [SMulCommClass R B B] [StarModule R B] (f : F) : + (⊥ : NonUnitalStarSubalgebra R A).map f = ⊥ := SetLike.coe_injective <| by simp [NonUnitalAlgebra.coe_bot, NonUnitalStarSubalgebra.coe_map] @[simp] -theorem comap_top (f : F) : (⊤ : NonUnitalStarSubalgebra R B).comap f = ⊤ := +theorem comap_top [IsScalarTower R B B] [SMulCommClass R B B] [StarModule R B] (f : F) : + (⊤ : NonUnitalStarSubalgebra R B).comap f = ⊤ := eq_top_iff.2 fun _x => mem_top /-- `NonUnitalStarAlgHom` to `⊤ : NonUnitalStarSubalgebra R A`. -/ def toTop : A →⋆ₙₐ[R] (⊤ : NonUnitalStarSubalgebra R A) := NonUnitalStarAlgHom.codRestrict (NonUnitalStarAlgHom.id R A) ⊤ fun _ => mem_top +end StarSubAlgebraA + +theorem range_top_iff_surjective [IsScalarTower R B B] [SMulCommClass R B B] [StarModule R B] + (f : F) : NonUnitalStarAlgHom.range f = (⊤ : NonUnitalStarSubalgebra R B) ↔ + Function.Surjective f := + NonUnitalStarAlgebra.eq_top_iff + +@[simp] +theorem map_top [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] (f : F) : + (⊤ : NonUnitalStarSubalgebra R A).map f = NonUnitalStarAlgHom.range f := + SetLike.coe_injective Set.image_univ + end NonUnitalStarAlgebra namespace NonUnitalStarSubalgebra open NonUnitalStarAlgebra -variable [CommSemiring R] [StarRing R] -variable [NonUnitalSemiring A] [StarRing A] -variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] -variable [NonUnitalSemiring B] [StarRing B] -variable [Module R B] [IsScalarTower R B B] [SMulCommClass R B B] [StarModule R B] +variable [CommSemiring R] +variable [NonUnitalSemiring A] [StarRing A] [Module R A] +variable [NonUnitalSemiring B] [StarRing B] [Module R B] variable [FunLike F A B] [NonUnitalAlgHomClass F R A B] [NonUnitalStarAlgHomClass F R A B] variable (S : NonUnitalStarSubalgebra R A) +section StarSubalgebra + +variable [StarRing R] +variable [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] +variable [IsScalarTower R B B] [SMulCommClass R B B] [StarModule R B] + lemma _root_.NonUnitalStarAlgHom.map_adjoin (f : F) (s : Set A) : map f (adjoin R s) = adjoin R (f '' s) := Set.image_preimage.l_comm_of_u_comm (gc_map_comap f) NonUnitalStarAlgebra.gi.gc @@ -848,9 +861,6 @@ instance _root_.NonUnitalStarAlgHom.subsingleton [Subsingleton (NonUnitalStarSub Subsingleton.elim (⊤ : NonUnitalStarSubalgebra R A) ⊥ ▸ mem_top (mem_bot.mp this).symm ▸ (map_zero f).trans (map_zero g).symm⟩ -theorem range_val : NonUnitalStarAlgHom.range (NonUnitalStarSubalgebraClass.subtype S) = S := - ext <| Set.ext_iff.1 <| (NonUnitalStarSubalgebraClass.subtype S).coe_range.trans Subtype.range_val - /-- The map `S → T` when `S` is a non-unital star subalgebra contained in the non-unital star algebra `T`. @@ -888,6 +898,12 @@ theorem inclusion_inclusion {S T U : NonUnitalStarSubalgebra R A} (hst : S ≤ T theorem val_inclusion {S T : NonUnitalStarSubalgebra R A} (h : S ≤ T) (s : S) : (inclusion h s : A) = s := rfl + +end StarSubalgebra + +theorem range_val : NonUnitalStarAlgHom.range (NonUnitalStarSubalgebraClass.subtype S) = S := + ext <| Set.ext_iff.1 <| (NonUnitalStarSubalgebraClass.subtype S).coe_range.trans Subtype.range_val + section Prod variable (S₁ : NonUnitalStarSubalgebra R B) @@ -911,6 +927,10 @@ theorem mem_prod {S : NonUnitalStarSubalgebra R A} {S₁ : NonUnitalStarSubalgeb x ∈ prod S S₁ ↔ x.1 ∈ S ∧ x.2 ∈ S₁ := Set.mem_prod +variable [StarRing R] +variable [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] +variable [IsScalarTower R B B] [SMulCommClass R B B] [StarModule R B] + @[simp] theorem prod_top : (prod ⊤ ⊤ : NonUnitalStarSubalgebra R (A × B)) = ⊤ := by ext; simp @@ -928,6 +948,11 @@ end Prod section iSupLift variable {ι : Type*} +variable [StarRing R] [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] + +section StarSubalgebraB + +variable [IsScalarTower R B B] [SMulCommClass R B B] [StarModule R B] theorem coe_iSup_of_directed [Nonempty ι] {S : ι → NonUnitalStarSubalgebra R A} (dir : Directed (· ≤ ·) S) : ↑(iSup S) = ⋃ i, (S i : Set A) := @@ -988,6 +1013,8 @@ noncomputable def iSupLift [Nonempty ι] (K : ι → NonUnitalStarSubalgebra R A (fun _ _ => rfl) all_goals simp [map_star] } +end StarSubalgebraB + variable [Nonempty ι] {K : ι → NonUnitalStarSubalgebra R A} {dir : Directed (· ≤ ·) K} {f : ∀ i, K i →⋆ₙₐ[R] B} {hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)} {T : NonUnitalStarSubalgebra R A} {hT : T = iSup K} @@ -1022,6 +1049,7 @@ end iSupLift section Center variable (R A) +variable [IsScalarTower R A A] [SMulCommClass R A A] /-- The center of a non-unital star algebra is the set of elements which commute with every element. They form a non-unital star subalgebra. -/ @@ -1038,7 +1066,7 @@ theorem center_toNonUnitalSubalgebra : rfl @[simp] -theorem center_eq_top (A : Type*) [NonUnitalCommSemiring A] [StarRing A] [Module R A] +theorem center_eq_top (A : Type*) [StarRing R] [NonUnitalCommSemiring A] [StarRing A] [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] : center R A = ⊤ := SetLike.coe_injective (Set.center_eq_univ A) @@ -1059,6 +1087,7 @@ end Center section Centralizer variable (R) +variable [IsScalarTower R A A] [SMulCommClass R A A] /-- The centralizer of the star-closure of a set as a non-unital star subalgebra. -/ def centralizer (s : Set A) : NonUnitalStarSubalgebra R A := diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 48cc3b0a7d48f..9e2cd3513685e 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -284,6 +284,7 @@ theorem map_fromSpec {V : X.Opens} (hV : IsAffineOpen V) (f : op U ⟶ op V) : rw [fromSpec, ← Spec.map_comp_assoc, ← Functor.map_comp] rfl +include hU in protected theorem isCompact : IsCompact (U : Set X) := by convert @IsCompact.image _ _ _ _ Set.univ hU.fromSpec.1.base PrimeSpectrum.compactSpace.1 @@ -291,6 +292,7 @@ protected theorem isCompact : convert hU.range_fromSpec.symm exact Set.image_univ +include hU in theorem image_of_isOpenImmersion (f : X ⟶ Y) [H : IsOpenImmersion f] : IsAffineOpen (f ''ᵁ U) := by have : IsAffine _ := hU @@ -390,6 +392,7 @@ theorem basicOpen_fromSpec_app : (Spec Γ(X, U)).basicOpen (hU.fromSpec.app U f) = PrimeSpectrum.basicOpen f := by rw [← hU.fromSpec_preimage_basicOpen, Scheme.preimage_basicOpen] +include hU in theorem basicOpen : IsAffineOpen (X.basicOpen f) := by rw [← hU.fromSpec_image_basicOpen, Scheme.Hom.isAffineOpen_iff_of_isOpenImmersion] @@ -400,6 +403,7 @@ theorem basicOpen : instance [IsAffine X] (r : Γ(X, ⊤)) : IsAffine (X.basicOpen r) := (isAffineOpen_top X).basicOpen _ +include hU in theorem ι_basicOpen_preimage (r : Γ(X, ⊤)) : IsAffineOpen ((X.basicOpen r).ι ⁻¹ᵁ U) := by apply (X.basicOpen r).ι.isAffineOpen_iff_of_isOpenImmersion.mp @@ -408,6 +412,7 @@ theorem ι_basicOpen_preimage (r : Γ(X, ⊤)) : ← Scheme.basicOpen_res _ _ (homOfLE le_top).op] exact hU.basicOpen _ +include hU in theorem exists_basicOpen_le {V : X.Opens} (x : V) (h : ↑x ∈ U) : ∃ f : Γ(X, U), X.basicOpen f ≤ V ∧ ↑x ∈ X.basicOpen f := by have : IsAffine _ := hU @@ -441,6 +446,7 @@ instance basicOpenSectionsToAffine_isIso : rw [hU.range_fromSpec] exact RingedSpace.basicOpen_le _ _ +include hU in theorem isLocalization_basicOpen : IsLocalization.Away f Γ(X, X.basicOpen f) := by apply @@ -476,6 +482,7 @@ lemma appLE_eq_away_map {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} (hU : IsA RingHom.algebraMap_toAlgebra, RingHom.algebraMap_toAlgebra, ← CommRingCat.comp_eq_ring_hom_comp, Scheme.Hom.appLE_map, Scheme.Hom.map_appLE] +include hU in theorem isLocalization_of_eq_basicOpen {V : X.Opens} (i : V ⟶ U) (e : V = X.basicOpen f) : @IsLocalization.Away _ _ f Γ(X, V) _ (X.presheaf.map i.op).toAlgebra := by subst e; convert isLocalization_basicOpen hU f using 3 @@ -485,6 +492,7 @@ instance _root_.AlgebraicGeometry.Γ_restrict_isLocalization IsLocalization.Away r Γ(X.basicOpen r, ⊤) := (isAffineOpen_top X).isLocalization_of_eq_basicOpen r _ (Opens.openEmbedding_obj_top _) +include hU in theorem basicOpen_basicOpen_is_basicOpen (g : Γ(X, X.basicOpen f)) : ∃ f' : Γ(X, U), X.basicOpen f' = X.basicOpen g := by have := isLocalization_basicOpen hU f @@ -501,6 +509,7 @@ theorem basicOpen_basicOpen_is_basicOpen (g : Γ(X, X.basicOpen f)) : (IsLocalization.toInvSubmonoid (Submonoid.powers f) (Γ(X, X.basicOpen f)) _).prop +include hU in theorem _root_.AlgebraicGeometry.exists_basicOpen_le_affine_inter {V : X.Opens} (hV : IsAffineOpen V) (x : X) (hx : x ∈ U ⊓ V) : ∃ (f : Γ(X, U)) (g : Γ(X, V)), X.basicOpen f = X.basicOpen g ∧ x ∈ X.basicOpen f := by @@ -566,6 +575,7 @@ def _root_.AlgebraicGeometry.Scheme.affineBasicOpen (X : Scheme) {U : X.affineOpens} (f : Γ(X, U)) : X.affineOpens := ⟨X.basicOpen f, U.prop.basicOpen f⟩ +include hU in /-- In an affine open set `U`, a family of basic open covers `U` iff the sections span `Γ(X, U)`. See `iSup_basicOpen_of_span_eq_top` for the inverse direction without the affine-ness assuption. @@ -598,6 +608,7 @@ theorem basicOpen_union_eq_self_iff (s : Set Γ(X, U)) : PrimeSpectrum.zeroLocus_empty_iff_eq_top, PrimeSpectrum.zeroLocus_span] simp only [Set.iUnion_singleton_eq_range, Subtype.range_val_subtype, Set.setOf_mem_eq] +include hU in theorem self_le_basicOpen_union_iff (s : Set Γ(X, U)) : (U ≤ ⨆ f : s, X.basicOpen f.1) ↔ Ideal.span s = ⊤ := by rw [← hU.basicOpen_union_eq_self_iff, @comm _ Eq] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 31ca529057810..855fcc83a8aa2 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -200,6 +200,7 @@ theorem app_top (H : P f) [IsAffine X] [IsAffine Y] : Q (f.app ⊤) := by rw [Scheme.Hom.app_eq_appLE] exact appLE P f H ⟨_, isAffineOpen_top _⟩ ⟨_, isAffineOpen_top _⟩ _ +include Q in theorem comp_of_isOpenImmersion [IsOpenImmersion f] (H : P g) : P (f ≫ g) := by rw [eq_affineLocally P, affineLocally_iff_affineOpens_le] at H ⊢ @@ -340,6 +341,7 @@ theorem of_comp instance : P.IsMultiplicative where +include Q in lemma of_isOpenImmersion [IsOpenImmersion f] : P f := IsLocalAtSource.of_isOpenImmersion f lemma stableUnderBaseChange (hP : RingHom.StableUnderBaseChange Q) : P.StableUnderBaseChange := by diff --git a/Mathlib/AlgebraicGeometry/Noetherian.lean b/Mathlib/AlgebraicGeometry/Noetherian.lean index ad7dc34669b3f..5b47926dddcf1 100644 --- a/Mathlib/AlgebraicGeometry/Noetherian.lean +++ b/Mathlib/AlgebraicGeometry/Noetherian.lean @@ -58,6 +58,7 @@ section localizationProps variable {R : Type u} [CommRing R] (S : Finset R) (hS : Ideal.span (α := R) S = ⊤) (hN : ∀ s : S, IsNoetherianRing (Away (M := R) s)) +include hS hN in /-- Let `R` be a ring, and `f i` a finite collection of elements of `R` generating the unit ideal. If the localization of `R` at each `f i` is noetherian, so is `R`. diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/TensorProduct.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/TensorProduct.lean index 05ca19ffa9555..c13845fdd8fb9 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/TensorProduct.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/TensorProduct.lean @@ -18,7 +18,7 @@ import Mathlib.AlgebraicGeometry.PrimeSpectrum.Basic -/ variable (R S T : Type*) [CommRing R] [CommRing S] [Algebra R S] -variable [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] +variable [CommRing T] [Algebra R T] open TensorProduct @@ -32,6 +32,7 @@ lemma PrimeSpectrum.continuous_tensorProductTo : Continuous (tensorProductTo R S (comap _).2.prod_mk (comap _).2 variable (hRT : (algebraMap R T).SurjectiveOnStalks) +include hRT lemma PrimeSpectrum.embedding_tensorProductTo_of_surjectiveOnStalks_aux (p₁ p₂ : PrimeSpectrum (S ⊗[R] T)) diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 3ed524534c561..5785868dc943e 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -377,6 +377,7 @@ theorem carrier.add_mem (q : Spec.T A⁰_ f) {a b : A} (ha : a ∈ carrier f_deg · simp_rw [pow_add]; rfl variable (hm : 0 < m) (q : Spec.T A⁰_ f) +include hm theorem carrier.zero_mem : (0 : A) ∈ carrier f_deg q := fun i => by convert Submodule.zero_mem q.1 using 1 @@ -522,6 +523,7 @@ end fromSpecToSpec namespace toSpec variable {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) +include hm variable {𝒜} in lemma image_basicOpen_eq_basicOpen (a : A) (i : ℕ) : diff --git a/Mathlib/AlgebraicGeometry/Pullbacks.lean b/Mathlib/AlgebraicGeometry/Pullbacks.lean index 27e3ad47bf323..36ae5c042ab8c 100644 --- a/Mathlib/AlgebraicGeometry/Pullbacks.lean +++ b/Mathlib/AlgebraicGeometry/Pullbacks.lean @@ -415,6 +415,7 @@ def gluedIsLimit : IsLimit (PullbackCone.mk _ _ (p_comm 𝒰 f g)) := by pullbackSymmetry_hom_comp_snd_assoc, pullback.lift_fst_assoc, Category.comp_id, pullbackRightPullbackFstIso_hom_fst_assoc, ← pullback.condition_assoc, h₂] +include 𝒰 in theorem hasPullback_of_cover : HasPullback f g := ⟨⟨⟨_, gluedIsLimit 𝒰 f g⟩⟩⟩ diff --git a/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean b/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean index 3745a6fd9c752..c47f77c73c988 100644 --- a/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean +++ b/Mathlib/AlgebraicTopology/FundamentalGroupoid/InducedMaps.lean @@ -83,6 +83,7 @@ theorem hcast_def {X : TopCat} {x₀ x₁ : X} (hx₀ : x₀ = x₁) : variable {X₁ X₂ Y : TopCat.{u}} {f : C(X₁, Y)} {g : C(X₂, Y)} {x₀ x₁ : X₁} {x₂ x₃ : X₂} {p : Path x₀ x₁} {q : Path x₂ x₃} (hfg : ∀ t, f (p t) = g (q t)) +include hfg /-- If `f(p(t) = g(q(t))` for two paths `p` and `q`, then the induced path homotopy classes `f(p)` and `g(p)` are the same as well, despite having a priori different types -/ diff --git a/Mathlib/Analysis/Analytic/Meromorphic.lean b/Mathlib/Analysis/Analytic/Meromorphic.lean index 5fc9bcc71b8be..e2dd5221e27e3 100644 --- a/Mathlib/Analysis/Analytic/Meromorphic.lean +++ b/Mathlib/Analysis/Analytic/Meromorphic.lean @@ -238,34 +238,45 @@ lemma const (e : E) {U : Set 𝕜} : MeromorphicOn (fun _ ↦ e) U := section arithmetic +include hf in lemma mono_set {V : Set 𝕜} (hv : V ⊆ U) : MeromorphicOn f V := fun x hx ↦ hf x (hv hx) +include hf hg in lemma add : MeromorphicOn (f + g) U := fun x hx ↦ (hf x hx).add (hg x hx) +include hf hg in lemma sub : MeromorphicOn (f - g) U := fun x hx ↦ (hf x hx).sub (hg x hx) +include hf in lemma neg : MeromorphicOn (-f) U := fun x hx ↦ (hf x hx).neg @[simp] lemma neg_iff : MeromorphicOn (-f) U ↔ MeromorphicOn f U := ⟨fun h ↦ by simpa only [neg_neg] using h.neg, neg⟩ +include hs hf in lemma smul : MeromorphicOn (s • f) U := fun x hx ↦ (hs x hx).smul (hf x hx) +include hs ht in lemma mul : MeromorphicOn (s * t) U := fun x hx ↦ (hs x hx).mul (ht x hx) +include hs in lemma inv : MeromorphicOn s⁻¹ U := fun x hx ↦ (hs x hx).inv @[simp] lemma inv_iff : MeromorphicOn s⁻¹ U ↔ MeromorphicOn s U := ⟨fun h ↦ by simpa only [inv_inv] using h.inv, inv⟩ +include hs ht in lemma div : MeromorphicOn (s / t) U := fun x hx ↦ (hs x hx).div (ht x hx) +include hs in lemma pow (n : ℕ) : MeromorphicOn (s ^ n) U := fun x hx ↦ (hs x hx).pow _ +include hs in lemma zpow (n : ℤ) : MeromorphicOn (s ^ n) U := fun x hx ↦ (hs x hx).zpow _ end arithmetic +include hf in lemma congr (h_eq : Set.EqOn f g U) (hu : IsOpen U) : MeromorphicOn g U := by refine fun x hx ↦ (hf x hx).congr (EventuallyEq.filter_mono ?_ nhdsWithin_le_nhds) exact eventually_of_mem (hu.mem_nhds hx) h_eq diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index 2e4ffbbc89320..c9d1ff9e06ab5 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -645,7 +645,7 @@ protected theorem edist_zero_of_eq_zero {f : α → E} {s : Set α} (hf : Locall edist (f a) (f b) = 0 := by wlog h' : a ≤ b · rw [edist_comm] - apply this f s hf hb ha _ (le_of_not_le h') + apply this hf hb ha _ (le_of_not_le h') rw [variationOnFromTo.eq_neg_swap, h, neg_zero] · apply le_antisymm _ (zero_le _) rw [← ENNReal.ofReal_zero, ← h, variationOnFromTo.eq_of_le f s h', diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean index bfbdfec1ce738..40711ef598fae 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean @@ -315,7 +315,7 @@ def toFilterDistortioniUnion (l : IntegrationParams) (I : Box ι) (c : ℝ≥0) /-- A set `s : Set (TaggedPrepartition I)` belongs to `l.toFilteriUnion I π₀` if for any `c : ℝ≥0` there exists a function `r : ℝⁿ → (0, ∞)` (or a constant `r` if `l.bRiemann = true`) such that `s` contains each prepartition `π` such that `l.MemBaseSet I c r π` and `π.iUnion = π₀.iUnion`. -/ -def toFilteriUnion (l : IntegrationParams) (I : Box ι) (π₀ : Prepartition I) := +def toFilteriUnion (I : Box ι) (π₀ : Prepartition I) := ⨆ c : ℝ≥0, l.toFilterDistortioniUnion I c π₀ theorem rCond_of_bRiemann_eq_false {ι} (l : IntegrationParams) (hl : l.bRiemann = false) @@ -348,10 +348,9 @@ theorem MemBaseSet.exists_common_compl (hU : π₁.iUnion = π₂.iUnion) : ∃ π : Prepartition I, π.iUnion = ↑I \ π₁.iUnion ∧ (l.bDistortion → π.distortion ≤ c₁) ∧ (l.bDistortion → π.distortion ≤ c₂) := by - clear l₁ l₂ wlog hc : c₁ ≤ c₂ with H · simpa [hU, _root_.and_comm] using - @H _ _ I J c c₂ c₁ l r₂ r₁ π π₂ π₁ _ h₂ h₁ hU.symm (le_of_not_le hc) + @H _ _ I c₂ c₁ r₂ r₁ π₂ π₁ _ h₂ h₁ hU.symm (le_of_not_le hc) by_cases hD : (l.bDistortion : Prop) · rcases h₁.4 hD with ⟨π, hπU, hπc⟩ exact ⟨π, hπU, fun _ => hπc, fun _ => hπc.trans hc⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index cd5df457cc5a8..cc87a12072105 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -120,6 +120,7 @@ lemma cfcₙAux_mem_range_inr (f : C(σₙ 𝕜 a, 𝕜)₀) : variable [CStarRing A] +include hp₁ in open Unitization NonUnitalStarAlgHom in theorem RCLike.nonUnitalContinuousFunctionalCalculus : NonUnitalContinuousFunctionalCalculus 𝕜 (p : A → Prop) where diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index cfbfc80013afc..dfb187b624e62 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -245,17 +245,21 @@ lemma cfcₙ_cases (P : A → Prop) (a : A) (f : R → R) (h₀ : P 0) · rwa [cfcₙ_apply_of_not_predicate _ h] variable (R) in +include ha in lemma cfcₙ_id : cfcₙ (id : R → R) a = a := cfcₙ_apply (id : R → R) a ▸ cfcₙHom_id (p := p) ha variable (R) in +include ha in lemma cfcₙ_id' : cfcₙ (fun x : R ↦ x) a = a := cfcₙ_id R a +include ha hf hf0 in /-- The **spectral mapping theorem** for the non-unital continuous functional calculus. -/ lemma cfcₙ_map_quasispectrum : σₙ R (cfcₙ f a) = f '' σₙ R a := by simp [cfcₙ_apply f a, cfcₙHom_map_quasispectrum (p := p)] variable (R) in +include R in lemma cfcₙ_predicate_zero : p 0 := NonUnitalContinuousFunctionalCalculus.predicate_zero (R := R) @@ -305,12 +309,14 @@ lemma cfcₙ_const_zero : cfcₙ (fun _ : R ↦ 0) a = 0 := cfcₙ_zero R a variable {R} +include hf hf0 hg hg0 in lemma cfcₙ_mul : cfcₙ (fun x ↦ f x * g x) a = cfcₙ f a * cfcₙ g a := by by_cases ha : p a · rw [cfcₙ_apply f a, cfcₙ_apply g a, ← map_mul, cfcₙ_apply _ a] congr · simp [cfcₙ_apply_of_not_predicate a ha] +include hf hf0 hg hg0 in lemma cfcₙ_add : cfcₙ (fun x ↦ f x + g x) a = cfcₙ f a + cfcₙ g a := by by_cases ha : p a · rw [cfcₙ_apply f a, cfcₙ_apply g a, cfcₙ_apply _ a] @@ -380,6 +386,7 @@ lemma cfcₙ_smul_id {S : Type*} [SMulZeroClass S R] [ContinuousConstSMul S R] lemma cfcₙ_const_mul_id (r : R) (a : A) (ha : p a := by cfc_tac) : cfcₙ (r * ·) a = r • a := cfcₙ_smul_id r a +include ha in lemma cfcₙ_star_id : cfcₙ (star · : R → R) a = star a := by rw [cfcₙ_star _ a, cfcₙ_id' R a] @@ -437,6 +444,7 @@ lemma CFC.eq_zero_of_quasispectrum_eq_zero (h_spec : σₙ R a ⊆ {0}) (ha : p a = 0 := by simpa [cfcₙ_id R a] using cfcₙ_congr (a := a) (f := id) (g := fun _ : R ↦ 0) fun x ↦ by simp_all +include instCFCₙ in lemma CFC.quasispectrum_zero_eq : σₙ R (0 : A) = {0} := by refine Set.eq_singleton_iff_unique_mem.mpr ⟨quasispectrum.zero_mem R 0, fun x hx ↦ ?_⟩ rw [← cfcₙ_zero R (0 : A), @@ -472,6 +480,7 @@ variable (f g : R → R) (a : A) variable (hf : ContinuousOn f (σₙ R a) := by cfc_cont_tac) (hf0 : f 0 = 0 := by cfc_zero_tac) variable (hg : ContinuousOn g (σₙ R a) := by cfc_cont_tac) (hg0 : g 0 = 0 := by cfc_zero_tac) +include hf hf0 hg hg0 in lemma cfcₙ_sub : cfcₙ (fun x ↦ f x - g x) a = cfcₙ f a - cfcₙ g a := by by_cases ha : p a · rw [cfcₙ_apply f a, cfcₙ_apply g a, ← map_sub, cfcₙ_apply ..] diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index ab9b6a6c87cf1..bb752066f3920 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -340,6 +340,7 @@ lemma cfc_const (r : R) (a : A) (ha : p a := by cfc_tac) : congr variable (R) in +include R in lemma cfc_predicate_zero : p 0 := ContinuousFunctionalCalculus.predicate_zero (R := R) @@ -350,6 +351,7 @@ lemma cfc_predicate_algebraMap (r : R) : p (algebraMap R A r) := cfc_const r (0 : A) (cfc_predicate_zero R) ▸ cfc_predicate (fun _ ↦ r) 0 variable (R) in +include R in lemma cfc_predicate_one : p 1 := map_one (algebraMap R A) ▸ cfc_predicate_algebraMap (1 : R) @@ -374,14 +376,17 @@ lemma eqOn_of_cfc_eq_cfc {f g : R → R} {a : A} (h : cfc f a = cfc g a) congrm($(this) ⟨x, hx⟩) variable {a f g} in +include ha hf hg in lemma cfc_eq_cfc_iff_eqOn : cfc f a = cfc g a ↔ (spectrum R a).EqOn f g := ⟨eqOn_of_cfc_eq_cfc, cfc_congr⟩ variable (R) +include ha in lemma cfc_one : cfc (1 : R → R) a = 1 := cfc_apply (1 : R → R) a ▸ map_one (cfcHom (show p a from ha)) +include ha in lemma cfc_const_one : cfc (fun _ : R ↦ 1) a = 1 := cfc_one R a @[simp] @@ -490,6 +495,7 @@ lemma cfc_smul_id {S : Type*} [SMul S R] [ContinuousConstSMul S R] lemma cfc_const_mul_id (r : R) (a : A) (ha : p a := by cfc_tac) : cfc (r * ·) a = r • a := cfc_smul_id r a +include ha in lemma cfc_star_id : cfc (star · : R → R) a = star a := by rw [cfc_star .., cfc_id' ..] @@ -588,23 +594,27 @@ lemma CFC.eq_one_of_spectrum_subset_one (h_spec : spectrum R a ⊆ {1}) (ha : p a = 1 := by simpa using eq_algebraMap_of_spectrum_subset_singleton a 1 h_spec +include instCFC in lemma CFC.spectrum_algebraMap_subset (r : R) : spectrum R (algebraMap R A r) ⊆ {r} := by rw [← cfc_const r 0 (cfc_predicate_zero R), cfc_map_spectrum (fun _ ↦ r) 0 (cfc_predicate_zero R)] rintro - ⟨x, -, rfl⟩ simp +include instCFC in lemma CFC.spectrum_algebraMap_eq [Nontrivial A] (r : R) : spectrum R (algebraMap R A r) = {r} := by have hp : p 0 := cfc_predicate_zero R rw [← cfc_const r 0 hp, cfc_map_spectrum (fun _ => r) 0 hp] exact Set.Nonempty.image_const (⟨0, spectrum.zero_mem (R := R) not_isUnit_zero⟩) _ +include instCFC in lemma CFC.spectrum_zero_eq [Nontrivial A] : spectrum R (0 : A) = {0} := by have : (0 : A) = algebraMap R A 0 := Eq.symm (RingHom.map_zero (algebraMap R A)) rw [this, spectrum_algebraMap_eq] +include instCFC in lemma CFC.spectrum_one_eq [Nontrivial A] : spectrum R (1 : A) = {1} := by have : (1 : A) = algebraMap R A 1 := Eq.symm (RingHom.map_one (algebraMap R A)) @@ -761,6 +771,7 @@ variable [Ring A] [StarRing A] [Algebra R A] [ContinuousFunctionalCalculus R p] variable (f g : R → R) (a : A) (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) variable (hg : ContinuousOn g (spectrum R a) := by cfc_cont_tac) +include hf hg in lemma cfc_sub : cfc (fun x ↦ f x - g x) a = cfc f a - cfc g a := by by_cases ha : p a · rw [cfc_apply f a, cfc_apply g a, ← map_sub, cfc_apply ..] diff --git a/Mathlib/Analysis/CStarAlgebra/Module.lean b/Mathlib/Analysis/CStarAlgebra/Module.lean index c524080d6161f..80dc5201159e3 100644 --- a/Mathlib/Analysis/CStarAlgebra/Module.lean +++ b/Mathlib/Analysis/CStarAlgebra/Module.lean @@ -154,6 +154,7 @@ noncomputable def norm : Norm E where lemma inner_self_eq_norm_sq {x : E} : ‖⟪x, x⟫‖ = ‖x‖ ^ 2 := by simp [norm_eq_sqrt_norm_inner_self] section +include A protected lemma norm_nonneg {x : E} : 0 ≤ ‖x‖ := by simp [norm_eq_sqrt_norm_inner_self]; positivity diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index b488789c0194b..8b829d21486fe 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -445,6 +445,7 @@ open FormalMultilinearSeries ENNReal Nat variable {p : FormalMultilinearSeries 𝕜 E F} {f : E → F} {x : E} {r : ℝ≥0∞} (h : HasFPowerSeriesOnBall f p x r) (y : E) +include h in theorem iteratedFDeriv_zero_apply_diag : iteratedFDeriv 𝕜 0 f x = p 0 := by ext convert (h.hasSum <| EMetric.mem_ball_self h.r_pos).tsum_eq.symm @@ -465,6 +466,7 @@ private theorem factorial_smul' {n : ℕ} : ∀ {F : Type max u v} [NormedAddCom rfl variable [CompleteSpace F] +include h theorem factorial_smul (n : ℕ) : n ! • p n (fun _ ↦ y) = iteratedFDeriv 𝕜 n f x (fun _ ↦ y) := by diff --git a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean index a3cef0f7858f8..fa52f1fbe6370 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean @@ -56,6 +56,9 @@ variable {E F : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [NormedAddComm {f'' : E →L[ℝ] E →L[ℝ] F} (hf : ∀ x ∈ interior s, HasFDerivAt f (f' x) x) {x : E} (xs : x ∈ s) (hx : HasFDerivWithinAt f' f'' (interior s) x) +section +include s_conv hf xs hx + /-- Assume that `f` is differentiable inside a convex set `s`, and that its derivative `f'` is differentiable at a point `x`. Then, given two vectors `v` and `w` pointing inside `s`, one can Taylor-expand to order two the function `f` on the segment `[x + h v, x + h (v + w)]`, giving a @@ -249,6 +252,8 @@ theorem Convex.second_derivative_within_at_symmetric_of_mem_interior {v w : E} field_simp [LT.lt.ne' hpos, SMul.smul] simpa only [sub_eq_zero] using isLittleO_const_const_iff.1 B +end + /-- If a function is differentiable inside a convex set with nonempty interior, and has a second derivative at a point of this convex set, then this second derivative is symmetric. -/ theorem Convex.second_derivative_within_at_symmetric {s : Set E} (s_conv : Convex ℝ s) diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean index f4db529395d41..48cbff9518b9a 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/Deriv.lean @@ -23,6 +23,7 @@ noncomputable section namespace HasStrictDerivAt variable (f' a : 𝕜) (hf : HasStrictDerivAt f f' a) (hf' : f' ≠ 0) +include hf hf' /-- A function that is inverse to `f` near `a`. -/ abbrev localInverse : 𝕜 → 𝕜 := diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean index 3dffaba885d4e..f81f092709a46 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean @@ -21,11 +21,8 @@ variable {R : Type*} [Semiring R] [Module R F] [SMulCommClass 𝕜 R F] [ContinuousConstSMul R F] {n : ℕ} {x : 𝕜} {s : Set 𝕜} (hx : x ∈ s) (h : UniqueDiffOn 𝕜 s) {f g : 𝕜 → F} -theorem iteratedDerivWithin_add (hf : ContDiffOn 𝕜 n f s) (hg : ContDiffOn 𝕜 n g s) : - iteratedDerivWithin n (f + g) s x = - iteratedDerivWithin n f s x + iteratedDerivWithin n g s x := by - simp_rw [iteratedDerivWithin, iteratedFDerivWithin_add_apply hf hg h hx, - ContinuousMultilinearMap.add_apply] +section +include h theorem iteratedDerivWithin_congr (hfg : Set.EqOn f g s) : Set.EqOn (iteratedDerivWithin n f s) (iteratedDerivWithin n g s) s := by @@ -37,6 +34,14 @@ theorem iteratedDerivWithin_congr (hfg : Set.EqOn f g s) : rw [iteratedDerivWithin_succ this, iteratedDerivWithin_succ this] exact derivWithin_congr (IH hfg) (IH hfg hy) +include hx + +theorem iteratedDerivWithin_add (hf : ContDiffOn 𝕜 n f s) (hg : ContDiffOn 𝕜 n g s) : + iteratedDerivWithin n (f + g) s x = + iteratedDerivWithin n f s x + iteratedDerivWithin n g s x := by + simp_rw [iteratedDerivWithin, iteratedFDerivWithin_add_apply hf hg h hx, + ContinuousMultilinearMap.add_apply] + theorem iteratedDerivWithin_const_add (hn : 0 < n) (c : F) : iteratedDerivWithin n (fun z => c + f z) s x = iteratedDerivWithin n f s x := by obtain ⟨n, rfl⟩ := n.exists_eq_succ_of_ne_zero hn.ne' @@ -82,6 +87,8 @@ theorem iteratedDerivWithin_sub (hf : ContDiffOn 𝕜 n f s) (hg : ContDiffOn rw [sub_eq_add_neg, sub_eq_add_neg, Pi.neg_def, iteratedDerivWithin_add hx h hf hg.neg, iteratedDerivWithin_neg' hx h] +end + theorem iteratedDeriv_const_smul {n : ℕ} {f : 𝕜 → F} (h : ContDiff 𝕜 n f) (c : 𝕜) : iteratedDeriv n (fun x => f (c * x)) = fun x => c ^ n • iteratedDeriv n f (c * x) := by induction n with diff --git a/Mathlib/Analysis/Calculus/MeanValue.lean b/Mathlib/Analysis/Calculus/MeanValue.lean index 0064eb24ae840..d28d3e874e5b0 100644 --- a/Mathlib/Analysis/Calculus/MeanValue.lean +++ b/Mathlib/Analysis/Calculus/MeanValue.lean @@ -678,6 +678,7 @@ variable (f f' : ℝ → ℝ) {a b : ℝ} (hab : a < b) (hfc : ContinuousOn f (I (g g' : ℝ → ℝ) (hgc : ContinuousOn g (Icc a b)) (hgg' : ∀ x ∈ Ioo a b, HasDerivAt g (g' x) x) (hgd : DifferentiableOn ℝ g (Ioo a b)) +include hab hfc hff' hgc hgg' in /-- Cauchy's **Mean Value Theorem**, `HasDerivAt` version. -/ theorem exists_ratio_hasDerivAt_eq_ratio_slope : ∃ c ∈ Ioo a b, (g b - g a) * f' c = (f b - f a) * g' c := by @@ -691,6 +692,7 @@ theorem exists_ratio_hasDerivAt_eq_ratio_slope : rcases exists_hasDerivAt_eq_zero hab hhc hI hhh' with ⟨c, cmem, hc⟩ exact ⟨c, cmem, sub_eq_zero.1 hc⟩ +include hab in /-- Cauchy's **Mean Value Theorem**, extended `HasDerivAt` version. -/ theorem exists_ratio_hasDerivAt_eq_ratio_slope' {lfa lga lfb lgb : ℝ} (hff' : ∀ x ∈ Ioo a b, HasDerivAt f (f' x) x) (hgg' : ∀ x ∈ Ioo a b, HasDerivAt g (g' x) x) @@ -715,6 +717,7 @@ theorem exists_ratio_hasDerivAt_eq_ratio_slope' {lfa lga lfb lgb : ℝ} rcases exists_hasDerivAt_eq_zero' hab hha hhb hhh' with ⟨c, cmem, hc⟩ exact ⟨c, cmem, sub_eq_zero.1 hc⟩ +include hab hfc hff' in /-- Lagrange's Mean Value Theorem, `HasDerivAt` version -/ theorem exists_hasDerivAt_eq_slope : ∃ c ∈ Ioo a b, f' c = (f b - f a) / (b - a) := by obtain ⟨c, cmem, hc⟩ : ∃ c ∈ Ioo a b, (b - a) * f' c = (f b - f a) * 1 := @@ -723,6 +726,7 @@ theorem exists_hasDerivAt_eq_slope : ∃ c ∈ Ioo a b, f' c = (f b - f a) / (b use c, cmem rwa [mul_one, mul_comm, ← eq_div_iff (sub_ne_zero.2 hab.ne')] at hc +include hab hfc hgc hgd hfd in /-- Cauchy's Mean Value Theorem, `deriv` version. -/ theorem exists_ratio_deriv_eq_ratio_slope : ∃ c ∈ Ioo a b, (g b - g a) * deriv f c = (f b - f a) * deriv g c := @@ -731,6 +735,7 @@ theorem exists_ratio_deriv_eq_ratio_slope : (deriv g) hgc fun x hx => ((hgd x hx).differentiableAt <| IsOpen.mem_nhds isOpen_Ioo hx).hasDerivAt +include hab in /-- Cauchy's Mean Value Theorem, extended `deriv` version. -/ theorem exists_ratio_deriv_eq_ratio_slope' {lfa lga lfb lgb : ℝ} (hdf : DifferentiableOn ℝ f <| Ioo a b) (hdg : DifferentiableOn ℝ g <| Ioo a b) @@ -741,11 +746,13 @@ theorem exists_ratio_deriv_eq_ratio_slope' {lfa lga lfb lgb : ℝ} (fun x hx => ((hdf x hx).differentiableAt <| Ioo_mem_nhds hx.1 hx.2).hasDerivAt) (fun x hx => ((hdg x hx).differentiableAt <| Ioo_mem_nhds hx.1 hx.2).hasDerivAt) hfa hga hfb hgb +include hab hfc hfd in /-- Lagrange's **Mean Value Theorem**, `deriv` version. -/ theorem exists_deriv_eq_slope : ∃ c ∈ Ioo a b, deriv f c = (f b - f a) / (b - a) := exists_hasDerivAt_eq_slope f (deriv f) hab hfc fun x hx => ((hfd x hx).differentiableAt <| IsOpen.mem_nhds isOpen_Ioo hx).hasDerivAt +include hab hfc hfd in /-- Lagrange's **Mean Value Theorem**, `deriv` version. -/ theorem exists_deriv_eq_slope' : ∃ c ∈ Ioo a b, deriv f c = slope f a b := by rw [slope_def_field] diff --git a/Mathlib/Analysis/Calculus/Rademacher.lean b/Mathlib/Analysis/Calculus/Rademacher.lean index 5d94dc5f85717..cbca99bb29e9a 100644 --- a/Mathlib/Analysis/Calculus/Rademacher.lean +++ b/Mathlib/Analysis/Calculus/Rademacher.lean @@ -46,10 +46,10 @@ open Filter MeasureTheory Measure FiniteDimensional Metric Set Asymptotics open scoped NNReal ENNReal Topology -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [MeasurableSpace E] [BorelSpace E] {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] {C D : ℝ≥0} {f g : E → ℝ} {s : Set E} - {μ : Measure E} [IsAddHaarMeasure μ] + {μ : Measure E} namespace LipschitzWith @@ -60,7 +60,16 @@ This follows from the one-dimensional result that a Lipschitz function on `ℝ` variation, and is therefore ae differentiable, together with a Fubini argument. -/ -theorem ae_lineDifferentiableAt (hf : LipschitzWith C f) (v : E) : + +theorem memℒp_lineDeriv (hf : LipschitzWith C f) (v : E) : + Memℒp (fun x ↦ lineDeriv ℝ f x v) ∞ μ := + memℒp_top_of_bound (aestronglyMeasurable_lineDeriv hf.continuous μ) + (C * ‖v‖) (eventually_of_forall (fun _x ↦ norm_lineDeriv_le_of_lipschitz ℝ hf)) + +variable [FiniteDimensional ℝ E] [IsAddHaarMeasure μ] + +theorem ae_lineDifferentiableAt + (hf : LipschitzWith C f) (v : E) : ∀ᵐ p ∂μ, LineDifferentiableAt ℝ f p v := by let L : ℝ →L[ℝ] E := ContinuousLinearMap.smulRight (1 : ℝ →L[ℝ] ℝ) v suffices A : ∀ p, ∀ᵐ (t : ℝ) ∂volume, LineDifferentiableAt ℝ f (p + t • v) v from @@ -76,11 +85,6 @@ theorem ae_lineDifferentiableAt (hf : LipschitzWith C f) (v : E) : convert h's.comp 0 this with _ t simp only [LineDifferentiableAt, add_assoc, Function.comp_apply, add_smul] -theorem memℒp_lineDeriv (hf : LipschitzWith C f) (v : E) : - Memℒp (fun x ↦ lineDeriv ℝ f x v) ∞ μ := - memℒp_top_of_bound (aestronglyMeasurable_lineDeriv hf.continuous μ) - (C * ‖v‖) (eventually_of_forall (fun _x ↦ norm_lineDeriv_le_of_lipschitz ℝ hf)) - theorem locallyIntegrable_lineDeriv (hf : LipschitzWith C f) (v : E) : LocallyIntegrable (fun x ↦ lineDeriv ℝ f x v) μ := (hf.memℒp_lineDeriv v).locallyIntegrable le_top @@ -254,7 +258,10 @@ theorem ae_exists_fderiv_of_countable /-- If a Lipschitz functions has line derivatives in a dense set of directions, all of them given by a single continuous linear map `L`, then it admits `L` as Fréchet derivative. -/ -theorem hasFderivAt_of_hasLineDerivAt_of_closure {f : E → F} +-- We redeclare `E` here as we do not need the `[MeasurableSpace E]` instance +-- available in the rest of the file. +theorem hasFderivAt_of_hasLineDerivAt_of_closure + {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [ProperSpace E] {f : E → F} (hf : LipschitzWith C f) {s : Set E} (hs : sphere 0 1 ⊆ closure s) {L : E →L[ℝ] F} {x : E} (hL : ∀ v ∈ s, HasLineDerivAt ℝ f (L v) x v) : HasFDerivAt f L x := by @@ -321,7 +328,7 @@ theorem ae_differentiableAt_of_real (hf : LipschitzWith C f) : end LipschitzWith -variable [FiniteDimensional ℝ F] +variable [FiniteDimensional ℝ E] [FiniteDimensional ℝ F] [IsAddHaarMeasure μ] namespace LipschitzOnWith diff --git a/Mathlib/Analysis/Complex/Angle.lean b/Mathlib/Analysis/Complex/Angle.lean index 063d48f81fcec..95b1c80bb289e 100644 --- a/Mathlib/Analysis/Complex/Angle.lean +++ b/Mathlib/Analysis/Complex/Angle.lean @@ -79,7 +79,6 @@ up to a factor of `π / 2`. /-- Chord-length is a multiple of arc-length up to constants. -/ lemma norm_sub_mem_Icc_angle (hx : ‖x‖ = 1) (hy : ‖y‖ = 1) : ‖x - y‖ ∈ Icc (2 / π * angle x y) (angle x y) := by - clear a wlog h : y = 1 · have := @this (x / y) 1 (by simp only [norm_div, hx, hy, div_one]) norm_one rfl rwa [angle_div_left_eq_angle_mul_right, div_sub_one, norm_div, hy, div_one, one_mul] diff --git a/Mathlib/Analysis/Complex/Hadamard.lean b/Mathlib/Analysis/Complex/Hadamard.lean index 18b2355e35748..663f0d998f29a 100644 --- a/Mathlib/Analysis/Complex/Hadamard.lean +++ b/Mathlib/Analysis/Complex/Hadamard.lean @@ -72,7 +72,7 @@ noncomputable def sSupNormIm {E : Type*} [NormedAddCommGroup E] section invInterpStrip -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] (f : ℂ → E) (z : ℂ) +variable {E : Type*} [NormedAddCommGroup E] (f : ℂ → E) (z : ℂ) /-- The inverse of the interpolation of `sSupNormIm` on the two boundaries. @@ -85,7 +85,7 @@ noncomputable def invInterpStrip (ε : ℝ) : ℂ := (ε + sSupNormIm f 0) ^ (z - 1) * (ε + sSupNormIm f 1) ^ (-z) /-- A function useful for the proofs steps. We will aim to show that it is bounded by 1. -/ -noncomputable def F (ε : ℝ) := fun z ↦ invInterpStrip f z ε • f z +noncomputable def F [NormedSpace ℂ E] (ε : ℝ) := fun z ↦ invInterpStrip f z ε • f z /-- `sSup` of `norm` is nonneg applied to the image of `f` on the vertical line `re z = x` -/ lemma sSupNormIm_nonneg (x : ℝ) : 0 ≤ sSupNormIm f x := by @@ -137,6 +137,8 @@ lemma norm_lt_sSupNormIm_eps (f : ℂ → E) (ε : ℝ) (hε : ε > 0) (z : ℂ) ‖f z‖ < ε + sSupNormIm f (z.re) := lt_add_of_pos_of_le hε (norm_le_sSupNormIm f z hD hB) +variable [NormedSpace ℂ E] + /-- When the function `f` is bounded above on a vertical strip, then so is `F`. -/ lemma F_BddAbove (f : ℂ → E) (ε : ℝ) (hε : ε > 0) (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) : @@ -227,7 +229,7 @@ end invInterpStrip ----- -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] (f : ℂ → E) +variable {E : Type*} [NormedAddCommGroup E] (f : ℂ → E) /-- The interpolation of `sSupNormIm` on the two boundaries. @@ -294,6 +296,8 @@ lemma diffContOnCl_interpStrip : · apply differentiableAt_id' · left; simp only [Ne, ofReal_eq_zero]; rwa [eq_comm] +variable [NormedSpace ℂ E] + lemma norm_le_interpStrip_of_mem_verticalClosedStrip_eps (ε : ℝ) (hε : ε > 0) (z : ℂ) (hB : BddAbove ((norm ∘ f) '' (verticalClosedStrip 0 1))) (hd : DiffContOnCl ℂ f (verticalStrip 0 1)) (hz : z ∈ verticalClosedStrip 0 1) : diff --git a/Mathlib/Analysis/Complex/TaylorSeries.lean b/Mathlib/Analysis/Complex/TaylorSeries.lean index a764eeee99d6f..90a8b74ad8c80 100644 --- a/Mathlib/Analysis/Complex/TaylorSeries.lean +++ b/Mathlib/Analysis/Complex/TaylorSeries.lean @@ -35,6 +35,7 @@ section ball variable ⦃c : ℂ⦄ ⦃r : ℝ⦄ (hf : DifferentiableOn ℂ f (Metric.ball c r)) variable ⦃z : ℂ⦄ (hz : z ∈ Metric.ball c r) +include hf hz in /-- A function that is complex differentiable on the open ball of radius `r` around `c` is given by evaluating its Taylor series at `c` on this open ball. -/ lemma hasSum_taylorSeries_on_ball : @@ -54,12 +55,14 @@ lemma hasSum_taylorSeries_on_ball : Finset.card_fin] using ((iteratedFDeriv ℂ n f c).map_smul_univ (fun _ ↦ z - c) (fun _ ↦ 1)).symm +include hf hz in /-- A function that is complex differentiable on the open ball of radius `r` around `c` is given by evaluating its Taylor series at `c` on this open ball. -/ lemma taylorSeries_eq_on_ball : ∑' n : ℕ, (n ! : ℂ)⁻¹ • (z - c) ^ n • iteratedDeriv n f c = f z := (hasSum_taylorSeries_on_ball hf hz).tsum_eq +include hz in /-- A function that is complex differentiable on the open ball of radius `r` around `c` is given by evaluating its Taylor series at `c` on this open ball. -/ lemma taylorSeries_eq_on_ball' {f : ℂ → ℂ} (hf : DifferentiableOn ℂ f (Metric.ball c r)) : @@ -74,6 +77,7 @@ section emetric variable ⦃c : ℂ⦄ ⦃r : ENNReal⦄ (hf : DifferentiableOn ℂ f (EMetric.ball c r)) variable ⦃z : ℂ⦄ (hz : z ∈ EMetric.ball c r) +include hf hz in /-- A function that is complex differentiable on the open ball of radius `r ≤ ∞` around `c` is given by evaluating its Taylor series at `c` on this open ball. -/ lemma hasSum_taylorSeries_on_emetric_ball : @@ -85,12 +89,14 @@ lemma hasSum_taylorSeries_on_emetric_ball : rw [← Metric.emetric_ball_nnreal] exact hf.mono <| EMetric.ball_subset_ball hr'.le +include hf hz in /-- A function that is complex differentiable on the open ball of radius `r ≤ ∞` around `c` is given by evaluating its Taylor series at `c` on this open ball. -/ lemma taylorSeries_eq_on_emetric_ball : ∑' n : ℕ, (n ! : ℂ)⁻¹ • (z - c) ^ n • iteratedDeriv n f c = f z := (hasSum_taylorSeries_on_emetric_ball hf hz).tsum_eq +include hz in /-- A function that is complex differentiable on the open ball of radius `r ≤ ∞` around `c` is given by evaluating its Taylor series at `c` on this open ball. -/ lemma taylorSeries_eq_on_emetric_ball' {f : ℂ → ℂ} (hf : DifferentiableOn ℂ f (EMetric.ball c r)) : @@ -104,6 +110,7 @@ section entire variable ⦃f : ℂ → E⦄ (hf : Differentiable ℂ f) (c z : ℂ) +include hf in /-- A function that is complex differentiable on the complex plane is given by evaluating its Taylor series at any point `c`. -/ lemma hasSum_taylorSeries_of_entire : @@ -111,6 +118,7 @@ lemma hasSum_taylorSeries_of_entire : hasSum_taylorSeries_on_emetric_ball hf.differentiableOn <| EMetric.mem_ball.mpr <| edist_lt_top .. +include hf in /-- A function that is complex differentiable on the complex plane is given by evaluating its Taylor series at any point `c`. -/ lemma taylorSeries_eq_of_entire : diff --git a/Mathlib/Analysis/Complex/Tietze.lean b/Mathlib/Analysis/Complex/Tietze.lean index a6c38c9091e65..09efc63e6ae39 100644 --- a/Mathlib/Analysis/Complex/Tietze.lean +++ b/Mathlib/Analysis/Complex/Tietze.lean @@ -93,11 +93,12 @@ theorem Metric.instTietzeExtensionClosedBall (𝕜 : Type v) [RCLike 𝕜] {E : exact (mul_le_iff_le_one_right hr).symm variable {X : Type u} [TopologicalSpace X] [NormalSpace X] {s : Set X} (hs : IsClosed s) -variable (𝕜 : Type v) [RCLike 𝕜] [TietzeExtension.{u, v} 𝕜] +variable (𝕜 : Type v) [RCLike 𝕜] variable {E : Type w} [NormedAddCommGroup E] [NormedSpace 𝕜 E] [FiniteDimensional 𝕜 E] namespace BoundedContinuousFunction +include 𝕜 hs in /-- **Tietze extension theorem** for real-valued bounded continuous maps, a version with a closed embedding and bundled composition. If `e : C(X, Y)` is a closed embedding of a topological space into a normal topological space and `f : X →ᵇ ℝ` is a bounded continuous function, then there exists diff --git a/Mathlib/Analysis/Convex/Function.lean b/Mathlib/Analysis/Convex/Function.lean index 261e230a5eb97..d3bd49537ee13 100644 --- a/Mathlib/Analysis/Convex/Function.lean +++ b/Mathlib/Analysis/Convex/Function.lean @@ -401,9 +401,6 @@ theorem LinearOrder.convexOn_of_lt (hs : Convex 𝕜 s) f (a • x + b • y) ≤ a • f x + b • f y) : ConvexOn 𝕜 s f := by refine convexOn_iff_pairwise_pos.2 ⟨hs, fun x hx y hy hxy a b ha hb hab => ?_⟩ - -- Porting note: without clearing the stray variables, `wlog` gives a bad term. - -- See https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/wlog.20.2316495 - clear! α F ι wlog h : x < y · rw [add_comm (a • x), add_comm (a • f x)] rw [add_comm] at hab @@ -429,9 +426,6 @@ theorem LinearOrder.strictConvexOn_of_lt (hs : Convex 𝕜 s) f (a • x + b • y) < a • f x + b • f y) : StrictConvexOn 𝕜 s f := by refine ⟨hs, fun x hx y hy hxy a b ha hb hab => ?_⟩ - -- Porting note: without clearing the stray variables, `wlog` gives a bad term. - -- See https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/wlog.20.2316495 - clear! α F ι wlog h : x < y · rw [add_comm (a • x), add_comm (a • f x)] rw [add_comm] at hab diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 5b93a0b901333..2d2ee1ed85185 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -573,8 +573,7 @@ lemma _root_.ContinuousLinearMap.hasTemperateGrowth (f : E →L[ℝ] F) : simpa [this] using .const _ · exact (f.le_opNorm x).trans (by simp [mul_add]) -variable [NormedAddCommGroup D] [NormedSpace ℝ D] -variable [MeasurableSpace D] [BorelSpace D] [SecondCountableTopology D] [FiniteDimensional ℝ D] +variable [NormedAddCommGroup D] [MeasurableSpace D] open MeasureTheory FiniteDimensional @@ -597,6 +596,7 @@ lemma integrable_pow_neg_integrablePower instance _root_.MeasureTheory.Measure.IsFiniteMeasure.instHasTemperateGrowth {μ : Measure D} [h : IsFiniteMeasure μ] : μ.HasTemperateGrowth := ⟨⟨0, by simp⟩⟩ +variable [NormedSpace ℝ D] [FiniteDimensional ℝ D] [BorelSpace D] in instance _root_.MeasureTheory.Measure.IsAddHaarMeasure.instHasTemperateGrowth {μ : Measure D} [h : μ.IsAddHaarMeasure] : μ.HasTemperateGrowth := ⟨⟨finrank ℝ D + 1, by apply integrable_one_add_norm; norm_num⟩⟩ @@ -631,11 +631,14 @@ lemma pow_mul_le_of_le_of_pow_mul_le {C₁ C₂ : ℝ} {k l : ℕ} {x f : ℝ} ( · exact Real.rpow_le_rpow_of_nonpos (by linarith) (by linarith) (by simp) · exact h₂.trans (by linarith) +variable [BorelSpace D] [SecondCountableTopology D] in /-- Given a function such that `f` and `x ^ (k + l) * f` are bounded for a suitable `l`, then `x ^ k * f` is integrable. The bounds are not relevant for the integrability conclusion, but they are relevant for bounding the integral in `integral_pow_mul_le_of_le_of_pow_mul_le`. We formulate the two lemmas with the same set of assumptions for ease of applications. -/ +-- We redeclare `E` here to avoid the `NormedSpace ℝ E` typeclass available throughout this file. lemma integrable_of_le_of_pow_mul_le + {E : Type*} [NormedAddCommGroup E] {μ : Measure D} [μ.HasTemperateGrowth] {f : D → E} {C₁ C₂ : ℝ} {k : ℕ} (hf : ∀ x, ‖f x‖ ≤ C₁) (h'f : ∀ x, ‖x‖ ^ (k + μ.integrablePower) * ‖f x‖ ≤ C₂) (h''f : AEStronglyMeasurable f μ) : @@ -648,7 +651,9 @@ lemma integrable_of_le_of_pow_mul_le /-- Given a function such that `f` and `x ^ (k + l) * f` are bounded for a suitable `l`, then one can bound explicitly the integral of `x ^ k * f`. -/ +-- We redeclare `E` here to avoid the `NormedSpace ℝ E` typeclass available throughout this file. lemma integral_pow_mul_le_of_le_of_pow_mul_le + {E : Type*} [NormedAddCommGroup E] {μ : Measure D} [μ.HasTemperateGrowth] {f : D → E} {C₁ C₂ : ℝ} {k : ℕ} (hf : ∀ x, ‖f x‖ ≤ C₁) (h'f : ∀ x, ‖x‖ ^ (k + μ.integrablePower) * ‖f x‖ ≤ C₂) : ∫ x, ‖x‖ ^ k * ‖f x‖ ∂μ ≤ 2 ^ μ.integrablePower * @@ -1034,14 +1039,25 @@ open Real Complex Filter MeasureTheory MeasureTheory.Measure FiniteDimensional variable [RCLike 𝕜] variable [NormedAddCommGroup D] [NormedSpace ℝ D] variable [NormedAddCommGroup V] [NormedSpace ℝ V] [NormedSpace 𝕜 V] -variable [MeasurableSpace D] [BorelSpace D] [SecondCountableTopology D] +variable [MeasurableSpace D] variable {μ : Measure D} [hμ : HasTemperateGrowth μ] attribute [local instance 101] secondCountableTopologyEither_of_left +variable (𝕜 μ) in +lemma integral_pow_mul_iteratedFDeriv_le (f : 𝓢(D, V)) (k n : ℕ) : + ∫ x, ‖x‖ ^ k * ‖iteratedFDeriv ℝ n f x‖ ∂μ ≤ 2 ^ μ.integrablePower * + (∫ x, (1 + ‖x‖) ^ (- (μ.integrablePower : ℝ)) ∂μ) * + (SchwartzMap.seminorm 𝕜 0 n f + SchwartzMap.seminorm 𝕜 (k + μ.integrablePower) n f) := + integral_pow_mul_le_of_le_of_pow_mul_le (norm_iteratedFDeriv_le_seminorm ℝ _ _) + (le_seminorm ℝ _ _ _) + +variable [BorelSpace D] [SecondCountableTopology D] + variable (μ) in -lemma integrable_pow_mul_iteratedFDeriv (f : 𝓢(D, V)) +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 @@ -1052,14 +1068,6 @@ lemma integrable_pow_mul (f : 𝓢(D, V)) convert integrable_pow_mul_iteratedFDeriv μ f k 0 with x simp -variable (𝕜 μ) in -lemma integral_pow_mul_iteratedFDeriv_le (f : 𝓢(D, V)) (k n : ℕ) : - ∫ x, ‖x‖ ^ k * ‖iteratedFDeriv ℝ n f x‖ ∂μ ≤ 2 ^ μ.integrablePower * - (∫ x, (1 + ‖x‖) ^ (- (μ.integrablePower : ℝ)) ∂μ) * - (SchwartzMap.seminorm 𝕜 0 n f + SchwartzMap.seminorm 𝕜 (k + μ.integrablePower) n f) := - integral_pow_mul_le_of_le_of_pow_mul_le (norm_iteratedFDeriv_le_seminorm ℝ _ _) - (le_seminorm ℝ _ _ _) - lemma integrable (f : 𝓢(D, V)) : Integrable f μ := (f.integrable_pow_mul μ 0).mono f.continuous.aestronglyMeasurable (eventually_of_forall (fun _ ↦ by simp)) diff --git a/Mathlib/Analysis/Fourier/AddCircle.lean b/Mathlib/Analysis/Fourier/AddCircle.lean index 45761b766817c..3de5271afa8a5 100644 --- a/Mathlib/Analysis/Fourier/AddCircle.lean +++ b/Mathlib/Analysis/Fourier/AddCircle.lean @@ -274,7 +274,7 @@ variable [hT : Fact (0 < T)] section fourierCoeff -variable {E : Type} [NormedAddCommGroup E] [NormedSpace ℂ E] [CompleteSpace E] +variable {E : Type} [NormedAddCommGroup E] [NormedSpace ℂ E] /-- The `n`-th Fourier coefficient of a function `AddCircle T → E`, for `E` a complete normed `ℂ`-vector space, defined as the integral over `AddCircle T` of `fourier (-n) t • f t`. -/ diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 6b2938e160a15..f5aa42ed27d81 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -1133,12 +1133,14 @@ theorem real_inner_le_norm (x y : F) : ⟪x, y⟫_ℝ ≤ ‖x‖ * ‖y‖ := variable (𝕜) +include 𝕜 in theorem parallelogram_law_with_norm (x y : E) : ‖x + y‖ * ‖x + y‖ + ‖x - y‖ * ‖x - y‖ = 2 * (‖x‖ * ‖x‖ + ‖y‖ * ‖y‖) := by simp only [← @inner_self_eq_norm_mul_norm 𝕜] rw [← re.map_add, parallelogram_law, two_mul, two_mul] simp only [re.map_add] +include 𝕜 in theorem parallelogram_law_with_nnnorm (x y : E) : ‖x + y‖₊ * ‖x + y‖₊ + ‖x - y‖₊ * ‖x - y‖₊ = 2 * (‖x‖₊ * ‖x‖₊ + ‖y‖₊ * ‖y‖₊) := Subtype.ext <| parallelogram_law_with_norm 𝕜 x y @@ -1993,19 +1995,24 @@ def OrthogonalFamily (G : ι → Type*) [∀ i, SeminormedAddCommGroup (G i)] variable {𝕜} 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)] + {V : ∀ i, G i →ₗᵢ[𝕜] E} theorem Orthonormal.orthogonalFamily {v : ι → E} (hv : Orthonormal 𝕜 v) : OrthogonalFamily 𝕜 (fun _i : ι => 𝕜) fun i => LinearIsometry.toSpanSingleton 𝕜 E (hv.1 i) := fun i j hij a b => by simp [inner_smul_left, inner_smul_right, hv.2 hij] +section +variable (hV : OrthogonalFamily 𝕜 G V) +include hV + theorem OrthogonalFamily.eq_ite [DecidableEq ι] {i j : ι} (v : G i) (w : G j) : ⟪V i v, V j w⟫ = ite (i = j) ⟪V i v, V j w⟫ 0 := by split_ifs with h · rfl · exact hV h v w -theorem OrthogonalFamily.inner_right_dfinsupp [DecidableEq ι] (l : ⨁ i, G i) (i : ι) (v : G i) : +theorem OrthogonalFamily.inner_right_dfinsupp + [∀ (i) (x : G i), Decidable (x ≠ 0)] [DecidableEq ι] (l : ⨁ i, G i) (i : ι) (v : G i) : ⟪V i v, l.sum fun j => V j⟫ = ⟪v, l i⟫ := calc ⟪V i v, l.sum fun j => V j⟫ = l.sum fun j => fun w => ⟪V i v, V j w⟫ := @@ -2137,6 +2144,8 @@ theorem OrthogonalFamily.summable_iff_norm_sq_summable [CompleteSpace E] (f : · exact fun i => sq_nonneg _ linarith +end + end OrthogonalFamily_Seminormed section OrthogonalFamily diff --git a/Mathlib/Analysis/InnerProductSpace/Calculus.lean b/Mathlib/Analysis/InnerProductSpace/Calculus.lean index eae31940da226..031592910f445 100644 --- a/Mathlib/Analysis/InnerProductSpace/Calculus.lean +++ b/Mathlib/Analysis/InnerProductSpace/Calculus.lean @@ -124,6 +124,9 @@ theorem deriv_inner_apply {f g : ℝ → E} {x : ℝ} (hf : DifferentiableAt ℝ deriv (fun t => ⟪f t, g t⟫) x = ⟪f x, deriv g x⟫ + ⟪deriv f x, g x⟫ := (hf.hasDerivAt.inner 𝕜 hg.hasDerivAt).deriv +section +include 𝕜 + theorem contDiff_norm_sq : ContDiff ℝ n fun x : E => ‖x‖ ^ 2 := by convert (reCLM : 𝕜 →L[ℝ] ℝ).contDiff.comp ((contDiff_id (E := E)).inner 𝕜 (contDiff_id (E := E))) exact (inner_self_eq_norm_sq _).symm @@ -176,6 +179,8 @@ theorem ContDiff.dist (hf : ContDiff ℝ n f) (hg : ContDiff ℝ n g) (hne : ∀ ContDiff ℝ n fun y => dist (f y) (g y) := contDiff_iff_contDiffAt.2 fun x => hf.contDiffAt.dist 𝕜 hg.contDiffAt (hne x) +end + -- Porting note: use `2 •` instead of `bit0` theorem hasStrictFDerivAt_norm_sq (x : F) : HasStrictFDerivAt (fun x => ‖x‖ ^ 2) (2 • (innerSL ℝ x)) x := by @@ -201,6 +206,9 @@ theorem HasDerivWithinAt.norm_sq {f : ℝ → F} {f' : F} {s : Set ℝ} {x : ℝ HasDerivWithinAt (‖f ·‖ ^ 2) (2 * Inner.inner (f x) f') s x := by simpa using hf.hasFDerivWithinAt.norm_sq.hasDerivWithinAt +section +include 𝕜 + theorem DifferentiableAt.norm_sq (hf : DifferentiableAt ℝ f x) : DifferentiableAt ℝ (fun y => ‖f y‖ ^ 2) x := ((contDiffAt_id.norm_sq 𝕜).differentiableAt le_rfl).comp x hf @@ -247,6 +255,8 @@ theorem DifferentiableOn.dist (hf : DifferentiableOn ℝ f s) (hg : Differentiab (hne : ∀ x ∈ s, f x ≠ g x) : DifferentiableOn ℝ (fun y => dist (f y) (g y)) s := fun x hx => (hf x hx).dist 𝕜 (hg x hx) (hne x hx) +end + end DerivInner section PiLike diff --git a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean index 1d532e503222f..9019b62655b0c 100644 --- a/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean +++ b/Mathlib/Analysis/InnerProductSpace/LinearPMap.lean @@ -170,14 +170,17 @@ theorem adjoint_apply_of_dense (y : T†.domain) : T† y = adjointAux hT y := b change (if hT : Dense (T.domain : Set E) then adjointAux hT else 0) y = _ simp only [hT, dif_pos, LinearMap.coe_mk] +include hT in theorem adjoint_apply_eq (y : T†.domain) {x₀ : E} (hx₀ : ∀ x : T.domain, ⟪x₀, x⟫ = ⟪(y : F), T x⟫) : T† y = x₀ := (adjoint_apply_of_dense hT y).symm ▸ adjointAux_unique hT _ hx₀ +include hT in /-- The fundamental property of the adjoint. -/ theorem adjoint_isFormalAdjoint : T†.IsFormalAdjoint T := fun x => (adjoint_apply_of_dense hT x).symm ▸ adjointAux_inner hT x +include hT in /-- The adjoint is maximal in the sense that it contains every formal adjoint. -/ theorem IsFormalAdjoint.le_adjoint (h : T.IsFormalAdjoint S) : S ≤ T† := ⟨-- Trivially, every `x : S.domain` is in `T.adjoint.domain` diff --git a/Mathlib/Analysis/InnerProductSpace/l2Space.lean b/Mathlib/Analysis/InnerProductSpace/l2Space.lean index a93c696be81f4..4552de407ddb4 100644 --- a/Mathlib/Analysis/InnerProductSpace/l2Space.lean +++ b/Mathlib/Analysis/InnerProductSpace/l2Space.lean @@ -85,7 +85,7 @@ open scoped NNReal ENNReal ComplexConjugate Topology noncomputable section variable {ι 𝕜 : Type*} [RCLike 𝕜] {E : Type*} -variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] [cplt : CompleteSpace E] +variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] variable {G : ι → Type*} [∀ i, NormedAddCommGroup (G i)] [∀ i, InnerProductSpace 𝕜 (G i)] local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y @@ -175,9 +175,11 @@ end lp namespace OrthogonalFamily -variable {V : ∀ i, G i →ₗᵢ[𝕜] E} (hV : OrthogonalFamily 𝕜 G V) +variable [CompleteSpace E] {V : ∀ i, G i →ₗᵢ[𝕜] E} (hV : OrthogonalFamily 𝕜 G V) +include hV -protected theorem summable_of_lp (f : lp G 2) : Summable fun i => V i (f i) := by +protected theorem summable_of_lp (f : lp G 2) : + Summable fun i => V i (f i) := by rw [hV.summable_iff_norm_sq_summable] convert (lp.memℓp f).summable _ · norm_cast @@ -254,7 +256,7 @@ end OrthogonalFamily section IsHilbertSum variable (𝕜 G) -variable (V : ∀ i, G i →ₗᵢ[𝕜] E) (F : ι → Submodule 𝕜 E) +variable [CompleteSpace E] (V : ∀ i, G i →ₗᵢ[𝕜] E) (F : ι → Submodule 𝕜 E) /-- Given a family of Hilbert spaces `G : ι → Type*`, a Hilbert sum of `G` consists of a Hilbert space `E` and an orthogonal family `V : Π i, G i →ₗᵢ[𝕜] E` such that the induced isometry @@ -492,7 +494,11 @@ theorem finite_spans_dense [DecidableEq E] (b : HilbertBasis ι 𝕜 E) : Set.mem_iUnion_of_mem {i} <| Finset.mem_coe.mpr <| Finset.mem_image_of_mem _ <| Finset.mem_singleton_self i)) +variable [CompleteSpace E] + +section variable {v : ι → E} (hv : Orthonormal 𝕜 v) +include hv /-- An orthonormal family of vectors whose span is dense in the whole module is a Hilbert basis. -/ protected def mk (hsp : ⊤ ≤ (span 𝕜 (Set.range v)).topologicalClosure) : HilbertBasis ι 𝕜 E := @@ -529,6 +535,8 @@ protected def _root_.OrthonormalBasis.toHilbertBasis [Fintype ι] (b : Orthonorm simpa only [← OrthonormalBasis.coe_toBasis, b.toBasis.span_eq, eq_top_iff] using @subset_closure E _ _ +end + @[simp] theorem _root_.OrthonormalBasis.coe_toHilbertBasis [Fintype ι] (b : OrthonormalBasis ι 𝕜 E) : (b.toHilbertBasis : ι → E) = b := diff --git a/Mathlib/Analysis/LocallyConvex/Barrelled.lean b/Mathlib/Analysis/LocallyConvex/Barrelled.lean index ea40aa6ca9fd7..18409cff2a2f8 100644 --- a/Mathlib/Analysis/LocallyConvex/Barrelled.lean +++ b/Mathlib/Analysis/LocallyConvex/Barrelled.lean @@ -146,8 +146,9 @@ instance BaireSpace.instBarrelledSpace [TopologicalSpace E] [TopologicalAddGroup namespace WithSeminorms variable [UniformSpace E] [UniformSpace F] [UniformAddGroup E] [UniformAddGroup F] - [ContinuousSMul 𝕜₁ E] [ContinuousSMul 𝕜₂ F] [BarrelledSpace 𝕜₁ E] {𝓕 : ι → E →SL[σ₁₂] F} + [ContinuousSMul 𝕜₁ E] [BarrelledSpace 𝕜₁ E] {𝓕 : ι → E →SL[σ₁₂] F} {q : SeminormFamily 𝕜₂ F κ} (hq : WithSeminorms q) +include hq /-- The **Banach-Steinhaus** theorem, or **Uniform Boundedness Principle**, for maps from a barrelled space to any space whose topology is generated by a family of seminorms. Use @@ -170,6 +171,8 @@ protected theorem banach_steinhaus (H : ∀ k x, BddAbove (range fun i ↦ q k ( exact ⟨this, Seminorm.continuous_iSup _ (fun i ↦ (hq.continuous_seminorm k).comp (𝓕 i).continuous) this⟩ +variable [ContinuousSMul 𝕜₂ F] + /-- Given a sequence of continuous linear maps which converges pointwise and for which the domain is barrelled, the Banach-Steinhaus theorem is used to guarantee that the limit map is a *continuous* linear map as well. diff --git a/Mathlib/Analysis/Normed/Affine/Isometry.lean b/Mathlib/Analysis/Normed/Affine/Isometry.lean index 8fac64e2bb06c..737521f2a9dec 100644 --- a/Mathlib/Analysis/Normed/Affine/Isometry.lean +++ b/Mathlib/Analysis/Normed/Affine/Isometry.lean @@ -636,6 +636,7 @@ theorem coe_constVAdd (v : V) : ⇑(constVAdd 𝕜 P v : P ≃ᵃⁱ[𝕜] P) = theorem constVAdd_zero : constVAdd 𝕜 P (0 : V) = refl 𝕜 P := ext <| zero_vadd V +include 𝕜 in /-- The map `g` from `V` to `V₂` corresponding to a map `f` from `P` to `P₂`, at a base point `p`, is an isometry if `f` is one. -/ theorem vadd_vsub {f : P → P₂} (hf : Isometry f) {p : P} {g : V → V₂} diff --git a/Mathlib/Analysis/Normed/Algebra/Exponential.lean b/Mathlib/Analysis/Normed/Algebra/Exponential.lean index dbed897cdb5f6..f9679a22cb191 100644 --- a/Mathlib/Analysis/Normed/Algebra/Exponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/Exponential.lean @@ -542,6 +542,7 @@ section DivisionAlgebra variable {𝕂 𝔸 : Type*} [RCLike 𝕂] [NormedDivisionRing 𝔸] [NormedAlgebra 𝕂 𝔸] variable (𝕂) +include 𝕂 theorem norm_expSeries_div_summable (x : 𝔸) : Summable fun n => ‖(x ^ n / n ! : 𝔸)‖ := norm_expSeries_div_summable_of_mem_ball 𝕂 x diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index cbe8ee3895f27..b09a207deb0d5 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -426,6 +426,7 @@ end NonemptySpectrum section GelfandMazurIsomorphism variable [NormedRing A] [NormedAlgebra ℂ A] (hA : ∀ {a : A}, IsUnit a ↔ a ≠ 0) +include hA local notation "σ" => spectrum ℂ @@ -689,7 +690,7 @@ lemma real_iff [Module ℂ A] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] {a end QuasispectrumRestricts -variable {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] [PartialOrder A] +variable {A : Type*} [Ring A] [PartialOrder A] lemma coe_mem_spectrum_real_of_nonneg [Algebra ℝ A] [NonnegSpectrumClass ℝ A] {a : A} {x : ℝ≥0} (ha : 0 ≤ a := by cfc_tac) : diff --git a/Mathlib/Analysis/Normed/Module/Basic.lean b/Mathlib/Analysis/Normed/Module/Basic.lean index fe58c989cdb6a..4541f328011b9 100644 --- a/Mathlib/Analysis/Normed/Module/Basic.lean +++ b/Mathlib/Analysis/Normed/Module/Basic.lean @@ -62,7 +62,7 @@ instance NormedField.to_boundedSMul : BoundedSMul 𝕜 𝕜 := NormedSpace.boundedSMul variable (𝕜) in -theorem norm_zsmul [NormedSpace 𝕜 E] (n : ℤ) (x : E) : ‖n • x‖ = ‖(n : 𝕜)‖ * ‖x‖ := by +theorem norm_zsmul (n : ℤ) (x : E) : ‖n • x‖ = ‖(n : 𝕜)‖ * ‖x‖ := by rw [← norm_smul, ← Int.smul_one_eq_cast, smul_assoc, one_smul] theorem eventually_nhds_norm_smul_sub_lt (c : 𝕜) (x : E) {ε : ℝ} (h : 0 < ε) : @@ -178,6 +178,7 @@ section NontriviallyNormedSpace variable (𝕜 E) variable [NontriviallyNormedField 𝕜] [NormedAddCommGroup E] [NormedSpace 𝕜 E] [Nontrivial E] +include 𝕜 /-- If `E` is a nontrivial normed space over a nontrivially normed field `𝕜`, then `E` is unbounded: for any `c : ℝ`, there exists a vector `x : E` with norm strictly greater than `c`. -/ @@ -212,6 +213,7 @@ section NormedSpace variable (𝕜 E) variable [NormedField 𝕜] [Infinite 𝕜] [NormedAddCommGroup E] [Nontrivial E] [NormedSpace 𝕜 E] +include 𝕜 /-- A normed vector space over an infinite normed field is a noncompact space. This cannot be an instance because in order to apply it, diff --git a/Mathlib/Analysis/Normed/Operator/Banach.lean b/Mathlib/Analysis/Normed/Operator/Banach.lean index bfc054cac483e..26a210ee9dcea 100644 --- a/Mathlib/Analysis/Normed/Operator/Banach.lean +++ b/Mathlib/Analysis/Normed/Operator/Banach.lean @@ -22,8 +22,6 @@ open Function Metric Set Filter Finset Topology NNReal open LinearMap (range ker) variable {𝕜 𝕜' : Type*} [NontriviallyNormedField 𝕜] [NontriviallyNormedField 𝕜'] {σ : 𝕜 →+* 𝕜'} - {σ' : 𝕜' →+* 𝕜} [RingHomInvPair σ σ'] [RingHomInvPair σ' σ] [RingHomIsometric σ] - [RingHomIsometric σ'] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜' F] (f : E →SL[σ] F) @@ -53,16 +51,19 @@ theorem NonlinearRightInverse.bound {f : E →SL[σ] F} (fsymm : NonlinearRightI end ContinuousLinearMap +variable {σ' : 𝕜' →+* 𝕜} [RingHomInvPair σ σ'] [RingHomIsometric σ] [RingHomIsometric σ'] + /-- Given a continuous linear equivalence, the inverse is in particular an instance of `ContinuousLinearMap.NonlinearRightInverse` (which turns out to be linear). -/ -noncomputable def ContinuousLinearEquiv.toNonlinearRightInverse (f : E ≃SL[σ] F) : +noncomputable def ContinuousLinearEquiv.toNonlinearRightInverse + [RingHomInvPair σ' σ] (f : E ≃SL[σ] F) : ContinuousLinearMap.NonlinearRightInverse (f : E →SL[σ] F) where toFun := f.invFun nnnorm := ‖(f.symm : F →SL[σ'] E)‖₊ bound' _ := ContinuousLinearMap.le_opNorm (f.symm : F →SL[σ'] E) _ right_inv' := f.apply_symm_apply -noncomputable instance (f : E ≃SL[σ] F) : +noncomputable instance [RingHomInvPair σ' σ] (f : E ≃SL[σ] F) : Inhabited (ContinuousLinearMap.NonlinearRightInverse (f : E →SL[σ] F)) := ⟨f.toNonlinearRightInverse⟩ @@ -73,6 +74,7 @@ variable [CompleteSpace F] namespace ContinuousLinearMap +include σ' in /-- First step of the proof of the Banach open mapping theorem (using completeness of `F`): by Baire's theorem, there exists a ball in `E` whose image closure has nonempty interior. Rescaling everything, it follows that any `y ∈ F` is arbitrarily well approached by @@ -152,6 +154,9 @@ theorem exists_approx_preimage_norm_le (surj : Surjective f) : variable [CompleteSpace E] +section +include σ' + /-- The Banach open mapping theorem: if a bounded linear map between Banach spaces is onto, then any point has a preimage with controlled norm. -/ theorem exists_preimage_norm_le (surj : Surjective f) : @@ -246,6 +251,8 @@ protected theorem isOpenMap (surj : Surjective f) : IsOpenMap f := by protected theorem quotientMap (surj : Surjective f) : QuotientMap f := (f.isOpenMap surj).to_quotientMap f.continuous surj +end + theorem _root_.AffineMap.isOpenMap {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] [CompleteSpace F] {P Q : Type*} [MetricSpace P] [NormedAddTorsor E P] [MetricSpace Q] [NormedAddTorsor F Q] (f : P →ᵃ[𝕜] Q) (hf : Continuous f) (surj : Surjective f) : @@ -256,6 +263,8 @@ theorem _root_.AffineMap.isOpenMap {F : Type*} [NormedAddCommGroup F] [NormedSpa /-! ### Applications of the Banach open mapping theorem -/ +section +include σ' theorem interior_preimage (hsurj : Surjective f) (s : Set F) : interior (f ⁻¹' s) = f ⁻¹' interior s := @@ -279,6 +288,8 @@ theorem exists_nonlinearRightInverse_of_surjective (f : E →SL[σ] F) right_inv' := fun y => (h y).1 } exact hC +end + /-- A surjective continuous linear map between Banach spaces admits a (possibly nonlinear) controlled right inverse. In general, it is not possible to ensure that such a right inverse is linear (take for instance the map from `E` to `E/F` where `F` is a closed subspace of `E` @@ -296,7 +307,7 @@ end ContinuousLinearMap namespace LinearEquiv -variable [CompleteSpace E] +variable [CompleteSpace E] [RingHomInvPair σ' σ] /-- If a bounded linear map is a bijection, then its inverse is also a bounded linear map. -/ @[continuity] @@ -329,7 +340,7 @@ end LinearEquiv namespace ContinuousLinearMap -variable [CompleteSpace E] +variable [CompleteSpace E] [RingHomInvPair σ' σ] /-- An injective continuous linear map with a closed range defines a continuous linear equivalence between its domain and its range. -/ @@ -353,7 +364,7 @@ end ContinuousLinearMap namespace ContinuousLinearEquiv -variable [CompleteSpace E] +variable [CompleteSpace E] [RingHomInvPair σ' σ] /-- Convert a bijective continuous linear map `f : E →SL[σ] F` from a Banach space to a normed space to a continuous linear equivalence. -/ @@ -519,12 +530,22 @@ section BijectivityCriteria namespace ContinuousLinearMap +variable {σ : 𝕜 →+* 𝕜'} {σ' : 𝕜' →+* 𝕜} [RingHomInvPair σ σ'] {f : E →SL[σ] F} +variable {F : Type u_4} [NormedAddCommGroup F] [NormedSpace 𝕜' F] variable [CompleteSpace E] lemma closed_range_of_antilipschitz {f : E →SL[σ] F} {c : ℝ≥0} (hf : AntilipschitzWith c f) : (LinearMap.range f).topologicalClosure = LinearMap.range f := SetLike.ext'_iff.mpr <| (hf.isClosed_range f.uniformContinuous).closure_eq +variable [CompleteSpace F] + +lemma _root_.AntilipschitzWith.completeSpace_range_clm {f : E →SL[σ] F} {c : ℝ≥0} + (hf : AntilipschitzWith c f) : CompleteSpace (LinearMap.range f) := + IsClosed.completeSpace_coe <| hf.isClosed_range f.uniformContinuous + +variable [RingHomInvPair σ' σ] [RingHomIsometric σ] [RingHomIsometric σ'] + open Function lemma bijective_iff_dense_range_and_antilipschitz (f : E →SL[σ] F) : Bijective f ↔ (LinearMap.range f).topologicalClosure = ⊤ ∧ ∃ c, AntilipschitzWith c f := by @@ -535,10 +556,6 @@ lemma bijective_iff_dense_range_and_antilipschitz (f : E →SL[σ] F) : simp only [LinearMap.range_eq_top, LinearMapClass.ker_eq_bot, h.1, h.2] case surj => rwa [← LinearMap.range_eq_top, ← closed_range_of_antilipschitz hf] -lemma _root_.AntilipschitzWith.completeSpace_range_clm {f : E →SL[σ] F} {c : ℝ≥0} - (hf : AntilipschitzWith c f) : CompleteSpace (LinearMap.range f) := - IsClosed.completeSpace_coe <| hf.isClosed_range f.uniformContinuous - end ContinuousLinearMap end BijectivityCriteria diff --git a/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean b/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean index bf8d8abdf6efe..7f54cef7e5812 100644 --- a/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean +++ b/Mathlib/Analysis/Normed/Ring/SeminormFromConst.lean @@ -44,8 +44,7 @@ open scoped Topology section Ring -variable {R : Type _} [CommRing R] (c : R) (f : RingSeminorm R) (hf1 : f 1 ≤ 1) (hc : f c ≠ 0) - (hpm : IsPowMul f) +variable {R : Type _} [CommRing R] (c : R) (f : RingSeminorm R) /-- For a ring seminorm `f` on `R` and `c ∈ R`, the sequence given by `(f (x * c^n))/((f c)^n)`. -/ def seminormFromConst_seq (x : R) : ℕ → ℝ := fun n ↦ f (x * c ^ n) / f c ^ n @@ -73,12 +72,16 @@ theorem seminormFromConst_seq_zero (hf : f 0 = 0) : seminormFromConst_seq c f 0 rw [zero_mul, hf, zero_div, Pi.zero_apply] variable {c} +variable (hf1 : f 1 ≤ 1) (hc : f c ≠ 0) (hpm : IsPowMul f) +include hpm hc /-- If `1 ≤ n`, then `seminormFromConst_seq c f 1 n = 1`. -/ theorem seminormFromConst_seq_one (n : ℕ) (hn : 1 ≤ n) : seminormFromConst_seq c f 1 n = 1 := by simp only [seminormFromConst_seq] rw [one_mul, hpm _ hn, div_self (pow_ne_zero n hc)] +include hf1 + /-- `seminormFromConst_seq c f x` is antitone. -/ theorem seminormFromConst_seq_antitone (x : R) : Antitone (seminormFromConst_seq c f x) := by intro m n hmn diff --git a/Mathlib/Analysis/NormedSpace/BallAction.lean b/Mathlib/Analysis/NormedSpace/BallAction.lean index 83b77c0ba1c7f..262d4babb606c 100644 --- a/Mathlib/Analysis/NormedSpace/BallAction.lean +++ b/Mathlib/Analysis/NormedSpace/BallAction.lean @@ -170,8 +170,10 @@ end SMulCommClass variable (𝕜) variable [CharZero 𝕜] +include 𝕜 in theorem ne_neg_of_mem_sphere {r : ℝ} (hr : r ≠ 0) (x : sphere (0 : E) r) : x ≠ -x := fun h => ne_zero_of_mem_sphere hr x ((self_eq_neg 𝕜 _).mp (by (conv_lhs => rw [h]); rfl)) +include 𝕜 in theorem ne_neg_of_mem_unit_sphere (x : sphere (0 : E) 1) : x ≠ -x := ne_neg_of_mem_sphere 𝕜 one_ne_zero x diff --git a/Mathlib/Analysis/NormedSpace/RCLike.lean b/Mathlib/Analysis/NormedSpace/RCLike.lean index 2237337083a4d..6b6b64a156b06 100644 --- a/Mathlib/Analysis/NormedSpace/RCLike.lean +++ b/Mathlib/Analysis/NormedSpace/RCLike.lean @@ -90,7 +90,7 @@ alias ContinuousLinearMap.op_norm_bound_of_ball_bound := ContinuousLinearMap.opNorm_bound_of_ball_bound variable (𝕜) - +include 𝕜 in theorem NormedSpace.sphere_nonempty_rclike [Nontrivial E] {r : ℝ} (hr : 0 ≤ r) : Nonempty (sphere (0 : E) r) := letI : NormedSpace ℝ E := NormedSpace.restrictScalars ℝ 𝕜 E diff --git a/Mathlib/Analysis/ODE/Gronwall.lean b/Mathlib/Analysis/ODE/Gronwall.lean index d09579fb63b54..243e6fdc3e392 100644 --- a/Mathlib/Analysis/ODE/Gronwall.lean +++ b/Mathlib/Analysis/ODE/Gronwall.lean @@ -130,6 +130,7 @@ theorem norm_le_gronwallBound_of_norm_deriv_right_le {f f' : ℝ → E} {δ K ε variable {v : ℝ → E → E} {s : ℝ → Set E} {K : ℝ≥0} {f g f' g' : ℝ → E} {a b t₀ : ℝ} {εf εg δ : ℝ} (hv : ∀ t, LipschitzOnWith K (v t) (s t)) +include hv in /-- If `f` and `g` are two approximate solutions of the same ODE, then the distance between them can't grow faster than exponentially. This is a simple corollary of Grönwall's inequality, and some people call this Grönwall's inequality too. @@ -177,6 +178,7 @@ theorem dist_le_of_approx_trajectories_ODE dist_le_of_approx_trajectories_ODE_of_mem (fun t => (hv t).lipschitzOnWith _) hf hf' f_bound hfs hg hg' g_bound (fun _ _ => trivial) ha +include hv in /-- If `f` and `g` are two exact solutions of the same ODE, then the distance between them can't grow faster than exponentially. This is a simple corollary of Grönwall's inequality, and some people call this Grönwall's inequality too. @@ -214,6 +216,7 @@ theorem dist_le_of_trajectories_ODE dist_le_of_trajectories_ODE_of_mem (fun t => (hv t).lipschitzOnWith _) hf hf' hfs hg hg' (fun _ _ => trivial) ha +include hv in /-- There exists only one solution of an ODE \(\dot x=v(t, x)\) in a set `s ⊆ ℝ × E` with a given initial value provided that the RHS is Lipschitz continuous in `x` within `s`, and we consider only solutions included in `s`. @@ -231,6 +234,7 @@ theorem ODE_solution_unique_of_mem_Icc_right have := dist_le_of_trajectories_ODE_of_mem hv hf hf' hfs hg hg' hgs (dist_le_zero.2 ha) t ht rwa [zero_mul, dist_le_zero] at this +include hv in /-- A time-reversed version of `ODE_solution_unique_of_mem_Icc_right`. Uniqueness is shown in a closed interval `Icc a b`, where `b` is the "initial" time. -/ theorem ODE_solution_unique_of_mem_Icc_left @@ -267,6 +271,7 @@ theorem ODE_solution_unique_of_mem_Icc_left (hasDerivAt_neg t).hasDerivWithinAt (hmt3 t) simp +include hv in /-- A version of `ODE_solution_unique_of_mem_Icc_right` for uniqueness in a closed interval whose interior contains the initial time. -/ theorem ODE_solution_unique_of_mem_Icc @@ -294,6 +299,7 @@ theorem ODE_solution_unique_of_mem_Icc (hg.mono <| Icc_subset_Icc_left <| le_of_lt ht.1) (fun _ ht' ↦ (hg' _ (hss ht')).hasDerivWithinAt) (fun _ ht' ↦ (hgs _ (hss ht'))) heq +include hv in /-- A version of `ODE_solution_unique_of_mem_Icc` for uniqueness in an open interval. -/ theorem ODE_solution_unique_of_mem_Ioo (ht : t₀ ∈ Ioo a b) @@ -324,6 +330,7 @@ theorem ODE_solution_unique_of_mem_Ioo (fun _ ht'' ↦ (hg _ <| hss <| Ico_subset_Icc_self ht'').2) heq ⟨h, le_rfl⟩ +include hv in /-- Local unqueness of ODE solutions. -/ theorem ODE_solution_unique_of_eventually (hf : ∀ᶠ t in 𝓝 t₀, HasDerivAt f (v t (f t)) t ∧ f t ∈ s t) diff --git a/Mathlib/Analysis/ODE/PicardLindelof.lean b/Mathlib/Analysis/ODE/PicardLindelof.lean index ca5ac0ac8a0a4..ed06739eb6415 100644 --- a/Mathlib/Analysis/ODE/PicardLindelof.lean +++ b/Mathlib/Analysis/ODE/PicardLindelof.lean @@ -227,8 +227,6 @@ instance [CompleteSpace E] : CompleteSpace v.FunSpace := by theorem intervalIntegrable_vComp (t₁ t₂ : ℝ) : IntervalIntegrable f.vComp volume t₁ t₂ := f.continuous_vComp.intervalIntegrable _ _ -variable [CompleteSpace E] - /-- The Picard-Lindelöf operator. This is a contracting map on `PicardLindelof.FunSpace v` such that the fixed point of this map is the solution of the corresponding ODE. @@ -244,20 +242,6 @@ def next (f : FunSpace v) : FunSpace v where theorem next_apply (t : Icc v.tMin v.tMax) : f.next t = v.x₀ + ∫ τ : ℝ in v.t₀..t, f.vComp τ := rfl -theorem hasDerivWithinAt_next (t : Icc v.tMin v.tMax) : - HasDerivWithinAt (f.next ∘ v.proj) (v t (f t)) (Icc v.tMin v.tMax) t := by - haveI : Fact ((t : ℝ) ∈ Icc v.tMin v.tMax) := ⟨t.2⟩ - simp only [(· ∘ ·), next_apply] - refine HasDerivWithinAt.const_add _ ?_ - have : HasDerivWithinAt (∫ τ in v.t₀..·, f.vComp τ) (f.vComp t) (Icc v.tMin v.tMax) t := - integral_hasDerivWithinAt_right (f.intervalIntegrable_vComp _ _) - (f.continuous_vComp.stronglyMeasurableAtFilter _ _) - f.continuous_vComp.continuousWithinAt - rw [vComp_apply_coe] at this - refine this.congr_of_eventuallyEq_of_mem ?_ t.coe_prop - filter_upwards [self_mem_nhdsWithin] with _ ht' - rw [v.proj_of_mem ht'] - theorem dist_next_apply_le_of_le {f₁ f₂ : FunSpace v} {n : ℕ} {d : ℝ} (h : ∀ t, dist (f₁ t) (f₂ t) ≤ (v.L * |t.1 - v.t₀|) ^ n / n ! * d) (t : Icc v.tMin v.tMax) : dist (next f₁ t) (next f₂ t) ≤ (v.L * |t.1 - v.t₀|) ^ (n + 1) / (n + 1)! * d := by @@ -293,10 +277,24 @@ theorem dist_iterate_next_le (f₁ f₂ : FunSpace v) (n : ℕ) : have : |(t - v.t₀ : ℝ)| ≤ v.tDist := v.dist_t₀_le t gcongr -end FunSpace - variable [CompleteSpace E] +theorem hasDerivWithinAt_next (t : Icc v.tMin v.tMax) : + HasDerivWithinAt (f.next ∘ v.proj) (v t (f t)) (Icc v.tMin v.tMax) t := by + haveI : Fact ((t : ℝ) ∈ Icc v.tMin v.tMax) := ⟨t.2⟩ + simp only [(· ∘ ·), next_apply] + refine HasDerivWithinAt.const_add _ ?_ + have : HasDerivWithinAt (∫ τ in v.t₀..·, f.vComp τ) (f.vComp t) (Icc v.tMin v.tMax) t := + integral_hasDerivWithinAt_right (f.intervalIntegrable_vComp _ _) + (f.continuous_vComp.stronglyMeasurableAtFilter _ _) + f.continuous_vComp.continuousWithinAt + rw [vComp_apply_coe] at this + refine this.congr_of_eventuallyEq_of_mem ?_ t.coe_prop + filter_upwards [self_mem_nhdsWithin] with _ ht' + rw [v.proj_of_mem ht'] + +end FunSpace + section theorem exists_contracting_iterate : @@ -308,7 +306,7 @@ theorem exists_contracting_iterate : exact ⟨N, ⟨_, this⟩, hN, LipschitzWith.of_dist_le_mul fun f g => FunSpace.dist_iterate_next_le f g N⟩ -theorem exists_fixed : ∃ f : v.FunSpace, f.next = f := +theorem exists_fixed [CompleteSpace E] : ∃ f : v.FunSpace, f.next = f := let ⟨_N, _K, hK⟩ := exists_contracting_iterate v ⟨_, hK.isFixedPt_fixedPoint_iterate⟩ @@ -316,7 +314,7 @@ end /-- Picard-Lindelöf (Cauchy-Lipschitz) theorem. Use `IsPicardLindelof.exists_forall_hasDerivWithinAt_Icc_eq` instead for the public API. -/ -theorem exists_solution : +theorem exists_solution [CompleteSpace E] : ∃ f : ℝ → E, f v.t₀ = v.x₀ ∧ ∀ t ∈ Icc v.tMin v.tMax, HasDerivWithinAt f (v t (f t)) (Icc v.tMin v.tMax) t := by rcases v.exists_fixed with ⟨f, hf⟩ diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean index 86c0558714218..5c8d9e5532473 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean @@ -104,6 +104,7 @@ theorem logb_pow {k : ℕ} (hx : 0 < x) : logb b (x ^ k) = k * logb b x := by section BPosAndNeOne variable (b_pos : 0 < b) (b_ne_one : b ≠ 1) +include b_pos b_ne_one private theorem log_b_ne_zero : log b ≠ 0 := by have b_ne_zero : b ≠ 0 := by linarith @@ -159,6 +160,7 @@ end BPosAndNeOne section OneLtB variable (hb : 1 < b) +include hb private theorem b_pos : 0 < b := by linarith @@ -253,9 +255,12 @@ end OneLtB section BPosAndBLtOne variable (b_pos : 0 < b) (b_lt_one : b < 1) +include b_lt_one private theorem b_ne_one : b ≠ 1 := by linarith +include b_pos + @[simp] theorem logb_le_logb_of_base_lt_one (h : 0 < x) (h₁ : 0 < y) : logb b x ≤ logb b y ↔ y ≤ x := by rw [logb, logb, div_le_div_right_of_neg (log_neg b_pos b_lt_one), log_le_log_iff h₁ h] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Integral.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Integral.lean index 2fced8eca57ba..c15cc1874154a 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Integral.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Integral.lean @@ -35,9 +35,11 @@ namespace MeasureTheory variable {α : Type*} [MeasurableSpace α] {f : α → ℝ} (μ : Measure α) (f_nn : 0 ≤ᵐ[μ] f) (f_mble : AEMeasurable f μ) {p : ℝ} (p_pos : 0 < p) +include f_nn f_mble p_pos section Layercake +include f_nn f_mble p_pos in /-- An application of the layer cake formula / Cavalieri's principle / tail probability formula: For a nonnegative function `f` on a measure space, the Lebesgue integral of `f` can @@ -73,6 +75,7 @@ end Layercake section LayercakeLT +include f_nn f_mble p_pos in /-- An application of the layer cake formula / Cavalieri's principle / tail probability formula: For a nonnegative function `f` on a measure space, the Lebesgue integral of `f` can diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 293a306a56774..903a92a213646 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -382,6 +382,7 @@ section EdistLeGeometric variable [PseudoEMetricSpace α] (r C : ℝ≥0∞) (hr : r < 1) (hC : C ≠ ⊤) {f : ℕ → α} (hu : ∀ n, edist (f n) (f (n + 1)) ≤ C * r ^ n) +include hr hC hu in /-- If `edist (f n) (f (n+1))` is bounded by `C * r^n`, `C ≠ ∞`, `r < 1`, then `f` is a Cauchy sequence. -/ theorem cauchySeq_of_edist_le_geometric : CauchySeq f := by @@ -390,6 +391,7 @@ theorem cauchySeq_of_edist_le_geometric : CauchySeq f := by refine ENNReal.mul_ne_top hC (ENNReal.inv_ne_top.2 ?_) exact (tsub_pos_iff_lt.2 hr).ne' +include hu in /-- If `edist (f n) (f (n+1))` is bounded by `C * r^n`, then the distance from `f n` to the limit of `f` is bounded above by `C * r^n / (1 - r)`. -/ theorem edist_le_of_edist_le_geometric_of_tendsto {a : α} (ha : Tendsto f atTop (𝓝 a)) (n : ℕ) : @@ -397,6 +399,7 @@ theorem edist_le_of_edist_le_geometric_of_tendsto {a : α} (ha : Tendsto f atTop convert edist_le_tsum_of_edist_le_of_tendsto _ hu ha _ simp only [pow_add, ENNReal.tsum_mul_left, ENNReal.tsum_geometric, div_eq_mul_inv, mul_assoc] +include hu in /-- If `edist (f n) (f (n+1))` is bounded by `C * r^n`, then the distance from `f 0` to the limit of `f` is bounded above by `C / (1 - r)`. -/ theorem edist_le_of_edist_le_geometric_of_tendsto₀ {a : α} (ha : Tendsto f atTop (𝓝 a)) : @@ -410,12 +413,14 @@ section EdistLeGeometricTwo variable [PseudoEMetricSpace α] (C : ℝ≥0∞) (hC : C ≠ ⊤) {f : ℕ → α} (hu : ∀ n, edist (f n) (f (n + 1)) ≤ C / 2 ^ n) {a : α} (ha : Tendsto f atTop (𝓝 a)) +include hC hu in /-- If `edist (f n) (f (n+1))` is bounded by `C * 2^-n`, then `f` is a Cauchy sequence. -/ theorem cauchySeq_of_edist_le_geometric_two : CauchySeq f := by simp only [div_eq_mul_inv, ENNReal.inv_pow] at hu refine cauchySeq_of_edist_le_geometric 2⁻¹ C ?_ hC hu simp [ENNReal.one_lt_two] +include hu ha in /-- If `edist (f n) (f (n+1))` is bounded by `C * 2^-n`, then the distance from `f n` to the limit of `f` is bounded above by `2 * C * 2^-n`. -/ theorem edist_le_of_edist_le_geometric_two_of_tendsto (n : ℕ) : edist (f n) a ≤ 2 * C / 2 ^ n := by @@ -424,6 +429,7 @@ theorem edist_le_of_edist_le_geometric_two_of_tendsto (n : ℕ) : edist (f n) a convert edist_le_of_edist_le_geometric_of_tendsto 2⁻¹ C hu ha n using 1 rw [ENNReal.one_sub_inv_two, div_eq_mul_inv, inv_inv] +include hu ha in /-- If `edist (f n) (f (n+1))` is bounded by `C * 2^-n`, then the distance from `f 0` to the limit of `f` is bounded above by `2 * C`. -/ theorem edist_le_of_edist_le_geometric_two_of_tendsto₀ : edist (f 0) a ≤ 2 * C := by @@ -434,9 +440,13 @@ end EdistLeGeometricTwo section LeGeometric -variable [PseudoMetricSpace α] {r C : ℝ} (hr : r < 1) {f : ℕ → α} - (hu : ∀ n, dist (f n) (f (n + 1)) ≤ C * r ^ n) +variable [PseudoMetricSpace α] {r C : ℝ} {f : ℕ → α} +section +variable (hr : r < 1) (hu : ∀ n, dist (f n) (f (n + 1)) ≤ C * r ^ n) +include hr hu + +/-- If `dist (f n) (f (n+1))` is bounded by `C * r^n`, `r < 1`, then `f` is a Cauchy sequence. -/ theorem aux_hasSum_of_le_geometric : HasSum (fun n : ℕ ↦ C * r ^ n) (C / (1 - r)) := by rcases sign_cases_of_C_mul_pow_nonneg fun n ↦ dist_nonneg.trans (hu n) with (rfl | ⟨_, r₀⟩) · simp [hasSum_zero] @@ -467,7 +477,10 @@ theorem dist_le_of_le_geometric_of_tendsto {a : α} (ha : Tendsto f atTop (𝓝 rw [mul_comm] exact (this.mul_left _).tsum_eq.symm +end + variable (hu₂ : ∀ n, dist (f n) (f (n + 1)) ≤ C / 2 / 2 ^ n) +include hu₂ /-- If `dist (f n) (f (n+1))` is bounded by `(C / 2) / 2^n`, then `f` is a Cauchy sequence. -/ theorem cauchySeq_of_le_geometric_two : CauchySeq f := diff --git a/Mathlib/Analysis/Subadditive.lean b/Mathlib/Analysis/Subadditive.lean index fb466ac626ee5..e3104ca58d187 100644 --- a/Mathlib/Analysis/Subadditive.lean +++ b/Mathlib/Analysis/Subadditive.lean @@ -43,6 +43,7 @@ theorem lim_le_div (hbdd : BddBelow (range fun n => u n / n)) {n : ℕ} (hn : n rw [Subadditive.lim] exact csInf_le (hbdd.mono <| image_subset_range _ _) ⟨n, hn.bot_lt, rfl⟩ +include h in theorem apply_mul_add_le (k n r) : u (k * n + r) ≤ k * u n + u r := by induction k with | zero => simp only [Nat.zero_eq, Nat.cast_zero, zero_mul, zero_add]; rfl @@ -53,6 +54,7 @@ theorem apply_mul_add_le (k n r) : u (k * n + r) ≤ k * u n + u r := by _ ≤ u n + (k * u n + u r) := add_le_add_left IH _ _ = (k + 1 : ℕ) * u n + u r := by simp; ring +include h in theorem eventually_div_lt_of_div_lt {L : ℝ} {n : ℕ} (hn : n ≠ 0) (hL : u n / n < L) : ∀ᶠ p in atTop, u p / p < L := by /- It suffices to prove the statement for each arithmetic progression `(n * · + r)`. -/ diff --git a/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean b/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean index 19eb362cc37b0..1919022e9a035 100644 --- a/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean +++ b/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean @@ -128,6 +128,7 @@ end Four section Five variable {R₁ R₂ : ComposableArrows C 4} (hR₁ : R₁.Exact) (hR₂ : R₂.Exact) (φ : R₁ ⟶ R₂) +include hR₁ hR₂ #adaptation_note /-- nightly-2024-03-11 We turn off simprocs here. diff --git a/Mathlib/CategoryTheory/Abelian/Exact.lean b/Mathlib/CategoryTheory/Abelian/Exact.lean index 887f80467a65d..cfc118215dd0f 100644 --- a/Mathlib/CategoryTheory/Abelian/Exact.lean +++ b/Mathlib/CategoryTheory/Abelian/Exact.lean @@ -231,6 +231,7 @@ section variable [L.PreservesZeroMorphisms] variable (hL : ∀ (S : ShortComplex A), S.Exact → (S.map L).Exact) +include hL open ZeroObject diff --git a/Mathlib/CategoryTheory/Abelian/ProjectiveResolution.lean b/Mathlib/CategoryTheory/Abelian/ProjectiveResolution.lean index b9dd2157df844..854dacb44ab3e 100644 --- a/Mathlib/CategoryTheory/Abelian/ProjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Abelian/ProjectiveResolution.lean @@ -197,12 +197,12 @@ abbrev projectiveResolution (Z : C) [HasZeroObject C] ProjectiveResolution Z := (HasProjectiveResolution.out (Z := Z)).some +variable (C) variable [Abelian C] section variable [HasProjectiveResolutions C] -variable (C) in /-- Taking projective resolutions is functorial, if considered with target the homotopy category (`ℕ`-indexed chain complexes and chain maps up to homotopy). @@ -219,6 +219,8 @@ def projectiveResolutions : C ⥤ HomotopyCategory C (ComplexShape.down ℕ) whe apply HomotopyCategory.eq_of_homotopy apply ProjectiveResolution.liftCompHomotopy +variable {C} + /-- If `P : ProjectiveResolution X`, then the chosen `(projectiveResolutions C).obj X` is isomorphic (in the homotopy category) to `P.complex`. -/ def ProjectiveResolution.iso {X : C} (P : ProjectiveResolution X) : @@ -250,6 +252,7 @@ end variable [EnoughProjectives C] +variable {C} in theorem exact_d_f {X Y : C} (f : X ⟶ Y) : (ShortComplex.mk (d f) f (by simp)).Exact := by let α : ShortComplex.mk (d f) f (by simp) ⟶ ShortComplex.mk (kernel.ι f) f (by simp) := @@ -274,6 +277,7 @@ applied to the previously constructed morphism, and the map from the `n`-th object as `Projective.d`. -/ +variable {C} variable (Z : C) -- The construction of the projective resolution `of` would be very, very slow @@ -296,12 +300,8 @@ lemma ofComplex_exactAt_succ (n : ℕ) : simp only [ChainComplex.of_d] -- TODO: this should just be apply exact_d_f so something is missing match n with - | 0 => - apply exact_d_f ((ChainComplex.mkAux _ _ _ (d (Projective.π Z)) (d (d (Projective.π Z))) _ _ - 0).g) - | n+1 => - apply exact_d_f ((ChainComplex.mkAux _ _ _ (d (Projective.π Z)) (d (d (Projective.π Z))) _ _ - (n+1)).g) + | 0 => apply exact_d_f + | n + 1 => apply exact_d_f instance (n : ℕ) : Projective ((ofComplex Z).X n) := by obtain (_ | _ | _ | n) := n <;> apply Projective.projective_over diff --git a/Mathlib/CategoryTheory/CofilteredSystem.lean b/Mathlib/CategoryTheory/CofilteredSystem.lean index 277d38409282c..eb69db1cc53a2 100644 --- a/Mathlib/CategoryTheory/CofilteredSystem.lean +++ b/Mathlib/CategoryTheory/CofilteredSystem.lean @@ -328,6 +328,7 @@ section FiniteCofilteredSystem variable [∀ j : J, Nonempty (F.obj j)] [∀ j : J, Finite (F.obj j)] (Fsur : ∀ ⦃i j : J⦄ (f : i ⟶ j), (F.map f).Surjective) +include Fsur theorem eval_section_surjective_of_surjective (i : J) : (fun s : F.sections => s.val i).Surjective := fun x => by diff --git a/Mathlib/CategoryTheory/ComposableArrows.lean b/Mathlib/CategoryTheory/ComposableArrows.lean index 8acbee9382883..d7b5a70f49f70 100644 --- a/Mathlib/CategoryTheory/ComposableArrows.lean +++ b/Mathlib/CategoryTheory/ComposableArrows.lean @@ -852,7 +852,6 @@ variable (obj : Fin (n + 1) → C) (mapSucc : ∀ (i : Fin n), obj i.castSucc lemma mkOfObjOfMapSucc_exists : ∃ (F : ComposableArrows C n) (e : ∀ i, F.obj i ≅ obj i), ∀ (i : ℕ) (hi : i < n), mapSucc ⟨i, hi⟩ = (e ⟨i, _⟩).inv ≫ F.map' i (i + 1) ≫ (e ⟨i + 1, _⟩).hom := by - clear F G revert obj mapSucc induction' n with n hn · intro obj _ diff --git a/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean b/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean index 23770f32d2e54..c926a1b32013f 100644 --- a/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean +++ b/Mathlib/CategoryTheory/Functor/Derived/RightDerived.lean @@ -72,7 +72,6 @@ lemma isRightDerivedFunctor_iff_of_iso (α' : F ⟶ L ⋙ RF') (W : MorphismProp section variable [L.IsLocalization W] [RF.IsRightDerivedFunctor α W] - [RF'.IsRightDerivedFunctor α' W] [RF''.IsRightDerivedFunctor α'' W] /-- Constructor for natural transformations from a right derived functor. -/ noncomputable def rightDerivedDesc (G : D ⥤ H) (β : F ⟶ L ⋙ G) : RF ⟶ G := @@ -91,6 +90,7 @@ lemma rightDerived_fac_app (G : D ⥤ H) (β : F ⟶ L ⋙ G) (X : C) : have := IsRightDerivedFunctor.isLeftKanExtension RF α W RF.descOfIsLeftKanExtension_fac_app α G β X +include W in lemma rightDerived_ext (G : D ⥤ H) (γ₁ γ₂ : RF ⟶ G) (hγ : α ≫ whiskerLeft L γ₁ = α ≫ whiskerLeft L γ₂) : γ₁ = γ₂ := have := IsRightDerivedFunctor.isLeftKanExtension RF α W @@ -119,6 +119,8 @@ lemma rightDerivedNatTrans_id : rightDerivedNatTrans RF RF α α W (𝟙 F) = 𝟙 RF := rightDerived_ext RF α W _ _ _ (by aesop_cat) +variable [RF'.IsRightDerivedFunctor α' W] + @[reassoc (attr := simp)] lemma rightDerivedNatTrans_comp (τ : F ⟶ F') (τ' : F' ⟶ F'') : rightDerivedNatTrans RF RF' α α' W τ ≫ rightDerivedNatTrans RF' RF'' α' α'' W τ' = @@ -165,6 +167,7 @@ lemma hasRightDerivedFunctor_iff : variable {F} +include e in lemma hasRightDerivedFunctor_iff_of_iso : HasRightDerivedFunctor F W ↔ HasRightDerivedFunctor F' W := by rw [hasRightDerivedFunctor_iff F W.Q W, hasRightDerivedFunctor_iff F' W.Q W, diff --git a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean index 8845826dfdb7a..a637719323dad 100644 --- a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean @@ -144,6 +144,8 @@ def id : (𝟭 C).FullyFaithful where section variable (hF : F.FullyFaithful) +include hF + /-- The equivalence `(X ⟶ Y) ≃ (F.obj X ⟶ F.obj Y)` given by `h : F.FullyFaithful`. -/ @[simps] def homEquiv {X Y : C} : (X ⟶ Y) ≃ (F.obj X ⟶ F.obj Y) where @@ -178,7 +180,7 @@ def preimageIso {X Y : C} (e : F.obj X ≅ F.obj Y) : X ≅ Y where hom_inv_id := hF.map_injective (by simp) inv_hom_id := hF.map_injective (by simp) -lemma isIso_of_isIso_map (hF : F.FullyFaithful) {X Y : C} (f : X ⟶ Y) [IsIso (F.map f)] : +lemma isIso_of_isIso_map {X Y : C} (f : X ⟶ Y) [IsIso (F.map f)] : IsIso f := by simpa using (hF.preimageIso (asIso (F.map f))).isIso_hom diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean index 2c6d890495c7c..32996f6d67fcb 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Basic.lean @@ -335,18 +335,18 @@ noncomputable instance (f : L ⋙ G ⟶ L') [IsIso f] (F : C ⥤ H) : IsEquivalence (RightExtension.postcomp₁ G f F) := by apply CostructuredArrow.isEquivalenceMap₂ -variable (e : L ⋙ G ≅ L') (F : C ⥤ H) - variable {G} in -lemma hasLeftExtension_iff_postcomp₁ : +lemma hasLeftExtension_iff_postcomp₁ (e : L ⋙ G ≅ L') (F : C ⥤ H) : HasLeftKanExtension L' F ↔ HasLeftKanExtension L F := (LeftExtension.postcomp₁ G e.inv F).asEquivalence.hasInitial_iff variable {G} in -lemma hasRightExtension_iff_postcomp₁ : +lemma hasRightExtension_iff_postcomp₁ (e : L ⋙ G ≅ L') (F : C ⥤ H) : HasRightKanExtension L' F ↔ HasRightKanExtension L F := (RightExtension.postcomp₁ G e.hom F).asEquivalence.hasTerminal_iff +variable (e : L ⋙ G ≅ L') (F : C ⥤ H) + /-- Given an isomorphism `e : L ⋙ G ≅ L'`, a left extension of `F` along `L'` is universal iff the corresponding left extension of `L` along `L` is. -/ noncomputable def LeftExtension.isUniversalPostcomp₁Equiv (ex : LeftExtension L' F) : @@ -458,6 +458,7 @@ a natural isomorphism `L ≅ L'`. -/ def rightExtensionEquivalenceOfIso₁ : RightExtension L F ≌ RightExtension L' F := CostructuredArrow.mapNatIso ((whiskeringLeft C D H).mapIso iso₁) +include iso₁ in lemma hasRightExtension_iff_of_iso₁ : HasRightKanExtension L F ↔ HasRightKanExtension L' F := (rightExtensionEquivalenceOfIso₁ iso₁ F).hasTerminal_iff @@ -466,6 +467,7 @@ a natural isomorphism `L ≅ L'`. -/ def leftExtensionEquivalenceOfIso₁ : LeftExtension L F ≌ LeftExtension L' F := StructuredArrow.mapNatIso ((whiskeringLeft C D H).mapIso iso₁) +include iso₁ in lemma hasLeftExtension_iff_of_iso₁ : HasLeftKanExtension L F ↔ HasLeftKanExtension L' F := (leftExtensionEquivalenceOfIso₁ iso₁ F).hasInitial_iff @@ -480,6 +482,7 @@ a natural isomorphism `F ≅ F'`. -/ def rightExtensionEquivalenceOfIso₂ : RightExtension L F ≌ RightExtension L F' := CostructuredArrow.mapIso iso₂ +include iso₂ in lemma hasRightExtension_iff_of_iso₂ : HasRightKanExtension L F ↔ HasRightKanExtension L F' := (rightExtensionEquivalenceOfIso₂ L iso₂).hasTerminal_iff @@ -488,6 +491,7 @@ a natural isomorphism `F ≅ F'`. -/ def leftExtensionEquivalenceOfIso₂ : LeftExtension L F ≌ LeftExtension L F' := StructuredArrow.mapIso iso₂ +include iso₂ in lemma hasLeftExtension_iff_of_iso₂ : HasLeftKanExtension L F ↔ HasLeftKanExtension L F' := (leftExtensionEquivalenceOfIso₂ L iso₂).hasInitial_iff diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean index c3bd4a60730ab..cf9297f47d979 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Pointwise.lean @@ -120,6 +120,7 @@ def isPointwiseLeftKanExtensionEquivOfIso (e : E ≅ E') : right_inv h := by aesop variable (h : E.IsPointwiseLeftKanExtension) +include h lemma IsPointwiseLeftKanExtension.hasPointwiseLeftKanExtension : HasPointwiseLeftKanExtension L F := @@ -136,7 +137,7 @@ def IsPointwiseLeftKanExtension.homFrom (G : LeftExtension L F) : E ⟶ G := ext X simpa using (h (L.obj X)).fac (LeftExtension.coconeAt G _) (CostructuredArrow.mk (𝟙 _))) -lemma IsPointwiseLeftKanExtension.hom_ext (h : E.IsPointwiseLeftKanExtension) +lemma IsPointwiseLeftKanExtension.hom_ext {G : LeftExtension L F} {f₁ f₂ : E ⟶ G} : f₁ = f₂ := by ext Y apply (h Y).hom_ext @@ -232,6 +233,7 @@ def isPointwiseRightKanExtensionEquivOfIso (e : E ≅ E') : right_inv h := by aesop variable (h : E.IsPointwiseRightKanExtension) +include h lemma IsPointwiseRightKanExtension.hasPointwiseRightKanExtension : HasPointwiseRightKanExtension L F := @@ -248,7 +250,7 @@ def IsPointwiseRightKanExtension.homTo (G : RightExtension L F) : G ⟶ E := ext X simpa using (h (L.obj X)).fac (RightExtension.coneAt G _) (StructuredArrow.mk (𝟙 _)) ) -lemma IsPointwiseRightKanExtension.hom_ext (h : E.IsPointwiseRightKanExtension) +lemma IsPointwiseRightKanExtension.hom_ext {G : RightExtension L F} {f₁ f₂ : G ⟶ E} : f₁ = f₂ := by ext Y apply (h Y).hom_ext diff --git a/Mathlib/CategoryTheory/Galois/Decomposition.lean b/Mathlib/CategoryTheory/Galois/Decomposition.lean index 77a9a29932f42..b9febf20bbfc8 100644 --- a/Mathlib/CategoryTheory/Galois/Decomposition.lean +++ b/Mathlib/CategoryTheory/Galois/Decomposition.lean @@ -40,7 +40,6 @@ variable {C : Type u₁} [Category.{u₂} C] namespace PreGaloisCategory -variable [GaloisCategory C] section Decomposition @@ -70,6 +69,8 @@ private lemma has_decomp_connected_components_aux_initial (X : C) (h : IsInitial (fun s m _ ↦ IsInitial.hom_ext h m _) exact ⟨by simp only [IsEmpty.forall_iff], inferInstance⟩ +variable [GaloisCategory C] + /- Show decomposition by inducting on `Nat.card (F.obj X)`. -/ private lemma has_decomp_connected_components_aux (F : C ⥤ FintypeCat.{w}) [FiberFunctor F] (n : ℕ) : ∀ (X : C), n = Nat.card (F.obj X) → ∃ (ι : Type) (f : ι → C) @@ -184,7 +185,7 @@ Reference: [lenstraGSchemes, 3.14] -/ -variable (F : C ⥤ FintypeCat.{w}) [FiberFunctor F] +variable [GaloisCategory C] (F : C ⥤ FintypeCat.{w}) [FiberFunctor F] section GaloisRepAux @@ -208,6 +209,7 @@ private lemma mkSelfProdFib_map_π (t : F.obj X) : F.map (Pi.π _ t) (mkSelfProd variable {X} {A : C} (u : A ⟶ selfProd F X) (a : F.obj A) (h : F.map u a = mkSelfProdFib F X) {F} +include h /-- For each `x : F.obj X`, this is the composition of `u` with the projection at `x`. -/ @[simp] @@ -300,6 +302,7 @@ lemma exists_hom_from_galois_of_fiber_nonempty (X : C) (h : Nonempty (F.obj X)) obtain ⟨A, f, a, h1, _⟩ := exists_hom_from_galois_of_fiber F X x exact ⟨A, f, h1⟩ +include F in /-- Any connected object admits a hom from a Galois object. -/ lemma exists_hom_from_galois_of_connected (X : C) [IsConnected X] : ∃ (A : C) (_ : A ⟶ X), IsGalois A := diff --git a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean index e02cc45fc92e2..f8261a14065e9 100644 --- a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean +++ b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean @@ -222,16 +222,16 @@ lemma AutGalois.ext {f g : AutGalois F} ext A exact h A +variable [FiberFunctor F] + /-- `autGalois.π` is surjective for every pointed Galois object. -/ -theorem AutGalois.π_surjective [FiberFunctor F] (A : PointedGaloisObject F) : +theorem AutGalois.π_surjective (A : PointedGaloisObject F) : Function.Surjective (AutGalois.π F A) := fun (σ : Aut A.obj) ↦ by have (i : PointedGaloisObject F) : Finite ((autGaloisSystem F ⋙ forget _).obj i) := inferInstanceAs <| Finite (Aut (i.obj)) exact eval_section_surjective_of_surjective (autGaloisSystem F ⋙ forget _) (autGaloisSystem_map_surjective F) A σ -variable [FiberFunctor F] - section EndAutGaloisIsomorphism /-! diff --git a/Mathlib/CategoryTheory/GradedObject.lean b/Mathlib/CategoryTheory/GradedObject.lean index f38f0041573d5..fcee74c33d22f 100644 --- a/Mathlib/CategoryTheory/GradedObject.lean +++ b/Mathlib/CategoryTheory/GradedObject.lean @@ -485,6 +485,7 @@ def isColimitCofanMapObjComp : dsimp rw [assoc]) +include hpqr in lemma hasMap_comp [(X.mapObj p).HasMap q] : X.HasMap r := fun k => ⟨_, isColimitCofanMapObjComp X p q r hpqr k _ (fun j _ => X.isColimitCofanMapObj p j) _ ((X.mapObj p).isColimitCofanMapObj q k)⟩ diff --git a/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean b/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean index 5568b3f19475a..fa00a95c23594 100644 --- a/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean +++ b/Mathlib/CategoryTheory/GradedObject/Trifunctor.lean @@ -350,11 +350,14 @@ noncomputable def isColimitCofan₃MapBifunctor₁₂BifunctorMapObj (j : J) : variable {F₁₂ G ρ₁₂ X₁ X₂ X₃} +include ρ₁₂ in lemma HasGoodTrifunctor₁₂Obj.hasMap : HasMap ((((mapTrifunctor (bifunctorComp₁₂ F₁₂ G) I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) r := fun j => ⟨_, isColimitCofan₃MapBifunctor₁₂BifunctorMapObj F₁₂ G ρ₁₂ X₁ X₂ X₃ j⟩ variable (F₁₂ G ρ₁₂ X₁ X₂ X₃) + +section variable [HasMap ((((mapTrifunctor (bifunctorComp₁₂ F₁₂ G) I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) r] /-- The action on graded objects of a trifunctor obtained by composition of two @@ -383,6 +386,8 @@ lemma ι_mapBifunctorComp₁₂MapObjIso_inv (i₁ : I₁) (i₂ : I₂) (i₃ : CofanMapObjFun.inj_iso_hom (isColimitCofan₃MapBifunctor₁₂BifunctorMapObj F₁₂ G ρ₁₂ X₁ X₂ X₃ j) _ h +end + variable {X₁ X₂ X₃ F₁₂ G ρ₁₂} variable {j : J} {A : C₄} @@ -523,11 +528,14 @@ noncomputable def isColimitCofan₃MapBifunctorBifunctor₂₃MapObj (j : J) : variable {F₁₂ G ρ₁₂ X₁ X₂ X₃} +include ρ₂₃ in lemma HasGoodTrifunctor₂₃Obj.hasMap : HasMap ((((mapTrifunctor (bifunctorComp₂₃ F G₂₃) I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) r := fun j => ⟨_, isColimitCofan₃MapBifunctorBifunctor₂₃MapObj F G₂₃ ρ₂₃ X₁ X₂ X₃ j⟩ variable (F₁₂ G ρ₁₂ X₁ X₂ X₃) + +section variable [HasMap ((((mapTrifunctor (bifunctorComp₂₃ F G₂₃) I₁ I₂ I₃).obj X₁).obj X₂).obj X₃) r] /-- The action on graded objects of a trifunctor obtained by composition of two @@ -556,6 +564,8 @@ lemma ι_mapBifunctorComp₂₃MapObjIso_inv (i₁ : I₁) (i₂ : I₂) (i₃ : CofanMapObjFun.inj_iso_hom (isColimitCofan₃MapBifunctorBifunctor₂₃MapObj F G₂₃ ρ₂₃ X₁ X₂ X₃ j) _ h +end + variable {X₁ X₂ X₃ F G₂₃ ρ₂₃} variable {j : J} {A : C₄} diff --git a/Mathlib/CategoryTheory/GradedObject/Unitor.lean b/Mathlib/CategoryTheory/GradedObject/Unitor.lean index eece8ef2deb7c..265f2b3635c67 100644 --- a/Mathlib/CategoryTheory/GradedObject/Unitor.lean +++ b/Mathlib/CategoryTheory/GradedObject/Unitor.lean @@ -84,6 +84,7 @@ noncomputable def mapBifunctorLeftUnitorCofanIsColimit (j : J) : exact mapBifunctorObjSingle₀ObjIsInitial _ _ _ _ hi) (fun s m hm => by simp [← hm ⟨⟨0, j⟩, hp j⟩]) +include e hp in lemma mapBifunctorLeftUnitor_hasMap : HasMap (((mapBifunctor F I J).obj ((single₀ I).obj X)).obj Y) p := CofanMapObjFun.hasMap _ _ _ (mapBifunctorLeftUnitorCofanIsColimit F X e p hp Y) @@ -200,6 +201,7 @@ noncomputable def mapBifunctorRightUnitorCofanIsColimit (j : J) : rw [← hm ⟨⟨j, 0⟩, hp j⟩, mapBifunctorRightUnitorCofan_inj, assoc, ← Functor.map_comp_assoc, Iso.inv_hom_id, Functor.map_id, id_comp, Iso.inv_hom_id_app_assoc]) +include e hp in lemma mapBifunctorRightUnitor_hasMap : HasMap (((mapBifunctor F J I).obj X).obj ((single₀ I).obj Y)) p := CofanMapObjFun.hasMap _ _ _ (mapBifunctorRightUnitorCofanIsColimit F Y e p hp X) @@ -269,8 +271,8 @@ structure TriangleIndexData (r : I₁ × I₂ × I₃ → J) (π : I₁ × I₃ h₁ (i₁ : I₁) : p₁₂ (i₁, 0) = i₁ h₃ (i₃ : I₃) : p₂₃ (0, i₃) = i₃ -variable {r : I₁ × I₂ × I₃ → J} {π : I₁ × I₃ → J} - (τ : TriangleIndexData r π) +variable {r : I₁ × I₂ × I₃ → J} {π : I₁ × I₃ → J} (τ : TriangleIndexData r π) +include τ namespace TriangleIndexData @@ -320,10 +322,10 @@ variable {C₁ C₂ C₃ D I₁ I₂ I₃ J : Type*} [Category C₁] [Category C [HasGoodTrifunctor₁₂Obj F₁ G τ.ρ₁₂ X₁ ((single₀ I₂).obj X₂) X₃] [HasGoodTrifunctor₂₃Obj G F₂ τ.ρ₂₃ X₁ ((single₀ I₂).obj X₂) X₃] [HasMap (((mapBifunctor G I₁ I₃).obj X₁).obj X₃) π] - (triangle : ∀ (X₁ : C₁) (X₃ : C₃), ((associator.hom.app X₁).app X₂).app X₃ ≫ - (G.obj X₁).map (e₂.hom.app X₃) = (G.map (e₁.hom.app X₁)).app X₃) -lemma mapBifunctor_triangle : +lemma mapBifunctor_triangle + (triangle : ∀ (X₁ : C₁) (X₃ : C₃), ((associator.hom.app X₁).app X₂).app X₃ ≫ + (G.obj X₁).map (e₂.hom.app X₃) = (G.map (e₁.hom.app X₁)).app X₃) : (mapBifunctorAssociator associator τ.ρ₁₂ τ.ρ₂₃ X₁ ((single₀ I₂).obj X₂) X₃).hom ≫ mapBifunctorMapMap G π (𝟙 X₁) (mapBifunctorLeftUnitor F₂ X₂ e₂ τ.p₂₃ τ.h₃ X₃).hom = mapBifunctorMapMap G π (mapBifunctorRightUnitor F₁ X₂ e₁ τ.p₁₂ τ.h₁ X₁).hom (𝟙 X₃) := by diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 6d2c629e1fcdb..88f2a7239bcc4 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -325,6 +325,7 @@ 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))) +include F in theorem hasColimitsOfShape_of_final [HasColimitsOfShape C E] : HasColimitsOfShape D E where has_colimit := fun _ => hasColimit_of_comp F @@ -596,6 +597,7 @@ 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))) +include F in theorem hasLimitsOfShape_of_initial [HasLimitsOfShape C E] : HasLimitsOfShape D E where has_limit := fun _ => hasLimit_of_comp F diff --git a/Mathlib/CategoryTheory/Limits/MonoCoprod.lean b/Mathlib/CategoryTheory/Limits/MonoCoprod.lean index 4f00dd6062184..9922662f9f26a 100644 --- a/Mathlib/CategoryTheory/Limits/MonoCoprod.lean +++ b/Mathlib/CategoryTheory/Limits/MonoCoprod.lean @@ -109,6 +109,7 @@ section variable {I₁ I₂ : Type*} {X : I₁ ⊕ I₂ → C} (c : Cofan X) (c₁ : Cofan (X ∘ Sum.inl)) (c₂ : Cofan (X ∘ Sum.inr)) (hc : IsColimit c) (hc₁ : IsColimit c₁) (hc₂ : IsColimit c₂) +include hc hc₁ hc₂ /-- Given a family of objects `X : I₁ ⊕ I₂ → C`, a cofan of `X`, and two colimit cofans of `X ∘ Sum.inl` and `X ∘ Sum.inr`, this is a cofan for `c₁.pt` and `c₂.pt` whose @@ -172,7 +173,9 @@ lemma mono_of_injective_aux (hι : Function.Injective ι) (c : Cofan X) (c₁ : variable (hι : Function.Injective ι) (c : Cofan X) (c₁ : Cofan (X ∘ ι)) (hc : IsColimit c) (hc₁ : IsColimit c₁) +include hι +include hc in lemma mono_of_injective [HasCoproduct (fun (k : ((Set.range ι)ᶜ : Set I)) => X k.1)] : Mono (Cofan.IsColimit.desc hc₁ (fun i => c.inj (ι i))) := mono_of_injective_aux X ι hι c c₁ hc hc₁ _ (colimit.isColimit _) diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Square.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Square.lean index fb95223d7a2e0..3ed3c0339f3a5 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Square.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Square.lean @@ -85,6 +85,7 @@ variable {sq₁ : Square (Type v)} {sq₂ : Square (Type u)} (comm₁₃ : e₃ ∘ sq₁.f₁₃ = sq₂.f₁₃ ∘ e₁) (comm₂₄ : e₄ ∘ sq₁.f₂₄ = sq₂.f₂₄ ∘ e₂) (comm₃₄ : e₄ ∘ sq₁.f₃₄ = sq₂.f₃₄ ∘ e₃) +include comm₁₂ comm₁₃ comm₂₄ comm₃₄ variable (sq₁ sq₂) in lemma IsPullback.iff_of_equiv : sq₁.IsPullback ↔ sq₂.IsPullback := by diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean index 6fe88cd02ae8f..80021b7b4364e 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean @@ -658,6 +658,7 @@ lemma equivPullbackObj_symm_apply_snd (x : Types.PullbackObj f g) : obtain ⟨x, rfl⟩ := (equivPullbackObj hc).surjective x simp +include hc in lemma type_ext {x y : c.pt} (h₁ : c.fst x = c.fst y) (h₂ : c.snd x = c.snd y) : x = y := (equivPullbackObj hc).injective (by ext <;> assumption) diff --git a/Mathlib/CategoryTheory/Limits/Types.lean b/Mathlib/CategoryTheory/Limits/Types.lean index 10095a65f2e14..1b69256ba80da 100644 --- a/Mathlib/CategoryTheory/Limits/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Types.lean @@ -642,6 +642,7 @@ private noncomputable def limitOfSurjectionsSurjective.preimage | 0 => a | n+1 => (hF n (preimage a n)).choose +include hF in open limitOfSurjectionsSurjective in /-- Auxiliary lemma. Use `limit_of_surjections_surjective` instead. -/ lemma surjective_π_app_zero_of_surjective_map_aux : diff --git a/Mathlib/CategoryTheory/Localization/Adjunction.lean b/Mathlib/CategoryTheory/Localization/Adjunction.lean index e3621c0fd3eb7..4c50c4d4c9e5e 100644 --- a/Mathlib/CategoryTheory/Localization/Adjunction.lean +++ b/Mathlib/CategoryTheory/Localization/Adjunction.lean @@ -127,6 +127,7 @@ lemma localization_counit_app (X₂ : C₂) : end +include adj in lemma isLocalization [F.Full] [F.Faithful] : G.IsLocalization ((MorphismProperty.isomorphisms C₂).inverseImage G) := by let W := ((MorphismProperty.isomorphisms C₂).inverseImage G) diff --git a/Mathlib/CategoryTheory/Localization/Bousfield.lean b/Mathlib/CategoryTheory/Localization/Bousfield.lean index 37a5a70ba851a..35fde70a59774 100644 --- a/Mathlib/CategoryTheory/Localization/Bousfield.lean +++ b/Mathlib/CategoryTheory/Localization/Bousfield.lean @@ -108,6 +108,7 @@ end section variable {F : C ⥤ D} {G : D ⥤ C} (adj : G ⊣ F) [F.Full] [F.Faithful] +include adj lemma W_adj_unit_app (X : D) : W (· ∈ Set.range F.obj) (adj.unit.app X) := by rintro _ ⟨Y, rfl⟩ diff --git a/Mathlib/CategoryTheory/Localization/CalculusOfFractions.lean b/Mathlib/CategoryTheory/Localization/CalculusOfFractions.lean index 1301779a6aec0..b6cc4f0377426 100644 --- a/Mathlib/CategoryTheory/Localization/CalculusOfFractions.lean +++ b/Mathlib/CategoryTheory/Localization/CalculusOfFractions.lean @@ -768,6 +768,7 @@ lemma MorphismProperty.map_eq_iff_postcomp {X Y : C} (f₁ f₂ : X ⟶ Y) : simp only [← cancel_mono (Localization.isoOfHom L W s hs).hom, Localization.isoOfHom_hom, ← L.map_comp, fac] +include W in lemma Localization.essSurj_mapArrow : L.mapArrow.EssSurj where mem_essImage f := by @@ -976,6 +977,7 @@ lemma MorphismProperty.map_eq_iff_precomp {Y Z : C} (f₁ f₂ : Y ⟶ Z) : simp only [← cancel_epi (Localization.isoOfHom L W s hs).hom, Localization.isoOfHom_hom, ← L.map_comp, fac] +include W in lemma Localization.essSurj_mapArrow_of_hasRightCalculusOfFractions : L.mapArrow.EssSurj where mem_essImage f := by diff --git a/Mathlib/CategoryTheory/Localization/CalculusOfFractions/Preadditive.lean b/Mathlib/CategoryTheory/Localization/CalculusOfFractions/Preadditive.lean index d7bc5297388e4..4658401ffef2f 100644 --- a/Mathlib/CategoryTheory/Localization/CalculusOfFractions/Preadditive.lean +++ b/Mathlib/CategoryTheory/Localization/CalculusOfFractions/Preadditive.lean @@ -36,7 +36,7 @@ namespace CategoryTheory open MorphismProperty Preadditive Limits Category variable {C D : Type*} [Category C] [Category D] [Preadditive C] (L : C ⥤ D) - {W : MorphismProperty C} [L.IsLocalization W] [W.HasLeftCalculusOfFractions] + {W : MorphismProperty C} [L.IsLocalization W] namespace MorphismProperty @@ -91,7 +91,7 @@ should only rely on the fact that the localization functor is additive, as this completely determines the preadditive structure on the localized category when there is a calculus of left fractions. -/ -variable {X Y Z : C} +variable [W.HasLeftCalculusOfFractions] {X Y Z : C} variable {L} /-- The opposite of a map `L.obj X ⟶ L.obj Y` when `L : C ⥤ D` is a localization @@ -302,6 +302,8 @@ end ImplementationDetails end Preadditive +variable [W.HasLeftCalculusOfFractions] + /-- The preadditive structure on `D`, when `L : C ⥤ D` is a localization functor, `C` is preadditive and there is a left calculus of fractions. -/ noncomputable def preadditive : Preadditive D where @@ -316,7 +318,7 @@ lemma functor_additive : ⟨by apply Preadditive.map_add⟩ attribute [irreducible] preadditive - +include W in lemma functor_additive_iff {E : Type*} [Category E] [Preadditive E] [Preadditive D] [L.Additive] (G : D ⥤ E) : G.Additive ↔ (L ⋙ G).Additive := by diff --git a/Mathlib/CategoryTheory/Localization/Equivalence.lean b/Mathlib/CategoryTheory/Localization/Equivalence.lean index cbe9e06a98b35..3f76b0aaa237e 100644 --- a/Mathlib/CategoryTheory/Localization/Equivalence.lean +++ b/Mathlib/CategoryTheory/Localization/Equivalence.lean @@ -51,6 +51,7 @@ lemma equivalence_counitIso_app (X : C₂) : dsimp [Lifting.iso] rw [comp_id] +include L₁ W₁ L₂ W₂ G F F' α β in /-- Basic constructor of an equivalence between localized categories -/ lemma isEquivalence : G'.IsEquivalence := (equivalence L₁ W₁ L₂ W₂ G G' F F' α β).isEquivalence_functor diff --git a/Mathlib/CategoryTheory/Localization/FiniteProducts.lean b/Mathlib/CategoryTheory/Localization/FiniteProducts.lean index e69a4d3c58363..7fe7c9bfe0e8a 100644 --- a/Mathlib/CategoryTheory/Localization/FiniteProducts.lean +++ b/Mathlib/CategoryTheory/Localization/FiniteProducts.lean @@ -28,17 +28,20 @@ open Limits namespace Localization variable {C : Type u₁} {D : Type u₂} [Category.{v₁} C] [Category.{v₂} D] (L : C ⥤ D) - {W : MorphismProperty C} [L.IsLocalization W] [W.ContainsIdentities] + {W : MorphismProperty C} [L.IsLocalization W] namespace HasProductsOfShapeAux -variable {J : Type} [Finite J] [HasProductsOfShape J C] +variable {J : Type} [HasProductsOfShape J C] (hW : W.IsStableUnderProductsOfShape J) +include hW lemma inverts : (W.functorCategory (Discrete J)).IsInvertedBy (lim ⋙ L) := fun _ _ f hf => Localization.inverts L W _ (hW.lim_map f hf) +variable [W.ContainsIdentities] [Finite J] + /-- The (candidate) limit functor for the localized category. It is induced by `lim ⋙ L : (Discrete J ⥤ C) ⥤ D`. -/ noncomputable abbrev limitFunctor : @@ -84,7 +87,9 @@ noncomputable def isLimitMapCone (F : Discrete J ⥤ C) : end HasProductsOfShapeAux variable (W) +variable [W.ContainsIdentities] +include L lemma hasProductsOfShape (J : Type) [Finite J] [HasProductsOfShape J C] (hW : W.IsStableUnderProductsOfShape J) : HasProductsOfShape J D := @@ -102,6 +107,7 @@ noncomputable def preservesProductsOfShape (J : Type) [Finite J] variable [HasFiniteProducts C] [W.IsStableUnderFiniteProducts] +include W in lemma hasFiniteProducts : HasFiniteProducts D := ⟨fun _ => hasProductsOfShape L W _ (W.isStableUnderProductsOfShape_of_isStableUnderFiniteProducts _)⟩ diff --git a/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean b/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean index 40d5dc9790705..c0f59f13389ca 100644 --- a/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean +++ b/Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean @@ -100,6 +100,7 @@ variable [CatCommSq Φ.functor L₁ L₂ G] [Category.{v₄'} D₁'] [Category.{v₅'} D₂'] (L₁' : C₁ ⥤ D₁') (L₂' : C₂ ⥤ D₂') [L₁'.IsLocalization W₁] [L₂'.IsLocalization W₂] (G' : D₁' ⥤ D₂') [CatCommSq Φ.functor L₁' L₂' G'] +include W₁ W₂ Φ L₁ L₂ L₁' L₂' /-- If a localizer morphism induces an equivalence on some choice of localized categories, it will be so for any choice of localized categoriees. -/ diff --git a/Mathlib/CategoryTheory/Localization/SmallHom.lean b/Mathlib/CategoryTheory/Localization/SmallHom.lean index ded70ab026400..4f7063c5ba725 100644 --- a/Mathlib/CategoryTheory/Localization/SmallHom.lean +++ b/Mathlib/CategoryTheory/Localization/SmallHom.lean @@ -55,6 +55,7 @@ lemma hasSmallLocalizedHom_iff : · intro h exact ⟨small_map (homEquiv W W.Q L)⟩ +include L in lemma hasSmallLocalizedHom_of_isLocalization : HasSmallLocalizedHom.{v₂} W X Y := by rw [hasSmallLocalizedHom_iff W L] diff --git a/Mathlib/CategoryTheory/Localization/SmallShiftedHom.lean b/Mathlib/CategoryTheory/Localization/SmallShiftedHom.lean index 712b539adf8cc..378b267f05134 100644 --- a/Mathlib/CategoryTheory/Localization/SmallShiftedHom.lean +++ b/Mathlib/CategoryTheory/Localization/SmallShiftedHom.lean @@ -30,7 +30,6 @@ open Category variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D] (W : MorphismProperty C) {M : Type w'} [AddMonoid M] [HasShift C M] [HasShift D M] - [W.IsCompatibleWithShift M] namespace Localization @@ -58,6 +57,7 @@ lemma hasSmallLocalizedShiftedHom_iff variable [HasSmallLocalizedShiftedHom.{w} W M X Y] +include M in variable (M) in lemma hasSmallLocalizedHom_of_hasSmallLocalizedShiftedHom₀ : HasSmallLocalizedHom.{w} W X Y := @@ -85,7 +85,7 @@ end namespace SmallHom variable {W} -variable (L : C ⥤ D) [L.IsLocalization W] [L.CommShift M] +variable [W.IsCompatibleWithShift M] (L : C ⥤ D) [L.IsLocalization W] [L.CommShift M] {X Y : C} [HasSmallLocalizedHom.{w} W X Y] (f : SmallHom.{w} W X Y) (a : M) [HasSmallLocalizedHom.{w} W (X⟦a⟧) (Y⟦a⟧)] @@ -111,6 +111,7 @@ namespace SmallShiftedHom section variable {W} +variable [W.IsCompatibleWithShift M] variable {X Y Z : C} /-- Given `f : SmallShiftedHom.{w} W X Y a`, this is the element in @@ -142,7 +143,7 @@ end section -variable (L : C ⥤ D) [L.IsLocalization W] [HasShift D M] [L.CommShift M] +variable (L : C ⥤ D) [L.IsLocalization W] [L.CommShift M] {X Y Z T : C} /-- The bijection `SmallShiftedHom.{w} W X Y m ≃ ShiftedHom (L.obj X) (L.obj Y) m` @@ -153,6 +154,9 @@ noncomputable def equiv [HasSmallLocalizedShiftedHom.{w} W M X Y] {m : M} : SmallShiftedHom.{w} W X Y m ≃ ShiftedHom (L.obj X) (L.obj Y) m := (SmallHom.equiv W L).trans ((L.commShiftIso m).app Y).homToEquiv +section +variable [W.IsCompatibleWithShift M] + lemma equiv_shift' {a : M} [HasSmallLocalizedShiftedHom.{w} W M X Y] [HasSmallLocalizedShiftedHom.{w} W M Y Y] (f : SmallShiftedHom.{w} W X Y a) (n a' : M) (h : a + n = a') : @@ -184,6 +188,8 @@ lemma equiv_comp [HasSmallLocalizedShiftedHom.{w} W M X Y] comp_id, Functor.map_comp] rfl +end + @[simp] lemma equiv_mk₀ [HasSmallLocalizedShiftedHom.{w} W M X Y] (m₀ : M) (hm₀ : m₀ = 0) (f : X ⟶ Y) : @@ -198,6 +204,8 @@ lemma equiv_mk₀ [HasSmallLocalizedShiftedHom.{w} W M X Y] end +variable [W.IsCompatibleWithShift M] + lemma comp_assoc {X Y Z T : C} {a₁ a₂ a₃ a₁₂ a₂₃ a : M} [HasSmallLocalizedShiftedHom.{w} W M X Y] [HasSmallLocalizedShiftedHom.{w} W M X Z] [HasSmallLocalizedShiftedHom.{w} W M X T] [HasSmallLocalizedShiftedHom.{w} W M Y Z] diff --git a/Mathlib/CategoryTheory/Localization/Triangulated.lean b/Mathlib/CategoryTheory/Localization/Triangulated.lean index d97d95429aa7f..1be2f5f65a3cd 100644 --- a/Mathlib/CategoryTheory/Localization/Triangulated.lean +++ b/Mathlib/CategoryTheory/Localization/Triangulated.lean @@ -120,6 +120,7 @@ namespace Localization variable (W : MorphismProperty C) [L.IsLocalization W] [W.HasLeftCalculusOfFractions] +include W in lemma distinguished_cocone_triangle {X Y : D} (f : X ⟶ Y) : ∃ (Z : D) (g : Y ⟶ Z) (h : Z ⟶ X⟦(1 : ℤ)⟧), Triangle.mk f g h ∈ L.essImageDistTriang := by @@ -138,6 +139,7 @@ lemma distinguished_cocone_triangle {X Y : D} (f : X ⟶ Y) : section variable [W.IsCompatibleWithTriangulation] +include W in lemma complete_distinguished_triangle_morphism (T₁ T₂ : Triangle D) (hT₁ : T₁ ∈ L.essImageDistTriang) (hT₂ : T₂ ∈ L.essImageDistTriang) (a : T₁.obj₁ ⟶ T₂.obj₁) (b : T₁.obj₂ ⟶ T₂.obj₂) (fac : T₁.mor₁ ≫ b = a ≫ T₂.mor₁) : @@ -200,6 +202,7 @@ end variable [HasZeroObject D] [Preadditive D] [∀ (n : ℤ), (shiftFunctor D n).Additive] +include W in lemma isTriangulated [Pretriangulated D] [L.IsTriangulated] [IsTriangulated C] : IsTriangulated D := by have := essSurj_mapComposableArrows L W 2 diff --git a/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean b/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean index 4d2244e5a0fa4..1a22be0cc03ff 100644 --- a/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean +++ b/Mathlib/CategoryTheory/Preadditive/AdditiveFunctor.lean @@ -120,6 +120,7 @@ lemma additive_of_comp_faithful rw [← Functor.comp_map, G.map_add, (F ⋙ G).map_add, Functor.comp_map, Functor.comp_map]) open ZeroObject Limits in +include F in lemma hasZeroObject_of_additive [HasZeroObject C] : HasZeroObject D where zero := ⟨F.obj 0, by rw [IsZero.iff_id_eq_zero, ← F.map_id, id_zero, F.map_zero]⟩ diff --git a/Mathlib/CategoryTheory/Sites/Coherent/ReflectsPrecoherent.lean b/Mathlib/CategoryTheory/Sites/Coherent/ReflectsPrecoherent.lean index ded20250ddf8c..a42e9b7bca302 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/ReflectsPrecoherent.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/ReflectsPrecoherent.lean @@ -22,6 +22,7 @@ variable {C D : Type*} [Category C] [Category D] (F : C ⥤ D) [F.EffectivelyEnough] [Precoherent D] [F.Full] [F.Faithful] +include F in lemma Functor.reflects_precoherent : Precoherent C where pullback {B₁ B₂} f α _ X₁ π₁ _ := by obtain ⟨β, _, Y₂, τ₂, H, i, ι, hh⟩ := Precoherent.pullback (F.map f) _ _ diff --git a/Mathlib/CategoryTheory/Sites/Coherent/ReflectsPreregular.lean b/Mathlib/CategoryTheory/Sites/Coherent/ReflectsPreregular.lean index 89d59708b6161..9b3784170505c 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/ReflectsPreregular.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/ReflectsPreregular.lean @@ -22,6 +22,7 @@ variable {C D : Type*} [Category C] [Category D] (F : C ⥤ D) [F.EffectivelyEnough] [Preregular D] [F.Full] [F.Faithful] +include F in lemma Functor.reflects_preregular : Preregular C where exists_fac f g _ := by obtain ⟨W, f', _, i, w⟩ := Preregular.exists_fac (F.map f) (F.map g) diff --git a/Mathlib/CategoryTheory/Sites/CoverLifting.lean b/Mathlib/CategoryTheory/Sites/CoverLifting.lean index 3755ebea1a0f2..fc2152f710887 100644 --- a/Mathlib/CategoryTheory/Sites/CoverLifting.lean +++ b/Mathlib/CategoryTheory/Sites/CoverLifting.lean @@ -184,6 +184,7 @@ lemma fac (i : S.Arrow) : lift hF hR s ≫ R.map i.f.op = s.ι i := by rw [Category.assoc, eq] simpa using liftAux_map hF α s (j.hom.unop ≫ i.f) (𝟙 _) i j.hom.unop (by simp) +include hR hF in variable (K) in lemma hom_ext {W : A} {f g : W ⟶ R.obj (op X)} (h : ∀ (i : S.Arrow), f ≫ R.map i.f.op = g ≫ R.map i.f.op) : f = g := by diff --git a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean index 722f976deaf36..49db6e1c4c551 100644 --- a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean +++ b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean @@ -80,8 +80,10 @@ structure CompatiblePreserving (K : GrothendieckTopology D) (G : C ⥤ D) : Prop {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 {T : Presieve Z} {x : FamilyOfElements (G.op ⋙ ℱ.val) T} (h : x.Compatible) +include hG h /-- `CompatiblePreserving` functors indeed preserve compatible families. -/ theorem Presieve.FamilyOfElements.Compatible.functorPushforward : @@ -103,6 +105,8 @@ theorem CompatiblePreserving.apply_map {Y : C} {f : Y ⟶ Z} (hf : T f) : ⟨X, g, f', hg, eq⟩ simpa using hG.compatible ℱ h f' (𝟙 _) hg hf (by simp [eq]) +end + open Limits.WalkingCospan theorem compatiblePreservingOfFlat {C : Type u₁} [Category.{v₁} C] {D : Type u₁} [Category.{v₁} D] @@ -150,9 +154,6 @@ theorem compatiblePreservingOfDownwardsClosed (F : C ⥤ D) [F.Full] [F.Faithful hx (F.preimage <| e.hom ≫ f₁) (F.preimage <| e.hom ≫ f₂) hg₁ hg₂ (F.map_injective <| by simpa using he) -variable (J K) - - variable {F J K} /-- If `F` is cover-preserving and compatible-preserving, diff --git a/Mathlib/CategoryTheory/Sites/CoversTop.lean b/Mathlib/CategoryTheory/Sites/CoversTop.lean index d36281d946589..fafa6768c9f15 100644 --- a/Mathlib/CategoryTheory/Sites/CoversTop.lean +++ b/Mathlib/CategoryTheory/Sites/CoversTop.lean @@ -51,6 +51,7 @@ namespace CoversTop variable {J} variable {I : Type*} {Y : I → C} (hY : J.CoversTop Y) +include hY /-- The cover of any object `W : C` attached to a family of objects `Y` that satisfy `J.CoversTop Y` -/ diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean index f5b5acba282c3..dba0ba9fb3362 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean @@ -112,10 +112,10 @@ namespace Functor namespace IsCoverDense variable {K} -variable {A : Type*} [Category A] (G : C ⥤ D) [G.IsCoverDense 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 (ℱ : SheafOfTypes K) (X : D) {s t : ℱ.val.obj (op X)} +theorem ext [G.IsCoverDense K] (ℱ : SheafOfTypes K) (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 rintro Y _ ⟨Z, f₁, f₂, ⟨rfl⟩⟩ @@ -123,7 +123,7 @@ theorem ext (ℱ : SheafOfTypes K) (X : D) {s t : ℱ.val.obj (op X)} variable {G} -theorem functorPullback_pushforward_covering [G.IsLocallyFull K] {X : C} +theorem functorPullback_pushforward_covering [G.IsCoverDense K] [G.IsLocallyFull K] {X : C} (T : K (G.obj X)) : (T.val.functorPullback G).functorPushforward G ∈ K (G.obj X) := by refine K.transitive T.2 _ fun Y iYX hiYX ↦ ?_ apply K.transitive (G.is_cover_of_isCoverDense _ _) _ @@ -154,13 +154,11 @@ theorem sheaf_eq_amalgamation (ℱ : Sheaf K A) {X : A} {U : D} {T : Sieve U} (h t = (ℱ.cond X T hT).amalgamate x hx := (ℱ.cond X T hT).isSeparatedFor x t _ h ((ℱ.cond X T hT).isAmalgamation hx) -variable [G.IsLocallyFull K] - namespace Types variable {ℱ : Dᵒᵖ ⥤ Type v} {ℱ' : SheafOfTypes.{v} K} (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) -theorem naturality_apply {X Y : C} (i : G.obj X ⟶ G.obj Y) (x) : +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 have {X Y} (i : X ⟶ Y) (x) : ℱ'.1.map (G.map i).op (α.app _ x) = α.app _ (ℱ.map (G.map i).op x) := by @@ -169,7 +167,7 @@ theorem naturality_apply {X Y : C} (i : G.obj X ⟶ G.obj Y) (x) : simp only [comp_obj, types_comp_apply, ← FunctorToTypes.map_comp_apply, ← op_comp, ← e, this] @[reassoc] -theorem naturality {X Y : C} (i : G.obj X ⟶ G.obj Y) : +theorem naturality [G.IsLocallyFull K] {X Y : C} (i : G.obj X ⟶ G.obj Y) : α.app _ ≫ ℱ'.1.map i.op = ℱ.map i.op ≫ α.app _ := types_ext _ _ (naturality_apply α i) /-- @@ -189,6 +187,17 @@ noncomputable def pushforwardFamily {X} (x : ℱ.obj (op X)) : pushforwardFamily α x = fun _ _ hf => ℱ'.val.map hf.some.lift.op <| α.app (op _) (ℱ.map hf.some.map.op x : _) := rfl +@[simp] +theorem pushforwardFamily_apply [G.IsLocallyFull K] + {X} (x : ℱ.obj (op X)) {Y : C} (f : G.obj Y ⟶ X) : + pushforwardFamily α x f (Presieve.in_coverByImage G f) = α.app (op Y) (ℱ.map f.op x) := by + simp only [pushforwardFamily_def, op_obj] + generalize Nonempty.some (Presieve.in_coverByImage G f) = l + obtain ⟨W, iYW, iWX, rfl⟩ := l + simp only [← op_comp, ← FunctorToTypes.map_comp_apply, naturality_apply] + +variable [G.IsCoverDense K] [G.IsLocallyFull K] + /-- (Implementation). The `pushforwardFamily` defined is compatible. -/ theorem pushforwardFamily_compatible {X} (x : ℱ.obj (op X)) : (pushforwardFamily α x).Compatible := by @@ -212,14 +221,6 @@ noncomputable def appHom (X : D) : ℱ.obj (op X) ⟶ ℱ'.val.obj (op X) := fun (ℱ'.cond _ (G.is_cover_of_isCoverDense _ X)).amalgamate (pushforwardFamily α x) (pushforwardFamily_compatible α x) -@[simp] -theorem pushforwardFamily_apply {X} (x : ℱ.obj (op X)) {Y : C} (f : G.obj Y ⟶ X) : - pushforwardFamily α x f (Presieve.in_coverByImage G f) = α.app (op Y) (ℱ.map f.op x) := by - simp only [pushforwardFamily_def, op_obj] - generalize Nonempty.some (Presieve.in_coverByImage G f) = l - obtain ⟨W, iYW, iWX, rfl⟩ := l - simp only [← op_comp, ← FunctorToTypes.map_comp_apply, naturality_apply] - @[simp] theorem appHom_restrict {X : D} {Y : C} (f : op X ⟶ op (G.obj Y)) (x) : ℱ'.val.map f (appHom α X x) = α.app (op Y) (ℱ.map f x) := @@ -298,7 +299,7 @@ end Types open Types -variable {ℱ : Dᵒᵖ ⥤ A} {ℱ' : Sheaf K A} +variable [G.IsCoverDense K] [G.IsLocallyFull K] {ℱ : Dᵒᵖ ⥤ A} {ℱ' : Sheaf K A} /-- (Implementation). The sheaf map given in `types.sheaf_hom` is natural in terms of `X`. -/ @[simps] @@ -499,8 +500,7 @@ lemma whiskerLeft_obj_map_bijective_of_isCoverDense (G : C ⥤ D) Function.Bijective (((whiskeringLeft Cᵒᵖ Dᵒᵖ A).obj G.op).map : (P ⟶ Q) → _) := (IsCoverDense.restrictHomEquivHom (ℱ' := ⟨Q, hQ⟩)).symm.bijective -variable {A : Type*} [Category A] -variable (J : GrothendieckTopology C) (K : GrothendieckTopology D) (G : C ⥤ D) +variable {A : Type*} [Category A] (G : C ⥤ D) /-- The functor `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)` if `G` is cover-dense, locally fully-faithful, @@ -518,44 +518,46 @@ namespace IsDenseSubsite variable [G.IsDenseSubsite J K] +include J K + lemma isCoverDense : G.IsCoverDense K := isCoverDense' J lemma isLocallyFull : G.IsLocallyFull K := isLocallyFull' J lemma isLocallyFaithful : G.IsLocallyFaithful K := isLocallyFaithful' J -lemma coverPreserving [G.IsDenseSubsite J K] : CoverPreserving J K G := +lemma coverPreserving : CoverPreserving J K G := ⟨functorPushforward_mem_iff.mpr⟩ -instance (priority := 900) [G.IsDenseSubsite J K] : G.IsContinuous J K := +instance (priority := 900) : G.IsContinuous J K := letI := IsDenseSubsite.isCoverDense J K G letI := IsDenseSubsite.isLocallyFull J K G letI := IsDenseSubsite.isLocallyFaithful J K G IsCoverDense.isContinuous J K G (IsDenseSubsite.coverPreserving J K G) -instance (priority := 900) [G.IsDenseSubsite J K] : G.IsCocontinuous J K where +instance (priority := 900) : G.IsCocontinuous J K where cover_lift hS := letI := IsDenseSubsite.isCoverDense J K G letI := IsDenseSubsite.isLocallyFull J K G IsDenseSubsite.functorPushforward_mem_iff.mp (IsCoverDense.functorPullback_pushforward_covering ⟨_, hS⟩) -instance full_sheafPushforwardContinuous [G.IsDenseSubsite J K] : +instance full_sheafPushforwardContinuous : Full (G.sheafPushforwardContinuous A J K) := letI := IsDenseSubsite.isCoverDense J K G letI := IsDenseSubsite.isLocallyFull J K G inferInstance -instance faithful_sheafPushforwardContinuous [G.IsDenseSubsite J K] : +instance faithful_sheafPushforwardContinuous : Faithful (G.sheafPushforwardContinuous A J K) := letI := IsDenseSubsite.isCoverDense J K G letI := IsDenseSubsite.isLocallyFull J K G inferInstance -lemma imageSieve_mem [G.IsDenseSubsite J K] {U V} (f : G.obj U ⟶ G.obj V) : +lemma imageSieve_mem {U V} (f : G.obj U ⟶ G.obj V) : G.imageSieve f ∈ J _ := letI := IsDenseSubsite.isLocallyFull J K G IsDenseSubsite.functorPushforward_mem_iff.mp (G.functorPushforward_imageSieve_mem K f) -lemma equalizer_mem [G.IsDenseSubsite J K] {U V} (f₁ f₂ : U ⟶ V) (e : G.map f₁ = G.map f₂) : +lemma equalizer_mem {U V} (f₁ f₂ : U ⟶ V) (e : G.map f₁ = G.map f₂) : Sieve.equalizer f₁ f₂ ∈ J _ := letI := IsDenseSubsite.isLocallyFaithful J K G IsDenseSubsite.functorPushforward_mem_iff.mp (G.functorPushforward_equalizer_mem K f₁ f₂ e) @@ -577,6 +579,7 @@ 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] diff --git a/Mathlib/CategoryTheory/Sites/Equivalence.lean b/Mathlib/CategoryTheory/Sites/Equivalence.lean index 6bfc05edd2dbc..742f549874445 100644 --- a/Mathlib/CategoryTheory/Sites/Equivalence.lean +++ b/Mathlib/CategoryTheory/Sites/Equivalence.lean @@ -147,6 +147,7 @@ def transportSheafificationAdjunction : transportAndSheafify J K e A ⊣ sheafTo noncomputable instance : PreservesFiniteLimits <| transportAndSheafify J K e A where preservesFiniteLimits _ := compPreservesLimitsOfShape _ _ +include K e in /-- Transport `HasSheafify` along an equivalence of sites. -/ theorem hasSheafify : HasSheafify J A := HasSheafify.mk' J A (transportSheafificationAdjunction J K e A) @@ -154,6 +155,7 @@ theorem hasSheafify : HasSheafify J A := variable {A : Type*} [Category A] {B : Type*} [Category B] (F : A ⥤ B) [K.HasSheafCompose F] +include K e in theorem hasSheafCompose : J.HasSheafCompose F where isSheaf P hP := by have hP' : Presheaf.IsSheaf K (e.inverse.op ⋙ P ⋙ F) := by @@ -169,8 +171,10 @@ variable {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) end Equivalence -variable [EssentiallySmall.{w} C] variable (B : Type u₄) [Category.{v₄} B] (F : A ⥤ B) + +section +variable [EssentiallySmall.{w} C] variable [HasSheafify ((equivSmallModel C).inverse.inducedTopology J) A] variable [((equivSmallModel C).inverse.inducedTopology J).HasSheafCompose F] @@ -205,11 +209,15 @@ instance hasColimitsEssentiallySmallSite Adjunction.has_colimits_of_equivalence ((equivSmallModel C).sheafCongr J ((equivSmallModel C).inverse.inducedTopology J) A).functor +end + namespace GrothendieckTopology variable {A} -variable [G.IsCoverDense J] [Functor.IsContinuous.{v₃} G K J] - [G.Full] [(G.sheafPushforwardContinuous A K J).EssSurj] +variable [G.IsCoverDense J] [G.Full] + +section +variable [Functor.IsContinuous.{v₃} G K J] [(G.sheafPushforwardContinuous A K J).EssSurj] open Localization @@ -249,9 +257,12 @@ lemma W_whiskerLeft_iff {P Q : Cᵒᵖ ⥤ A} (f : P ⟶ Q) : rw [← W_inverseImage_whiskeringLeft J K G] rfl -variable [Functor.IsContinuous.{v₄} G K J] [Functor.IsContinuous.{v₃} G K J] +end -lemma PreservesSheafification.transport [(G.sheafPushforwardContinuous B K J).EssSurj] +lemma PreservesSheafification.transport + [Functor.IsContinuous.{v₄} G K J] [Functor.IsContinuous.{v₃} G K J] + [(G.sheafPushforwardContinuous B K J).EssSurj] + [(G.sheafPushforwardContinuous A K J).EssSurj] [K.PreservesSheafification F] : J.PreservesSheafification F where le P Q f hf := by rw [← J.W_whiskerLeft_iff (G := G) (K := K)] at hf @@ -259,15 +270,18 @@ lemma PreservesSheafification.transport [(G.sheafPushforwardContinuous B K J).Es rwa [whiskerRight_left, K.W_whiskerLeft_iff (G := G) (J := J) (f := whiskerRight f F)] at this -variable [G.IsCocontinuous K J] (hG : CoverPreserving K J G) [ConcreteCategory A] +variable [Functor.IsContinuous.{v₃} G K J] [(G.sheafPushforwardContinuous A K J).EssSurj] +variable [G.IsCocontinuous K J] [ConcreteCategory A] [K.WEqualsLocallyBijective A] -lemma WEqualsLocallyBijective.transport : J.WEqualsLocallyBijective A where +lemma WEqualsLocallyBijective.transport (hG : CoverPreserving K J G) : + J.WEqualsLocallyBijective A where iff f := by rw [← W_whiskerLeft_iff J K G f, ← Presheaf.isLocallyInjective_whisker_iff K J G f hG, ← Presheaf.isLocallySurjective_whisker_iff K J G f hG, W_iff_isLocallyBijective] -variable [∀ (X : Cᵒᵖ), HasLimitsOfShape (StructuredArrow X (equivSmallModel C).inverse.op) A] +variable [EssentiallySmall.{w} C] + [∀ (X : Cᵒᵖ), HasLimitsOfShape (StructuredArrow X (equivSmallModel C).inverse.op) A] instance [((equivSmallModel C).inverse.inducedTopology J).WEqualsLocallyBijective A] : J.WEqualsLocallyBijective A := diff --git a/Mathlib/CategoryTheory/Sites/IsSheafOneHypercover.lean b/Mathlib/CategoryTheory/Sites/IsSheafOneHypercover.lean index 8a14ae283832b..2dee640b3fe0b 100644 --- a/Mathlib/CategoryTheory/Sites/IsSheafOneHypercover.lean +++ b/Mathlib/CategoryTheory/Sites/IsSheafOneHypercover.lean @@ -59,9 +59,7 @@ class IsGenerating : Prop where le {X : C} (S : Sieve X) (hS : S ∈ J X) : ∃ (E : J.OneHypercover X) (_ : H E), E.sieve₀ ≤ S -variable [H.IsGenerating] - -lemma exists_oneHypercover {X : C} (S : Sieve X) (hS : S ∈ J X) : +lemma exists_oneHypercover [H.IsGenerating] {X : C} (S : Sieve X) (hS : S ∈ J X) : ∃ (E : J.OneHypercover X) (_ : H E), E.sieve₀ ≤ S := IsGenerating.le _ hS @@ -71,7 +69,8 @@ namespace IsSheafIff variable (hP : ∀ ⦃X : C⦄ (E : J.OneHypercover X) (_ : H E), Nonempty (IsLimit (E.multifork P))) -lemma hom_ext {X : C} (S : Sieve X) (hS : S ∈ J X) {T : A} +include hP in +lemma hom_ext [H.IsGenerating] {X : C} (S : Sieve X) (hS : S ∈ J X) {T : A} {x y : T ⟶ P.obj (Opposite.op X)} (h : ∀ ⦃Y : C⦄ (f : Y ⟶ X) (_ : S f), x ≫ P.map f.op = y ≫ P.map f.op) : x = y := by @@ -99,7 +98,7 @@ lemma fac' (i : E.I₀) : F.ι ⟨_, E.f i, le _ (Sieve.ofArrows_mk _ _ _)⟩ := Multifork.IsLimit.fac (hP E hE).some _ _ i -lemma fac {Y : C} (f : Y ⟶ X) (hf : S f) : +lemma fac [H.IsGenerating] {Y : C} (f : Y ⟶ X) (hf : S f) : lift hP hE le F ≫ P.map f.op = F.ι ⟨Y, f, hf⟩ := by apply hom_ext H P hP _ (J.pullback_stable f E.mem₀) intro Z g @@ -113,7 +112,7 @@ lemma fac {Y : C} (f : Y ⟶ X) (hf : S f) : end /-- Auxiliary definition for `OneHypercoverFamily.isSheaf_iff`. -/ -noncomputable def isLimit : +noncomputable def isLimit [H.IsGenerating] : IsLimit (Cover.multifork ⟨S, J.superset_covering le E.mem₀⟩ P) := Multifork.IsLimit.mk _ (fun F => lift hP hE le F) @@ -128,7 +127,7 @@ noncomputable def isLimit : end IsSheafIff -lemma isSheaf_iff : +lemma isSheaf_iff [H.IsGenerating] : Presheaf.IsSheaf J P ↔ ∀ ⦃X : C⦄ (E : J.OneHypercover X) (_ : H E), Nonempty (IsLimit (E.multifork P)) := by constructor diff --git a/Mathlib/CategoryTheory/Sites/LocallyInjective.lean b/Mathlib/CategoryTheory/Sites/LocallyInjective.lean index 94740e0e9b99d..a537edefc713a 100644 --- a/Mathlib/CategoryTheory/Sites/LocallyInjective.lean +++ b/Mathlib/CategoryTheory/Sites/LocallyInjective.lean @@ -219,6 +219,11 @@ instance isLocallyInjective_of_iso [IsIso φ] : IsLocallyInjective φ := by change Presheaf.IsLocallyInjective J ((sheafToPresheaf _ _).map φ) infer_instance +lemma mono_of_injective + (hφ : ∀ (X : Cᵒᵖ), Function.Injective (φ.val.app X)) : Mono φ := + have := fun X ↦ ConcreteCategory.mono_of_injective _ (hφ X) + (sheafToPresheaf _ _).mono_of_mono_map (NatTrans.mono_of_mono_app φ.1) + variable [J.HasSheafCompose (forget D)] instance isLocallyInjective_forget [IsLocallyInjective φ] : @@ -232,11 +237,6 @@ lemma isLocallyInjective_iff_injective : rw [← isSheaf_iff_isSheaf_of_type] exact ((sheafCompose J (forget D)).obj F₁).2) -lemma mono_of_injective - (hφ : ∀ (X : Cᵒᵖ), Function.Injective (φ.val.app X)) : Mono φ := - have := fun X ↦ ConcreteCategory.mono_of_injective _ (hφ X) - (sheafToPresheaf _ _).mono_of_mono_map (NatTrans.mono_of_mono_app φ.1) - lemma mono_of_isLocallyInjective [IsLocallyInjective φ] : Mono φ := by apply mono_of_injective rw [← isLocallyInjective_iff_injective] diff --git a/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean b/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean index d4273af81064a..c0bacd71db3d6 100644 --- a/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean +++ b/Mathlib/CategoryTheory/Sites/MayerVietorisSquare.lean @@ -96,6 +96,7 @@ namespace SheafCondition variable {S} variable {P : Cᵒᵖ ⥤ Type v'} (h : S.SheafCondition P) +include h lemma bijective_toPullbackObj : Function.Bijective (S.toPullbackObj P) := by rwa [← sheafCondition_iff_bijective_toPullbackObj] diff --git a/Mathlib/CategoryTheory/Sites/Preserves.lean b/Mathlib/CategoryTheory/Sites/Preserves.lean index 4ef0bcad8b573..1fdb884f4aeb3 100644 --- a/Mathlib/CategoryTheory/Sites/Preserves.lean +++ b/Mathlib/CategoryTheory/Sites/Preserves.lean @@ -72,27 +72,7 @@ section Product variable (hI : IsInitial I) -- This is the data of a particular disjoint coproduct in `C`. -variable {α : Type} {X : α → C} (c : Cofan X) (hc : IsColimit c) [(ofArrows X c.inj).hasPullbacks] - [HasInitial C] [∀ i, Mono (c.inj i)] - (hd : Pairwise fun i j => IsPullback (initial.to _) (initial.to _) (c.inj i) (c.inj j)) - -/-- -The two parallel maps in the equalizer diagram for the sheaf condition corresponding to the -inclusion maps in a disjoint coproduct are equal. --/ -theorem firstMap_eq_secondMap : Equalizer.Presieve.Arrows.firstMap F X c.inj = - Equalizer.Presieve.Arrows.secondMap F X c.inj := by - ext a ⟨i, j⟩ - simp only [Equalizer.Presieve.Arrows.firstMap, Types.pi_lift_π_apply, types_comp_apply, - Equalizer.Presieve.Arrows.secondMap] - by_cases hi : i = j - · rw [hi, Mono.right_cancellation _ _ pullback.condition] - · have := preservesTerminalOfIsSheafForEmpty F hF hI - apply_fun (F.mapIso ((hd hi).isoPullback).op ≪≫ F.mapIso (terminalIsoIsTerminal - (terminalOpOfInitial initialIsInitial)).symm ≪≫ (PreservesTerminal.iso F)).hom using - injective_of_mono _ - ext ⟨i⟩ - exact i.elim +variable {α : Type} {X : α → C} (c : Cofan X) (hc : IsColimit c) theorem piComparison_fac : have : HasCoproduct X := ⟨⟨c, hc⟩⟩ @@ -108,21 +88,9 @@ theorem piComparison_fac : rw [hh, ← desc_op_comp_opCoproductIsoProduct'_hom hc] simp -/-- -If `F` is a presheaf which `IsSheafFor` a presieve of arrows and the empty presieve, then it -preserves the product corresponding to the presieve of arrows. --/ -noncomputable -def preservesProductOfIsSheafFor (hF' : (ofArrows X c.inj).IsSheafFor F) : - PreservesLimit (Discrete.functor (fun x ↦ op (X x))) F := by - have : HasCoproduct X := ⟨⟨c, hc⟩⟩ - refine @PreservesProduct.ofIsoComparison _ _ _ _ F _ (fun x ↦ op (X x)) _ _ ?_ - rw [piComparison_fac (hc := hc)] - refine @IsIso.comp_isIso _ _ _ _ _ _ _ inferInstance ?_ - rw [isIso_iff_bijective, Function.bijective_iff_existsUnique] - rw [Equalizer.Presieve.Arrows.sheaf_condition, Limits.Types.type_equalizer_iff_unique] at hF' - exact fun b ↦ hF' b (congr_fun (firstMap_eq_secondMap F hF hI c hd) b) +variable [(ofArrows X c.inj).hasPullbacks] +include hc in /-- If `F` preserves a particular product, then it `IsSheafFor` the corresponging presieve of arrows. -/ @@ -144,6 +112,48 @@ theorem isSheafFor_of_preservesProduct [PreservesLimit (Discrete.functor (fun x congr simp [Iso.eq_inv_comp, ← Category.assoc, ← op_comp, eq_comm, ← Iso.eq_comp_inv] +variable [HasInitial C] [∀ i, Mono (c.inj i)] + (hd : Pairwise fun i j => IsPullback (initial.to _) (initial.to _) (c.inj i) (c.inj j)) +include hd + +include hF hI in +/-- +The two parallel maps in the equalizer diagram for the sheaf condition corresponding to the +inclusion maps in a disjoint coproduct are equal. +-/ +theorem firstMap_eq_secondMap : + Equalizer.Presieve.Arrows.firstMap F X c.inj = + Equalizer.Presieve.Arrows.secondMap F X c.inj := by + ext a ⟨i, j⟩ + simp only [Equalizer.Presieve.Arrows.firstMap, Types.pi_lift_π_apply, types_comp_apply, + Equalizer.Presieve.Arrows.secondMap] + by_cases hi : i = j + · rw [hi, Mono.right_cancellation _ _ pullback.condition] + · have := preservesTerminalOfIsSheafForEmpty F hF hI + apply_fun (F.mapIso ((hd hi).isoPullback).op ≪≫ F.mapIso (terminalIsoIsTerminal + (terminalOpOfInitial initialIsInitial)).symm ≪≫ (PreservesTerminal.iso F)).hom using + injective_of_mono _ + ext ⟨i⟩ + exact i.elim + +/-- +If `F` is a presheaf which `IsSheafFor` a presieve of arrows and the empty presieve, then it +preserves the product corresponding to the presieve of arrows. +-/ +noncomputable +def preservesProductOfIsSheafFor + (hd : Pairwise fun i j => IsPullback (initial.to _) (initial.to _) (c.inj i) (c.inj j)) + (hF' : (ofArrows X c.inj).IsSheafFor F) : + PreservesLimit (Discrete.functor (fun x ↦ op (X x))) F := by + have : HasCoproduct X := ⟨⟨c, hc⟩⟩ + refine @PreservesProduct.ofIsoComparison _ _ _ _ F _ (fun x ↦ op (X x)) _ _ ?_ + rw [piComparison_fac (hc := hc)] + refine @IsIso.comp_isIso _ _ _ _ _ _ _ inferInstance ?_ + rw [isIso_iff_bijective, Function.bijective_iff_existsUnique] + rw [Equalizer.Presieve.Arrows.sheaf_condition, Limits.Types.type_equalizer_iff_unique] at hF' + exact fun b ↦ hF' b (congr_fun (firstMap_eq_secondMap F hF hI c hd) b) + +include hF hI hc in theorem isSheafFor_iff_preservesProduct : (ofArrows X c.inj).IsSheafFor F ↔ Nonempty (PreservesLimit (Discrete.functor (fun x ↦ op (X x))) F) := by refine ⟨fun hF' ↦ ⟨preservesProductOfIsSheafFor _ hF hI c hc hd hF'⟩, fun hF' ↦ ?_⟩ diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index bc886bdc1fcf9..2de13cf6759ad 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -265,6 +265,7 @@ variable {P : Cᵒᵖ ⥤ A} (hP : Presheaf.IsSheaf J P) {I : Type*} {S : C} {X (x : ∀ i, E ⟶ P.obj (op (X i))) (hx : ∀ ⦃W : C⦄ ⦃i j : I⦄ (a : W ⟶ X i) (b : W ⟶ X j), a ≫ f i = b ≫ f j → x i ≫ P.map a.op = x j ≫ P.map b.op) +include hP hf hx lemma IsSheaf.exists_unique_amalgamation_ofArrows : ∃! (g : E ⟶ P.obj (op S)), ∀ (i : I), g ≫ P.map (f i).op = x i := diff --git a/Mathlib/CategoryTheory/Sites/SheafHom.lean b/Mathlib/CategoryTheory/Sites/SheafHom.lean index ff0ab98ee6de9..aeec15f07d717 100644 --- a/Mathlib/CategoryTheory/Sites/SheafHom.lean +++ b/Mathlib/CategoryTheory/Sites/SheafHom.lean @@ -121,10 +121,10 @@ variable {X : C} {S : Sieve X} namespace PresheafHom.IsSheafFor -variable (x : Presieve.FamilyOfElements (presheafHom F G) S.arrows) (hx : x.Compatible) - {Y : C} (g : Y ⟶ X) +variable (x : Presieve.FamilyOfElements (presheafHom F G) S.arrows) {Y : C} -lemma exists_app : +include hG in +lemma exists_app (hx : x.Compatible) (g : Y ⟶ X) : ∃ (φ : F.obj (op Y) ⟶ G.obj (op Y)), ∀ {Z : C} (p : Z ⟶ Y) (hp : S (p ≫ g)), φ ≫ G.map p.op = F.map p.op ≫ (x (p ≫ g) hp).app ⟨Over.mk (𝟙 Z)⟩ := by @@ -149,9 +149,10 @@ lemma exists_app : exact ((hG g).fac c ⟨Over.mk p, hp⟩) /-- Auxiliary definition for `presheafHom_isSheafFor`. -/ -noncomputable def app : F.obj (op Y) ⟶ G.obj (op Y) := (exists_app hG x hx g).choose +noncomputable def app (hx : x.Compatible) (g : Y ⟶ X) : F.obj (op Y) ⟶ G.obj (op Y) := + (exists_app hG x hx g).choose -lemma app_cond {Z : C} (p : Z ⟶ Y) (hp : S (p ≫ g)) : +lemma app_cond (hx : x.Compatible) (g : Y ⟶ X) {Z : C} (p : Z ⟶ Y) (hp : S (p ≫ g)) : app hG x hx g ≫ G.map p.op = F.map p.op ≫ (x (p ≫ g) hp).app ⟨Over.mk (𝟙 Z)⟩ := (exists_app hG x hx g).choose_spec p hp @@ -159,6 +160,7 @@ end PresheafHom.IsSheafFor variable (F G S) +include hG in open PresheafHom.IsSheafFor in lemma presheafHom_isSheafFor : Presieve.IsSheafFor (presheafHom F G) S.arrows := by diff --git a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean index f5584c6bfe3e0..76cdea5e2bc17 100644 --- a/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean +++ b/Mathlib/CategoryTheory/Triangulated/HomologicalFunctor.lean @@ -182,6 +182,8 @@ variable [HasZeroObject C] [Preadditive C] [∀ (n : ℤ), (CategoryTheory.shift variable [F.ShiftSequence ℤ] (T T' : Triangle C) (hT : T ∈ distTriang C) (hT' : T' ∈ distTriang C) (φ : T ⟶ T') (n₀ n₁ : ℤ) (h : n₀ + 1 = n₁) +section +include hT @[reassoc] lemma comp_homologySequenceδ : (F.shift n₀).map T.mor₂ ≫ F.homologySequenceδ T n₀ n₁ h = 0 := by @@ -238,6 +240,7 @@ lemma homologySequence_epi_shift_map_mor₂_iff : lemma homologySequence_mono_shift_map_mor₂_iff : Mono ((F.shift n₀).map T.mor₂) ↔ (F.shift n₀).map T.mor₁ = 0 := (F.homologySequence_exact₂ T hT n₀).mono_g_iff +end lemma mem_homologicalKernel_W_iff {X Y : C} (f : X ⟶ Y) : F.homologicalKernel.W f ↔ ∀ (n : ℤ), IsIso ((F.shift n).map f) := by @@ -265,6 +268,7 @@ open ComposableArrows mk₅ ((F.shift n₀).map T.mor₁) ((F.shift n₀).map T.mor₂) (F.homologySequenceδ T n₀ n₁ h) ((F.shift n₁).map T.mor₁) ((F.shift n₁).map T.mor₂) +include hT in lemma homologySequenceComposableArrows₅_exact : (F.homologySequenceComposableArrows₅ T n₀ n₁ h).Exact := exact_of_δ₀ (F.homologySequence_exact₂ T hT n₀).exact_toComposableArrows diff --git a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean index 077faaf31d4b1..091de90c8c3c1 100644 --- a/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean +++ b/Mathlib/CategoryTheory/Triangulated/Pretriangulated.lean @@ -233,6 +233,7 @@ lemma contractible_distinguished₂ (X : C) : namespace Triangle variable (T : Triangle C) (hT : T ∈ distTriang C) +include hT lemma yoneda_exact₂ {X : C} (f : T.obj₂ ⟶ X) (hf : T.mor₁ ≫ f = 0) : ∃ (g : T.obj₃ ⟶ X), f = T.mor₂ ≫ g := by diff --git a/Mathlib/CategoryTheory/Triangulated/Subcategory.lean b/Mathlib/CategoryTheory/Triangulated/Subcategory.lean index a1bd5112f2909..f4c688ba57647 100644 --- a/Mathlib/CategoryTheory/Triangulated/Subcategory.lean +++ b/Mathlib/CategoryTheory/Triangulated/Subcategory.lean @@ -260,6 +260,8 @@ section variable (T : Triangle C) (hT : T ∈ distTriang C) +include hT + lemma ext₁ [ClosedUnderIsomorphisms S.P] (h₂ : S.P T.obj₂) (h₃ : S.P T.obj₃) : S.P T.obj₁ := S.ext₂ _ (inv_rot_of_distTriang _ hT) (S.shift _ _ h₃) h₂ diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean index 27372279d8fdb..ad0a9bc39e113 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Behrend.lean @@ -153,7 +153,6 @@ theorem map_eq_iff {x₁ x₂ : Fin n.succ → ℕ} (hx₁ : ∀ i, x₁ i < d) exact ⟨this, h.resolve_right (pos_of_gt (hx₁ 0)).ne'⟩ theorem map_injOn : {x : Fin n → ℕ | ∀ i, x i < d}.InjOn (map d) := by - clear x intro x₁ hx₁ x₂ hx₂ h induction' n with n ih · simp [eq_iff_true_of_subsingleton] diff --git a/Mathlib/Combinatorics/SimpleGraph/Maps.lean b/Mathlib/Combinatorics/SimpleGraph/Maps.lean index 8ccd1909b18de..93f6348a4d897 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Maps.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Maps.lean @@ -501,6 +501,7 @@ def mapNeighborSet (v : V) : G.neighborSet v ≃ G'.neighborSet (f v) where left_inv w := by simp right_inv w := by simp +include f in theorem card_eq [Fintype V] [Fintype W] : Fintype.card V = Fintype.card W := by rw [← Fintype.ofEquiv_card f.toEquiv] convert rfl diff --git a/Mathlib/Combinatorics/SimpleGraph/Operations.lean b/Mathlib/Combinatorics/SimpleGraph/Operations.lean index 253482d421220..814a10a5cc143 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Operations.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Operations.lean @@ -32,6 +32,7 @@ namespace Iso variable {G} {W : Type*} {G' : SimpleGraph W} (f : G ≃g G') +include f in theorem card_edgeFinset_eq [Fintype G.edgeSet] [Fintype G'.edgeSet] : G.edgeFinset.card = G'.edgeFinset.card := by apply Finset.card_eq_of_equiv diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean index fdcf18b1b5ddd..dac9907861a7d 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean @@ -186,7 +186,6 @@ variable (t) lemma map_toTriangle_disjoint [ExplicitDisjoint t] : (t.map toTriangle : Set (Finset (α ⊕ β ⊕ γ))).Pairwise fun x y ↦ (x ∩ y : Set (α ⊕ β ⊕ γ)).Subsingleton := by - clear x intro simp only [Finset.coe_map, Set.mem_image, Finset.mem_coe, Prod.exists, Ne, forall_exists_index, and_imp] diff --git a/Mathlib/Combinatorics/SimpleGraph/Turan.lean b/Mathlib/Combinatorics/SimpleGraph/Turan.lean index d97bad0845370..e83caeb3b7f52 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Turan.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Turan.lean @@ -81,9 +81,7 @@ theorem turanGraph_eq_top : turanGraph n r = ⊤ ↔ r = 0 ∨ n ≤ r := by · simp [Fin.val_inj] · rw [Nat.mod_eq_of_lt (a.2.trans_le h), Nat.mod_eq_of_lt (b.2.trans_le h), Fin.val_inj] -variable (hr : 0 < r) - -theorem turanGraph_cliqueFree : (turanGraph n r).CliqueFree (r + 1) := by +theorem turanGraph_cliqueFree (hr : 0 < r) : (turanGraph n r).CliqueFree (r + 1) := by rw [cliqueFree_iff] by_contra h rw [not_isEmpty_iff] at h @@ -106,7 +104,7 @@ theorem not_cliqueFree_of_isTuranMaximal (hn : r ≤ Fintype.card V) (hG : G.IsT exact hGab <| le_sup_right.trans_eq ((hG.le_iff_eq <| h.sup_edge _ _).1 le_sup_left).symm <| (edge_adj ..).2 ⟨Or.inl ⟨rfl, rfl⟩, hab⟩ -lemma exists_isTuranMaximal : +lemma exists_isTuranMaximal (hr : 0 < r): ∃ H : SimpleGraph V, ∃ _ : DecidableRel H.Adj, H.IsTuranMaximal r := by classical let c := {H : SimpleGraph V | H.CliqueFree (r + 1)} @@ -126,7 +124,7 @@ end Defs namespace IsTuranMaximal -variable {s t u : V} +variable {s t u : V} [DecidableEq V] /-- In a Turán-maximal graph, non-adjacent vertices have the same degree. -/ lemma degree_eq_of_not_adj (h : G.IsTuranMaximal r) (hn : ¬G.Adj s t) : @@ -170,6 +168,7 @@ lemma not_adj_trans (h : G.IsTuranMaximal r) (hts : ¬G.Adj t s) (hsu : ¬G.Adj omega variable (h : G.IsTuranMaximal r) +include h /-- In a Turán-maximal graph, non-adjacency is an equivalence relation. -/ theorem equivalence_not_adj : Equivalence (¬G.Adj · ·) where @@ -187,8 +186,6 @@ instance : DecidableRel h.setoid.r := /-- The finpartition derived from `h.setoid`. -/ def finpartition [DecidableEq V] : Finpartition (univ : Finset V) := Finpartition.ofSetoid h.setoid -variable [DecidableEq V] - lemma not_adj_iff_part_eq : ¬G.Adj s t ↔ h.finpartition.part s = h.finpartition.part t := by change h.setoid.r s t ↔ _ @@ -265,7 +262,7 @@ theorem card_parts : h.finpartition.parts.card = min (Fintype.card V) r := by /-- **Turán's theorem**, forward direction. Any `r + 1`-cliquefree Turán-maximal graph on `n` vertices is isomorphic to `turanGraph n r`. -/ -theorem nonempty_iso_turanGraph (h : G.IsTuranMaximal r) : +theorem nonempty_iso_turanGraph : Nonempty (G ≃g turanGraph (Fintype.card V) r) := by classical obtain ⟨zm, zp⟩ := h.isEquipartition.exists_partPreservingEquiv diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index a72ceadbd3573..52c4b7eeca310 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -901,7 +901,7 @@ theorem count_edges_takeUntil_le_one {u v w : V} (p : G.Walk v w) (h : u ∈ p.s simp · rw [edges_cons, List.count_cons] split_ifs with h'' - · simp only [beq_iff_eq, Sym2.eq, Sym2.rel_iff', Prod.mk.injEq, Prod.swap_prod_mk] at h'' + · simp only [beq_iff_eq, Sym2.eq, Sym2.rel_iff'] at h'' obtain ⟨rfl, rfl⟩ | ⟨rfl, rfl⟩ := h'' · exact (h' rfl).elim · cases p' <;> simp! @@ -1071,7 +1071,7 @@ theorem map_copy (hu : u = u') (hv : v = v') : theorem map_id (p : G.Walk u v) : p.map Hom.id = p := by induction p with | nil => rfl - | cons _ p' ih => simp [ih p'] + | cons _ p' ih => simp [ih] @[simp] theorem map_map : (p.map f).map f' = p.map (f'.comp f) := by diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index f324a7c9a4781..d25df9001edd4 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -127,14 +127,6 @@ lemma max_bi_le {b : α → ℝ} (i : α) : b i ≤ b (max_bi b) := end min_max -variable {α : Type*} [Fintype α] [Nonempty α] {T : ℕ → ℝ} {g : ℝ → ℝ} {a b : α → ℝ} {r : α → ℕ → ℕ} - (R : AkraBazziRecurrence T g a b r) - -lemma dist_r_b' : ∀ᶠ n in atTop, ∀ i, ‖(r i n : ℝ) - b i * n‖ ≤ n / log n ^ 2 := by - rw [Filter.eventually_all] - intro i - simpa using IsLittleO.eventuallyLE (R.dist_r_b i) - lemma isLittleO_self_div_log_id : (fun (n : ℕ) => n / log n ^ 2) =o[atTop] (fun (n : ℕ) => (n : ℝ)) := by calc (fun (n : ℕ) => (n : ℝ) / log n ^ 2) = fun (n : ℕ) => (n : ℝ) * ((log n) ^ 2)⁻¹ := by @@ -150,6 +142,16 @@ lemma isLittleO_self_div_log_id : <| isLittleO_const_log_atTop) (by norm_num) _ = (fun (n : ℕ) => (n : ℝ)) := by ext; simp +variable {α : Type*} [Fintype α] {T : ℕ → ℝ} {g : ℝ → ℝ} {a b : α → ℝ} {r : α → ℕ → ℕ} +variable [Nonempty α] (R : AkraBazziRecurrence T g a b r) +section +include R + +lemma dist_r_b' : ∀ᶠ n in atTop, ∀ i, ‖(r i n : ℝ) - b i * n‖ ≤ n / log n ^ 2 := by + rw [Filter.eventually_all] + intro i + simpa using IsLittleO.eventuallyLE (R.dist_r_b i) + lemma eventually_b_le_r : ∀ᶠ (n : ℕ) in atTop, ∀ i, (b i : ℝ) * n - (n / log n ^ 2) ≤ r i n := by filter_upwards [R.dist_r_b'] with n hn intro i @@ -277,6 +279,8 @@ lemma eventually_log_b_mul_pos : ∀ᶠ (n : ℕ) in atTop, ∀ i, 0 < log (b i @[aesop safe apply] lemma T_nonneg (n : ℕ) : 0 ≤ T n := le_of_lt <| R.T_pos n +end + /-! #### Smoothing function @@ -349,6 +353,7 @@ lemma eventually_one_sub_smoothingFn_pos : ∀ᶠ (n : ℕ) in atTop, 0 < 1 - ε lemma eventually_one_sub_smoothingFn_nonneg : ∀ᶠ (n : ℕ) in atTop, 0 ≤ 1 - ε n := by filter_upwards [eventually_one_sub_smoothingFn_pos] with n hn; exact le_of_lt hn +include R in lemma eventually_one_sub_smoothingFn_r_pos : ∀ᶠ (n : ℕ) in atTop, ∀ i, 0 < 1 - ε (r i n) := by rw [Filter.eventually_all] exact fun i => (R.tendsto_atTop_r_real i).eventually eventually_one_sub_smoothingFn_pos_real @@ -434,6 +439,7 @@ lemma eventually_one_add_smoothingFn_pos : ∀ᶠ (n : ℕ) in atTop, 0 < 1 + ε show 0 < 1 + 1 / log x positivity +include R in lemma eventually_one_add_smoothingFn_r_pos : ∀ᶠ (n : ℕ) in atTop, ∀ i, 0 < 1 + ε (r i n) := by rw [Filter.eventually_all] exact fun i => (R.tendsto_atTop_r i).eventually (f := r i) eventually_one_add_smoothingFn_pos @@ -456,6 +462,9 @@ lemma strictMonoOn_one_sub_smoothingFn : lemma strictAntiOn_one_add_smoothingFn : StrictAntiOn (fun (x : ℝ) => (1 : ℝ) + ε x) (Set.Ioi 1) := StrictAntiOn.const_add strictAntiOn_smoothingFn 1 +section +include R + lemma isEquivalent_smoothingFn_sub_self (i : α) : (fun (n : ℕ) => ε (b i * n) - ε n) ~[atTop] fun n => -log (b i) / (log n)^2 := by calc (fun (n : ℕ) => 1 / log (b i * n) - 1 / log n) @@ -543,10 +552,13 @@ lemma one_mem_range_sumCoeffsExp : 1 ∈ Set.range (fun (p : ℝ) => ∑ i, a i lemma injective_sumCoeffsExp : Function.Injective (fun (p : ℝ) => ∑ i, a i * (b i) ^ p) := R.strictAnti_sumCoeffsExp.injective +end + variable (a b) in /-- The exponent `p` associated with a particular Akra-Bazzi recurrence. -/ noncomputable irreducible_def p : ℝ := Function.invFun (fun (p : ℝ) => ∑ i, a i * (b i) ^ p) 1 +include R in @[simp] lemma sumCoeffsExp_p_eq_one : ∑ i, a i * (b i) ^ p a b = 1 := by simp only [p] @@ -576,14 +588,18 @@ variable (g) (a) (b) `n^p (1 + ∑_{u < n} g(u) / u^(p+1))`. -/ noncomputable def asympBound (n : ℕ) : ℝ := n ^ p a b + sumTransform (p a b) g 0 n -lemma asympBound_def {n : ℕ} : asympBound g a b n = n ^ p a b + sumTransform (p a b) g 0 n := rfl +lemma asympBound_def {α} [Fintype α] (a b : α → ℝ) {n : ℕ} : + asympBound g a b n = n ^ p a b + sumTransform (p a b) g 0 n := rfl variable {g} {a} {b} -lemma asympBound_def' {n : ℕ} : +lemma asympBound_def' {α} [Fintype α] (a b : α → ℝ) {n : ℕ} : asympBound g a b n = n ^ p a b * (1 + (∑ u ∈ range n, g u / u ^ (p a b + 1))) := by simp [asympBound_def, sumTransform, mul_add, mul_one, Finset.sum_Ico_eq_sum_range] +section +include R + lemma asympBound_pos (n : ℕ) (hn : 0 < n) : 0 < asympBound g a b n := by calc 0 < (n : ℝ) ^ p a b * (1 + 0) := by aesop (add safe Real.rpow_pos_of_pos) _ ≤ asympBound g a b n := by @@ -769,6 +785,8 @@ lemma eventually_atTop_sumTransform_ge : _ ≥ min (c₂ * (1 - c₃)) ((1 - c₃) * c₂ / c₁ ^ ((p a b) + 1)) * g n := by gcongr; exact min_le_right _ _ +end + /-! #### Technical lemmas @@ -776,39 +794,6 @@ The next several lemmas are technical lemmas leading up to `rpow_p_mul_one_sub_s `rpow_p_mul_one_add_smoothingFn_ge`, which are key steps in the main proof. -/ -lemma isBigO_apply_r_sub_b (q : ℝ → ℝ) (hq_diff : DifferentiableOn ℝ q (Set.Ioi 1)) - (hq_poly : GrowsPolynomially fun x => ‖deriv q x‖) (i : α) : - (fun n => q (r i n) - q (b i * n)) =O[atTop] fun n => (deriv q n) * (r i n - b i * n) := by - let b' := b (min_bi b) / 2 - have hb_pos : 0 < b' := by have := R.b_pos (min_bi b); positivity - have hb_lt_one : b' < 1 := calc - b (min_bi b) / 2 < b (min_bi b) := by exact div_two_lt_of_pos (R.b_pos (min_bi b)) - _ < 1 := R.b_lt_one (min_bi b) - have hb : b' ∈ Set.Ioo 0 1 := ⟨hb_pos, hb_lt_one⟩ - have hb' : ∀ i, b' ≤ b i := fun i => calc - b (min_bi b) / 2 ≤ b i / 2 := by gcongr; aesop - _ ≤ b i := by exact le_of_lt <| div_two_lt_of_pos (R.b_pos i) - obtain ⟨c₁, _, c₂, _, hq_poly⟩ := hq_poly b' hb - rw [isBigO_iff] - refine ⟨c₂, ?_⟩ - have h_tendsto : Tendsto (fun x => b' * x) atTop atTop := - Tendsto.const_mul_atTop hb_pos tendsto_id - filter_upwards [hq_poly.natCast_atTop, R.eventually_bi_mul_le_r, eventually_ge_atTop R.n₀, - eventually_gt_atTop 0, (h_tendsto.eventually_gt_atTop 1).natCast_atTop] with - n hn h_bi_le_r h_ge_n₀ h_n_pos h_bn - rw [norm_mul, ← mul_assoc] - refine Convex.norm_image_sub_le_of_norm_deriv_le - (s := Set.Icc (b'*n) n) (fun z hz => ?diff) (fun z hz => (hn z hz).2) - (convex_Icc _ _) ?mem_Icc <| ⟨h_bi_le_r i, by exact_mod_cast (le_of_lt (R.r_lt_n i n h_ge_n₀))⟩ - case diff => - refine hq_diff.differentiableAt (Ioi_mem_nhds ?_) - calc 1 < b' * n := by exact h_bn - _ ≤ z := hz.1 - case mem_Icc => - refine ⟨by gcongr; exact hb' i, ?_⟩ - calc b i * n ≤ 1 * n := by gcongr; exact le_of_lt <| R.b_lt_one i - _ = n := by simp - lemma eventually_deriv_rpow_p_mul_one_sub_smoothingFn (p : ℝ) : deriv (fun z => z ^ p * (1 - ε z)) =ᶠ[atTop] fun z => p * z ^ (p-1) * (1 - ε z) + z ^ (p-1) / (log z ^ 2) := calc @@ -952,6 +937,41 @@ lemma growsPolynomially_deriv_rpow_p_mul_one_add_smoothingFn (p : ℝ) : filter_upwards [eventually_gt_atTop 0] with _ _ positivity +include R + +lemma isBigO_apply_r_sub_b (q : ℝ → ℝ) (hq_diff : DifferentiableOn ℝ q (Set.Ioi 1)) + (hq_poly : GrowsPolynomially fun x => ‖deriv q x‖) (i : α) : + (fun n => q (r i n) - q (b i * n)) =O[atTop] fun n => (deriv q n) * (r i n - b i * n) := by + let b' := b (min_bi b) / 2 + have hb_pos : 0 < b' := by have := R.b_pos (min_bi b); positivity + have hb_lt_one : b' < 1 := calc + b (min_bi b) / 2 < b (min_bi b) := by exact div_two_lt_of_pos (R.b_pos (min_bi b)) + _ < 1 := R.b_lt_one (min_bi b) + have hb : b' ∈ Set.Ioo 0 1 := ⟨hb_pos, hb_lt_one⟩ + have hb' : ∀ i, b' ≤ b i := fun i => calc + b (min_bi b) / 2 ≤ b i / 2 := by gcongr; aesop + _ ≤ b i := by exact le_of_lt <| div_two_lt_of_pos (R.b_pos i) + obtain ⟨c₁, _, c₂, _, hq_poly⟩ := hq_poly b' hb + rw [isBigO_iff] + refine ⟨c₂, ?_⟩ + have h_tendsto : Tendsto (fun x => b' * x) atTop atTop := + Tendsto.const_mul_atTop hb_pos tendsto_id + filter_upwards [hq_poly.natCast_atTop, R.eventually_bi_mul_le_r, eventually_ge_atTop R.n₀, + eventually_gt_atTop 0, (h_tendsto.eventually_gt_atTop 1).natCast_atTop] with + n hn h_bi_le_r h_ge_n₀ h_n_pos h_bn + rw [norm_mul, ← mul_assoc] + refine Convex.norm_image_sub_le_of_norm_deriv_le + (s := Set.Icc (b'*n) n) (fun z hz => ?diff) (fun z hz => (hn z hz).2) + (convex_Icc _ _) ?mem_Icc <| ⟨h_bi_le_r i, by exact_mod_cast (le_of_lt (R.r_lt_n i n h_ge_n₀))⟩ + case diff => + refine hq_diff.differentiableAt (Ioi_mem_nhds ?_) + calc 1 < b' * n := by exact h_bn + _ ≤ z := hz.1 + case mem_Icc => + refine ⟨by gcongr; exact hb' i, ?_⟩ + calc b i * n ≤ 1 * n := by gcongr; exact le_of_lt <| R.b_lt_one i + _ = n := by simp + lemma rpow_p_mul_one_sub_smoothingFn_le : ∀ᶠ (n : ℕ) in atTop, ∀ i, (r i n) ^ (p a b) * (1 - ε (r i n)) ≤ (b i) ^ (p a b) * n ^ (p a b) * (1 - ε n) := by diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 8523938c673cd..020bd9ea089ba 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -1620,6 +1620,7 @@ theorem stepAux_write (q : Stmt'₁) (v : σ) (a b : Γ) (L R : ListBlank Γ) : ListBlank.tail_cons] variable (encdec : ∀ a, dec (enc a) = a) +include encdec theorem stepAux_read (f : Γ → Stmt'₁) (v : σ) (L R : ListBlank Γ) : stepAux (read dec f) v (trTape' enc0 L R) = stepAux (f R.head) v (trTape' enc0 L R) := by @@ -1651,8 +1652,9 @@ theorem stepAux_read (f : Γ → Stmt'₁) (v : σ) (L R : ListBlank Γ) : rw [← ListBlank.append, IH] rfl -theorem tr_respects {enc₀} : - Respects (step M) (step (tr enc dec M)) fun c₁ c₂ ↦ trCfg enc enc₀ c₁ = c₂ := +variable {enc0} in +theorem tr_respects : + Respects (step M) (step (tr enc dec M)) fun c₁ c₂ ↦ trCfg enc enc0 c₁ = c₂ := fun_respects.2 fun ⟨l₁, v, T⟩ ↦ by obtain ⟨L, R, rfl⟩ := T.exists_mk' cases' l₁ with l₁ @@ -2433,7 +2435,7 @@ theorem tr_respects_aux {q v T k} {S : ∀ k, List (Γ k)} (TM1.stepAux (trNormal (stRun o q)) v (Tape.mk' ∅ (addBottom T))) b := by simp only [trNormal_run, step_run] have hgo := tr_respects_aux₁ M o q v (hT k) _ le_rfl - obtain ⟨T', hT', hrun⟩ := tr_respects_aux₂ hT o + obtain ⟨T', hT', hrun⟩ := tr_respects_aux₂ (Λ := Λ) hT o have := hgo.tail' rfl rw [tr, TM1.stepAux, Tape.move_right_n_head, Tape.mk'_nth_nat, addBottom_nth_snd, stk_nth_val _ (hT k), List.get?_len_le (le_of_eq (List.length_reverse _)), Option.isNone, cond, diff --git a/Mathlib/Data/FunLike/Equiv.lean b/Mathlib/Data/FunLike/Equiv.lean index f99cdf26ac4fd..d2a640eddc593 100644 --- a/Mathlib/Data/FunLike/Equiv.lean +++ b/Mathlib/Data/FunLike/Equiv.lean @@ -218,7 +218,7 @@ theorem comp_bijective (f : α → β) (e : F) : Function.Bijective (e ∘ f) (EquivLike.bijective e).of_comp_iff' f /-- This is not an instance to avoid slowing down every single `Subsingleton` typeclass search. -/ -lemma subsingleton_dom [Subsingleton β] : Subsingleton F := +lemma subsingleton_dom [FunLike F β γ] [Subsingleton β] : Subsingleton F := ⟨fun f g ↦ DFunLike.ext f g fun _ ↦ (right_inv f).injective <| Subsingleton.elim _ _⟩ end EquivLike diff --git a/Mathlib/Data/Int/CardIntervalMod.lean b/Mathlib/Data/Int/CardIntervalMod.lean index b4c80f203c4db..8b80171fc546a 100644 --- a/Mathlib/Data/Int/CardIntervalMod.lean +++ b/Mathlib/Data/Int/CardIntervalMod.lean @@ -38,6 +38,7 @@ lemma Ioc_filter_modEq_eq (v : ℤ) : (Ioc a b).filter (· ≡ v [ZMOD r]) = exists_eq_right, modEq_comm, modEq_iff_dvd, sub_lt_sub_iff_right, sub_le_sub_iff_right] variable (hr : 0 < r) +include hr lemma Ico_filter_dvd_eq : (Ico a b).filter (r ∣ ·) = (Ico ⌈a / (r : ℚ)⌉ ⌈b / (r : ℚ)⌉).map ⟨(· * r), mul_left_injective₀ hr.ne'⟩ := by @@ -98,6 +99,7 @@ lemma Ioc_filter_modEq_cast {v : ℕ} : ((Ioc a b).filter (· ≡ v [MOD r])).ma · intro h; lift x to ℕ using (by linarith); exact ⟨x, by simp_all [natCast_modEq_iff]⟩ variable (hr : 0 < r) +include hr /-- There are `⌈(b - v) / r⌉ - ⌈(a - v) / r⌉` numbers congruent to `v` mod `r` in `[a, b)`, if `a ≤ b`. `Nat` version of `Int.Ico_filter_modEq_card`. -/ diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index ff79ec56c9722..22a70729a558d 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1530,6 +1530,8 @@ section FoldlEqFoldlr' variable {f : α → β → α} variable (hf : ∀ a b c, f (f a b) c = f (f a c) b) +include hf + theorem foldl_eq_of_comm' : ∀ a b l, foldl f a (b :: l) = f (foldl f a l) b | a, b, [] => rfl | a, b, c :: l => by rw [foldl, foldl, foldl, ← foldl_eq_of_comm' .., foldl, hf] diff --git a/Mathlib/Data/Real/ConjExponents.lean b/Mathlib/Data/Real/ConjExponents.lean index 214fffe15846f..6b6586723f5c1 100644 --- a/Mathlib/Data/Real/ConjExponents.lean +++ b/Mathlib/Data/Real/ConjExponents.lean @@ -50,6 +50,10 @@ namespace IsConjExponent `q`: many computations using these exponents require clearing out denominators, which can be done with `field_simp` given a proof that these denominators are non-zero, so we record the most usual ones. -/ + +section +include h + theorem pos : 0 < p := lt_trans zero_lt_one h.one_lt theorem nonneg : 0 ≤ p := le_of_lt h.pos @@ -70,7 +74,7 @@ theorem one_div_nonneg : 0 ≤ 1 / p := le_of_lt h.one_div_pos theorem one_div_ne_zero : 1 / p ≠ 0 := ne_of_gt h.one_div_pos -theorem conj_eq (h : p.IsConjExponent q) : q = p / (p - 1) := by +theorem conj_eq : q = p / (p - 1) := by have := h.inv_add_inv_conj rw [← eq_sub_iff_add_eq', inv_eq_iff_eq_inv] at this field_simp [this, h.ne_zero] @@ -86,7 +90,7 @@ theorem sub_one_mul_conj : (p - 1) * q = p := theorem mul_eq_add : p * q = p + q := by simpa only [sub_mul, sub_eq_iff_eq_add, one_mul] using h.sub_one_mul_conj -@[symm] protected lemma symm (h : p.IsConjExponent q) : q.IsConjExponent p where +@[symm] protected lemma symm : q.IsConjExponent p where one_lt := by simpa only [h.conj_eq] using (one_lt_div h.sub_one_pos).mpr (sub_one_lt p) inv_add_inv_conj := by simpa [add_comm] using h.inv_add_inv_conj @@ -99,6 +103,8 @@ theorem inv_add_inv_conj_ennreal : (ENNReal.ofReal p)⁻¹ + (ENNReal.ofReal q) ← ENNReal.ofReal_inv_of_pos h.symm.pos, ← ENNReal.ofReal_add h.inv_nonneg h.symm.inv_nonneg, h.inv_add_inv_conj] +end + protected lemma inv_inv (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) : a⁻¹.IsConjExponent b⁻¹ := ⟨one_lt_inv ha $ by linarith, by simpa only [inv_inv]⟩ @@ -147,6 +153,10 @@ namespace IsConjExponent `q`: many computations using these exponents require clearing out denominators, which can be done with `field_simp` given a proof that these denominators are non-zero, so we record the most usual ones. -/ + +section +include h + lemma one_le : 1 ≤ p := h.one_lt.le lemma pos : 0 < p := zero_lt_one.trans h.one_lt lemma ne_zero : p ≠ 0 := h.pos.ne' @@ -159,7 +169,7 @@ lemma inv_ne_zero : p⁻¹ ≠ 0 := h.inv_pos.ne' lemma one_sub_inv : 1 - p⁻¹ = q⁻¹ := tsub_eq_of_eq_add_rev h.inv_add_inv_conj.symm -lemma conj_eq (h : p.IsConjExponent q) : q = p / (p - 1) := by +lemma conj_eq : q = p / (p - 1) := by simpa only [← coe_one, ← NNReal.coe_sub h.one_le, ← NNReal.coe_div, coe_inj] using h.coe.conj_eq lemma conjExponent_eq : conjExponent p = q := h.conj_eq.symm @@ -171,7 +181,7 @@ lemma mul_eq_add : p * q = p + q := by simpa only [← NNReal.coe_mul, ← NNReal.coe_add, NNReal.coe_inj] using h.coe.mul_eq_add @[symm] -protected lemma symm (h : p.IsConjExponent q) : q.IsConjExponent p where +protected lemma symm : q.IsConjExponent p where one_lt := by rw [h.conj_eq] exact (one_lt_div h.sub_one_pos).mpr (tsub_lt_self h.pos zero_lt_one) @@ -181,6 +191,8 @@ lemma div_conj_eq_sub_one : p / q = p - 1 := by field_simp [h.symm.ne_zero]; rw lemma inv_add_inv_conj_ennreal : (p⁻¹ + q⁻¹ : ℝ≥0∞) = 1 := by norm_cast; exact h.inv_add_inv_conj +end + protected lemma inv_inv (ha : a ≠ 0) (hb : b ≠ 0) (hab : a + b = 1) : a⁻¹.IsConjExponent b⁻¹ := ⟨one_lt_inv ha.bot_lt $ by rw [← hab]; exact lt_add_of_pos_right _ hb.bot_lt, by diff --git a/Mathlib/Data/Set/Lattice.lean b/Mathlib/Data/Set/Lattice.lean index 277bfe9ec88af..c994bb0f5366b 100644 --- a/Mathlib/Data/Set/Lattice.lean +++ b/Mathlib/Data/Set/Lattice.lean @@ -1247,6 +1247,7 @@ section open Function variable {f : α → β} {U : ι → Set β} (hU : iUnion U = univ) +include hU theorem injective_iff_injective_of_iUnion_eq_univ : Injective f ↔ ∀ i, Injective ((U i).restrictPreimage f) := by diff --git a/Mathlib/Data/Set/Sigma.lean b/Mathlib/Data/Set/Sigma.lean index a1805649dbc09..d8a73c0f8e40f 100644 --- a/Mathlib/Data/Set/Sigma.lean +++ b/Mathlib/Data/Set/Sigma.lean @@ -153,7 +153,6 @@ variable {β : ι → Type*} theorem insert_sigma : (insert i s).sigma t = Sigma.mk i '' t i ∪ s.sigma t := by rw [insert_eq, union_sigma, singleton_sigma] - exact a theorem sigma_insert {a : ∀ i, α i} : s.sigma (fun i ↦ insert (a i) (t i)) = (fun i ↦ ⟨i, a i⟩) '' s ∪ s.sigma t := by diff --git a/Mathlib/Data/Set/UnionLift.lean b/Mathlib/Data/Set/UnionLift.lean index f673e2f82b811..9b8184a95693c 100644 --- a/Mathlib/Data/Set/UnionLift.lean +++ b/Mathlib/Data/Set/UnionLift.lean @@ -104,7 +104,6 @@ theorem iUnionLift_unary (u : T → T) (ui : ∀ i, S i → S i) Set.inclusion (show S i ⊆ T from hT'.symm ▸ Set.subset_iUnion S i) (ui i x)) (uβ : β → β) (h : ∀ (i) (x : S i), f i (ui i x) = uβ (f i x)) (x : T) : iUnionLift S f hf T (le_of_eq hT') (u x) = uβ (iUnionLift S f hf T (le_of_eq hT') x) := by - clear hT -- this prevents the argument from getting inserted by accident. subst hT' cases' Set.mem_iUnion.1 x.prop with i hi rw [iUnionLift_of_mem x hi, ← h i] @@ -126,7 +125,6 @@ theorem iUnionLift_binary (dir : Directed (· ≤ ·) S) (op : T → T → T) (o (opβ : β → β → β) (h : ∀ (i) (x y : S i), f i (opi i x y) = opβ (f i x) (f i y)) (x y : T) : iUnionLift S f hf T (le_of_eq hT') (op x y) = opβ (iUnionLift S f hf T (le_of_eq hT') x) (iUnionLift S f hf T (le_of_eq hT') y) := by - clear hT -- this prevents the argument from getting inserted by accident. subst hT' cases' Set.mem_iUnion.1 x.prop with i hi cases' Set.mem_iUnion.1 y.prop with j hj diff --git a/Mathlib/Data/Setoid/Partition.lean b/Mathlib/Data/Setoid/Partition.lean index 53a4b3e898dfd..ecf06ea7c220d 100644 --- a/Mathlib/Data/Setoid/Partition.lean +++ b/Mathlib/Data/Setoid/Partition.lean @@ -324,7 +324,7 @@ namespace IndexedPartition open Set -variable {ι α : Type*} {s : ι → Set α} (hs : IndexedPartition s) +variable {ι α : Type*} {s : ι → Set α} /-- On a unique index set there is the obvious trivial partition -/ instance [Unique ι] [Inhabited α] : Inhabited (IndexedPartition fun _i : ι => (Set.univ : Set α)) := @@ -337,13 +337,18 @@ instance [Unique ι] [Inhabited α] : Inhabited (IndexedPartition fun _i : ι => -- Porting note: `simpNF` complains about `mem_index` attribute [simp] some_mem --mem_index +variable (hs : IndexedPartition s) + +include hs in theorem exists_mem (x : α) : ∃ i, x ∈ s i := ⟨hs.index x, hs.mem_index x⟩ +include hs in theorem iUnion : ⋃ i, s i = univ := by ext x simp [hs.exists_mem x] +include hs in theorem disjoint : Pairwise fun i j => Disjoint (s i) (s j) := fun {_i _j} h => disjoint_left.mpr fun {_x} hxi hxj => h (hs.eq_of_mem hxi hxj) diff --git a/Mathlib/Data/TwoPointing.lean b/Mathlib/Data/TwoPointing.lean index 59143f1bf6b2c..a4473d81ef91c 100644 --- a/Mathlib/Data/TwoPointing.lean +++ b/Mathlib/Data/TwoPointing.lean @@ -55,6 +55,7 @@ theorem swap_snd : p.swap.snd = p.fst := rfl @[simp] theorem swap_swap : p.swap.swap = p := rfl +include p in theorem to_nontrivial : Nontrivial α := ⟨⟨p.fst, p.snd, p.fst_ne_snd⟩⟩ diff --git a/Mathlib/Data/Vector/MapLemmas.lean b/Mathlib/Data/Vector/MapLemmas.lean index ba8f98bc718bf..3c67c95406853 100644 --- a/Mathlib/Data/Vector/MapLemmas.lean +++ b/Mathlib/Data/Vector/MapLemmas.lean @@ -224,7 +224,6 @@ variable {xs : Vector α n} {ys : Vector β n} protected theorem map_eq_mapAccumr : map f xs = (mapAccumr (fun x (_ : Unit) ↦ ((), f x)) xs ()).snd := by - clear ys induction xs using Vector.revInductionOn <;> simp_all /-- @@ -266,7 +265,6 @@ theorem mapAccumr₂_eq_map₂ {f : α → β → σ → σ × γ} {s₀ : σ} ( @[simp] theorem mapAccumr_eq_map_of_constant_state (f : α → σ → σ × β) (s : σ) (h : ∀ a, (f a s).fst = s) : mapAccumr f xs s = (s, (map (fun x => (f x s).snd) xs)) := by - clear ys induction xs using revInductionOn <;> simp_all /-- diff --git a/Mathlib/Data/Vector/Snoc.lean b/Mathlib/Data/Vector/Snoc.lean index 5af561ba61449..41b4c9e7e94c8 100644 --- a/Mathlib/Data/Vector/Snoc.lean +++ b/Mathlib/Data/Vector/Snoc.lean @@ -55,7 +55,6 @@ theorem reverse_snoc : reverse (xs.snoc x) = x ::ᵥ (reverse xs) := by theorem replicate_succ_to_snoc (val : α) : replicate (n+1) val = (replicate n val).snoc val := by - clear xs induction n with | zero => rfl | succ n ih => diff --git a/Mathlib/Deprecated/Group.lean b/Mathlib/Deprecated/Group.lean index bd990fbce0326..d2d59cec9f5b0 100644 --- a/Mathlib/Deprecated/Group.lean +++ b/Mathlib/Deprecated/Group.lean @@ -137,11 +137,10 @@ end MulEquiv namespace IsMonoidHom -variable [MulOneClass α] [MulOneClass β] {f : α → β} (hf : IsMonoidHom f) - /-- A monoid homomorphism preserves multiplication. -/ @[to_additive "An additive monoid homomorphism preserves addition."] -theorem map_mul' (x y) : f (x * y) = f x * f y := +theorem map_mul' [MulOneClass α] [MulOneClass β] {f : α → β} (hf : IsMonoidHom f) (x y) : + f (x * y) = f x * f y := hf.map_mul x y /-- The inverse of a map which preserves multiplication, @@ -228,6 +227,14 @@ namespace IsGroupHom variable [Group α] [Group β] {f : α → β} (hf : IsGroupHom f) +/-- The identity is a group homomorphism. -/ +@[to_additive "The identity is an additive group homomorphism."] +theorem id : IsGroupHom (@id α) := + { map_mul := fun _ _ => rfl } + +section +include hf + open IsMulHom (map_mul) theorem map_mul' : ∀ x y, f (x * y) = f x * f y := @@ -245,33 +252,29 @@ theorem map_one : f 1 = 1 := /-- A group homomorphism sends inverses to inverses. -/ @[to_additive "An additive group homomorphism sends negations to negations."] -theorem map_inv (hf : IsGroupHom f) (a : α) : f a⁻¹ = (f a)⁻¹ := +theorem map_inv (a : α) : f a⁻¹ = (f a)⁻¹ := eq_inv_of_mul_eq_one_left <| by rw [← hf.map_mul, inv_mul_cancel, hf.map_one] @[to_additive] -theorem map_div (hf : IsGroupHom f) (a b : α) : f (a / b) = f a / f b := by +theorem map_div (a b : α) : f (a / b) = f a / f b := by simp_rw [div_eq_mul_inv, hf.map_mul, hf.map_inv] -/-- The identity is a group homomorphism. -/ -@[to_additive "The identity is an additive group homomorphism."] -theorem id : IsGroupHom (@id α) := - { map_mul := fun _ _ => rfl } - /-- The composition of two group homomorphisms is a group homomorphism. -/ @[to_additive "The composition of two additive group homomorphisms is an additive group homomorphism."] -theorem comp (hf : IsGroupHom f) {γ} [Group γ] {g : β → γ} (hg : IsGroupHom g) : +theorem comp {γ} [Group γ] {g : β → γ} (hg : IsGroupHom g) : IsGroupHom (g ∘ f) := { IsMulHom.comp hf.toIsMulHom hg.toIsMulHom with } /-- A group homomorphism is injective iff its kernel is trivial. -/ @[to_additive "An additive group homomorphism is injective if its kernel is trivial."] -theorem injective_iff {f : α → β} (hf : IsGroupHom f) : - Function.Injective f ↔ ∀ a, f a = 1 → a = 1 := +theorem injective_iff : Function.Injective f ↔ ∀ a, f a = 1 → a = 1 := ⟨fun h _ => by rw [← hf.map_one]; exact @h _ _, fun h x y hxy => eq_of_div_eq_one <| h _ <| by rwa [hf.map_div, div_eq_one]⟩ +end + /-- The product of group homomorphisms is a group homomorphism if the target is commutative. -/ @[to_additive "The sum of two additive group homomorphisms is an additive group homomorphism diff --git a/Mathlib/Deprecated/Subgroup.lean b/Mathlib/Deprecated/Subgroup.lean index 566783c7bb5b5..17691db27d6d4 100644 --- a/Mathlib/Deprecated/Subgroup.lean +++ b/Mathlib/Deprecated/Subgroup.lean @@ -117,6 +117,7 @@ namespace IsSubgroup open IsSubmonoid variable [Group G] {s : Set G} (hs : IsSubgroup s) +include hs @[to_additive] theorem inv_mem_iff : a⁻¹ ∈ s ↔ a ∈ s := diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index 644d722bfefa4..7a0cf2877f279 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -586,11 +586,11 @@ theorem transnumAuxSeq_dist_lt (n : ℕ) : pow_mul, sq, mul_apply] theorem tendsto_translationNumber_aux : Tendsto f.transnumAuxSeq atTop (𝓝 <| τ f) := - (cauchySeq_of_le_geometric_two 1 fun n => le_of_lt <| f.transnumAuxSeq_dist_lt n).tendsto_limUnder + (cauchySeq_of_le_geometric_two fun n => le_of_lt <| f.transnumAuxSeq_dist_lt n).tendsto_limUnder theorem dist_map_zero_translationNumber_le : dist (f 0) (τ f) ≤ 1 := f.transnumAuxSeq_zero ▸ - dist_le_of_le_geometric_two_of_tendsto₀ 1 (fun n => le_of_lt <| f.transnumAuxSeq_dist_lt n) + dist_le_of_le_geometric_two_of_tendsto₀ (fun n => le_of_lt <| f.transnumAuxSeq_dist_lt n) f.tendsto_translationNumber_aux theorem tendsto_translationNumber_of_dist_bounded_aux (x : ℕ → ℝ) (C : ℝ) diff --git a/Mathlib/FieldTheory/Extension.lean b/Mathlib/FieldTheory/Extension.lean index 2ded21368a221..4fca6eb496228 100644 --- a/Mathlib/FieldTheory/Extension.lean +++ b/Mathlib/FieldTheory/Extension.lean @@ -112,6 +112,7 @@ private theorem exists_algHom_adjoin_of_splits'' {L : IntermediateField F E} variable {L : Type*} [Field L] [Algebra F L] [Algebra L E] [IsScalarTower F L E] (f : L →ₐ[F] K) (hK : ∀ s ∈ S, IsIntegral L s ∧ (minpoly L s).Splits f.toRingHom) +include hK in theorem exists_algHom_adjoin_of_splits' : ∃ φ : adjoin L S →ₐ[F] K, φ.comp (IsScalarTower.toAlgHom F L _) = f := by let L' := (IsScalarTower.toAlgHom F L E).fieldRange @@ -129,6 +130,7 @@ theorem exists_algHom_adjoin_of_splits' : refine ⟨(hK s hs).1.tower_top, (hK s hs).1.minpoly_splits_tower_top' ?_⟩ convert (hK s hs).2; ext; exact congr_arg f (AlgEquiv.symm_apply_apply _ _) +include hK in theorem exists_algHom_of_adjoin_splits' (hS : adjoin L S = ⊤) : ∃ φ : E →ₐ[F] K, φ.comp (IsScalarTower.toAlgHom F L E) = f := have ⟨φ, hφ⟩ := exists_algHom_adjoin_of_splits' f hK @@ -141,10 +143,11 @@ theorem exists_algHom_of_splits' (hK : ∀ s : E, IsIntegral L s ∧ (minpoly L end variable (hK : ∀ s ∈ S, IsIntegral F s ∧ (minpoly F s).Splits (algebraMap F K)) - (hK' : ∀ s : E, IsIntegral F s ∧ (minpoly F s).Splits (algebraMap F K)) - {L : IntermediateField F E} (f : L →ₐ[F] K) (hL : L ≤ adjoin F S) + (hK' : ∀ s : E, IsIntegral F s ∧ (minpoly F s).Splits (algebraMap F K)) + {L : IntermediateField F E} (f : L →ₐ[F] K) (hL : L ≤ adjoin F S) {x : E} {y : K} --- The following uses the hypothesis `hK`. +section +include hK theorem exists_algHom_adjoin_of_splits : ∃ φ : adjoin F S →ₐ[F] K, φ.comp (inclusion hL) = f := by obtain ⟨φ, hfφ, hφ⟩ := zorn_le_nonempty_Ici₀ _ @@ -159,14 +162,17 @@ theorem nonempty_algHom_adjoin_of_splits : Nonempty (adjoin F S →ₐ[F] K) := variable (hS : adjoin F S = ⊤) +include hS in theorem exists_algHom_of_adjoin_splits : ∃ φ : E →ₐ[F] K, φ.comp L.val = f := have ⟨φ, hφ⟩ := exists_algHom_adjoin_of_splits hK f (hS.symm ▸ le_top) ⟨φ.comp ((equivOfEq hS).trans topEquiv).symm.toAlgHom, hφ⟩ +include hS in theorem nonempty_algHom_of_adjoin_splits : Nonempty (E →ₐ[F] K) := have ⟨φ, _⟩ := exists_algHom_of_adjoin_splits hK (⊥ : Lifts F E K).emb hS; ⟨φ⟩ -variable {x : E} (hx : x ∈ adjoin F S) {y : K} (hy : aeval y (minpoly F x) = 0) +variable (hx : x ∈ adjoin F S) (hy : aeval y (minpoly F x) = 0) +include hy theorem exists_algHom_adjoin_of_splits_of_aeval : ∃ φ : adjoin F S →ₐ[F] K, φ ⟨x, hx⟩ = y := by have := isAlgebraic_adjoin (fun s hs ↦ (hK s hs).1) @@ -177,12 +183,18 @@ theorem exists_algHom_adjoin_of_splits_of_aeval : ∃ φ : adjoin F S →ₐ[F] exact ⟨φ, (DFunLike.congr_fun hφ <| AdjoinSimple.gen F x).trans <| algHomAdjoinIntegralEquiv_symm_apply_gen F ix _⟩ +include hS in theorem exists_algHom_of_adjoin_splits_of_aeval : ∃ φ : E →ₐ[F] K, φ x = y := have ⟨φ, hφ⟩ := exists_algHom_adjoin_of_splits_of_aeval hK (hS ▸ mem_top) hy ⟨φ.comp ((equivOfEq hS).trans topEquiv).symm.toAlgHom, hφ⟩ -/- The following uses the hypothesis - (hK' : ∀ s : E, IsIntegral F s ∧ (minpoly F s).Splits (algebraMap F K)) -/ + +include hK' + +end + +section +include hK' theorem exists_algHom_of_splits : ∃ φ : E →ₐ[F] K, φ.comp L.val = f := exists_algHom_of_adjoin_splits (fun x _ ↦ hK' x) f (adjoin_univ F E) @@ -190,9 +202,12 @@ theorem exists_algHom_of_splits : ∃ φ : E →ₐ[F] K, φ.comp L.val = f := theorem nonempty_algHom_of_splits : Nonempty (E →ₐ[F] K) := nonempty_algHom_of_adjoin_splits (fun x _ ↦ hK' x) (adjoin_univ F E) -theorem exists_algHom_of_splits_of_aeval : ∃ φ : E →ₐ[F] K, φ x = y := +theorem exists_algHom_of_splits_of_aeval (hy : aeval y (minpoly F x) = 0) : + ∃ φ : E →ₐ[F] K, φ x = y := exists_algHom_of_adjoin_splits_of_aeval (fun x _ ↦ hK' x) (adjoin_univ F E) hy +end + end IntermediateField section Algebra.IsAlgebraic diff --git a/Mathlib/FieldTheory/KummerExtension.lean b/Mathlib/FieldTheory/KummerExtension.lean index 8e4904f42f1a7..c36f3f23e6e98 100644 --- a/Mathlib/FieldTheory/KummerExtension.lean +++ b/Mathlib/FieldTheory/KummerExtension.lean @@ -258,6 +258,7 @@ open scoped KummerExtension section AdjoinRoot +include hζ H in /-- Also see `Polynomial.separable_X_pow_sub_C_unit` -/ theorem Polynomial.separable_X_pow_sub_C_of_irreducible : (X ^ n - C a).Separable := by letI := Fact.mk H @@ -386,6 +387,7 @@ section IsSplittingField variable {a} variable {L : Type*} [Field L] [Algebra K L] [IsSplittingField K L (X ^ n - C a)] +include hζ in lemma isSplittingField_AdjoinRoot_X_pow_sub_C : haveI := Fact.mk H letI : Algebra K K[n√a] := inferInstance @@ -429,6 +431,7 @@ lemma adjoinRootXPowSubCEquiv_symm_eq_root : apply (adjoinRootXPowSubCEquiv hζ H hα).injective rw [(adjoinRootXPowSubCEquiv hζ H hα).apply_symm_apply, adjoinRootXPowSubCEquiv_root] +include hζ H hα in lemma Algebra.adjoin_root_eq_top_of_isSplittingField : Algebra.adjoin K {α} = ⊤ := by apply Subalgebra.map_injective (f := (adjoinRootXPowSubCEquiv hζ H hα).symm) @@ -437,6 +440,7 @@ lemma Algebra.adjoin_root_eq_top_of_isSplittingField : (adjoinRootXPowSubCEquiv hζ H hα).symm.surjective, AlgHom.map_adjoin, Set.image_singleton, AlgHom.coe_coe, adjoinRootXPowSubCEquiv_symm_eq_root, adjoinRoot_eq_top] +include hζ H hα in lemma IntermediateField.adjoin_root_eq_top_of_isSplittingField : K⟮α⟯ = ⊤ := by refine (IntermediateField.eq_adjoin_of_eq_algebra_adjoin _ _ _ ?_).symm @@ -478,6 +482,7 @@ lemma autEquivRootsOfUnity_apply_rootOfSplit (σ : L ≃ₐ[K] L) : adjoinRootXPowSubCEquiv_root] rfl +include hα in lemma autEquivRootsOfUnity_smul (σ : L ≃ₐ[K] L) : autEquivRootsOfUnity hζ hn H L σ • α = σ α := by have ⟨ζ, hζ'⟩ := hζ @@ -501,24 +506,29 @@ def autEquivZmod {ζ : K} (hζ : IsPrimitiveRoot ζ n) : (hζ.isUnit_unit' hn)).symm).trans (AddEquiv.toMultiplicative' (hζ.isUnit_unit' hn).zmodEquivZPowers.symm)) +include hα in lemma autEquivZmod_symm_apply_intCast {ζ : K} (hζ : IsPrimitiveRoot ζ n) (m : ℤ) : (autEquivZmod H L hζ).symm (Multiplicative.ofAdd (m : ZMod n)) α = ζ ^ m • α := by have hn := Nat.pos_iff_ne_zero.mpr (ne_zero_of_irreducible_X_pow_sub_C H) rw [← autEquivRootsOfUnity_smul ⟨ζ, (mem_primitiveRoots hn).mpr hζ⟩ hn H L hα] simp [MulEquiv.subgroupCongr_symm_apply, Subgroup.smul_def, Units.smul_def, autEquivZmod] +include hα in lemma autEquivZmod_symm_apply_natCast {ζ : K} (hζ : IsPrimitiveRoot ζ n) (m : ℕ) : (autEquivZmod H L hζ).symm (Multiplicative.ofAdd (m : ZMod n)) α = ζ ^ m • α := by simpa only [Int.cast_natCast, zpow_natCast] using autEquivZmod_symm_apply_intCast H L hα hζ m +include hζ H in lemma isCyclic_of_isSplittingField_X_pow_sub_C : IsCyclic (L ≃ₐ[K] L) := have hn := Nat.pos_iff_ne_zero.mpr (ne_zero_of_irreducible_X_pow_sub_C H) isCyclic_of_surjective _ (autEquivZmod H _ <| (mem_primitiveRoots hn).mp hζ.choose_spec).symm.surjective +include hζ H in lemma isGalois_of_isSplittingField_X_pow_sub_C : IsGalois K L := IsGalois.of_separable_splitting_field (separable_X_pow_sub_C_of_irreducible hζ a H) +include hζ H in lemma finrank_of_isSplittingField_X_pow_sub_C : FiniteDimensional.finrank K L = n := by have := Polynomial.IsSplittingField.finiteDimensional L (X ^ n - C a) have := isGalois_of_isSplittingField_X_pow_sub_C hζ H L @@ -533,15 +543,16 @@ end IsSplittingField section IsCyclic -variable {L} [Field L] [Algebra K L] [IsGalois K L] [FiniteDimensional K L] [IsCyclic (L ≃ₐ[K] L)] +variable {L} [Field L] [Algebra K L] [FiniteDimensional K L] variable (hK : (primitiveRoots (FiniteDimensional.finrank K L) K).Nonempty) open FiniteDimensional variable (K L) +include hK in /-- If `L/K` is a cyclic extension of degree `n`, and `K` contains all `n`-th roots of unity, then `L = K[α]` for some `α ^ n ∈ K`. -/ -lemma exists_root_adjoin_eq_top_of_isCyclic : +lemma exists_root_adjoin_eq_top_of_isCyclic [IsGalois K L] [IsCyclic (L ≃ₐ[K] L)] : ∃ (α : L), α ^ (finrank K L) ∈ Set.range (algebraMap K L) ∧ K⟮α⟯ = ⊤ := by -- Let `ζ` be an `n`-th root of unity, and `σ` be a generator of `L ≃ₐ[K] L`. have ⟨ζ, hζ⟩ := hK @@ -593,6 +604,7 @@ lemma irreducible_X_pow_sub_C_of_root_adjoin_eq_top exact (finrank_top K L).ge exact this ▸ minpoly.irreducible (IsIntegral.of_finite K α) +include hK in lemma isSplittingField_X_pow_sub_C_of_root_adjoin_eq_top {a : K} {α : L} (ha : α ^ (finrank K L) = algebraMap K L a) (hα : K⟮α⟯ = ⊤) : IsSplittingField K L (X ^ (finrank K L) - C a) := by diff --git a/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean b/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean index 78394928e9f8a..317ae7f335968 100644 --- a/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean +++ b/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean @@ -37,25 +37,12 @@ lemma coeff_minpolyDiv (i) : coeff (minpolyDiv R x) i = algebraMap R S (coeff (minpoly R x) (i + 1)) + coeff (minpolyDiv R x) (i + 1) * x := by rw [← coeff_map, ← minpolyDiv_spec R x]; simp [mul_sub] -variable (hx : IsIntegral R x) {R x} - -lemma minpolyDiv_ne_zero [Nontrivial S] : minpolyDiv R x ≠ 0 := by - intro e - have := minpolyDiv_spec R x - rw [e, zero_mul] at this - exact ((minpoly.monic hx).map (algebraMap R S)).ne_zero this.symm +variable {R x} lemma minpolyDiv_eq_zero (hx : ¬IsIntegral R x) : minpolyDiv R x = 0 := by delta minpolyDiv minpoly rw [dif_neg hx, Polynomial.map_zero, zero_divByMonic] -lemma minpolyDiv_monic : Monic (minpolyDiv R x) := by - nontriviality S - have := congr_arg leadingCoeff (minpolyDiv_spec R x) - rw [leadingCoeff_mul', ((minpoly.monic hx).map (algebraMap R S)).leadingCoeff] at this - · simpa using this - · simpa using minpolyDiv_ne_zero hx - lemma eval_minpolyDiv_self : (minpolyDiv R x).eval x = aeval x (derivative <| minpoly R x) := by rw [aeval_def, ← eval_map, ← derivative_map, ← minpolyDiv_spec R x]; simp @@ -89,23 +76,6 @@ lemma eval_minpolyDiv_of_aeval_eq_zero [IsDomain S] [DecidableEq S] rw [eval, eval₂_minpolyDiv_of_eval₂_eq_zero, RingHom.id_apply, RingHom.id_apply] simpa [aeval_def] using hy -lemma natDegree_minpolyDiv_succ [Nontrivial S] : - natDegree (minpolyDiv R x) + 1 = natDegree (minpoly R x) := by - rw [← (minpoly.monic hx).natDegree_map (algebraMap R S), ← minpolyDiv_spec, natDegree_mul'] - · simp - · simpa using minpolyDiv_ne_zero hx - -lemma natDegree_minpolyDiv : - natDegree (minpolyDiv R x) = natDegree (minpoly R x) - 1 := by - nontriviality S - by_cases hx : IsIntegral R x - · rw [← natDegree_minpolyDiv_succ hx]; rfl - · rw [minpolyDiv_eq_zero hx, minpoly.eq_zero hx]; rfl - -lemma natDegree_minpolyDiv_lt [Nontrivial S] : - natDegree (minpolyDiv R x) < natDegree (minpoly R x) := by - rw [← natDegree_minpolyDiv_succ hx] - exact Nat.lt_succ_self _ lemma coeff_minpolyDiv_mem_adjoin (x : S) (i) : coeff (minpolyDiv R x) i ∈ Algebra.adjoin R {x} := by @@ -125,6 +95,34 @@ lemma coeff_minpolyDiv_mem_adjoin (x : S) (i) : rw [← add_assoc] exact Nat.lt_succ_self _ +section IsIntegral +variable (hx : IsIntegral R x) +include hx + +lemma minpolyDiv_ne_zero [Nontrivial S] : minpolyDiv R x ≠ 0 := by + intro e + have := minpolyDiv_spec R x + rw [e, zero_mul] at this + exact ((minpoly.monic hx).map (algebraMap R S)).ne_zero this.symm + +lemma minpolyDiv_monic : Monic (minpolyDiv R x) := by + nontriviality S + have := congr_arg leadingCoeff (minpolyDiv_spec R x) + rw [leadingCoeff_mul', ((minpoly.monic hx).map (algebraMap R S)).leadingCoeff] at this + · simpa using this + · simpa using minpolyDiv_ne_zero hx + +lemma natDegree_minpolyDiv_succ [Nontrivial S] : + natDegree (minpolyDiv R x) + 1 = natDegree (minpoly R x) := by + rw [← (minpoly.monic hx).natDegree_map (algebraMap R S), ← minpolyDiv_spec, natDegree_mul'] + · simp + · simpa using minpolyDiv_ne_zero hx + +lemma natDegree_minpolyDiv_lt [Nontrivial S] : + natDegree (minpolyDiv R x) < natDegree (minpoly R x) := by + rw [← natDegree_minpolyDiv_succ hx] + exact Nat.lt_succ_self _ + lemma minpolyDiv_eq_of_isIntegrallyClosed [IsDomain R] [IsIntegrallyClosed R] [IsDomain S] [Algebra R K] [Algebra K S] [IsScalarTower R K S] [IsFractionRing R K] : minpolyDiv R x = minpolyDiv K x := by @@ -178,6 +176,16 @@ lemma span_coeff_minpolyDiv : exact hi j hj (lt_trans hj hi') · rwa [← natDegree_minpolyDiv_succ hx, Set.mem_Iio, Nat.lt_succ_iff] at hi' +end IsIntegral + +lemma natDegree_minpolyDiv : + natDegree (minpolyDiv R x) = natDegree (minpoly R x) - 1 := by + nontriviality S + by_cases hx : IsIntegral R x + · rw [← natDegree_minpolyDiv_succ hx]; rfl + · rw [minpolyDiv_eq_zero hx, minpoly.eq_zero hx]; rfl + + section PowerBasis variable {K} diff --git a/Mathlib/FieldTheory/NormalClosure.lean b/Mathlib/FieldTheory/NormalClosure.lean index eff047d5d9efd..a1d7bf881f7dd 100644 --- a/Mathlib/FieldTheory/NormalClosure.lean +++ b/Mathlib/FieldTheory/NormalClosure.lean @@ -71,6 +71,7 @@ lemma normalClosure_le_iSup_adjoin : variable (splits : ∀ x : K, (minpoly F x).Splits (algebraMap F L)) +include splits in lemma normalClosure_eq_iSup_adjoin_of_splits : normalClosure F K L = ⨆ x : K, IntermediateField.adjoin F ((minpoly F x).rootSet L) := normalClosure_le_iSup_adjoin.antisymm <| @@ -87,6 +88,7 @@ lemma isNormalClosure_iff : IsNormalClosure F K L ↔ simpa only [normalClosure_eq_iSup_adjoin_of_splits splits] using h -- TODO: IntermediateField.isNormalClosure_iff similar to IntermediateField.isSplittingField_iff +include splits in /-- `normalClosure F K L` is a valid normal closure if `K/F` is algebraic and all minimal polynomials of `K/F` splits in `L/F`. -/ lemma isNormalClosure_normalClosure : IsNormalClosure F K (normalClosure F K L) := by diff --git a/Mathlib/FieldTheory/PrimitiveElement.lean b/Mathlib/FieldTheory/PrimitiveElement.lean index 74874763ddba2..65b6d1707f19d 100644 --- a/Mathlib/FieldTheory/PrimitiveElement.lean +++ b/Mathlib/FieldTheory/PrimitiveElement.lean @@ -384,6 +384,7 @@ theorem primitive_element_iff_minpoly_degree_eq (α : E) : variable [Algebra.IsSeparable F E] (A : Type*) [Field A] [Algebra F A] (hA : ∀ x : E, (minpoly F x).Splits (algebraMap F A)) +include hA theorem primitive_element_iff_algHom_eq_of_eval' (α : E) : F⟮α⟯ = ⊤ ↔ Function.Injective fun φ : E →ₐ[F] A ↦ φ α := by diff --git a/Mathlib/FieldTheory/Separable.lean b/Mathlib/FieldTheory/Separable.lean index 58a4f2be710fe..fddc5fbc1dcee 100644 --- a/Mathlib/FieldTheory/Separable.lean +++ b/Mathlib/FieldTheory/Separable.lean @@ -382,8 +382,6 @@ theorem unique_separable_of_irreducible {f : F[X]} (hf : Irreducible f) (hp : 0 (g₁ : F[X]) (hg₁ : g₁.Separable) (hgf₁ : expand F (p ^ n₁) g₁ = f) (n₂ : ℕ) (g₂ : F[X]) (hg₂ : g₂.Separable) (hgf₂ : expand F (p ^ n₂) g₂ = f) : n₁ = n₂ ∧ g₁ = g₂ := by revert g₁ g₂ - -- Porting note: the variable `K` affects the `wlog` tactic. - clear! K wlog hn : n₁ ≤ n₂ · intro g₁ hg₁ Hg₁ g₂ hg₂ Hg₂ simpa only [eq_comm] using this p hf hp n₂ n₁ (le_of_not_le hn) g₂ hg₂ Hg₂ g₁ hg₁ Hg₁ @@ -557,6 +555,7 @@ theorem Algebra.isSeparable_iff : fun h => ⟨fun x => (h x).2⟩⟩ variable {E : Type*} [Ring E] [Algebra F E] (e : K ≃ₐ[F] E) +include e /-- Transfer `Algebra.IsSeparable` across an `AlgEquiv`. -/ theorem AlgEquiv.isSeparable [Algebra.IsSeparable F K] : Algebra.IsSeparable F E := diff --git a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean index ea3592eaa2d83..d819c1008538f 100644 --- a/Mathlib/Geometry/Euclidean/Inversion/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Inversion/Basic.lean @@ -174,6 +174,7 @@ theorem dist_inversion_mul_dist_center_eq (hx : x ≠ c) (hy : y ≠ c) : ### Ptolemy's inequality -/ +include V in /-- **Ptolemy's inequality**: in a quadrangle `ABCD`, `|AC| * |BD| ≤ |AB| * |CD| + |BC| * |AD|`. If `ABCD` is a convex cyclic polygon, then this inequality becomes an equality, see `EuclideanGeometry.mul_dist_add_mul_dist_eq_mul_dist_of_cospherical`. -/ diff --git a/Mathlib/Geometry/Euclidean/Sphere/Basic.lean b/Mathlib/Geometry/Euclidean/Sphere/Basic.lean index 8b1bfe714bc4f..0b37ec458a9f3 100644 --- a/Mathlib/Geometry/Euclidean/Sphere/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Sphere/Basic.lean @@ -170,6 +170,7 @@ section NormedSpace variable [NormedAddCommGroup V] [NormedSpace ℝ V] [MetricSpace P] [NormedAddTorsor V P] +include V in /-- Two points are cospherical. -/ theorem cospherical_pair (p₁ p₂ : P) : Cospherical ({p₁, p₂} : Set P) := ⟨midpoint ℝ p₁ p₂, ‖(2 : ℝ)‖⁻¹ * dist p₁ p₂, by diff --git a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean index b1f65c2f05087..47d6a6e3e4f01 100644 --- a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean +++ b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean @@ -97,6 +97,7 @@ variable (I) theorem smooth_inv : Smooth I I fun x : G => x⁻¹ := LieGroup.smooth_inv +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 @@ -240,6 +241,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalS theorem smoothAt_inv₀ {x : G} (hx : x ≠ 0) : SmoothAt I I (fun y ↦ y⁻¹) x := SmoothInv₀.smoothAt_inv₀ hx +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]. -/ diff --git a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean index b7c7224c17e18..9202595ec5d94 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean @@ -73,6 +73,7 @@ variable (I) theorem smooth_mul : Smooth (I.prod I) I fun p : G × G => p.1 * p.2 := SmoothMul.smooth_mul +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 diff --git a/Mathlib/Geometry/Manifold/Complex.lean b/Mathlib/Geometry/Manifold/Complex.lean index 706c9fd5659b9..a0110600277bb 100644 --- a/Mathlib/Geometry/Manifold/Complex.lean +++ b/Mathlib/Geometry/Manifold/Complex.lean @@ -42,7 +42,7 @@ open Function Set Filter Complex variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℂ F] variable {H : Type*} [TopologicalSpace H] {I : ModelWithCorners ℂ E H} [I.Boundaryless] -variable {M : Type*} [TopologicalSpace M] [CompactSpace M] [ChartedSpace H M] +variable {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] /-- **Maximum modulus principle**: if `f : M → F` is complex differentiable in a neighborhood of `c` @@ -146,6 +146,8 @@ model so that it works, e.g., on a product of two manifolds without a boundary. namespace MDifferentiable +variable [CompactSpace M] + /-- A holomorphic function on a compact complex manifold is locally constant. -/ protected theorem isLocallyConstant {f : M → F} (hf : MDifferentiable I 𝓘(ℂ, F) f) : IsLocallyConstant f := diff --git a/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean b/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean index 97804c5a82769..3f612e90bddc7 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean @@ -167,9 +167,6 @@ theorem ContMDiff.clm_postcomp {f : M → F₂ →L[𝕜] F₃} (hf : ContMDiff (fun y ↦ (f y).postcomp F₁ : M → (F₁ →L[𝕜] F₂) →L[𝕜] (F₁ →L[𝕜] F₃)) := fun x ↦ (hf x).clm_postcomp --- Now make `M` a smooth manifold. -variable [SmoothManifoldWithCorners I M] - theorem ContMDiffWithinAt.clm_comp {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ →L[𝕜] F₁} {s : Set M} {x : M} (hg : ContMDiffWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₃) n g s x) (hf : ContMDiffWithinAt I 𝓘(𝕜, F₂ →L[𝕜] F₁) n f s x) : diff --git a/Mathlib/Geometry/Manifold/IntegralCurve.lean b/Mathlib/Geometry/Manifold/IntegralCurve.lean index ce8f118abfbcf..beb8c885abb24 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve.lean @@ -63,7 +63,7 @@ open scoped Manifold Topology open Function Set variable - {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] + {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] @@ -308,7 +308,7 @@ variable (t₀) {x₀ : M} /-- Existence of local integral curves for a $C^1$ vector field at interior points of a smooth manifold. -/ -theorem exists_isIntegralCurveAt_of_contMDiffAt +theorem exists_isIntegralCurveAt_of_contMDiffAt [CompleteSpace E] (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) x₀) (hx : I.IsInteriorPoint x₀) : ∃ γ : ℝ → M, γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := by @@ -362,7 +362,8 @@ theorem exists_isIntegralCurveAt_of_contMDiffAt /-- Existence of local integral curves for a $C^1$ vector field on a smooth manifold without boundary. -/ -lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless [BoundarylessManifold I M] +lemma exists_isIntegralCurveAt_of_contMDiffAt_boundaryless + [CompleteSpace E] [BoundarylessManifold I M] (hv : ContMDiffAt I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M)) x₀) : ∃ γ : ℝ → M, γ t₀ = x₀ ∧ IsIntegralCurveAt γ v t₀ := exists_isIntegralCurveAt_of_contMDiffAt t₀ hv (BoundarylessManifold.isInteriorPoint I) diff --git a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean index 6f17294b2d254..13b8450ed5ea7 100644 --- a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean +++ b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean @@ -73,6 +73,7 @@ structure LocalInvariantProp (P : (H → H') → Set H → H → Prop) : Prop wh variable {G G'} {P : (H → H') → Set H → H → Prop} {s t u : Set H} {x : H} variable (hG : G.LocalInvariantProp G' P) +include hG namespace LocalInvariantProp @@ -214,6 +215,7 @@ namespace LocalInvariantProp section variable (hG : G.LocalInvariantProp G' P) +include hG /-- `LiftPropWithinAt P f s x` is equivalent to a definition where we restrict the set we are considering to the domain of the charts at `x` and `f x`. -/ diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean index 7fdb2ed6ae39d..cd4d98910c206 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean @@ -166,6 +166,7 @@ end Charts namespace PartialHomeomorph.MDifferentiable variable {I I' I''} variable {e : PartialHomeomorph M M'} (he : e.MDifferentiable I I') {e' : PartialHomeomorph M' M''} +include he nonrec theorem symm : e.symm.MDifferentiable I' I := he.symm diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index 2cfef4d3ff400..29f6ccec46202 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -87,6 +87,175 @@ theorem IsOpen.uniqueMDiffOn (hs : IsOpen s) : UniqueMDiffOn I s := theorem uniqueMDiffOn_univ : UniqueMDiffOn I (univ : Set M) := isOpen_univ.uniqueMDiffOn +nonrec theorem UniqueMDiffWithinAt.prod {x : M} {y : M'} {s t} (hs : UniqueMDiffWithinAt I s x) + (ht : UniqueMDiffWithinAt I' t y) : UniqueMDiffWithinAt (I.prod I') (s ×ˢ t) (x, y) := by + refine (hs.prod ht).mono ?_ + rw [ModelWithCorners.range_prod, ← prod_inter_prod] + rfl + +theorem UniqueMDiffOn.prod {s : Set M} {t : Set M'} (hs : UniqueMDiffOn I s) + (ht : UniqueMDiffOn I' t) : UniqueMDiffOn (I.prod I') (s ×ˢ t) := fun x h ↦ + (hs x.1 h.1).prod (ht x.2 h.2) + +theorem mdifferentiableWithinAt_iff {f : M → M'} {s : Set M} {x : M} : + MDifferentiableWithinAt I I' f s x ↔ + ContinuousWithinAt f s x ∧ + DifferentiableWithinAt 𝕜 (writtenInExtChartAt I I' x f) + ((extChartAt I x).target ∩ (extChartAt I x).symm ⁻¹' s) ((extChartAt I x) x) := by + rw [mdifferentiableWithinAt_iff'] + refine and_congr Iff.rfl (exists_congr fun f' => ?_) + rw [inter_comm] + simp only [HasFDerivWithinAt, nhdsWithin_inter, nhdsWithin_extChartAt_target_eq] + +theorem MDifferentiableWithinAt.mono (hst : s ⊆ t) (h : MDifferentiableWithinAt I I' f t x) : + MDifferentiableWithinAt I I' f s x := + ⟨ContinuousWithinAt.mono h.1 hst, DifferentiableWithinAt.mono + h.differentiableWithinAt_writtenInExtChartAt + (inter_subset_inter_left _ (preimage_mono hst))⟩ + +theorem mdifferentiableWithinAt_univ : + MDifferentiableWithinAt I I' f univ x ↔ MDifferentiableAt I I' f x := by + simp_rw [MDifferentiableWithinAt, MDifferentiableAt, ChartedSpace.LiftPropAt] + +theorem mdifferentiableWithinAt_inter (ht : t ∈ 𝓝 x) : + MDifferentiableWithinAt I I' f (s ∩ t) x ↔ MDifferentiableWithinAt I I' f s x := by + rw [MDifferentiableWithinAt, MDifferentiableWithinAt, + (differentiable_within_at_localInvariantProp I I').liftPropWithinAt_inter ht] + +theorem mdifferentiableWithinAt_inter' (ht : t ∈ 𝓝[s] x) : + MDifferentiableWithinAt I I' f (s ∩ t) x ↔ MDifferentiableWithinAt I I' f s x := by + rw [MDifferentiableWithinAt, MDifferentiableWithinAt, + (differentiable_within_at_localInvariantProp I I').liftPropWithinAt_inter' ht] + +theorem MDifferentiableAt.mdifferentiableWithinAt (h : MDifferentiableAt I I' f x) : + MDifferentiableWithinAt I I' f s x := + MDifferentiableWithinAt.mono (subset_univ _) (mdifferentiableWithinAt_univ.2 h) + +theorem MDifferentiableWithinAt.mdifferentiableAt (h : MDifferentiableWithinAt I I' f s x) + (hs : s ∈ 𝓝 x) : MDifferentiableAt I I' f x := by + have : s = univ ∩ s := by rw [univ_inter] + rwa [this, mdifferentiableWithinAt_inter hs, mdifferentiableWithinAt_univ] at h + +theorem MDifferentiableOn.mono (h : MDifferentiableOn I I' f t) (st : s ⊆ t) : + MDifferentiableOn I I' f s := fun x hx => (h x (st hx)).mono st + +theorem mdifferentiableOn_univ : MDifferentiableOn I I' f univ ↔ MDifferentiable I I' f := by + simp only [MDifferentiableOn, mdifferentiableWithinAt_univ, mfld_simps]; rfl + +theorem MDifferentiableOn.mdifferentiableAt (h : MDifferentiableOn I I' f s) (hx : s ∈ 𝓝 x) : + MDifferentiableAt I I' f x := + (h x (mem_of_mem_nhds hx)).mdifferentiableAt hx + +theorem MDifferentiable.mdifferentiableOn (h : MDifferentiable I I' f) : + MDifferentiableOn I I' f s := + (mdifferentiableOn_univ.2 h).mono (subset_univ _) + +theorem mdifferentiableOn_of_locally_mdifferentiableOn + (h : ∀ x ∈ s, ∃ u, IsOpen u ∧ x ∈ u ∧ MDifferentiableOn I I' f (s ∩ u)) : + MDifferentiableOn I I' f s := by + intro x xs + rcases h x xs with ⟨t, t_open, xt, ht⟩ + exact (mdifferentiableWithinAt_inter (t_open.mem_nhds xt)).1 (ht x ⟨xs, xt⟩) + +/-! ### Deducing differentiability from smoothness -/ + +variable {n : ℕ∞} + +theorem ContMDiffWithinAt.mdifferentiableWithinAt (hf : ContMDiffWithinAt I I' n f s x) + (hn : 1 ≤ n) : MDifferentiableWithinAt I I' f s x := by + suffices h : MDifferentiableWithinAt I I' f (s ∩ f ⁻¹' (extChartAt I' (f x)).source) x by + rwa [mdifferentiableWithinAt_inter'] at h + apply hf.1.preimage_mem_nhdsWithin + exact extChartAt_source_mem_nhds I' (f x) + rw [mdifferentiableWithinAt_iff] + exact ⟨hf.1.mono inter_subset_left, (hf.2.differentiableWithinAt hn).mono (by mfld_set_tac)⟩ + +theorem ContMDiffAt.mdifferentiableAt (hf : ContMDiffAt I I' n f x) (hn : 1 ≤ n) : + MDifferentiableAt I I' f x := + mdifferentiableWithinAt_univ.1 <| ContMDiffWithinAt.mdifferentiableWithinAt hf hn + +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 + +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 + +nonrec theorem SmoothOn.mdifferentiableOn (hf : SmoothOn I I' f s) : MDifferentiableOn I I' f s := + hf.mdifferentiableOn le_top + +theorem Smooth.mdifferentiable (hf : Smooth I I' f) : MDifferentiable I I' f := + ContMDiff.mdifferentiable hf le_top + +theorem Smooth.mdifferentiableAt (hf : Smooth I I' f) : MDifferentiableAt I I' f x := + hf.mdifferentiable x + +theorem MDifferentiableOn.continuousOn (h : MDifferentiableOn I I' f s) : ContinuousOn f s := + fun x hx => (h x hx).continuousWithinAt + +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 + +/-! ### Deriving continuity from differentiability on manifolds -/ + +theorem MDifferentiableWithinAt.prod_mk {f : M → M'} {g : M → M''} + (hf : MDifferentiableWithinAt I I' f s x) (hg : MDifferentiableWithinAt I I'' g s x) : + MDifferentiableWithinAt I (I'.prod I'') (fun x => (f x, g x)) s x := + ⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ + +theorem MDifferentiableAt.prod_mk {f : M → M'} {g : M → M''} (hf : MDifferentiableAt I I' f x) + (hg : MDifferentiableAt I I'' g x) : + MDifferentiableAt I (I'.prod I'') (fun x => (f x, g x)) x := + ⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ + +theorem MDifferentiableWithinAt.prod_mk_space {f : M → E'} {g : M → E''} + (hf : MDifferentiableWithinAt I 𝓘(𝕜, E') f s x) + (hg : MDifferentiableWithinAt I 𝓘(𝕜, E'') g s x) : + MDifferentiableWithinAt I 𝓘(𝕜, E' × E'') (fun x => (f x, g x)) s x := + ⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ + +theorem MDifferentiableAt.prod_mk_space {f : M → E'} {g : M → E''} + (hf : MDifferentiableAt I 𝓘(𝕜, E') f x) (hg : MDifferentiableAt I 𝓘(𝕜, E'') g x) : + MDifferentiableAt I 𝓘(𝕜, E' × E'') (fun x => (f x, g x)) x := + ⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ + +theorem MDifferentiableOn.prod_mk {f : M → M'} {g : M → M''} (hf : MDifferentiableOn I I' f s) + (hg : MDifferentiableOn I I'' g s) : + MDifferentiableOn I (I'.prod I'') (fun x => (f x, g x)) s := fun x hx => + (hf x hx).prod_mk (hg x hx) + +theorem MDifferentiable.prod_mk {f : M → M'} {g : M → M''} (hf : MDifferentiable I I' f) + (hg : MDifferentiable I I'' g) : MDifferentiable I (I'.prod I'') fun x => (f x, g x) := fun x => + (hf x).prod_mk (hg x) + +theorem MDifferentiableOn.prod_mk_space {f : M → E'} {g : M → E''} + (hf : MDifferentiableOn I 𝓘(𝕜, E') f s) (hg : MDifferentiableOn I 𝓘(𝕜, E'') g s) : + MDifferentiableOn I 𝓘(𝕜, E' × E'') (fun x => (f x, g x)) s := fun x hx => + (hf x hx).prod_mk_space (hg x hx) + +theorem MDifferentiable.prod_mk_space {f : M → E'} {g : M → E''} (hf : MDifferentiable I 𝓘(𝕜, E') f) + (hg : MDifferentiable I 𝓘(𝕜, E'') g) : MDifferentiable I 𝓘(𝕜, E' × E'') fun x => (f x, g x) := + fun x => (hf x).prod_mk_space (hg x) + +theorem writtenInExtChartAt_comp (h : ContinuousWithinAt f s x) : + {y | writtenInExtChartAt I I'' x (g ∘ f) y = + (writtenInExtChartAt I' I'' (f x) g ∘ writtenInExtChartAt I I' x f) y} ∈ + 𝓝[(extChartAt I x).symm ⁻¹' s ∩ range I] (extChartAt I x) x := by + apply + @Filter.mem_of_superset _ _ (f ∘ (extChartAt I x).symm ⁻¹' (extChartAt I' (f x)).source) _ + (extChartAt_preimage_mem_nhdsWithin I + (h.preimage_mem_nhdsWithin (extChartAt_source_mem_nhds _ _))) + mfld_set_tac + /- We name the typeclass variables related to `SmoothManifoldWithCorners` structure as they are necessary in lemmas mentioning the derivative, but not in lemmas about differentiability, so we want to include them or omit them when necessary. -/ @@ -105,32 +274,12 @@ theorem UniqueMDiffOn.eq (U : UniqueMDiffOn I s) (hx : x ∈ s) (h : HasMFDerivW (h₁ : HasMFDerivWithinAt I I' f s x f₁') : f' = f₁' := UniqueMDiffWithinAt.eq (U _ hx) h h₁ -nonrec theorem UniqueMDiffWithinAt.prod {x : M} {y : M'} {s t} (hs : UniqueMDiffWithinAt I s x) - (ht : UniqueMDiffWithinAt I' t y) : UniqueMDiffWithinAt (I.prod I') (s ×ˢ t) (x, y) := by - refine (hs.prod ht).mono ?_ - rw [ModelWithCorners.range_prod, ← prod_inter_prod] - rfl - -theorem UniqueMDiffOn.prod {s : Set M} {t : Set M'} (hs : UniqueMDiffOn I s) - (ht : UniqueMDiffOn I' t) : UniqueMDiffOn (I.prod I') (s ×ˢ t) := fun x h ↦ - (hs x.1 h.1).prod (ht x.2 h.2) - /-! ### General lemmas on derivatives of functions between manifolds We mimick the API for functions between vector spaces -/ -theorem mdifferentiableWithinAt_iff {f : M → M'} {s : Set M} {x : M} : - MDifferentiableWithinAt I I' f s x ↔ - ContinuousWithinAt f s x ∧ - DifferentiableWithinAt 𝕜 (writtenInExtChartAt I I' x f) - ((extChartAt I x).target ∩ (extChartAt I x).symm ⁻¹' s) ((extChartAt I x) x) := by - rw [mdifferentiableWithinAt_iff'] - refine and_congr Iff.rfl (exists_congr fun f' => ?_) - rw [inter_comm] - simp only [HasFDerivWithinAt, nhdsWithin_inter, nhdsWithin_extChartAt_target_eq] - /-- One can reformulate differentiability within a set at a point as continuity within this set at this point, and differentiability in any chart containing that point. -/ theorem mdifferentiableWithinAt_iff_of_mem_source {x' : M} {y : M'} @@ -246,56 +395,6 @@ theorem mfderivWithin_subset (st : s ⊆ t) (hs : UniqueMDiffWithinAt I s x) mfderivWithin I I' f s x = mfderivWithin I I' f t x := ((MDifferentiableWithinAt.hasMFDerivWithinAt h).mono st).mfderivWithin hs -theorem MDifferentiableWithinAt.mono (hst : s ⊆ t) (h : MDifferentiableWithinAt I I' f t x) : - MDifferentiableWithinAt I I' f s x := - ⟨ContinuousWithinAt.mono h.1 hst, DifferentiableWithinAt.mono - h.differentiableWithinAt_writtenInExtChartAt - (inter_subset_inter_left _ (preimage_mono hst))⟩ - -theorem mdifferentiableWithinAt_univ : - MDifferentiableWithinAt I I' f univ x ↔ MDifferentiableAt I I' f x := by - simp_rw [MDifferentiableWithinAt, MDifferentiableAt, ChartedSpace.LiftPropAt] - -theorem mdifferentiableWithinAt_inter (ht : t ∈ 𝓝 x) : - MDifferentiableWithinAt I I' f (s ∩ t) x ↔ MDifferentiableWithinAt I I' f s x := by - rw [MDifferentiableWithinAt, MDifferentiableWithinAt, - (differentiable_within_at_localInvariantProp I I').liftPropWithinAt_inter ht] - -theorem mdifferentiableWithinAt_inter' (ht : t ∈ 𝓝[s] x) : - MDifferentiableWithinAt I I' f (s ∩ t) x ↔ MDifferentiableWithinAt I I' f s x := by - rw [MDifferentiableWithinAt, MDifferentiableWithinAt, - (differentiable_within_at_localInvariantProp I I').liftPropWithinAt_inter' ht] - -theorem MDifferentiableAt.mdifferentiableWithinAt (h : MDifferentiableAt I I' f x) : - MDifferentiableWithinAt I I' f s x := - MDifferentiableWithinAt.mono (subset_univ _) (mdifferentiableWithinAt_univ.2 h) - -theorem MDifferentiableWithinAt.mdifferentiableAt (h : MDifferentiableWithinAt I I' f s x) - (hs : s ∈ 𝓝 x) : MDifferentiableAt I I' f x := by - have : s = univ ∩ s := by rw [univ_inter] - rwa [this, mdifferentiableWithinAt_inter hs, mdifferentiableWithinAt_univ] at h - -theorem MDifferentiableOn.mdifferentiableAt (h : MDifferentiableOn I I' f s) (hx : s ∈ 𝓝 x) : - MDifferentiableAt I I' f x := - (h x (mem_of_mem_nhds hx)).mdifferentiableAt hx - -theorem MDifferentiableOn.mono (h : MDifferentiableOn I I' f t) (st : s ⊆ t) : - MDifferentiableOn I I' f s := fun x hx => (h x (st hx)).mono st - -theorem mdifferentiableOn_univ : MDifferentiableOn I I' f univ ↔ MDifferentiable I I' f := by - simp only [MDifferentiableOn, mdifferentiableWithinAt_univ, mfld_simps]; rfl - -theorem MDifferentiable.mdifferentiableOn (h : MDifferentiable I I' f) : - MDifferentiableOn I I' f s := - (mdifferentiableOn_univ.2 h).mono (subset_univ _) - -theorem mdifferentiableOn_of_locally_mdifferentiableOn - (h : ∀ x ∈ s, ∃ u, IsOpen u ∧ x ∈ u ∧ MDifferentiableOn I I' f (s ∩ u)) : - MDifferentiableOn I I' f s := by - intro x xs - rcases h x xs with ⟨t, t_open, xt, ht⟩ - exact (mdifferentiableWithinAt_inter (t_open.mem_nhds xt)).1 (ht x ⟨xs, xt⟩) - @[simp, mfld_simps] theorem mfderivWithin_univ : mfderivWithin I I' f univ = mfderiv I I' f := by ext x : 1 @@ -329,49 +428,6 @@ theorem mdifferentiableAt_iff_of_mem_source {x' : M} {y : M'} (mdifferentiableWithinAt_iff_of_mem_source hx hy).trans <| by rw [continuousWithinAt_univ, Set.preimage_univ, Set.univ_inter] -/-! ### Deducing differentiability from smoothness -/ - --- Porting note: moved from `ContMDiffMFDeriv` -variable {n : ℕ∞} - -theorem ContMDiffWithinAt.mdifferentiableWithinAt (hf : ContMDiffWithinAt I I' n f s x) - (hn : 1 ≤ n) : MDifferentiableWithinAt I I' f s x := by - suffices h : MDifferentiableWithinAt I I' f (s ∩ f ⁻¹' (extChartAt I' (f x)).source) x by - rwa [mdifferentiableWithinAt_inter'] at h - apply hf.1.preimage_mem_nhdsWithin - exact extChartAt_source_mem_nhds I' (f x) - rw [mdifferentiableWithinAt_iff] - exact ⟨hf.1.mono inter_subset_left, (hf.2.differentiableWithinAt hn).mono (by mfld_set_tac)⟩ - -theorem ContMDiffAt.mdifferentiableAt (hf : ContMDiffAt I I' n f x) (hn : 1 ≤ n) : - MDifferentiableAt I I' f x := - mdifferentiableWithinAt_univ.1 <| ContMDiffWithinAt.mdifferentiableWithinAt hf hn - -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 - -theorem ContMDiff.mdifferentiable (hf : ContMDiff I I' n f) (hn : 1 ≤ n) : MDifferentiable I I' f := - fun x => (hf x).mdifferentiableAt hn - -nonrec theorem SmoothWithinAt.mdifferentiableWithinAt (hf : SmoothWithinAt I I' f s x) : - MDifferentiableWithinAt I I' f s x := - hf.mdifferentiableWithinAt le_top - -nonrec theorem SmoothAt.mdifferentiableAt (hf : SmoothAt I I' f x) : MDifferentiableAt I I' f x := - hf.mdifferentiableAt le_top - -nonrec theorem SmoothOn.mdifferentiableOn (hf : SmoothOn I I' f s) : MDifferentiableOn I I' f s := - hf.mdifferentiableOn le_top - -theorem Smooth.mdifferentiable (hf : Smooth I I' f) : MDifferentiable I I' f := - ContMDiff.mdifferentiable hf le_top - -theorem Smooth.mdifferentiableAt (hf : Smooth I I' f) : MDifferentiableAt I I' f x := - hf.mdifferentiable x - -theorem Smooth.mdifferentiableWithinAt (hf : Smooth I I' f) : MDifferentiableWithinAt I I' f s x := - hf.mdifferentiableAt.mdifferentiableWithinAt - /-! ### Deriving continuity from differentiability on manifolds -/ theorem HasMFDerivWithinAt.continuousWithinAt (h : HasMFDerivWithinAt I I' f s x f') : @@ -381,12 +437,6 @@ theorem HasMFDerivWithinAt.continuousWithinAt (h : HasMFDerivWithinAt I I' f s x theorem HasMFDerivAt.continuousAt (h : HasMFDerivAt I I' f x f') : ContinuousAt f x := h.1 -theorem MDifferentiableOn.continuousOn (h : MDifferentiableOn I I' f s) : ContinuousOn f s := - fun x hx => (h x hx).continuousWithinAt - -theorem MDifferentiable.continuous (h : MDifferentiable I I' f) : Continuous f := - continuous_iff_continuousAt.2 fun x => (h x).continuousAt - theorem tangentMapWithin_subset {p : TangentBundle I M} (st : s ⊆ t) (hs : UniqueMDiffWithinAt I s p.1) (h : MDifferentiableWithinAt I I' f t p.1) : tangentMapWithin I I' f s p = tangentMapWithin I I' f t p := by @@ -412,45 +462,6 @@ theorem tangentMapWithin_proj {p : TangentBundle I M} : theorem tangentMap_proj {p : TangentBundle I M} : (tangentMap I I' f p).proj = f p.proj := rfl -theorem MDifferentiableWithinAt.prod_mk {f : M → M'} {g : M → M''} - (hf : MDifferentiableWithinAt I I' f s x) (hg : MDifferentiableWithinAt I I'' g s x) : - MDifferentiableWithinAt I (I'.prod I'') (fun x => (f x, g x)) s x := - ⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ - -theorem MDifferentiableAt.prod_mk {f : M → M'} {g : M → M''} (hf : MDifferentiableAt I I' f x) - (hg : MDifferentiableAt I I'' g x) : - MDifferentiableAt I (I'.prod I'') (fun x => (f x, g x)) x := - ⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ - -theorem MDifferentiableOn.prod_mk {f : M → M'} {g : M → M''} (hf : MDifferentiableOn I I' f s) - (hg : MDifferentiableOn I I'' g s) : - MDifferentiableOn I (I'.prod I'') (fun x => (f x, g x)) s := fun x hx => - (hf x hx).prod_mk (hg x hx) - -theorem MDifferentiable.prod_mk {f : M → M'} {g : M → M''} (hf : MDifferentiable I I' f) - (hg : MDifferentiable I I'' g) : MDifferentiable I (I'.prod I'') fun x => (f x, g x) := fun x => - (hf x).prod_mk (hg x) - -theorem MDifferentiableWithinAt.prod_mk_space {f : M → E'} {g : M → E''} - (hf : MDifferentiableWithinAt I 𝓘(𝕜, E') f s x) - (hg : MDifferentiableWithinAt I 𝓘(𝕜, E'') g s x) : - MDifferentiableWithinAt I 𝓘(𝕜, E' × E'') (fun x => (f x, g x)) s x := - ⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ - -theorem MDifferentiableAt.prod_mk_space {f : M → E'} {g : M → E''} - (hf : MDifferentiableAt I 𝓘(𝕜, E') f x) (hg : MDifferentiableAt I 𝓘(𝕜, E'') g x) : - MDifferentiableAt I 𝓘(𝕜, E' × E'') (fun x => (f x, g x)) x := - ⟨hf.1.prod hg.1, hf.2.prod hg.2⟩ - -theorem MDifferentiableOn.prod_mk_space {f : M → E'} {g : M → E''} - (hf : MDifferentiableOn I 𝓘(𝕜, E') f s) (hg : MDifferentiableOn I 𝓘(𝕜, E'') g s) : - MDifferentiableOn I 𝓘(𝕜, E' × E'') (fun x => (f x, g x)) s := fun x hx => - (hf x hx).prod_mk_space (hg x hx) - -theorem MDifferentiable.prod_mk_space {f : M → E'} {g : M → E''} (hf : MDifferentiable I 𝓘(𝕜, E') f) - (hg : MDifferentiable I 𝓘(𝕜, E'') g) : MDifferentiable I 𝓘(𝕜, E' × E'') fun x => (f x, g x) := - fun x => (hf x).prod_mk_space (hg x) - /-! ### Congruence lemmas for derivatives on manifolds -/ theorem HasMFDerivAt.congr_mfderiv (h : HasMFDerivAt I I' f x f') (h' : f' = f₁') : @@ -562,16 +573,6 @@ theorem mfderiv_congr {f' : M → M'} (h : f = f') : /-! ### Composition lemmas -/ -theorem writtenInExtChartAt_comp (h : ContinuousWithinAt f s x) : - {y | writtenInExtChartAt I I'' x (g ∘ f) y = - (writtenInExtChartAt I' I'' (f x) g ∘ writtenInExtChartAt I I' x f) y} ∈ - 𝓝[(extChartAt I x).symm ⁻¹' s ∩ range I] (extChartAt I x) x := by - apply - @Filter.mem_of_superset _ _ (f ∘ (extChartAt I x).symm ⁻¹' (extChartAt I' (f x)).source) _ - (extChartAt_preimage_mem_nhdsWithin I - (h.preimage_mem_nhdsWithin (extChartAt_source_mem_nhds _ _))) - mfld_set_tac - variable (x) theorem HasMFDerivWithinAt.comp (hg : HasMFDerivWithinAt I' I'' g u (f x) g') diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index e4fa08c639677..7b50a99e3073a 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -175,9 +175,6 @@ theorem finsum_smul_mem_convex {g : ι → M → F} {t : Set F} {x : M} (hx : x (hg : ∀ i, f i x ≠ 0 → g i x ∈ t) (ht : Convex ℝ t) : ∑ᶠ i, f i x • g i x ∈ t := ht.finsum_mem (fun _ => f.nonneg _ _) (f.sum_eq_one hx) hg -section SmoothManifoldWithCorners -variable [SmoothManifoldWithCorners I M] - theorem contMDiff_smul {g : M → F} {i} (hg : ∀ x ∈ tsupport (f i), ContMDiffAt I 𝓘(ℝ, F) n g x) : ContMDiff I 𝓘(ℝ, F) n fun x => f i x • g x := contMDiff_of_tsupport fun x hx => @@ -219,8 +216,6 @@ theorem contDiffAt_finsum {s : Set E} (f : SmoothPartitionOfUnity ι 𝓘(ℝ, E simp only [← contMDiffAt_iff_contDiffAt] at * exact f.contMDiffAt_finsum hφ -end SmoothManifoldWithCorners - section finsupport variable {s : Set M} (ρ : SmoothPartitionOfUnity ι I M s) (x₀ : M) @@ -293,8 +288,6 @@ theorem isSubordinate_toPartitionOfUnity : alias ⟨_, IsSubordinate.toPartitionOfUnity⟩ := isSubordinate_toPartitionOfUnity -variable [SmoothManifoldWithCorners I M] - /-- 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 $C^n$ smooth on `U i`, then the sum `fun x ↦ ∑ᶠ i, f i x • g i x` is $C^n$ smooth on the whole manifold. -/ @@ -380,7 +373,7 @@ theorem IsSubordinate.support_subset {fs : SmoothBumpCovering ι I M s} {U : M Subset.trans subset_closure (h i) variable (I) in -variable [SmoothManifoldWithCorners I M] in + /-- Let `M` be a smooth manifold with corners modelled on a finite dimensional real vector space. Suppose also that `M` is a Hausdorff `σ`-compact topological space. Let `s` be a closed set in `M` and `U : M → Set M` be a collection of sets such that `U x ∈ 𝓝 x` for every `x ∈ s`. @@ -426,8 +419,6 @@ theorem apply_ind (x : M) (hx : x ∈ s) : fs (fs.ind x hx) x = 1 := theorem mem_support_ind (x : M) (hx : x ∈ s) : x ∈ support (fs <| fs.ind x hx) := by simp [fs.apply_ind x hx] -variable [SmoothManifoldWithCorners I M] - theorem mem_chartAt_source_of_eq_one {i : ι} {x : M} (h : fs i x = 1) : x ∈ (chartAt H (fs.c i)).source := (fs i).support_subset_source <| by simp [h] @@ -448,6 +439,7 @@ protected def fintype [CompactSpace M] : Fintype ι := fs.locallyFinite.fintypeOfCompact fun i => (fs i).nonempty_support variable [T2Space M] +variable [SmoothManifoldWithCorners I M] /-- Reinterpret a `SmoothBumpCovering` as a continuous `BumpCovering`. Note that not every `f : BumpCovering ι M s` with smooth functions `f i` is a `SmoothBumpCovering`. -/ diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index 2f4709813f8fd..ff5840676376f 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -32,10 +32,8 @@ variable (F : Type*) [NormedAddCommGroup F] [NormedSpace 𝕜 F] (n : ℕ∞) (V : M → Type*) [TopologicalSpace (TotalSpace F V)] -- `V` vector bundle - [∀ x : M, TopologicalSpace (V x)] [FiberBundle F V] - /-- Bundled `n` times continuously differentiable sections of a vector bundle. -/ structure ContMDiffSection where /-- the underlying function of this section -/ diff --git a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean index e40360732b4b1..09569f2558908 100644 --- a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean +++ b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean @@ -47,7 +47,7 @@ In this section we prove a version of the Whitney embedding theorem: for any com `M`, for sufficiently large `n` there exists a smooth embedding `M → ℝ^n`. -/ -variable [T2Space M] [hi : Fintype ι] {s : Set M} (f : SmoothBumpCovering ι I M s) +variable [T2Space M] [Fintype ι] {s : Set M} (f : SmoothBumpCovering ι I M s) /-- Smooth embedding of `M` into `(E × ℝ) ^ ι`. -/ def embeddingPiTangent : C^∞⟮I, M; 𝓘(ℝ, ι → E × ℝ), ι → E × ℝ⟯ where @@ -106,7 +106,7 @@ theorem embeddingPiTangent_injective_mfderiv (x : M) (hx : x ∈ s) : /-- Baby version of the **Whitney weak embedding theorem**: if `M` admits a finite covering by supports of bump functions, then for some `n` it can be immersed into the `n`-dimensional Euclidean space. -/ -theorem exists_immersion_euclidean [Finite ι] (f : SmoothBumpCovering ι I M) : +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 cases nonempty_fintype ι diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index 6d942b7c97ac8..547d6a74acec3 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -314,7 +314,6 @@ instance (i : ι) : Inhabited (Pair M i) := ⟨⟨1, empty, by tauto⟩⟩ variable {M} -variable [∀ i, DecidableEq (M i)] /-- Construct a new `Word` without any reduction. The underlying list of `cons m w _ _` is `⟨_, m⟩::w` -/ @@ -328,6 +327,18 @@ def cons {i} (m : M i) (w : Word M) (hmw : w.fstIdx ≠ some i) (h1 : m ≠ 1) : · exact w.ne_one l hl chain_ne := w.chain_ne.cons' (fstIdx_ne_iff.mp hmw) } +@[simp] +theorem fstIdx_cons {i} (m : M i) (w : Word M) (hmw : w.fstIdx ≠ some i) (h1 : m ≠ 1) : + fstIdx (cons m w hmw h1) = some i := by simp [cons, fstIdx] + +@[simp] +theorem prod_cons (i) (m : M i) (w : Word M) (h1 : m ≠ 1) (h2 : w.fstIdx ≠ some i) : + prod (cons m w h2 h1) = of m * prod w := by + simp [cons, prod, List.map_cons, List.prod_cons] + +section +variable [∀ i, DecidableEq (M i)] + /-- Given a pair `(head, tail)`, we can form a word by prepending `head` to `tail`, except if `head` is `1 : M i` then we have to just return `Word` since we need the result to be reduced. -/ def rcons {i} (p : Pair M i) : Word M := @@ -371,14 +382,7 @@ theorem mem_rcons_iff {i j : ι} (p : Pair M i) (m : M j) : · split_ifs <;> simp_all · split_ifs <;> simp_all [Ne.symm hij] -@[simp] -theorem fstIdx_cons {i} (m : M i) (w : Word M) (hmw : w.fstIdx ≠ some i) (h1 : m ≠ 1) : - fstIdx (cons m w hmw h1) = some i := by simp [cons, fstIdx] - -@[simp] -theorem prod_cons (i) (m : M i) (w : Word M) (h1 : m ≠ 1) (h2 : w.fstIdx ≠ some i) : - prod (cons m w h2 h1) = of m * prod w := by - simp [cons, prod, List.map_cons, List.prod_cons] +end /-- Induct on a word by adding letters one at a time without reduction, effectively inducting on the underlying `List`. -/ @@ -410,7 +414,7 @@ theorem consRecOn_cons {motive : Word M → Sort*} (i) (m : M i) (w : Word M) h1 consRecOn (cons m w h1 h2) h_empty h_cons = h_cons i m w h1 h2 (consRecOn w h_empty h_cons) := rfl -variable [DecidableEq ι] +variable [DecidableEq ι] [∀ i, DecidableEq (M i)] -- This definition is computable but not very nice to look at. Thankfully we don't have to inspect -- it, since `rcons` is known to be injective. @@ -807,7 +811,6 @@ open Pointwise open Cardinal -variable [hnontriv : Nontrivial ι] variable {G : Type*} [Group G] variable {H : ι → Type*} [∀ i, Group (H i)] variable (f : ∀ i, H i →* G) @@ -821,6 +824,7 @@ variable (X : ι → Set α) variable (hXnonempty : ∀ i, (X i).Nonempty) variable (hXdisj : Pairwise fun i j => Disjoint (X i) (X j)) variable (hpp : Pairwise fun i j => ∀ h : H i, h ≠ 1 → f i h • X j ⊆ X i) +include hpp theorem lift_word_ping_pong {i j k} (w : NeWord H i j) (hk : j ≠ k) : lift f w.prod • X k ⊆ X i := by @@ -832,6 +836,8 @@ theorem lift_word_ping_pong {i j k} (w : NeWord H i j) (hk : j ≠ k) : _ ⊆ lift f w₁.prod • X _ := set_smul_subset_set_smul_iff.mpr (hIw₂ hk) _ ⊆ X i := hIw₁ hne +include hXnonempty hXdisj + theorem lift_word_prod_nontrivial_of_other_i {i j k} (w : NeWord H i j) (hhead : k ≠ i) (hlast : k ≠ j) : lift f w.prod ≠ 1 := by intro heq1 @@ -839,12 +845,15 @@ theorem lift_word_prod_nontrivial_of_other_i {i j k} (w : NeWord H i j) (hhead : obtain ⟨x, hx⟩ := hXnonempty k exact (hXdisj hhead).le_bot ⟨hx, this hx⟩ -theorem lift_word_prod_nontrivial_of_head_eq_last {i} (w : NeWord H i i) : lift f w.prod ≠ 1 := by +variable [Nontrivial ι] + +theorem lift_word_prod_nontrivial_of_head_eq_last {i} (w : NeWord H i i) : + lift f w.prod ≠ 1 := by obtain ⟨k, hk⟩ := exists_ne i exact lift_word_prod_nontrivial_of_other_i f X hXnonempty hXdisj hpp w hk hk -theorem lift_word_prod_nontrivial_of_head_card {i j} (w : NeWord H i j) (hcard : 3 ≤ #(H i)) - (hheadtail : i ≠ j) : lift f w.prod ≠ 1 := by +theorem lift_word_prod_nontrivial_of_head_card {i j} (w : NeWord H i j) + (hcard : 3 ≤ #(H i)) (hheadtail : i ≠ j) : lift f w.prod ≠ 1 := by obtain ⟨h, hn1, hnh⟩ := Cardinal.three_le hcard 1 w.head⁻¹ have hnot1 : h * w.head ≠ 1 := by rw [← div_inv_eq_mul] @@ -858,7 +867,9 @@ theorem lift_word_prod_nontrivial_of_head_card {i j} (w : NeWord H i j) (hcard : apply hw' simp [w', heq1] -theorem lift_word_prod_nontrivial_of_not_empty {i j} (w : NeWord H i j) : lift f w.prod ≠ 1 := by +include hcard in +theorem lift_word_prod_nontrivial_of_not_empty {i j} (w : NeWord H i j) : + lift f w.prod ≠ 1 := by classical cases' hcard with hcard hcard · obtain ⟨i, h1, h2⟩ := Cardinal.three_le hcard i j @@ -890,11 +901,14 @@ theorem lift_word_prod_nontrivial_of_not_empty {i j} (w : NeWord H i j) : lift f apply hw' simp [w', heq1] -theorem empty_of_word_prod_eq_one {w : Word H} (h : lift f w.prod = 1) : w = Word.empty := by +include hcard in +theorem empty_of_word_prod_eq_one {w : Word H} (h : lift f w.prod = 1) : + w = Word.empty := by by_contra hnotempty obtain ⟨i, j, w, rfl⟩ := NeWord.of_word w hnotempty exact lift_word_prod_nontrivial_of_not_empty f hcard X hXnonempty hXdisj hpp w h +include hcard in /-- The **Ping-Pong-Lemma**. Given a group action of `G` on `X` so that the `H i` acts in a specific way on disjoint subsets @@ -964,6 +978,7 @@ variable (hXYdisj : ∀ i j, Disjoint (X i) (Y j)) variable (hX : ∀ i, a i • (Y i)ᶜ ⊆ X i) variable (hY : ∀ i, a⁻¹ i • (X i)ᶜ ⊆ Y i) +include hXnonempty hXdisj hYdisj hXYdisj hX hY in /-- The Ping-Pong-Lemma. Given a group action of `G` on `X` so that the generators of the free groups act in specific diff --git a/Mathlib/GroupTheory/CosetCover.lean b/Mathlib/GroupTheory/CosetCover.lean index 2accbbd944243..c9cefdcb06558 100644 --- a/Mathlib/GroupTheory/CosetCover.lean +++ b/Mathlib/GroupTheory/CosetCover.lean @@ -77,6 +77,7 @@ theorem leftCoset_cover_const_iff_surjOn : QuotientGroup.forall_mk, QuotientGroup.eq] variable (hcovers : ⋃ i ∈ s, g i • (H : Set G) = Set.univ) +include hcovers /-- If `H` is a subgroup of `G` and `G` is the union of a finite family of left cosets of `H` then `H` has finite index. -/ @@ -119,6 +120,7 @@ section variable {ι : Type*} {H : ι → Subgroup G} {g : ι → G} {s : Finset ι} (hcovers : ⋃ i ∈ s, (g i) • (H i : Set G) = Set.univ) +include hcovers -- Inductive inner part of `Subgroup.exists_finiteIndex_of_leftCoset_cover` @[to_additive] @@ -353,9 +355,8 @@ section Submodule variable {R M ι : Type*} [Ring R] [AddCommGroup M] [Module R M] {p : ι → Submodule R M} {s : Finset ι} - (hcovers : ⋃ i ∈ s, (p i : Set M) = Set.univ) -theorem Submodule.exists_finiteIndex_of_cover : +theorem Submodule.exists_finiteIndex_of_cover (hcovers : ⋃ i ∈ s, (p i : Set M) = Set.univ) : ∃ k ∈ s, (p k).toAddSubgroup.FiniteIndex := have hcovers' : ⋃ i ∈ s, (0 : M) +ᵥ ((p i).toAddSubgroup : Set M) = Set.univ := by simpa only [zero_vadd] using hcovers diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index 8ca8f8ad2b96a..ff587ef4dcc2e 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -65,6 +65,7 @@ namespace IsReflection variable {cs} variable {t : W} (ht : cs.IsReflection t) +include ht theorem pow_two : t ^ 2 = 1 := by rcases ht with ⟨w, i, rfl⟩ @@ -138,6 +139,7 @@ namespace IsReflection variable {cs} variable {t : W} (ht : cs.IsReflection t) +include ht theorem isRightInversion_mul_left_iff {w : W} : cs.IsRightInversion (w * t) t ↔ ¬cs.IsRightInversion w t := by diff --git a/Mathlib/GroupTheory/EckmannHilton.lean b/Mathlib/GroupTheory/EckmannHilton.lean index 9c6b5b653e7f9..9fb670e07acbf 100644 --- a/Mathlib/GroupTheory/EckmannHilton.lean +++ b/Mathlib/GroupTheory/EckmannHilton.lean @@ -43,6 +43,8 @@ variable {m₁ m₂ : X → X → X} {e₁ e₂ : X} variable (h₁ : IsUnital m₁ e₁) (h₂ : IsUnital m₂ e₂) variable (distrib : ∀ a b c d, ((a b) c d) = (a c) b d) +include h₁ h₂ distrib + /-- If a type carries two unital binary operations that distribute over each other, then they have the same unit elements. diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index 6a0ff7eea9fcc..1b48a93653134 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -57,6 +57,7 @@ alias ⟨exists_card_eq, _⟩ := iff_card section GIsPGroup variable (hG : IsPGroup p G) +include hG theorem of_injective {H : Type*} [Group H] (ϕ : H →* G) (hϕ : Function.Injective ϕ) : IsPGroup p H := by diff --git a/Mathlib/GroupTheory/SchurZassenhaus.lean b/Mathlib/GroupTheory/SchurZassenhaus.lean index dfa79ebca9a3b..1cb524d82f333 100644 --- a/Mathlib/GroupTheory/SchurZassenhaus.lean +++ b/Mathlib/GroupTheory/SchurZassenhaus.lean @@ -130,12 +130,13 @@ The proof is by contradiction. We assume that `G` is a minimal counterexample to -/ -variable {G : Type u} [Group G] [Finite G] {N : Subgroup G} [Normal N] +variable {G : Type u} [Group G] {N : Subgroup G} [Normal N] (h1 : Nat.Coprime (Nat.card N) N.index) (h2 : ∀ (G' : Type u) [Group G'] [Finite G'], Nat.card G' < Nat.card G → ∀ {N' : Subgroup G'} [N'.Normal], Nat.Coprime (Nat.card N') N'.index → ∃ H' : Subgroup G', IsComplement' N' H') (h3 : ∀ H : Subgroup G, ¬IsComplement' N H) +include h1 h3 /-! We will arrive at a contradiction via the following steps: * step 0: `N` (the normal Hall subgroup) is nontrivial. @@ -154,6 +155,9 @@ private theorem step0 : N ≠ ⊥ := by rintro rfl exact h3 ⊤ isComplement'_bot_top +variable [Finite G] + +include h2 in /-- Do not use this lemma: It is made obsolete by `exists_right_complement'_of_coprime` -/ private theorem step1 (K : Subgroup G) (hK : K ⊔ N = ⊤) : K = ⊤ := by contrapose! h3 @@ -177,6 +181,7 @@ private theorem step1 (K : Subgroup G) (hK : K ⊔ N = ⊤) : K = ⊤ := by rwa [hH] exact ⟨H.map K.subtype, isComplement'_of_coprime h7 h8⟩ +include h2 in /-- Do not use this lemma: It is made obsolete by `exists_right_complement'_of_coprime` -/ private theorem step2 (K : Subgroup G) [K.Normal] (hK : K ≤ N) : K = ⊥ ∨ K = N := by have : Function.Surjective (QuotientGroup.mk' K) := Quotient.surjective_Quotient_mk'' @@ -207,6 +212,7 @@ private theorem step2 (K : Subgroup G) [K.Normal] (hK : K ≤ N) : K = ⊥ ∨ K QuotientGroup.ker_mk'] at hH exact h4.2 (le_antisymm hK hH) +include h2 in /-- Do not use this lemma: It is made obsolete by `exists_right_complement'_of_coprime` -/ private theorem step3 (K : Subgroup N) [(K.map N.subtype).Normal] : K = ⊥ ∨ K = ⊤ := by have key := step2 h1 h2 h3 (K.map N.subtype) (map_subtype_le K) @@ -228,6 +234,7 @@ private theorem step5 {P : Sylow (Nat.card N).minFac N} : P.1 ≠ ⊥ := by apply P.ne_bot_of_dvd_card exact (Nat.card N).minFac_dvd +include h2 in /-- Do not use this lemma: It is made obsolete by `exists_right_complement'_of_coprime` -/ private theorem step6 : IsPGroup (Nat.card N).minFac N := by haveI : Fact (Nat.card N).minFac.Prime := ⟨step4 h1 h3⟩ @@ -237,6 +244,7 @@ private theorem step6 : IsPGroup (Nat.card N).minFac N := by 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) +include h2 in /-- Do not use this lemma: It is made obsolete by `exists_right_complement'_of_coprime` -/ theorem step7 : IsCommutative N := by haveI := N.bot_or_nontrivial.resolve_left (step0 h1 h3) diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 3be24db573211..965253cc75dc7 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -342,18 +342,7 @@ protected alias IsCyclic.iff_exists_ofOrder_eq_natCard_of_Fintype := section -variable [DecidableEq α] [Fintype α] - -@[to_additive] -theorem IsCyclic.image_range_orderOf (ha : ∀ x : α, x ∈ zpowers a) : - Finset.image (fun i => a ^ i) (range (orderOf a)) = univ := by - simp_rw [← SetLike.mem_coe] at ha - simp only [_root_.image_range_orderOf, Set.eq_univ_iff_forall.mpr ha, Set.toFinset_univ] - -@[to_additive] -theorem IsCyclic.image_range_card (ha : ∀ x : α, x ∈ zpowers a) : - Finset.image (fun i => a ^ i) (range (Fintype.card α)) = univ := by - rw [← orderOf_eq_card_of_forall_mem_zpowers ha, IsCyclic.image_range_orderOf ha] +variable [Fintype α] @[to_additive] theorem IsCyclic.unique_zpow_zmod (ha : ∀ x : α, x ∈ zpowers a) (x : α) : @@ -366,6 +355,19 @@ theorem IsCyclic.unique_zpow_zmod (ha : ∀ x : α, x ∈ zpowers a) (x : α) : ← ZMod.intCast_eq_intCast_iff] at hy simp [hy] +variable [DecidableEq α] + +@[to_additive] +theorem IsCyclic.image_range_orderOf (ha : ∀ x : α, x ∈ zpowers a) : + Finset.image (fun i => a ^ i) (range (orderOf a)) = univ := by + simp_rw [← SetLike.mem_coe] at ha + simp only [_root_.image_range_orderOf, Set.eq_univ_iff_forall.mpr ha, Set.toFinset_univ] + +@[to_additive] +theorem IsCyclic.image_range_card (ha : ∀ x : α, x ∈ zpowers a) : + Finset.image (fun i => a ^ i) (range (Fintype.card α)) = univ := by + rw [← orderOf_eq_card_of_forall_mem_zpowers ha, IsCyclic.image_range_orderOf ha] + @[to_additive] lemma IsCyclic.ext {G : Type*} [Group G] [Fintype G] [IsCyclic G] {d : ℕ} {a b : ZMod d} (hGcard : Fintype.card G = d) (h : ∀ t : G, t ^ a.val = t ^ b.val) : a = b := by @@ -382,6 +384,7 @@ section Totient variable [DecidableEq α] [Fintype α] (hn : ∀ n : ℕ, 0 < n → (univ.filter fun a : α => a ^ n = 1).card ≤ n) +include hn @[to_additive] private theorem card_pow_eq_one_eq_orderOf_aux (a : α) : @@ -786,6 +789,8 @@ lemma monoidHomOfForallMemZpowers_apply_gen : end monoidHom +include hg + /-- Two group homomorphisms `G →* G'` are equal if and only if they agree on a generator of `G`. -/ @[to_additive "Two homomorphisms `G →+ G'` of additive groups are equal if and only if they agree diff --git a/Mathlib/GroupTheory/Transfer.lean b/Mathlib/GroupTheory/Transfer.lean index f1d47f133261b..2ade7550cbcff 100644 --- a/Mathlib/GroupTheory/Transfer.lean +++ b/Mathlib/GroupTheory/Transfer.lean @@ -185,6 +185,7 @@ theorem transferCenterPow_apply [FiniteIndex (center G)] (g : G) : section BurnsideTransfer variable {p : ℕ} (P : Sylow p G) (hP : (P : Subgroup G).normalizer ≤ centralizer (P : Set G)) +include hP /-- The homomorphism `G →* P` in Burnside's transfer theorem. -/ noncomputable def transferSylow [FiniteIndex (P : Subgroup G)] : G →* (P : Subgroup G) := diff --git a/Mathlib/LinearAlgebra/AffineSpace/Basis.lean b/Mathlib/LinearAlgebra/AffineSpace/Basis.lean index b3c2e7715e81e..7d6f380b7c9dc 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Basis.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Basis.lean @@ -78,6 +78,7 @@ theorem ind : AffineIndependent k b := theorem tot : affineSpan k (range b) = ⊤ := b.tot' +include b in protected theorem nonempty : Nonempty ι := not_isEmpty_iff.mp fun hι => by simpa only [@range_eq_empty _ _ hι, AffineSubspace.span_empty, bot_ne_top] using b.tot diff --git a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean index 3b04b67dd0bca..6d283deb55741 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean @@ -413,6 +413,9 @@ theorem dualBasis_repr_apply (B.dualBasis hB b).repr x i = B x (b i) := by rw [dualBasis, Basis.map_repr, LinearEquiv.symm_symm, LinearEquiv.trans_apply, Basis.dualBasis_repr] + #adaptation_note + /-- Before leanprover/lean4#4814, we had a final `toDual_def` in the `rw`, + and no need for the `rfl`. I'm confused! -/ rfl theorem apply_dualBasis_left (B : BilinForm K V) (hB : B.Nondegenerate) (b : Basis ι K V) (i j) : diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean index b1e419160e7e2..978452899a1a5 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Contraction.lean @@ -231,7 +231,6 @@ theorem changeFormAux_changeFormAux (B : BilinForm R M) (v : M) (x : CliffordAlg variable {Q} variable {Q' Q'' : QuadraticForm R M} {B B' : BilinForm R M} -variable (h : B.toQuadraticMap = Q' - Q) (h' : B'.toQuadraticMap = Q'' - Q') /-- Convert between two algebras of different quadratic form, sending vector to vectors, scalars to scalars, and adjusting products by a contraction term. @@ -249,10 +248,14 @@ def changeForm (h : B.toQuadraticMap = Q' - Q) : CliffordAlgebra Q →ₗ[R] Cli theorem changeForm.zero_proof : (0 : BilinForm R M).toQuadraticMap = Q - Q := (sub_self _).symm +variable (h : B.toQuadraticMap = Q' - Q) (h' : B'.toQuadraticMap = Q'' - Q') + +include h h' in /-- Auxiliary lemma used as an argument to `CliffordAlgebra.changeForm` -/ theorem changeForm.add_proof : (B + B').toQuadraticMap = Q'' - Q := (congr_arg₂ (· + ·) h h').trans <| sub_add_sub_cancel' _ _ _ +include h in /-- Auxiliary lemma used as an argument to `CliffordAlgebra.changeForm` -/ theorem changeForm.neg_proof : (-B).toQuadraticMap = Q - Q' := (congr_arg Neg.neg h).trans <| neg_sub _ _ diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Prod.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Prod.lean index 457c0421d2c35..fab6f0c7bfae4 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Prod.lean @@ -41,6 +41,7 @@ section map_mul_map variable {Q₁ Q₂ Qₙ} variable (f₁ : Q₁ →qᵢ Qₙ) (f₂ : Q₂ →qᵢ Qₙ) (hf : ∀ x y, Qₙ.IsOrtho (f₁ x) (f₂ y)) variable (m₁ : CliffordAlgebra Q₁) (m₂ : CliffordAlgebra Q₂) +include hf /-- If `m₁` and `m₂` are both homogenous, and the quadratic spaces `Q₁` and `Q₂` map into diff --git a/Mathlib/LinearAlgebra/Dimension/Localization.lean b/Mathlib/LinearAlgebra/Dimension/Localization.lean index 5ffa00e14abf1..215eca3c7a657 100644 --- a/Mathlib/LinearAlgebra/Dimension/Localization.lean +++ b/Mathlib/LinearAlgebra/Dimension/Localization.lean @@ -29,6 +29,12 @@ variable [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S variable (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] variable (hp : p ≤ R⁰) +section +include hp + +section +include f + variable {S} in lemma IsLocalizedModule.linearIndependent_lift {ι} {v : ι → N} (hf : LinearIndependent S v) : ∃ w : ι → M, LinearIndependent R w := by @@ -77,10 +83,14 @@ lemma IsLocalizedModule.lift_rank_eq : apply hp (c * u i).prop exact hs t _ hc _ hit +end + lemma IsLocalizedModule.rank_eq {N : Type v} [AddCommGroup N] [Module R N] [Module S N] [IsScalarTower R S N] (f : M →ₗ[R] N) [IsLocalizedModule p f] : Module.rank S N = Module.rank R M := by simpa using IsLocalizedModule.lift_rank_eq S p f hp +end + variable (R M) in theorem exists_set_linearIndependent_of_isDomain [IsDomain R] : ∃ s : Set M, #s = Module.rank R M ∧ LinearIndependent (ι := s) R Subtype.val := by diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 535f8379922d6..1a73c224e14a7 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -718,6 +718,7 @@ theorem lc_def (e : ι → M) (l : ι →₀ R) : lc e l = Finsupp.total _ _ R e open Module variable [DecidableEq ι] (h : DualBases e ε) +include h theorem dual_lc (l : ι →₀ R) (i : ι) : ε i (DualBases.lc e l) = l i := by rw [lc, _root_.map_finsupp_sum, Finsupp.sum_eq_single i (g := fun a b ↦ (ε i) (b • e a))] @@ -1308,9 +1309,11 @@ variable [AddCommGroup V₁] [Module K V₁] [AddCommGroup V₂] [Module K V₂] namespace Module.Dual -variable [FiniteDimensional K V₁] {f : Module.Dual K V₁} (hf : f ≠ 0) +variable {f : Module.Dual K V₁} -open FiniteDimensional +section +variable (hf : f ≠ 0) +include hf lemma range_eq_top_of_ne_zero : LinearMap.range f = ⊤ := by @@ -1318,6 +1321,9 @@ lemma range_eq_top_of_ne_zero : rw [eq_top_iff] exact fun x _ ↦ ⟨x • (f v)⁻¹ • v, by simp [inv_mul_cancel₀ hv]⟩ +open FiniteDimensional +variable [FiniteDimensional K V₁] + lemma finrank_ker_add_one_of_ne_zero : finrank K (LinearMap.ker f) + 1 = finrank K V₁ := by suffices finrank K (LinearMap.range f) = 1 by @@ -1334,7 +1340,9 @@ lemma isCompl_ker_of_disjoint_of_ne_bot {p : Submodule K V₁} rwa [finrank_top, this, ← finrank_ker_add_one_of_ne_zero hf, add_le_add_iff_left, Submodule.one_le_finrank_iff] -lemma eq_of_ker_eq_of_apply_eq {f g : Module.Dual K V₁} (x : V₁) +end + +lemma eq_of_ker_eq_of_apply_eq [FiniteDimensional K V₁] {f g : Module.Dual K V₁} (x : V₁) (h : LinearMap.ker f = LinearMap.ker g) (h' : f x = g x) (hx : f x ≠ 0) : f = g := by let p := K ∙ x diff --git a/Mathlib/LinearAlgebra/LinearDisjoint.lean b/Mathlib/LinearAlgebra/LinearDisjoint.lean index 61752e40f40d3..725d409379f59 100644 --- a/Mathlib/LinearAlgebra/LinearDisjoint.lean +++ b/Mathlib/LinearAlgebra/LinearDisjoint.lean @@ -457,7 +457,9 @@ section not_linearIndependent_pair variable {M N} +section variable (H : M.LinearDisjoint N) +include H section @@ -537,6 +539,8 @@ theorem rank_inf_le_one_of_commute_of_flat_right [Module.Flat R N] (hc : ∀ (m n : ↥(M ⊓ N)), Commute m.1 n.1) : Module.rank R ↥(M ⊓ N) ≤ 1 := H.rank_inf_le_one_of_commute_of_flat (Or.inr ‹_›) hc +end + /-- If `M` and itself are linearly disjoint, if `M` is flat, if any two elements of `M` are commutative, then the rank of `M` is at most one. -/ theorem rank_le_one_of_commute_of_flat_of_self (H : M.LinearDisjoint M) [Module.Flat R M] @@ -562,7 +566,9 @@ section not_linearIndependent_pair variable {M N} +section variable (H : M.LinearDisjoint N) +include H section @@ -604,6 +610,8 @@ for commutative rings. -/ theorem rank_inf_le_one_of_flat_right [Module.Flat R N] : Module.rank R ↥(M ⊓ N) ≤ 1 := H.rank_inf_le_one_of_commute_of_flat_right fun _ _ ↦ mul_comm _ _ +end + /-- The `Submodule.LinearDisjoint.rank_le_one_of_commute_of_flat_of_self` for commutative rings. -/ theorem rank_le_one_of_flat_of_self (H : M.LinearDisjoint M) [Module.Flat R M] : diff --git a/Mathlib/LinearAlgebra/LinearPMap.lean b/Mathlib/LinearAlgebra/LinearPMap.lean index 1438fc123e93f..a0c9f4fde5fa5 100644 --- a/Mathlib/LinearAlgebra/LinearPMap.lean +++ b/Mathlib/LinearAlgebra/LinearPMap.lean @@ -985,6 +985,7 @@ theorem inverse_domain : (inverse f).domain = LinearMap.range f.toFun := by rfl variable (hf : LinearMap.ker f.toFun = ⊥) +include hf /-- The graph of the inverse generates a `LinearPMap`. -/ theorem mem_inverse_graph_snd_eq_zero (x : F × E) diff --git a/Mathlib/LinearAlgebra/Matrix/Block.lean b/Mathlib/LinearAlgebra/Matrix/Block.lean index bb41c66dd5941..c998bad25a879 100644 --- a/Mathlib/LinearAlgebra/Matrix/Block.lean +++ b/Mathlib/LinearAlgebra/Matrix/Block.lean @@ -198,7 +198,6 @@ theorem twoBlockTriangular_det' (M : Matrix m m R) (p : m → Prop) [DecidablePr protected theorem BlockTriangular.det [DecidableEq α] [LinearOrder α] (hM : BlockTriangular M b) : M.det = ∏ a ∈ univ.image b, (M.toSquareBlock b a).det := by - clear N induction' hs : univ.image b using Finset.strongInduction with s ih generalizing m subst hs cases isEmpty_or_nonempty m @@ -317,7 +316,6 @@ theorem toBlock_inverse_eq_zero [LinearOrder α] [Invertible M] (hM : BlockTrian /-- The inverse of a block-triangular matrix is block-triangular. -/ theorem blockTriangular_inv_of_blockTriangular [LinearOrder α] [Invertible M] (hM : BlockTriangular M b) : BlockTriangular M⁻¹ b := by - clear N induction' hs : univ.image b using Finset.strongInduction with s ih generalizing m subst hs intro i j hij diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean index 66566985607f4..45a4644bc6fc5 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/LinearMap.lean @@ -23,7 +23,7 @@ This is used to conclude the Cayley-Hamilton theorem for f.g. modules over arbit variable {ι : Type*} [Fintype ι] variable {M : Type*} [AddCommGroup M] (R : Type*) [CommRing R] [Module R M] (I : Ideal R) -variable (b : ι → M) (hb : Submodule.span R (Set.range b) = ⊤) +variable (b : ι → M) open Polynomial Matrix @@ -130,7 +130,8 @@ theorem Matrix.Represents.algebraMap (r : R) : (algebraMap _ (Matrix ι ι R) r).Represents b (algebraMap _ (Module.End R M) r) := by simpa only [Algebra.algebraMap_eq_smul_one] using Matrix.Represents.one.smul r -theorem Matrix.Represents.eq {A : Matrix ι ι R} {f f' : Module.End R M} (h : A.Represents b f) +theorem Matrix.Represents.eq (hb : Submodule.span R (Set.range b) = ⊤) + {A : Matrix ι ι R} {f f' : Module.End R M} (h : A.Represents b f) (h' : A.Represents b f') : f = f' := PiToModule.fromEnd_injective R b hb (h.symm.trans h') @@ -146,6 +147,9 @@ def Matrix.isRepresentation : Subalgebra R (Matrix ι ι R) where zero_mem' := ⟨0, Matrix.Represents.zero⟩ algebraMap_mem' r := ⟨algebraMap _ _ r, .algebraMap _⟩ +variable (hb : Submodule.span R (Set.range b) = ⊤) +include hb + /-- The map sending a matrix to the endomorphism it represents. This is an `R`-algebra morphism. -/ noncomputable def Matrix.isRepresentation.toEnd : Matrix.isRepresentation R b →ₐ[R] Module.End R M where diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index 2d222bd312cb6..fe70520844931 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -32,7 +32,7 @@ namespace Matrix variable {m n R 𝕜 : Type*} variable [Fintype m] [Fintype n] -variable [CommRing R] [PartialOrder R] [StarRing R] [StarOrderedRing R] +variable [CommRing R] [PartialOrder R] [StarRing R] variable [RCLike 𝕜] open scoped Matrix @@ -46,7 +46,7 @@ def PosSemidef (M : Matrix n n R) := M.IsHermitian ∧ ∀ x : n → R, 0 ≤ dotProduct (star x) (M *ᵥ x) /-- A diagonal matrix is positive semidefinite iff its diagonal entries are nonnegative. -/ -lemma posSemidef_diagonal_iff [DecidableEq n] {d : n → R} : +lemma posSemidef_diagonal_iff [StarOrderedRing R] [DecidableEq n] {d : n → R} : PosSemidef (diagonal d) ↔ (∀ i : n, 0 ≤ d i) := by refine ⟨fun ⟨_, hP⟩ i ↦ by simpa using hP (Pi.single i 1), ?_⟩ refine fun hd ↦ ⟨isHermitian_diagonal_iff.2 fun i ↦ IsSelfAdjoint.of_nonneg (hd i), ?_⟩ @@ -93,11 +93,12 @@ theorem conjTranspose {M : Matrix n n R} (hM : M.PosSemidef) : Mᴴ.PosSemidef : protected lemma zero : PosSemidef (0 : Matrix n n R) := ⟨isHermitian_zero, by simp⟩ -protected lemma one [DecidableEq n] : PosSemidef (1 : Matrix n n R) := +protected lemma one [StarOrderedRing R] [DecidableEq n] : PosSemidef (1 : Matrix n n R) := ⟨isHermitian_one, fun x => by rw [one_mulVec]; exact Fintype.sum_nonneg fun i => star_mul_self_nonneg _⟩ -protected lemma pow [DecidableEq n] {M : Matrix n n R} (hM : M.PosSemidef) (k : ℕ) : +protected lemma pow [StarOrderedRing R] [DecidableEq n] + {M : Matrix n n R} (hM : M.PosSemidef) (k : ℕ) : PosSemidef (M ^ k) := match k with | 0 => .one @@ -113,13 +114,14 @@ protected lemma inv [DecidableEq n] {M : Matrix n n R} (hM : M.PosSemidef) : M · rw [nonsing_inv_apply_not_isUnit _ h] exact .zero -protected lemma zpow [DecidableEq n] {M : Matrix n n R} (hM : M.PosSemidef) (z : ℤ) : +protected lemma zpow [StarOrderedRing R] [DecidableEq n] + {M : Matrix n n R} (hM : M.PosSemidef) (z : ℤ) : (M ^ z).PosSemidef := by obtain ⟨n, rfl | rfl⟩ := z.eq_nat_or_neg · simpa using hM.pow n · simpa using (hM.pow n).inv -protected lemma add {A : Matrix m m R} {B : Matrix m m R} +protected lemma add [CovariantClass R R (· + ·) (· ≤ · )] {A : Matrix m m R} {B : Matrix m m R} (hA : A.PosSemidef) (hB : B.PosSemidef) : (A + B).PosSemidef := ⟨hA.isHermitian.add hB.isHermitian, fun x => by rw [add_mulVec, dotProduct_add] @@ -178,6 +180,7 @@ lemma sq_sqrt : hA.sqrt ^ 2 = A := by @[simp] lemma sqrt_mul_self : hA.sqrt * hA.sqrt = A := by rw [← pow_two, sq_sqrt] +include hA in lemma eq_of_sq_eq_sq {B : Matrix n n 𝕜} (hB : PosSemidef B) (hAB : A ^ 2 = B ^ 2) : A = B := by /- This is deceptively hard, much more difficult than the positive *definite* case. We follow a clever proof due to Koeber and Schäfer. The idea is that if `A ≠ B`, then `A - B` has a nonzero @@ -217,6 +220,7 @@ lemma eq_of_sq_eq_sq {B : Matrix n n 𝕜} (hB : PosSemidef B) (hAB : A ^ 2 = B lemma sqrt_sq : (hA.pow 2 : PosSemidef (A ^ 2)).sqrt = A := (hA.pow 2).posSemidef_sqrt.eq_of_sq_eq_sq hA (hA.pow 2).sq_sqrt +include hA in lemma eq_sqrt_of_sq_eq {B : Matrix n n 𝕜} (hB : PosSemidef B) (hAB : A ^ 2 = B) : A = hB.sqrt := by subst B rw [hA.sqrt_sq] @@ -231,13 +235,15 @@ theorem posSemidef_submatrix_equiv {M : Matrix n n R} (e : m ≃ n) : ⟨fun h => by simpa using h.submatrix e.symm, fun h => h.submatrix _⟩ /-- The conjugate transpose of a matrix mulitplied by the matrix is positive semidefinite -/ -theorem posSemidef_conjTranspose_mul_self (A : Matrix m n R) : PosSemidef (Aᴴ * A) := by +theorem posSemidef_conjTranspose_mul_self [StarOrderedRing R] (A : Matrix m n R) : + PosSemidef (Aᴴ * A) := by refine ⟨isHermitian_transpose_mul_self _, fun x => ?_⟩ rw [← mulVec_mulVec, dotProduct_mulVec, vecMul_conjTranspose, star_star] exact Finset.sum_nonneg fun i _ => star_mul_self_nonneg _ /-- A matrix multiplied by its conjugate transpose is positive semidefinite -/ -theorem posSemidef_self_mul_conjTranspose (A : Matrix m n R) : PosSemidef (A * Aᴴ) := by +theorem posSemidef_self_mul_conjTranspose [StarOrderedRing R] (A : Matrix m n R) : + PosSemidef (A * Aᴴ) := by simpa only [conjTranspose_conjTranspose] using posSemidef_conjTranspose_mul_self Aᴴ lemma eigenvalues_conjTranspose_mul_self_nonneg (A : Matrix m n 𝕜) [DecidableEq n] (i : n) : @@ -312,19 +318,21 @@ theorem transpose {M : Matrix n n R} (hM : M.PosDef) : Mᵀ.PosDef := by convert hM.2 (star x) (star_ne_zero.2 hx) using 1 rw [mulVec_transpose, Matrix.dotProduct_mulVec, star_star, dotProduct_comm] -protected lemma add_posSemidef {A : Matrix m m R} {B : Matrix m m R} +protected lemma add_posSemidef [CovariantClass R R (· + ·) (· ≤ · )] + {A : Matrix m m R} {B : Matrix m m R} (hA : A.PosDef) (hB : B.PosSemidef) : (A + B).PosDef := ⟨hA.isHermitian.add hB.isHermitian, fun x hx => by rw [add_mulVec, dotProduct_add] exact add_pos_of_pos_of_nonneg (hA.2 x hx) (hB.2 x)⟩ -protected lemma posSemidef_add {A : Matrix m m R} {B : Matrix m m R} +protected lemma posSemidef_add [CovariantClass R R (· + ·) (· ≤ · )] + {A : Matrix m m R} {B : Matrix m m R} (hA : A.PosSemidef) (hB : B.PosDef) : (A + B).PosDef := ⟨hA.isHermitian.add hB.isHermitian, fun x hx => by rw [add_mulVec, dotProduct_add] exact add_pos_of_nonneg_of_pos (hA.2 x) (hB.2 x hx)⟩ -protected lemma add {A : Matrix m m R} {B : Matrix m m R} +protected lemma add [CovariantClass R R (· + ·) (· ≤ · )] {A : Matrix m m R} {B : Matrix m m R} (hA : A.PosDef) (hB : B.PosDef) : (A + B).PosDef := hA.add_posSemidef hB.posSemidef diff --git a/Mathlib/LinearAlgebra/PerfectPairing.lean b/Mathlib/LinearAlgebra/PerfectPairing.lean index 207133d7af90e..74a47d2060113 100644 --- a/Mathlib/LinearAlgebra/PerfectPairing.lean +++ b/Mathlib/LinearAlgebra/PerfectPairing.lean @@ -109,11 +109,13 @@ theorem bijective_toDualRight_symm_toDualLeft : Bijective.comp (LinearEquiv.bijective p.toDualRight.symm.dualMap) (LinearEquiv.bijective p.toDualLeft) +include p in theorem reflexive_left : IsReflexive R M where bijective_dual_eval' := by rw [← p.toDualRight_symm_comp_toDualLeft] exact p.bijective_toDualRight_symm_toDualLeft +include p in theorem reflexive_right : IsReflexive R N := p.flip.reflexive_left @@ -146,6 +148,7 @@ lemma symm_flip : e.flip.symm = e.symm.dualMap.trans (evalEquiv R M).symm := rfl lemma trans_dualMap_symm_flip : e.trans e.flip.symm.dualMap = Dual.eval R N := by ext; simp [symm_flip] +include e in /-- If `N` is in perfect pairing with `M`, then it is reflexive. -/ lemma isReflexive_of_equiv_dual_of_isReflexive : IsReflexive R N := by constructor diff --git a/Mathlib/LinearAlgebra/Pi.lean b/Mathlib/LinearAlgebra/Pi.lean index a6c4da9560237..002a80ea59395 100644 --- a/Mathlib/LinearAlgebra/Pi.lean +++ b/Mathlib/LinearAlgebra/Pi.lean @@ -212,7 +212,7 @@ def lsum (S) [AddCommMonoid M] [Module R M] [Fintype ι] [DecidableEq ι] [Semir rw [Finset.univ_sum_single] @[simp] -theorem lsum_apply (S) [AddCommMonoid M] [Module R M] [Fintype ι] [DecidableEq ι] [Semiring S] +theorem lsum_apply (S) [AddCommMonoid M] [Module R M] [Fintype ι] [Semiring S] [Module S M] [SMulCommClass R S M] (f : (i : ι) → φ i →ₗ[R] M) : lsum R φ S f = ∑ i : ι, (f i).comp (proj i) := rfl @@ -226,7 +226,7 @@ variable {R φ} section Ext -variable [Finite ι] [DecidableEq ι] [AddCommMonoid M] [Module R M] {f g : ((i : ι) → φ i) →ₗ[R] M} +variable [Finite ι] [AddCommMonoid M] [Module R M] {f g : ((i : ι) → φ i) →ₗ[R] M} theorem pi_ext (h : ∀ i x, f (Pi.single i x) = g (Pi.single i x)) : f = g := toAddMonoidHom_injective <| AddMonoidHom.functions_ext _ _ _ h @@ -281,8 +281,6 @@ end section -variable [DecidableEq ι] - /-- `diag i j` is the identity map if `i = j`. Otherwise it is the constant 0 map. -/ def diag (i j : ι) : φ i →ₗ[R] φ j := @Function.update ι (fun j => φ i →ₗ[R] φ j) _ 0 i id j @@ -311,7 +309,7 @@ end /-- A linear map `f` applied to `x : ι → R` can be computed using the image under `f` of elements of the canonical basis. -/ -theorem pi_apply_eq_sum_univ [Fintype ι] [DecidableEq ι] (f : (ι → R) →ₗ[R] M₂) (x : ι → R) : +theorem pi_apply_eq_sum_univ [Fintype ι] (f : (ι → R) →ₗ[R] M₂) (x : ι → R) : f x = ∑ i, x i • f fun j => if i = j then 1 else 0 := by conv_lhs => rw [pi_eq_sum_univ x, map_sum] refine Finset.sum_congr rfl (fun _ _ => ?_) diff --git a/Mathlib/LinearAlgebra/RootSystem/Basic.lean b/Mathlib/LinearAlgebra/RootSystem/Basic.lean index 9c4a0351f4b44..4b4ba33c9fa5a 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Basic.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Basic.lean @@ -45,12 +45,14 @@ section reflection_perm variable (p : PerfectPairing R M N) (root : ι ↪ M) (coroot : ι ↪ N) (i j : ι) (h : ∀ i, MapsTo (preReflection (root i) (p.toLin.flip (coroot i))) (range root) (range root)) +include h private theorem exist_eq_reflection_of_mapsTo : ∃ k, root k = (preReflection (root i) (p.toLin.flip (coroot i))) (root j) := h i (mem_range_self j) variable (hp : ∀ i, p.toLin (root i) (coroot i) = 2) +include hp private theorem choose_choose_eq_of_mapsTo : (exist_eq_reflection_of_mapsTo p root coroot i diff --git a/Mathlib/LinearAlgebra/RootSystem/RootPositive.lean b/Mathlib/LinearAlgebra/RootSystem/RootPositive.lean index 351254d0fea1c..ccf50330e9a81 100644 --- a/Mathlib/LinearAlgebra/RootSystem/RootPositive.lean +++ b/Mathlib/LinearAlgebra/RootSystem/RootPositive.lean @@ -49,6 +49,7 @@ class IsRootPositive (P : RootPairing ι R M N) (B : M →ₗ[R] M →ₗ[R] R) apply_reflection_eq : ∀ i x y, B (P.reflection i x) (P.reflection i y) = B x y variable {P : RootPairing ι R M N} (B : M →ₗ[R] M →ₗ[R] R) [IsRootPositive P B] (i j : ι) +include B lemma two_mul_apply_root_root : 2 * B (P.root i) (P.root j) = P.pairing i j * B (P.root j) (P.root j) := by diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index 32882cdc8f657..ea02d0bd4374e 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -136,6 +136,7 @@ variable [FiniteDimensional K M] section variable (hf : f.IsSemisimple) +include hf /-- The minimal polynomial of a semisimple endomorphism is square free -/ theorem IsSemisimple.minpoly_squarefree : Squarefree (minpoly K f) := @@ -166,6 +167,7 @@ end section PerfectField variable [PerfectField K] (comm : Commute f g) (hf : f.IsSemisimple) (hg : g.IsSemisimple) +include comm hf hg attribute [local simp] Submodule.Quotient.quot_mk_eq_mk in theorem IsSemisimple.of_mem_adjoin_pair {a : End K M} (ha : a ∈ Algebra.adjoin K {f, g}) : diff --git a/Mathlib/LinearAlgebra/SesquilinearForm.lean b/Mathlib/LinearAlgebra/SesquilinearForm.lean index 0226a9d6fc42b..8a2b3862b074a 100644 --- a/Mathlib/LinearAlgebra/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/SesquilinearForm.lean @@ -149,18 +149,20 @@ def IsRefl (B : M₁ →ₛₗ[I₁] M₁ →ₛₗ[I₂] M) : Prop := namespace IsRefl +section variable (H : B.IsRefl) +include H theorem eq_zero : ∀ {x y}, B x y = 0 → B y x = 0 := fun {x y} ↦ H x y theorem ortho_comm {x y} : IsOrtho B x y ↔ IsOrtho B y x := ⟨eq_zero H, eq_zero H⟩ -theorem domRestrict (H : B.IsRefl) (p : Submodule R₁ M₁) : (B.domRestrict₁₂ p p).IsRefl := +theorem domRestrict (p : Submodule R₁ M₁) : (B.domRestrict₁₂ p p).IsRefl := fun _ _ ↦ by simp_rw [domRestrict₁₂_apply] exact H _ _ - +end @[simp] theorem flip_isRefl_iff : B.flip.IsRefl ↔ B.IsRefl := ⟨fun h x y H ↦ h y x ((B.flip_apply _ _).trans H), fun h x y ↦ h y x⟩ @@ -239,6 +241,7 @@ def IsAlt (B : M₁ →ₛₗ[I₁] M₁ →ₛₗ[I₂] M) : Prop := ∀ x, B x x = 0 variable (H : B.IsAlt) +include H theorem IsAlt.self_eq_zero (x : M₁) : B x x = 0 := H x diff --git a/Mathlib/LinearAlgebra/Span.lean b/Mathlib/LinearAlgebra/Span.lean index 27a0a2d497a0c..c560119afc158 100644 --- a/Mathlib/LinearAlgebra/Span.lean +++ b/Mathlib/LinearAlgebra/Span.lean @@ -107,11 +107,13 @@ theorem span_coe_eq_restrictScalars [Semiring S] [SMul S R] [Module S M] [IsScal span S (p : Set M) = p.restrictScalars S := span_eq (p.restrictScalars S) +include σ₁₂ in /-- A version of `Submodule.map_span_le` that does not require the `RingHomSurjective` assumption. -/ theorem image_span_subset (f : F) (s : Set M) (N : Submodule R₂ M₂) : f '' span R s ⊆ N ↔ ∀ m ∈ s, f m ∈ N := image_subset_iff.trans <| span_le (p := N.comap f) +include σ₁₂ in theorem image_span_subset_span (f : F) (s : Set M) : f '' span R s ⊆ span R₂ (f '' s) := (image_span_subset f s _).2 fun x hx ↦ subset_span ⟨x, hx, rfl⟩ @@ -968,6 +970,7 @@ section AddCommMonoid variable [Semiring R] [AddCommMonoid M] [Module R M] variable [Semiring R₂] [AddCommMonoid M₂] [Module R₂ M₂] variable {F : Type*} {σ₁₂ : R →+* R₂} [FunLike F M M₂] [SemilinearMapClass F σ₁₂ M M₂] +include σ₁₂ /-- Two linear maps are equal on `Submodule.span s` iff they are equal on `s`. -/ theorem eqOn_span_iff {s : Set M} {f g : F} : Set.EqOn f g (span R s) ↔ Set.EqOn f g s := by diff --git a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean index 8abf71e5347e1..3cdcdf07810e4 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean @@ -259,6 +259,7 @@ noncomputable def lTensor.equiv : ((Q ⊗[R] N) ⧸ (LinearMap.range (lTensor Q f))) ≃ₗ[R] (Q ⊗[R] P) := lTensor.linearEquiv_of_rightInverse Q hfg (Function.rightInverse_surjInv hg) +include hfg hg in /-- Tensoring an exact pair on the left gives an exact pair -/ theorem lTensor_exact : Exact (lTensor Q f) (lTensor Q g) := by rw [exact_iff, ← Submodule.ker_mkQ (p := range (lTensor Q f)), @@ -364,6 +365,7 @@ noncomputable def rTensor.equiv : ((N ⊗[R] Q) ⧸ (LinearMap.range (rTensor Q f))) ≃ₗ[R] (P ⊗[R] Q) := rTensor.linearEquiv_of_rightInverse Q hfg (Function.rightInverse_surjInv hg) +include hfg hg in /-- Tensoring an exact pair on the right gives an exact pair -/ theorem rTensor_exact : Exact (rTensor Q f) (rTensor Q g) := by rw [exact_iff, ← Submodule.ker_mkQ (p := range (rTensor Q f)), @@ -385,10 +387,12 @@ variable {M' N' P' : Type*} {f' : M' →ₗ[R] N'} {g' : N' →ₗ[R] P'} (hfg' : Exact f' g') (hg' : Function.Surjective g') +include hg hg' in theorem TensorProduct.map_surjective : Function.Surjective (TensorProduct.map g g') := by rw [← lTensor_comp_rTensor, coe_comp] exact Function.Surjective.comp (lTensor_surjective _ hg') (rTensor_surjective _ hg) +include hg hg' hfg hfg' in /-- Kernel of a product map (right-exactness of tensor product) -/ theorem TensorProduct.map_ker : ker (TensorProduct.map g g') = range (lTensor N f') ⊔ range (rTensor N' f) := by diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 3c7429f258290..88fe378a09004 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -845,7 +845,7 @@ end BoundedQuantifiers section ite -variable {α : Sort*} {σ : α → Sort*} {P Q R : Prop} [Decidable P] [Decidable Q] +variable {α : Sort*} {σ : α → Sort*} {P Q R : Prop} [Decidable P] {a b c : α} {A : P → α} {B : ¬P → α} theorem dite_eq_iff : dite P A B = c ↔ (∃ h, A h = c) ∨ ∃ h, B h = c := by @@ -956,7 +956,7 @@ variable {P Q} theorem ite_prop_iff_or : (if P then Q else R) ↔ (P ∧ Q ∨ ¬ P ∧ R) := by by_cases p : P <;> simp [p] -theorem dite_prop_iff_or {Q : P → Prop} {R : ¬P → Prop} [Decidable P] : +theorem dite_prop_iff_or {Q : P → Prop} {R : ¬P → Prop} : dite P Q R ↔ (∃ p, Q p) ∨ (∃ p, R p) := by by_cases h : P <;> simp [h, exists_prop_of_false, exists_prop_of_true] @@ -964,7 +964,7 @@ theorem dite_prop_iff_or {Q : P → Prop} {R : ¬P → Prop} [Decidable P] : theorem ite_prop_iff_and : (if P then Q else R) ↔ ((P → Q) ∧ (¬ P → R)) := by by_cases p : P <;> simp [p] -theorem dite_prop_iff_and {Q : P → Prop} {R : ¬P → Prop} [Decidable P] : +theorem dite_prop_iff_and {Q : P → Prop} {R : ¬P → Prop} : dite P Q R ↔ (∀ h, Q h) ∧ (∀ h, R h) := by by_cases h : P <;> simp [h, forall_prop_of_false, forall_prop_of_true] diff --git a/Mathlib/Logic/Equiv/TransferInstance.lean b/Mathlib/Logic/Equiv/TransferInstance.lean index 8f6b53e102e82..4c39c97683377 100644 --- a/Mathlib/Logic/Equiv/TransferInstance.lean +++ b/Mathlib/Logic/Equiv/TransferInstance.lean @@ -462,6 +462,7 @@ protected abbrev commRing [CommRing β] : CommRing α := by noncomputable instance [Small.{v} α] [CommRing α] : CommRing (Shrink.{v} α) := (equivShrink α).symm.commRing +include e in /-- Transfer `Nontrivial` across an `Equiv` -/ protected theorem nontrivial [Nontrivial β] : Nontrivial α := e.surjective.nontrivial diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index abebf7a7f96df..23b06d62b1a9b 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -755,6 +755,8 @@ namespace Involutive variable {α : Sort u} {f : α → α} (h : Involutive f) +include h + @[simp] theorem comp_self : f ∘ f = id := funext h diff --git a/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean b/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean index d45ac29c66449..97bd8507cbce3 100644 --- a/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean +++ b/Mathlib/MeasureTheory/Constructions/HaarToSphere.lean @@ -32,8 +32,8 @@ local notation "dim" => FiniteDimensional.finrank ℝ noncomputable section namespace MeasureTheory -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] - [MeasurableSpace E] [BorelSpace E] +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] + [MeasurableSpace E] namespace Measure @@ -52,6 +52,8 @@ theorem toSphere_apply_aux (s : Set (sphere (0 : E) 1)) (r : Ioi (0 : ℝ)) : ← image_subtype_val_Ioi_Iio, image2_image_left, image2_swap, ← image_prod] rfl +variable [BorelSpace E] + theorem toSphere_apply' {s : Set (sphere (0 : E) 1)} (hs : MeasurableSet s) : μ.toSphere s = dim E * μ (Ioo (0 : ℝ) 1 • ((↑) '' s)) := by rw [toSphere, smul_apply, fst_apply hs, restrict_apply (measurable_fst hs), @@ -62,7 +64,7 @@ theorem toSphere_apply' {s : Set (sphere (0 : E) 1)} (hs : MeasurableSet s) : theorem toSphere_apply_univ' : μ.toSphere univ = dim E * μ (ball 0 1 \ {0}) := by rw [μ.toSphere_apply' .univ, image_univ, Subtype.range_coe, Ioo_smul_sphere_zero] <;> simp -variable [μ.IsAddHaarMeasure] +variable [FiniteDimensional ℝ E] [μ.IsAddHaarMeasure] @[simp] theorem toSphere_apply_univ : μ.toSphere univ = dim E * μ (ball 0 1) := by @@ -127,7 +129,7 @@ theorem measurePreserving_homeomorphUnitSphereProd : end Measure variable {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] - [Nontrivial E] (μ : Measure E) [μ.IsAddHaarMeasure] + [Nontrivial E] (μ : Measure E) [FiniteDimensional ℝ E] [BorelSpace E] [μ.IsAddHaarMeasure] lemma integral_fun_norm_addHaar (f : ℝ → F) : ∫ x, f (‖x‖) ∂μ = dim E • (μ (ball 0 1)).toReal • ∫ y in Ioi (0 : ℝ), y ^ (dim E - 1) • f y := diff --git a/Mathlib/MeasureTheory/Covering/Differentiation.lean b/Mathlib/MeasureTheory/Covering/Differentiation.lean index 9889a9e2e7aa6..cb24a16d2a1c8 100644 --- a/Mathlib/MeasureTheory/Covering/Differentiation.lean +++ b/Mathlib/MeasureTheory/Covering/Differentiation.lean @@ -203,6 +203,7 @@ theorem ae_eventually_measure_zero_of_singular (hρ : ρ ⟂ₘ μ) : section AbsolutelyContinuous variable (hρ : ρ ≪ μ) +include hρ /-- A set of points `s` satisfying both `ρ a ≤ c * μ a` and `ρ a ≥ d * μ a` at arbitrarily small sets in a Vitali family has measure `0` if `c < d`. Indeed, the first inequality should imply diff --git a/Mathlib/MeasureTheory/Covering/VitaliFamily.lean b/Mathlib/MeasureTheory/Covering/VitaliFamily.lean index e4c3a72dd39a7..a584639b9f6e9 100644 --- a/Mathlib/MeasureTheory/Covering/VitaliFamily.lean +++ b/Mathlib/MeasureTheory/Covering/VitaliFamily.lean @@ -103,6 +103,7 @@ def FineSubfamilyOn (v : VitaliFamily μ) (f : X → Set (Set X)) (s : Set X) : namespace FineSubfamilyOn variable {v : VitaliFamily μ} {f : X → Set (Set X)} {s : Set X} (h : v.FineSubfamilyOn f s) +include h theorem exists_disjoint_covering_ae : ∃ t : Set (X × Set X), diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean index 8efce39ac58ef..ef56edebde96d 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Unique.lean @@ -69,6 +69,7 @@ alias lpMeas.ae_eq_zero_of_forall_set_integral_eq_zero := variable (𝕜) +include 𝕜 in theorem Lp.ae_eq_zero_of_forall_setIntegral_eq_zero' (hm : m ≤ m0) (f : Lp E' p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf_int_finite : ∀ s, MeasurableSet[m] s → μ s < ∞ → IntegrableOn f s μ) @@ -92,6 +93,7 @@ theorem Lp.ae_eq_zero_of_forall_setIntegral_eq_zero' (hm : m ≤ m0) (f : Lp E' alias Lp.ae_eq_zero_of_forall_set_integral_eq_zero' := Lp.ae_eq_zero_of_forall_setIntegral_eq_zero' +include 𝕜 in /-- **Uniqueness of the conditional expectation** -/ theorem Lp.ae_eq_of_forall_setIntegral_eq' (hm : m ≤ m0) (f g : Lp E' p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) (hf_int_finite : ∀ s, MeasurableSet[m] s → μ s < ∞ → IntegrableOn f s μ) diff --git a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean index 43e55d7c27774..b866f1253eceb 100644 --- a/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean +++ b/Mathlib/MeasureTheory/Function/SpecialFunctions/Basic.lean @@ -133,6 +133,7 @@ section RealComposition open Real variable {α : Type*} {m : MeasurableSpace α} {f : α → ℝ} (hf : Measurable f) +include hf @[measurability] theorem Measurable.exp : Measurable fun x => Real.exp (f x) := @@ -169,6 +170,7 @@ section ComplexComposition open Complex variable {α : Type*} {m : MeasurableSpace α} {f : α → ℂ} (hf : Measurable f) +include hf @[measurability] theorem Measurable.cexp : Measurable fun x => Complex.exp (f x) := diff --git a/Mathlib/MeasureTheory/Group/Action.lean b/Mathlib/MeasureTheory/Group/Action.lean index 3238971b09711..159a5bff0f4ff 100644 --- a/Mathlib/MeasureTheory/Group/Action.lean +++ b/Mathlib/MeasureTheory/Group/Action.lean @@ -57,7 +57,7 @@ end SMulInvariantMeasure section AE_smul -variable {m : MeasurableSpace α} [MeasurableSpace G] [SMul G α] +variable {m : MeasurableSpace α} [SMul G α] (μ : Measure α) [SMulInvariantMeasure G α μ] {s : Set α} /-- See also `measure_preimage_smul_of_nullMeasurableSet` and `measure_preimage_smul`. -/ @@ -88,7 +88,7 @@ end AE_smul section AE -variable {m : MeasurableSpace α} [MeasurableSpace G] [Group G] [MulAction G α] +variable {m : MeasurableSpace α} [Group G] [MulAction G α] (μ : Measure α) [SMulInvariantMeasure G α μ] @[to_additive (attr := simp)] @@ -198,9 +198,9 @@ instance smulInvariantMeasure_map_smul [SMul M α] [SMul N α] [SMulCommClass N end SMulHomClass -variable (G) {m : MeasurableSpace α} [Group G] [MulAction G α] [MeasurableSpace G] - [MeasurableSMul G α] (c : G) (μ : Measure α) +variable (G) {m : MeasurableSpace α} [Group G] [MulAction G α] (c : G) (μ : Measure α) +variable [MeasurableSpace G] [MeasurableSMul G α] in /-- Equivalent definitions of a measure invariant under a multiplicative action of a group. - 0: `SMulInvariantMeasure G α μ`; @@ -268,7 +268,7 @@ variable {G} variable [SMulInvariantMeasure G α μ] variable {μ} - +variable [MeasurableSpace G] [MeasurableSMul G α] in @[to_additive] theorem NullMeasurableSet.smul {s} (hs : NullMeasurableSet s μ) (c : G) : NullMeasurableSet (c • s) μ := by @@ -280,6 +280,7 @@ section IsMinimal variable (G) variable [TopologicalSpace α] [ContinuousConstSMul G α] [MulAction.IsMinimal G α] {K U : Set α} +include G in /-- If measure `μ` is invariant under a group action and is nonzero on a compact set `K`, then it is positive on any nonempty open set. In case of a regular measure, one can assume `μ ≠ 0` instead of `μ K ≠ 0`, see `MeasureTheory.measure_isOpen_pos_of_smulInvariant_of_ne_zero`. -/ @@ -297,6 +298,8 @@ then it is positive on any nonempty open set. In case of a regular measure, one instead of `μ K ≠ 0`, see `MeasureTheory.measure_isOpen_pos_of_vaddInvariant_of_ne_zero`. -/ add_decl_doc measure_isOpen_pos_of_vaddInvariant_of_compact_ne_zero +include G + @[to_additive] theorem isLocallyFiniteMeasure_of_smulInvariant (hU : IsOpen U) (hne : U.Nonempty) (hμU : μ U ≠ ∞) : IsLocallyFiniteMeasure μ := diff --git a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean index 181c922b7b868..5b5f6208fd138 100644 --- a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean +++ b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean @@ -364,7 +364,7 @@ protected theorem integrableOn_iff (hs : IsFundamentalDomain G s μ) (ht : IsFun {f : α → E} (hf : ∀ (g : G) (x), f (g • x) = f x) : IntegrableOn f s μ ↔ IntegrableOn f t μ := and_congr (hs.aEStronglyMeasurable_on_iff ht hf) (hs.hasFiniteIntegral_on_iff ht hf) -variable [NormedSpace ℝ E] [CompleteSpace E] +variable [NormedSpace ℝ E] @[to_additive] theorem integral_eq_tsum_of_ac (h : IsFundamentalDomain G s μ) (hν : ν ≪ μ) (f : α → E) @@ -587,10 +587,12 @@ end MeasurableSpace namespace IsFundamentalDomain -section Group - variable [Countable G] [Group G] [MulAction G α] [MeasurableSpace α] {μ : Measure α} {s : Set α} (hs : IsFundamentalDomain G s μ) +include hs + +section Group + @[to_additive MeasureTheory.IsAddFundamentalDomain.measure_addFundamentalFrontier] theorem measure_fundamentalFrontier : μ (fundamentalFrontier G s) = 0 := by @@ -603,9 +605,7 @@ theorem measure_fundamentalInterior : μ (fundamentalInterior G s) = μ s := end Group -variable [Countable G] [Group G] [MulAction G α] [MeasurableSpace α] {μ : Measure α} {s : Set α} - (hs : IsFundamentalDomain G s μ) [MeasurableSpace G] [MeasurableSMul G α] - [SMulInvariantMeasure G α μ] +variable [MeasurableSpace G] [MeasurableSMul G α] [SMulInvariantMeasure G α μ] protected theorem fundamentalInterior : IsFundamentalDomain G (fundamentalInterior G s) μ where nullMeasurableSet := hs.nullMeasurableSet.fundamentalInterior _ _ @@ -698,7 +698,7 @@ noncomputable def covolume (G α : Type*) [One G] [SMul G α] [MeasurableSpace (ν : Measure α := by volume_tac) : ℝ≥0∞ := if funDom : HasFundamentalDomain G α ν then ν funDom.ExistsIsFundamentalDomain.choose else 0 -variable [Group G] [MulAction G α] [MeasurableSpace G] [MeasurableSpace α] +variable [Group G] [MulAction G α] [MeasurableSpace α] /-- If there is a fundamental domain `s`, then `HasFundamentalDomain` holds. -/ @[to_additive] @@ -709,7 +709,7 @@ lemma IsFundamentalDomain.hasFundamentalDomain (ν : Measure α) {s : Set α} /-- The `covolume` can be computed by taking the `volume` of any given fundamental domain `s`. -/ @[to_additive] lemma IsFundamentalDomain.covolume_eq_volume (ν : Measure α) [Countable G] - [MeasurableSMul G α] [SMulInvariantMeasure G α ν] {s : Set α} + [MeasurableSpace G] [MeasurableSMul G α] [SMulInvariantMeasure G α ν] {s : Set α} (fund_dom_s : IsFundamentalDomain G s ν) : covolume G α ν = ν s := by dsimp [covolume] simp only [(fund_dom_s.hasFundamentalDomain ν), ↓reduceDIte] @@ -780,25 +780,7 @@ lemma IsFundamentalDomain.projection_respects_measure_apply {ν : Measure α} (meas_U : MeasurableSet U) : μ U = ν (π ⁻¹' U ∩ t) := by rw [fund_dom_t.projection_respects_measure (μ := μ), measure_map_restrict_apply ν t meas_U] -variable {ν : Measure α} [Countable G] [MeasurableSpace G] - [SMulInvariantMeasure G α ν] [MeasurableSMul G α] - -/-- Given a measure upstairs (i.e., on `α`), and a choice `s` of fundamental domain, there's always -an artificial way to generate a measure downstairs such that the pair satisfies the -`QuotientMeasureEqMeasurePreimage` typeclass. -/ -@[to_additive] -lemma IsFundamentalDomain.quotientMeasureEqMeasurePreimage_quotientMeasure - {s : Set α} (fund_dom_s : IsFundamentalDomain G s ν) : - QuotientMeasureEqMeasurePreimage ν ((ν.restrict s).map π) where - projection_respects_measure' t fund_dom_t := by rw [fund_dom_s.quotientMeasure_eq _ fund_dom_t] - -/-- One can prove `QuotientMeasureEqMeasurePreimage` by checking behavior with respect to a single -fundamental domain. -/ -@[to_additive] -lemma IsFundamentalDomain.quotientMeasureEqMeasurePreimage {μ : Measure (Quotient α_mod_G)} - {s : Set α} (fund_dom_s : IsFundamentalDomain G s ν) (h : μ = (ν.restrict s).map π) : - QuotientMeasureEqMeasurePreimage ν μ := by - simpa [h] using fund_dom_s.quotientMeasureEqMeasurePreimage_quotientMeasure +variable {ν : Measure α} /-- Any two measures satisfying `QuotientMeasureEqMeasurePreimage` are equal. -/ @[to_additive] @@ -821,6 +803,26 @@ theorem IsFundamentalDomain.measurePreserving_quotient_mk haveI : HasFundamentalDomain G α ν := ⟨𝓕, h𝓕⟩ rw [h𝓕.projection_respects_measure (μ := μ)] +variable [SMulInvariantMeasure G α ν] [Countable G] [MeasurableSpace G] [MeasurableSMul G α] + +/-- Given a measure upstairs (i.e., on `α`), and a choice `s` of fundamental domain, there's always +an artificial way to generate a measure downstairs such that the pair satisfies the +`QuotientMeasureEqMeasurePreimage` typeclass. -/ +@[to_additive] +lemma IsFundamentalDomain.quotientMeasureEqMeasurePreimage_quotientMeasure + {s : Set α} (fund_dom_s : IsFundamentalDomain G s ν) : + QuotientMeasureEqMeasurePreimage ν ((ν.restrict s).map π) where + projection_respects_measure' t fund_dom_t := by rw [fund_dom_s.quotientMeasure_eq _ fund_dom_t] + +/-- One can prove `QuotientMeasureEqMeasurePreimage` by checking behavior with respect to a single +fundamental domain. -/ +@[to_additive] +lemma IsFundamentalDomain.quotientMeasureEqMeasurePreimage {μ : Measure (Quotient α_mod_G)} + {s : Set α} (fund_dom_s : IsFundamentalDomain G s ν) (h : μ = (ν.restrict s).map π) : + QuotientMeasureEqMeasurePreimage ν μ := by + simpa [h] using fund_dom_s.quotientMeasureEqMeasurePreimage_quotientMeasure + + /-- If a fundamental domain has volume 0, then `QuotientMeasureEqMeasurePreimage` holds. -/ @[to_additive] theorem IsFundamentalDomain.quotientMeasureEqMeasurePreimage_of_zero diff --git a/Mathlib/MeasureTheory/Integral/CircleIntegral.lean b/Mathlib/MeasureTheory/Integral/CircleIntegral.lean index 1a25d84a8bf81..a2409b3431d6b 100644 --- a/Mathlib/MeasureTheory/Integral/CircleIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/CircleIntegral.lean @@ -298,7 +298,7 @@ theorem circleIntegrable_sub_inv_iff {c w : ℂ} {R : ℝ} : CircleIntegrable (fun z => (z - w)⁻¹) c R ↔ R = 0 ∨ w ∉ sphere c |R| := by simp only [← zpow_neg_one, circleIntegrable_sub_zpow_iff]; norm_num -variable [NormedSpace ℂ E] [CompleteSpace E] +variable [NormedSpace ℂ E] /-- Definition for $\oint_{|z-c|=R} f(z)\,dz$. -/ def circleIntegral (f : ℂ → E) (c : ℂ) (R : ℝ) : E := @@ -393,7 +393,7 @@ theorem integral_smul {𝕜 : Type*} [RCLike 𝕜] [NormedSpace 𝕜 E] [SMulCom simp only [circleIntegral, ← smul_comm a (_ : ℂ) (_ : E), intervalIntegral.integral_smul] @[simp] -theorem integral_smul_const (f : ℂ → ℂ) (a : E) (c : ℂ) (R : ℝ) : +theorem integral_smul_const [CompleteSpace E] (f : ℂ → ℂ) (a : E) (c : ℂ) (R : ℝ) : (∮ z in C(c, R), f z • a) = (∮ z in C(c, R), f z) • a := by simp only [circleIntegral, intervalIntegral.integral_smul_const, ← smul_assoc] @@ -411,7 +411,7 @@ theorem integral_sub_center_inv (c : ℂ) {R : ℝ} (hR : R ≠ 0) : /-- If `f' : ℂ → E` is a derivative of a complex differentiable function on the circle `Metric.sphere c |R|`, then `∮ z in C(c, R), f' z = 0`. -/ -theorem integral_eq_zero_of_hasDerivWithinAt' {f f' : ℂ → E} {c : ℂ} {R : ℝ} +theorem integral_eq_zero_of_hasDerivWithinAt' [CompleteSpace E] {f f' : ℂ → E} {c : ℂ} {R : ℝ} (h : ∀ z ∈ sphere c |R|, HasDerivWithinAt f (f' z) (sphere c |R|) z) : (∮ z in C(c, R), f' z) = 0 := by by_cases hi : CircleIntegrable f' c R @@ -423,7 +423,8 @@ theorem integral_eq_zero_of_hasDerivWithinAt' {f f' : ℂ → E} {c : ℂ} {R : /-- If `f' : ℂ → E` is a derivative of a complex differentiable function on the circle `Metric.sphere c R`, then `∮ z in C(c, R), f' z = 0`. -/ -theorem integral_eq_zero_of_hasDerivWithinAt {f f' : ℂ → E} {c : ℂ} {R : ℝ} (hR : 0 ≤ R) +theorem integral_eq_zero_of_hasDerivWithinAt [CompleteSpace E] + {f f' : ℂ → E} {c : ℂ} {R : ℝ} (hR : 0 ≤ R) (h : ∀ z ∈ sphere c R, HasDerivWithinAt f (f' z) (sphere c R) z) : (∮ z in C(c, R), f' z) = 0 := integral_eq_zero_of_hasDerivWithinAt' <| (_root_.abs_of_nonneg hR).symm ▸ h diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index 5f5bef893c3af..cbb80865f8332 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -150,15 +150,17 @@ end MetricSpace section Preorderα variable [Preorder α] [TopologicalSpace α] [OrderClosedTopology α] [OpensMeasurableSpace α] - {a b : ι → α} (ha : Tendsto a l atBot) (hb : Tendsto b l atTop) + {a b : ι → α} -theorem aecover_Ici : AECover μ l fun i => Ici (a i) where +theorem aecover_Ici (ha : Tendsto a l atBot) : AECover μ l fun i => Ici (a i) where ae_eventually_mem := ae_of_all μ ha.eventually_le_atBot measurableSet _ := measurableSet_Ici -theorem aecover_Iic : AECover μ l fun i => Iic <| b i := aecover_Ici (α := αᵒᵈ) hb +theorem aecover_Iic (hb : Tendsto b l atTop) : AECover μ l fun i => Iic <| b i := + aecover_Ici (α := αᵒᵈ) hb -theorem aecover_Icc : AECover μ l fun i => Icc (a i) (b i) := +theorem aecover_Icc (ha : Tendsto a l atBot) (hb : Tendsto b l atTop) : + AECover μ l fun i => Icc (a i) (b i) := (aecover_Ici ha).inter (aecover_Iic hb) end Preorderα @@ -168,12 +170,16 @@ section LinearOrderα variable [LinearOrder α] [TopologicalSpace α] [OrderClosedTopology α] [OpensMeasurableSpace α] {a b : ι → α} (ha : Tendsto a l atBot) (hb : Tendsto b l atTop) +include ha in theorem aecover_Ioi [NoMinOrder α] : AECover μ l fun i => Ioi (a i) where ae_eventually_mem := ae_of_all μ ha.eventually_lt_atBot measurableSet _ := measurableSet_Ioi +include hb in theorem aecover_Iio [NoMaxOrder α] : AECover μ l fun i => Iio (b i) := aecover_Ioi (α := αᵒᵈ) hb +include ha hb + theorem aecover_Ioo [NoMinOrder α] [NoMaxOrder α] : AECover μ l fun i => Ioo (a i) (b i) := (aecover_Ioi ha).inter (aecover_Iio hb) @@ -190,30 +196,38 @@ section FiniteIntervals variable [LinearOrder α] [TopologicalSpace α] [OrderClosedTopology α] [OpensMeasurableSpace α] {a b : ι → α} {A B : α} (ha : Tendsto a l (𝓝 A)) (hb : Tendsto b l (𝓝 B)) +include ha in theorem aecover_Ioi_of_Ioi : AECover (μ.restrict (Ioi A)) l fun i ↦ Ioi (a i) where ae_eventually_mem := (ae_restrict_mem measurableSet_Ioi).mono fun _x hx ↦ ha.eventually <| eventually_lt_nhds hx measurableSet _ := measurableSet_Ioi +include hb in theorem aecover_Iio_of_Iio : AECover (μ.restrict (Iio B)) l fun i ↦ Iio (b i) := aecover_Ioi_of_Ioi (α := αᵒᵈ) hb +include ha in theorem aecover_Ioi_of_Ici : AECover (μ.restrict (Ioi A)) l fun i ↦ Ici (a i) := (aecover_Ioi_of_Ioi ha).superset (fun _ ↦ Ioi_subset_Ici_self) fun _ ↦ measurableSet_Ici +include hb in theorem aecover_Iio_of_Iic : AECover (μ.restrict (Iio B)) l fun i ↦ Iic (b i) := aecover_Ioi_of_Ici (α := αᵒᵈ) hb +include ha hb in theorem aecover_Ioo_of_Ioo : AECover (μ.restrict <| Ioo A B) l fun i => Ioo (a i) (b i) := ((aecover_Ioi_of_Ioi ha).mono <| Measure.restrict_mono Ioo_subset_Ioi_self le_rfl).inter ((aecover_Iio_of_Iio hb).mono <| Measure.restrict_mono Ioo_subset_Iio_self le_rfl) +include ha hb in theorem aecover_Ioo_of_Icc : AECover (μ.restrict <| Ioo A B) l fun i => Icc (a i) (b i) := (aecover_Ioo_of_Ioo ha hb).superset (fun _ ↦ Ioo_subset_Icc_self) fun _ ↦ measurableSet_Icc +include ha hb in theorem aecover_Ioo_of_Ico : AECover (μ.restrict <| Ioo A B) l fun i => Ico (a i) (b i) := (aecover_Ioo_of_Ioo ha hb).superset (fun _ ↦ Ioo_subset_Ico_self) fun _ ↦ measurableSet_Ico +include ha hb in theorem aecover_Ioo_of_Ioc : AECover (μ.restrict <| Ioo A B) l fun i => Ioc (a i) (b i) := (aecover_Ioo_of_Ioo ha hb).superset (fun _ ↦ Ioo_subset_Ioc_self) fun _ ↦ measurableSet_Ioc @@ -452,7 +466,7 @@ end Integrable section Integral variable {α ι E : Type*} [MeasurableSpace α] {μ : Measure α} {l : Filter ι} [NormedAddCommGroup E] - [NormedSpace ℝ E] [CompleteSpace E] + [NormedSpace ℝ E] theorem AECover.integral_tendsto_of_countably_generated [l.IsCountablyGenerated] {φ : ι → Set α} (hφ : AECover μ l φ) {f : α → E} (hfi : Integrable f μ) : @@ -588,7 +602,7 @@ end IntegrableOfIntervalIntegral section IntegralOfIntervalIntegral variable {ι E : Type*} {μ : Measure ℝ} {l : Filter ι} [IsCountablyGenerated l] - [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] {a b : ι → ℝ} {f : ℝ → E} + [NormedAddCommGroup E] [NormedSpace ℝ E] {a b : ι → ℝ} {f : ℝ → E} theorem intervalIntegral_tendsto_integral (hfi : Integrable f μ) (ha : Tendsto a l atBot) (hb : Tendsto b l atTop) : Tendsto (fun i => ∫ x in a i..b i, f x ∂μ) l (𝓝 <| ∫ x, f x ∂μ) := by diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 5ca5b383e51ad..53d163b674b03 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -406,7 +406,7 @@ In this section we define `∫ x in a..b, f x ∂μ` as `∫ x in Ioc a b, f x and prove some basic properties. -/ -variable [CompleteSpace E] [NormedSpace ℝ E] +variable [NormedSpace ℝ E] /-- The interval integral `∫ x in a..b, f x ∂μ` is defined as `∫ x in Ioc a b, f x ∂μ - ∫ x in Ioc b a, f x ∂μ`. If `a ≤ b`, then it equals @@ -545,7 +545,8 @@ nonrec theorem integral_smul {𝕜 : Type*} [NontriviallyNormedField 𝕜] [Norm simp only [intervalIntegral, integral_smul, smul_sub] @[simp] -nonrec theorem integral_smul_const {𝕜 : Type*} [RCLike 𝕜] [NormedSpace 𝕜 E] (f : ℝ → 𝕜) (c : E) : +nonrec theorem integral_smul_const [CompleteSpace E] + {𝕜 : Type*} [RCLike 𝕜] [NormedSpace 𝕜 E] (f : ℝ → 𝕜) (c : E) : ∫ x in a..b, f x • c ∂μ = (∫ x in a..b, f x ∂μ) • c := by simp only [intervalIntegral_eq_integral_uIoc, integral_smul_const, smul_assoc] @@ -564,12 +565,12 @@ theorem integral_div {𝕜 : Type*} [RCLike 𝕜] (r : 𝕜) (f : ℝ → 𝕜) ∫ x in a..b, f x / r ∂μ = (∫ x in a..b, f x ∂μ) / r := by simpa only [div_eq_mul_inv] using integral_mul_const r⁻¹ f -theorem integral_const' (c : E) : +theorem integral_const' [CompleteSpace E] (c : E) : ∫ _ in a..b, c ∂μ = ((μ <| Ioc a b).toReal - (μ <| Ioc b a).toReal) • c := by simp only [intervalIntegral, setIntegral_const, sub_smul] @[simp] -theorem integral_const (c : E) : ∫ _ in a..b, c = (b - a) • c := by +theorem integral_const [CompleteSpace E] (c : E) : ∫ _ in a..b, c = (b - a) • c := by simp only [integral_const', Real.volume_Ioc, ENNReal.toReal_ofReal', ← neg_sub b, max_zero_sub_eq_self] @@ -605,7 +606,7 @@ theorem _root_.ContinuousLinearMap.intervalIntegral_apply {a b : ℝ} {φ : ℝ variable [NormedSpace ℝ F] [CompleteSpace F] -theorem _root_.ContinuousLinearMap.intervalIntegral_comp_comm (L : E →L[𝕜] F) +theorem _root_.ContinuousLinearMap.intervalIntegral_comp_comm [CompleteSpace E] (L : E →L[𝕜] F) (hf : IntervalIntegrable f μ a b) : (∫ x in a..b, L (f x) ∂μ) = L (∫ x in a..b, f x ∂μ) := by simp_rw [intervalIntegral, L.integral_comp_comm hf.1, L.integral_comp_comm hf.2, L.map_sub] @@ -878,7 +879,7 @@ theorem integral_Iio_add_Ici (h_left : IntegrableOn f (Iio b) μ) rw [Iio_union_Ici, Measure.restrict_univ] /-- If `μ` is a finite measure then `∫ x in a..b, c ∂μ = (μ (Iic b) - μ (Iic a)) • c`. -/ -theorem integral_const_of_cdf [IsFiniteMeasure μ] (c : E) : +theorem integral_const_of_cdf [CompleteSpace E] [IsFiniteMeasure μ] (c : E) : ∫ _ in a..b, c ∂μ = ((μ (Iic b)).toReal - (μ (Iic a)).toReal) • c := by simp only [sub_smul, ← setIntegral_const] refine (integral_Iic_sub_Iic ?_ ?_).symm <;> @@ -1027,7 +1028,24 @@ theorem abs_integral_le_integral_abs (hab : a ≤ b) : section Mono +theorem integral_mono_interval {c d} (hca : c ≤ a) (hab : a ≤ b) (hbd : b ≤ d) + (hf : 0 ≤ᵐ[μ.restrict (Ioc c d)] f) (hfi : IntervalIntegrable f μ c d) : + (∫ x in a..b, f x ∂μ) ≤ ∫ x in c..d, f x ∂μ := by + rw [integral_of_le hab, integral_of_le (hca.trans (hab.trans hbd))] + exact setIntegral_mono_set hfi.1 hf (Ioc_subset_Ioc hca hbd).eventuallyLE + +theorem abs_integral_mono_interval {c d} (h : Ι a b ⊆ Ι c d) (hf : 0 ≤ᵐ[μ.restrict (Ι c d)] f) + (hfi : IntervalIntegrable f μ c d) : |∫ x in a..b, f x ∂μ| ≤ |∫ x in c..d, f x ∂μ| := + have hf' : 0 ≤ᵐ[μ.restrict (Ι a b)] f := ae_mono (Measure.restrict_mono h le_rfl) hf + calc + |∫ x in a..b, f x ∂μ| = |∫ x in Ι a b, f x ∂μ| := abs_integral_eq_abs_integral_uIoc f + _ = ∫ x in Ι a b, f x ∂μ := abs_of_nonneg (MeasureTheory.integral_nonneg_of_ae hf') + _ ≤ ∫ x in Ι c d, f x ∂μ := setIntegral_mono_set hfi.def' hf h.eventuallyLE + _ ≤ |∫ x in Ι c d, f x ∂μ| := le_abs_self _ + _ = |∫ x in c..d, f x ∂μ| := (abs_integral_eq_abs_integral_uIoc f).symm + variable (hab : a ≤ b) (hf : IntervalIntegrable f μ a b) (hg : IntervalIntegrable g μ a b) +include hab hf hg theorem integral_mono_ae_restrict (h : f ≤ᵐ[μ.restrict (Icc a b)] g) : (∫ u in a..b, f u ∂μ) ≤ ∫ u in a..b, g u ∂μ := by @@ -1045,22 +1063,6 @@ theorem integral_mono_on (h : ∀ x ∈ Icc a b, f x ≤ g x) : theorem integral_mono (h : f ≤ g) : (∫ u in a..b, f u ∂μ) ≤ ∫ u in a..b, g u ∂μ := integral_mono_ae hab hf hg <| ae_of_all _ h -theorem integral_mono_interval {c d} (hca : c ≤ a) (hab : a ≤ b) (hbd : b ≤ d) - (hf : 0 ≤ᵐ[μ.restrict (Ioc c d)] f) (hfi : IntervalIntegrable f μ c d) : - (∫ x in a..b, f x ∂μ) ≤ ∫ x in c..d, f x ∂μ := by - rw [integral_of_le hab, integral_of_le (hca.trans (hab.trans hbd))] - exact setIntegral_mono_set hfi.1 hf (Ioc_subset_Ioc hca hbd).eventuallyLE - -theorem abs_integral_mono_interval {c d} (h : Ι a b ⊆ Ι c d) (hf : 0 ≤ᵐ[μ.restrict (Ι c d)] f) - (hfi : IntervalIntegrable f μ c d) : |∫ x in a..b, f x ∂μ| ≤ |∫ x in c..d, f x ∂μ| := - have hf' : 0 ≤ᵐ[μ.restrict (Ι a b)] f := ae_mono (Measure.restrict_mono h le_rfl) hf - calc - |∫ x in a..b, f x ∂μ| = |∫ x in Ι a b, f x ∂μ| := abs_integral_eq_abs_integral_uIoc f - _ = ∫ x in Ι a b, f x ∂μ := abs_of_nonneg (MeasureTheory.integral_nonneg_of_ae hf') - _ ≤ ∫ x in Ι c d, f x ∂μ := setIntegral_mono_set hfi.def' hf h.eventuallyLE - _ ≤ |∫ x in Ι c d, f x ∂μ| := le_abs_self _ - _ = |∫ x in c..d, f x ∂μ| := (abs_integral_eq_abs_integral_uIoc f).symm - end Mono end diff --git a/Mathlib/MeasureTheory/Integral/Periodic.lean b/Mathlib/MeasureTheory/Integral/Periodic.lean index 5b378bc95c27c..00da4d383fa2e 100644 --- a/Mathlib/MeasureTheory/Integral/Periodic.lean +++ b/Mathlib/MeasureTheory/Integral/Periodic.lean @@ -157,7 +157,7 @@ protected theorem lintegral_preimage (t : ℝ) (f : AddCircle T → ℝ≥0∞) rw [← map_map AddCircle.measurable_mk' measurable_subtype_coe, ← map_comap_subtype_coe m] rfl -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] attribute [local instance] Subtype.measureSpace in /-- The integral of an almost-everywhere strongly measurable function over `AddCircle T` is equal @@ -205,7 +205,7 @@ protected theorem lintegral_preimage (t : ℝ) (f : UnitAddCircle → ℝ≥0∞ (∫⁻ a in Ioc t (t + 1), f a) = ∫⁻ b : UnitAddCircle, f b := AddCircle.lintegral_preimage 1 t f -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] /-- The integral of an almost-everywhere strongly measurable function over `UnitAddCircle` is equal to the integral over an interval (t, t + 1] in `ℝ` of its lift to `ℝ`. -/ @@ -221,7 +221,7 @@ protected theorem intervalIntegral_preimage (t : ℝ) (f : UnitAddCircle → E) end UnitAddCircle -variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] namespace Function @@ -285,6 +285,7 @@ open Filter variable {g : ℝ → ℝ} variable (hg : Periodic g T) (h_int : ∀ t₁ t₂, IntervalIntegrable g MeasureSpace.volume t₁ t₂) +include hg h_int /-- If `g : ℝ → ℝ` is periodic with period `T > 0`, then for any `t : ℝ`, the function `t ↦ ∫ x in 0..t, g x` is bounded below by `t ↦ X + ⌊t/T⌋ • Y` for appropriate constants `X` and diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index b8cfddf9ab91b..44f6063be9b79 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -730,8 +730,11 @@ end NormedAddCommGroup section Mono -variable {μ : Measure X} {f g : X → ℝ} {s t : Set X} (hf : IntegrableOn f s μ) - (hg : IntegrableOn g s μ) +variable {μ : Measure X} {f g : X → ℝ} {s t : Set X} + +section +variable (hf : IntegrableOn f s μ) (hg : IntegrableOn g s μ) +include hf hg theorem setIntegral_mono_ae_restrict (h : f ≤ᵐ[μ.restrict s] g) : ∫ x in s, f x ∂μ ≤ ∫ x in s, g x ∂μ := @@ -767,6 +770,8 @@ theorem setIntegral_mono (h : f ≤ g) : ∫ x in s, f x ∂μ ≤ ∫ x in s, g @[deprecated (since := "2024-04-17")] alias set_integral_mono := setIntegral_mono +end + theorem setIntegral_mono_set (hfi : IntegrableOn f t μ) (hf : 0 ≤ᵐ[μ.restrict t] f) (hst : s ≤ᵐ[μ] t) : ∫ x in s, f x ∂μ ≤ ∫ x in t, f x ∂μ := integral_mono_measure (Measure.restrict_mono_ae hst) hf hfi @@ -1263,7 +1268,7 @@ end ContinuousLinearEquiv namespace ContinuousMap -lemma integral_apply [TopologicalSpace Y] [CompactSpace Y] [NormedAddCommGroup E] [NormedSpace ℝ E] +lemma integral_apply [TopologicalSpace Y] [CompactSpace Y] [NormedSpace ℝ E] [CompleteSpace E] {f : X → C(Y, E)} (hf : Integrable f μ) (y : Y) : (∫ x, f x ∂μ) y = ∫ x, f x y ∂μ := by calc (∫ x, f x ∂μ) y = ContinuousMap.evalCLM ℝ y (∫ x, f x ∂μ) := rfl @@ -1454,11 +1459,15 @@ theorem measure_le_lintegral_thickenedIndicator (μ : Measure X) {E : Set X} end thickenedIndicator +-- We declare a new `{X : Type*}` to discard the instance `[MeasureableSpace X]` +-- which has been in scope for the entire file up to this point. +variable {X : Type*} + section BilinearMap namespace MeasureTheory -variable {f : X → ℝ} {m m0 : MeasurableSpace X} {μ : Measure X} +variable {X : Type*} {f : X → ℝ} {m m0 : MeasurableSpace X} {μ : Measure X} theorem Integrable.simpleFunc_mul (g : SimpleFunc X ℝ) (hf : Integrable f μ) : Integrable (⇑g * f) μ := by @@ -1499,7 +1508,7 @@ open Metric ContinuousLinearMap under mild assumptions on the topologies involved. -/ theorem continuous_parametric_integral_of_continuous [FirstCountableTopology X] [LocallyCompactSpace X] - [OpensMeasurableSpace Y] [SecondCountableTopologyEither Y E] [IsLocallyFiniteMeasure μ] + [SecondCountableTopologyEither Y E] [IsLocallyFiniteMeasure μ] {f : X → Y → E} (hf : Continuous f.uncurry) {s : Set Y} (hs : IsCompact s) : Continuous (∫ y in s, f · y ∂μ) := by rw [continuous_iff_continuousAt] diff --git a/Mathlib/MeasureTheory/Integral/TorusIntegral.lean b/Mathlib/MeasureTheory/Integral/TorusIntegral.lean index ee80053903107..3398c9ead73c2 100644 --- a/Mathlib/MeasureTheory/Integral/TorusIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/TorusIntegral.lean @@ -134,7 +134,7 @@ theorem function_integrable [NormedSpace ℂ E] (hf : TorusIntegrable f c R) : end TorusIntegrable -variable [NormedSpace ℂ E] [CompleteSpace E] {f g : (Fin n → ℂ) → E} {c : Fin n → ℂ} {R : Fin n → ℝ} +variable [NormedSpace ℂ E] {f g : (Fin n → ℂ) → E} {c : Fin n → ℂ} {R : Fin n → ℝ} /-- The integral over a generalized torus with center `c ∈ ℂⁿ` and radius `R ∈ ℝⁿ`, defined as the `•`-product of the derivative of `torusMap` and `f (torusMap c R θ)`-/ @@ -185,7 +185,8 @@ theorem norm_torusIntegral_le_of_norm_le_const {C : ℝ} (hf : ∀ θ, ‖f (tor Fin.prod_const, mul_assoc, mul_comm ((2 * π) ^ (n : ℕ))] @[simp] -theorem torusIntegral_dim0 (f : ℂ⁰ → E) (c : ℂ⁰) (R : ℝ⁰) : (∯ x in T(c, R), f x) = f c := by +theorem torusIntegral_dim0 [CompleteSpace E] + (f : ℂ⁰ → E) (c : ℂ⁰) (R : ℝ⁰) : (∯ x in T(c, R), f x) = f c := by simp only [torusIntegral, Fin.prod_univ_zero, one_smul, Subsingleton.elim (fun _ : Fin 0 => 2 * π) 0, Icc_self, Measure.restrict_singleton, volume_pi, integral_smul_measure, integral_dirac, Measure.pi_of_empty (fun _ : Fin 0 ↦ volume) 0, @@ -207,7 +208,8 @@ theorem torusIntegral_dim1 (f : ℂ¹ → E) (c : ℂ¹) (R : ℝ¹) : simp [circleMap_zero] /-- Recurrent formula for `torusIntegral`, see also `torusIntegral_succ`. -/ -theorem torusIntegral_succAbove {f : ℂⁿ⁺¹ → E} {c : ℂⁿ⁺¹} {R : ℝⁿ⁺¹} (hf : TorusIntegrable f c R) +theorem torusIntegral_succAbove + {f : ℂⁿ⁺¹ → E} {c : ℂⁿ⁺¹} {R : ℝⁿ⁺¹} (hf : TorusIntegrable f c R) (i : Fin (n + 1)) : (∯ x in T(c, R), f x) = ∮ x in C(c i, R i), ∯ y in T(c ∘ i.succAbove, R ∘ i.succAbove), f (i.insertNth x y) := by @@ -231,7 +233,8 @@ theorem torusIntegral_succAbove {f : ℂⁿ⁺¹ → E} {c : ℂⁿ⁺¹} {R : rwa [← hem.integrableOn_comp_preimage e.measurableEmbedding, heπ] at this /-- Recurrent formula for `torusIntegral`, see also `torusIntegral_succAbove`. -/ -theorem torusIntegral_succ {f : ℂⁿ⁺¹ → E} {c : ℂⁿ⁺¹} {R : ℝⁿ⁺¹} (hf : TorusIntegrable f c R) : +theorem torusIntegral_succ + {f : ℂⁿ⁺¹ → E} {c : ℂⁿ⁺¹} {R : ℝⁿ⁺¹} (hf : TorusIntegrable f c R) : (∯ x in T(c, R), f x) = ∮ x in C(c 0, R 0), ∯ y in T(c ∘ Fin.succ, R ∘ Fin.succ), f (Fin.cons x y) := by simpa using torusIntegral_succAbove hf 0 diff --git a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean index 50342e831be4d..31dbff7cf6c84 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean @@ -63,13 +63,25 @@ end section smulInvariantMeasure -variable {G : Type*} [Group G] [MeasurableSpace G] [TopologicalSpace G] [TopologicalGroup G] - [BorelSpace G] [PolishSpace G] (ν : Measure G) {Γ : Subgroup G} [Countable Γ] - [T2Space (G ⧸ Γ)] [SecondCountableTopology (G ⧸ Γ)] {μ : Measure (G ⧸ Γ)} +variable {G : Type*} [Group G] [MeasurableSpace G] (ν : Measure G) {Γ : Subgroup G} + {μ : Measure (G ⧸ Γ)} [QuotientMeasureEqMeasurePreimage ν μ] +/-- Given a subgroup `Γ` of a topological group `G` with measure `ν`, and a measure 'μ' on the + quotient `G ⧸ Γ` satisfying `QuotientMeasureEqMeasurePreimage`, the restriction + of `ν` to a fundamental domain is measure-preserving with respect to `μ`. -/ +@[to_additive] +theorem measurePreserving_quotientGroup_mk_of_QuotientMeasureEqMeasurePreimage + {𝓕 : Set G} (h𝓕 : IsFundamentalDomain Γ.op 𝓕 ν) (μ : Measure (G ⧸ Γ)) + [QuotientMeasureEqMeasurePreimage ν μ] : + MeasurePreserving (@QuotientGroup.mk G _ Γ) (ν.restrict 𝓕) μ := + h𝓕.measurePreserving_quotient_mk μ + local notation "π" => @QuotientGroup.mk G _ Γ +variable [TopologicalSpace G] [TopologicalGroup G] [BorelSpace G] [PolishSpace G] + [T2Space (G ⧸ Γ)] [SecondCountableTopology (G ⧸ Γ)] + /-- If `μ` satisfies `QuotientMeasureEqMeasurePreimage` relative to a both left- and right- invariant measure `ν` on `G`, then it is a `G` invariant measure on `G ⧸ Γ`. -/ @[to_additive] @@ -96,28 +108,17 @@ lemma MeasureTheory.QuotientMeasureEqMeasurePreimage.smulInvariantMeasure_quotie rw [measure_preimage_mul] rw [this, ← preimage_smul_inv]; rfl -/-- Given a subgroup `Γ` of a topological group `G` with measure `ν`, and a measure 'μ' on the - quotient `G ⧸ Γ` satisfying `QuotientMeasureEqMeasurePreimage`, the restriction - of `ν` to a fundamental domain is measure-preserving with respect to `μ`. -/ -@[to_additive] -theorem measurePreserving_quotientGroup_mk_of_QuotientMeasureEqMeasurePreimage - {𝓕 : Set G} (h𝓕 : IsFundamentalDomain Γ.op 𝓕 ν) (μ : Measure (G ⧸ Γ)) - [QuotientMeasureEqMeasurePreimage ν μ] : - MeasurePreserving (@QuotientGroup.mk G _ Γ) (ν.restrict 𝓕) μ := - h𝓕.measurePreserving_quotient_mk μ - end smulInvariantMeasure section normal variable {G : Type*} [Group G] [MeasurableSpace G] [TopologicalSpace G] [TopologicalGroup G] - [BorelSpace G] [PolishSpace G] {Γ : Subgroup G} [Countable Γ] [Subgroup.Normal Γ] + [BorelSpace G] [PolishSpace G] {Γ : Subgroup G} [Subgroup.Normal Γ] [T2Space (G ⧸ Γ)] [SecondCountableTopology (G ⧸ Γ)] {μ : Measure (G ⧸ Γ)} section mulInvariantMeasure -variable (ν : Measure G) [IsMulLeftInvariant ν] [IsMulRightInvariant ν] - [SigmaFinite ν] +variable (ν : Measure G) [IsMulLeftInvariant ν] /-- If `μ` on `G ⧸ Γ` satisfies `QuotientMeasureEqMeasurePreimage` relative to a both left- and right-invariant measure on `G` and `Γ` is a normal subgroup, then `μ` is a left-invariant @@ -136,7 +137,8 @@ lemma MeasureTheory.QuotientMeasureEqMeasurePreimage.mulInvariantMeasure_quotien simp [← MulAction.Quotient.coe_smul_out', ← Quotient.mk''_eq_mk] exact smulInvariantMeasure_quotient ν -variable [IsMulLeftInvariant μ] [SigmaFinite μ] +variable [Countable Γ] [IsMulRightInvariant ν] [SigmaFinite ν] + [IsMulLeftInvariant μ] [SigmaFinite μ] local notation "π" => @QuotientGroup.mk G _ Γ @@ -204,7 +206,7 @@ end mulInvariantMeasure section haarMeasure -variable (ν : Measure G) [SigmaFinite ν] [IsHaarMeasure ν] [IsMulRightInvariant ν] +variable [Countable Γ] (ν : Measure G) [IsHaarMeasure ν] [IsMulRightInvariant ν] local notation "π" => @QuotientGroup.mk G _ Γ @@ -250,6 +252,8 @@ theorem MeasureTheory.QuotientMeasureEqMeasurePreimage.haarMeasure_quotient [Loc apply measure_mono exact inter_subset_right +variable [SigmaFinite ν] + /-- Given a normal subgroup `Γ` of a topological group `G` with Haar measure `μ`, which is also right-invariant, and a finite volume fundamental domain `𝓕`, the quotient map to `G ⧸ Γ`, properly normalized, satisfies `QuotientMeasureEqMeasurePreimage`. -/ @@ -319,6 +323,7 @@ variable {G : Type*} [Group G] [MeasurableSpace G] [TopologicalSpace G] [Topolog [BorelSpace G] {μ : Measure G} {Γ : Subgroup G} variable {𝓕 : Set G} (h𝓕 : IsFundamentalDomain Γ.op 𝓕 μ) +include h𝓕 variable [Countable Γ] [MeasurableSpace (G ⧸ Γ)] [BorelSpace (G ⧸ Γ)] diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean index 4d5668275bbf9..a43d6a7133501 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean @@ -69,6 +69,7 @@ variable {E : Type*} [AddCommGroup E] [Module ℝ E] [FiniteDimensional ℝ E] [ (μ : Measure E) [IsAddHaarMeasure μ] {g : E → ℝ} (h1 : g 0 = 0) (h2 : ∀ x, g (- x) = g x) (h3 : ∀ x y, g (x + y) ≤ g x + g y) (h4 : ∀ {x}, g x = 0 → x = 0) (h5 : ∀ r x, g (r • x) ≤ |r| * (g x)) +include h1 h2 h3 h4 h5 theorem MeasureTheory.measure_lt_one_eq_integral_div_gamma {p : ℝ} (hp : 0 < p) : μ {x : E | g x < 1} = @@ -161,9 +162,9 @@ section LpSpace open Real Fintype ENNReal FiniteDimensional MeasureTheory MeasureTheory.Measure -variable (ι : Type*) [Fintype ι] {p : ℝ} (hp : 1 ≤ p) +variable (ι : Type*) [Fintype ι] {p : ℝ} -theorem MeasureTheory.volume_sum_rpow_lt_one : +theorem MeasureTheory.volume_sum_rpow_lt_one (hp : 1 ≤ p) : volume {x : ι → ℝ | ∑ i, |x i| ^ p < 1} = .ofReal ((2 * Gamma (1 / p + 1)) ^ card ι / Gamma (card ι / p + 1)) := by have h₁ : 0 < p := by linarith diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index f1612b1c061fe..0acae1a750f6c 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -771,7 +771,11 @@ open MeasureTheory Measure namespace MeasurableEmbedding -variable {m0 : MeasurableSpace α} {m1 : MeasurableSpace β} {f : α → β} (hf : MeasurableEmbedding f) +variable {m0 : MeasurableSpace α} {m1 : MeasurableSpace β} {f : α → β} + +section +variable (hf : MeasurableEmbedding f) +include hf theorem map_comap (μ : Measure β) : (comap f μ).map f = μ.restrict (range f) := by ext1 t ht @@ -811,6 +815,8 @@ lemma restrict_comap (μ : Measure β) (s : Set α) : (μ.comap f).restrict s = (μ.restrict (f '' s)).comap f := by rw [comap_restrict hf, preimage_image_eq _ hf.injective] +end + theorem _root_.MeasurableEquiv.restrict_map (e : α ≃ᵐ β) (μ : Measure α) (s : Set β) : (μ.map e).restrict s = (μ.restrict <| e ⁻¹' s).map e := e.measurableEmbedding.restrict_map _ _ diff --git a/Mathlib/MeasureTheory/OuterMeasure/Induced.lean b/Mathlib/MeasureTheory/OuterMeasure/Induced.lean index b18e208fe28c9..b1dd27d6be85b 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Induced.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Induced.lean @@ -88,9 +88,6 @@ variable variable (msU : ∀ ⦃f : ℕ → Set α⦄ (hm : ∀ i, P (f i)), m (⋃ i, f i) (PU hm) ≤ ∑' i, m (f i) (hm i)) variable (m_mono : ∀ ⦃s₁ s₂ : Set α⦄ (hs₁ : P s₁) (hs₂ : P s₂), s₁ ⊆ s₂ → m s₁ hs₁ ≤ m s₂ hs₂) -theorem extend_empty : extend m ∅ = 0 := - (extend_eq _ P0).trans m0 - theorem extend_iUnion_nat {f : ℕ → Set α} (hm : ∀ i, P (f i)) (mU : m (⋃ i, f i) (PU hm) = ∑' i, m (f i) (hm i)) : extend m (⋃ i, f i) = ∑' i, extend m (f i) := @@ -99,8 +96,13 @@ theorem extend_iUnion_nat {f : ℕ → Set α} (hm : ∀ i, P (f i)) congr with i rw [extend_eq] +include P0 m0 in +theorem extend_empty : extend m ∅ = 0 := + (extend_eq _ P0).trans m0 + section Subadditive +include PU msU in theorem extend_iUnion_le_tsum_nat' (s : ℕ → Set α) : extend m (⋃ i, s i) ≤ ∑' i, extend m (s i) := by by_cases h : ∀ i, P (s i) @@ -115,6 +117,7 @@ end Subadditive section Mono +include m_mono in theorem extend_mono' ⦃s₁ s₂ : Set α⦄ (h₁ : P s₁) (hs : s₁ ⊆ s₂) : extend m s₁ ≤ extend m s₂ := by refine le_iInf ?_ intro h₂ @@ -125,6 +128,7 @@ end Mono section Unions +include P0 m0 PU mU in theorem extend_iUnion {β} [Countable β] {f : β → Set α} (hd : Pairwise (Disjoint on f)) (hm : ∀ i, P (f i)) : extend m (⋃ i, f i) = ∑' i, extend m (f i) := by cases nonempty_encodable β @@ -134,6 +138,7 @@ theorem extend_iUnion {β} [Countable β] {f : β → Set α} (hd : Pairwise (Di (mU _ (Encodable.iUnion_decode₂_disjoint_on hd)) · exact extend_empty P0 m0 +include P0 m0 PU mU in theorem extend_union {s₁ s₂ : Set α} (hd : Disjoint s₁ s₂) (h₁ : P s₁) (h₂ : P s₂) : extend m (s₁ ∪ s₂) = extend m s₁ + extend m s₂ := by rw [union_eq_iUnion, @@ -167,6 +172,8 @@ theorem inducedOuterMeasure_union_of_false_of_nonempty_inter {s t : Set α} inducedOuterMeasure m P0 m0 s + inducedOuterMeasure m P0 m0 t := ofFunction_union_of_top_of_nonempty_inter fun u hsu htu => @iInf_of_empty _ _ _ ⟨h u hsu htu⟩ _ +include PU msU m_mono + theorem inducedOuterMeasure_eq_extend' {s : Set α} (hs : P s) : inducedOuterMeasure m P0 m0 s = extend m s := ofFunction_eq s (fun _t => extend_mono' m_mono hs) (extend_iUnion_le_tsum_nat' PU msU) @@ -252,6 +259,7 @@ variable (mU : ∀ ⦃f : ℕ → Set α⦄ (hm : ∀ i, MeasurableSet (f i)), Pairwise (Disjoint on f) → m (⋃ i, f i) (MeasurableSet.iUnion hm) = ∑' i, m (f i) (hm i)) +include m0 mU theorem extend_mono {s₁ s₂ : Set α} (h₁ : MeasurableSet s₁) (hs : s₁ ⊆ s₂) : extend m s₁ ≤ extend m s₂ := by @@ -311,7 +319,7 @@ theorem trim_mono : Monotone (trim : OuterMeasure α → OuterMeasure α) := fun iInf₂_mono fun _f _hs => ENNReal.tsum_le_tsum fun _b => iInf_mono fun _hf => H _ /-- `OuterMeasure.trim` is antitone in the σ-algebra. -/ -theorem trim_anti_measurableSpace (m : OuterMeasure α) {m0 m1 : MeasurableSpace α} +theorem trim_anti_measurableSpace {α} (m : OuterMeasure α) {m0 m1 : MeasurableSpace α} (h : m0 ≤ m1) : @trim _ m1 m ≤ @trim _ m0 m := by simp only [le_trim_iff] intro s hs diff --git a/Mathlib/ModelTheory/Substructures.lean b/Mathlib/ModelTheory/Substructures.lean index 31517d42b7c94..b08cea3eb212d 100644 --- a/Mathlib/ModelTheory/Substructures.lean +++ b/Mathlib/ModelTheory/Substructures.lean @@ -478,6 +478,7 @@ def gciMapComap (hf : Function.Injective f) : GaloisCoinsertion (map f) (comap f (gc_map_comap f).toGaloisCoinsertion fun S x => by simp [mem_comap, mem_map, hf.eq_iff] variable (hf : Function.Injective f) +include hf theorem comap_map_eq_of_injective (S : L.Substructure M) : (S.map f).comap f = S := (gciMapComap hf).u_l_eq _ @@ -513,6 +514,7 @@ end GaloisCoinsertion section GaloisInsertion variable {ι : Type*} {f : M →[L] N} (hf : Function.Surjective f) +include hf /-- `map f` and `comap f` form a `GaloisInsertion` when `f` is surjective. -/ def giMapComap : GaloisInsertion (map f) (comap f) := diff --git a/Mathlib/NumberTheory/Cyclotomic/Basic.lean b/Mathlib/NumberTheory/Cyclotomic/Basic.lean index e066e17c51f07..ed4273a3b101f 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Basic.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Basic.lean @@ -451,6 +451,7 @@ def algEquiv (L' : Type*) [Field L'] [Algebra K L'] [IsCyclotomicExtension {n} K scoped[Cyclotomic] attribute [instance] IsCyclotomicExtension.isSplittingField_X_pow_sub_one +include n in theorem isGalois : IsGalois K L := letI := isSplittingField_X_pow_sub_one n K L IsGalois.of_separable_splitting_field (X_pow_sub_one_separable_iff.2 diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index 7dec6765c845c..14303f635d275 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -210,7 +210,7 @@ variable {K} [Field K] [NumberField K] variable (n) in /-- If a `n`-th cyclotomic extension of `ℚ` contains a primitive `l`-th root of unity, then `l ∣ 2 * n`. -/ -theorem dvd_of_isCyclotomicExtension [NumberField K] [IsCyclotomicExtension {n} ℚ K] {ζ : K} +theorem dvd_of_isCyclotomicExtension [IsCyclotomicExtension {n} ℚ K] {ζ : K} {l : ℕ} (hζ : IsPrimitiveRoot ζ l) (hl : l ≠ 0) : l ∣ 2 * n := by have hl : NeZero l := ⟨hl⟩ have hroot := IsCyclotomicExtension.zeta_spec n ℚ K @@ -236,7 +236,7 @@ theorem dvd_of_isCyclotomicExtension [NumberField K] [IsCyclotomicExtension {n} /-- If `x` is a root of unity (spelled as `IsOfFinOrder x`) in an `n`-th cyclotomic extension of `ℚ`, where `n` is odd, and `ζ` is a primitive `n`-th root of unity, then there exist `r` such that `x = (-ζ)^r`. -/ -theorem exists_neg_pow_of_isOfFinOrder [NumberField K] [IsCyclotomicExtension {n} ℚ K] +theorem exists_neg_pow_of_isOfFinOrder [IsCyclotomicExtension {n} ℚ K] (hno : Odd (n : ℕ)) {ζ x : K} (hζ : IsPrimitiveRoot ζ n) (hx : IsOfFinOrder x) : ∃ r : ℕ, x = (-ζ) ^ r := by have hnegζ : IsPrimitiveRoot (-ζ) (2 * n) := by @@ -257,7 +257,7 @@ theorem exists_neg_pow_of_isOfFinOrder [NumberField K] [IsCyclotomicExtension {n /-- If `x` is a root of unity (spelled as `IsOfFinOrder x`) in an `n`-th cyclotomic extension of `ℚ`, where `n` is odd, and `ζ` is a primitive `n`-th root of unity, then there exists `r < n` such that `x = ζ^r` or `x = -ζ^r`. -/ -theorem exists_pow_or_neg_mul_pow_of_isOfFinOrder [NumberField K] [IsCyclotomicExtension {n} ℚ K] +theorem exists_pow_or_neg_mul_pow_of_isOfFinOrder [IsCyclotomicExtension {n} ℚ K] (hno : Odd (n : ℕ)) {ζ x : K} (hζ : IsPrimitiveRoot ζ n) (hx : IsOfFinOrder x) : ∃ r : ℕ, r < n ∧ (x = ζ ^ r ∨ x = -ζ ^ r) := by obtain ⟨r, hr⟩ := hζ.exists_neg_pow_of_isOfFinOrder hno hx @@ -269,7 +269,7 @@ end Field section CommRing -variable [CommRing L] {ζ : L} (hζ : IsPrimitiveRoot ζ n) +variable [CommRing L] {ζ : L} variable {K} [Field K] [Algebra K L] /-- This mathematically trivial result is complementary to `norm_eq_one` below. -/ @@ -277,6 +277,9 @@ theorem norm_eq_neg_one_pow (hζ : IsPrimitiveRoot ζ 2) [IsDomain L] : norm K ζ = (-1 : K) ^ finrank K L := by rw [hζ.eq_neg_one_of_two_right, show -1 = algebraMap K L (-1) by simp, Algebra.norm_algebraMap] +variable (hζ : IsPrimitiveRoot ζ n) +include hζ + /-- If `Irreducible (cyclotomic n K)` (in particular for `K = ℚ`), the norm of a primitive root is `1` if `n ≠ 2`. -/ theorem norm_eq_one [IsDomain L] [IsCyclotomicExtension {n} K L] (hn : n ≠ 2) @@ -314,9 +317,13 @@ end CommRing section Field -variable [Field L] {ζ : L} (hζ : IsPrimitiveRoot ζ n) +variable [Field L] {ζ : L} variable {K} [Field K] [Algebra K L] +section +variable (hζ : IsPrimitiveRoot ζ n) +include hζ + /-- If `Irreducible (cyclotomic n K)` (in particular for `K = ℚ`), then the norm of `ζ - 1` is `eval 1 (cyclotomic n ℤ)`. -/ theorem sub_one_norm_eq_eval_cyclotomic [IsCyclotomicExtension {n} K L] (h : 2 < (n : ℕ)) @@ -362,6 +369,8 @@ theorem sub_one_norm_isPrimePow (hn : IsPrimePow (n : ℕ)) [IsCyclotomicExtensi (mem_primeFactorsList (IsPrimePow.ne_zero hn)).2 ⟨hprime.out, minFac_dvd _⟩) simp [hk, sub_one_norm_eq_eval_cyclotomic hζ this hirr] +end + variable {A} theorem minpoly_sub_one_eq_cyclotomic_comp [Algebra K A] [IsDomain A] {ζ : A} diff --git a/Mathlib/NumberTheory/Dioph.lean b/Mathlib/NumberTheory/Dioph.lean index d0e7977dfd15d..b6602777b1e87 100644 --- a/Mathlib/NumberTheory/Dioph.lean +++ b/Mathlib/NumberTheory/Dioph.lean @@ -521,7 +521,9 @@ theorem const_dioph (n : ℕ) : DiophFn (const (α → ℕ) n) := scoped prefix:100 "D." => Dioph.const_dioph +section variable {f g : (α → ℕ) → ℕ} (df : DiophFn f) (dg : DiophFn g) +include df dg theorem dioph_comp2 {S : ℕ → ℕ → Prop} (d : Dioph fun v : Vector3 ℕ 2 => S (v &0) (v &1)) : Dioph fun v => S (f v) (g v) := dioph_comp d [f, g] ⟨df, dg⟩ @@ -627,6 +629,8 @@ theorem div_dioph : DiophFn fun v => f v / g v := (le_antisymm_iff.trans <| and_congr (Nat.le_div_iff_mul_le ypos) <| Iff.trans ⟨lt_succ_of_le, le_of_lt_succ⟩ (div_lt_iff_lt_mul ypos)).symm +end + scoped infixl:80 " D/ " => Dioph.div_dioph open Pell @@ -666,7 +670,8 @@ theorem xn_dioph : DiophPFun fun v : Vector3 ℕ 2 => ⟨1 < v &0, fun h => xn h Dioph.ext this fun _ => ⟨fun ⟨_, h, xe, _⟩ => ⟨h, xe⟩, fun ⟨h, xe⟩ => ⟨_, h, xe, rfl⟩⟩ /-- A version of **Matiyasevic's theorem** -/ -theorem pow_dioph : DiophFn fun v => f v ^ g v := by +theorem pow_dioph {f g : (α → ℕ) → ℕ} (df : DiophFn f) (dg : DiophFn g) : + DiophFn fun v => f v ^ g v := by have proof := let D_pell := pell_dioph.reindex_dioph (Fin2 9) [&4, &8, &1, &0] (D&2 D= D.0 D∧ D&0 D= D.1) D∨ (D.0 D< D&2 D∧ diff --git a/Mathlib/NumberTheory/DiophantineApproximation.lean b/Mathlib/NumberTheory/DiophantineApproximation.lean index 4a523ae2f1007..0e5d09c4d8fa2 100644 --- a/Mathlib/NumberTheory/DiophantineApproximation.lean +++ b/Mathlib/NumberTheory/DiophantineApproximation.lean @@ -392,7 +392,11 @@ private theorem aux₀ {v : ℤ} (hv : 0 < v) : (0 : ℝ) < v ∧ (0 : ℝ) < 2 ⟨cast_pos.mpr hv, by norm_cast; omega⟩ -- In the following, we assume that `ass ξ u v` holds and `v ≥ 2`. -variable {ξ : ℝ} {u v : ℤ} (hv : 2 ≤ v) (h : ContfracLegendre.Ass ξ u v) +variable {ξ : ℝ} {u v : ℤ} + +section +variable (hv : 2 ≤ v) (h : ContfracLegendre.Ass ξ u v) +include hv h -- The fractional part of `ξ` is positive. private theorem aux₁ : 0 < fract ξ := by @@ -504,16 +508,16 @@ private theorem invariant : ContfracLegendre.Ass (fract ξ)⁻¹ v (u - ⌊ξ⌋ rwa [lt_sub_iff_add_lt', (by ring : (v : ℝ) + -(1 / 2) = (2 * v - 1) / 2), lt_inv (div_pos hv₀' zero_lt_two) (aux₁ hv h), inv_div] +end + /-! ### The main result -/ /-- The technical version of *Legendre's Theorem*. -/ -theorem exists_rat_eq_convergent' {v : ℕ} (h' : ContfracLegendre.Ass ξ u v) : +theorem exists_rat_eq_convergent' {v : ℕ} (h : ContfracLegendre.Ass ξ u v) : ∃ n, (u / v : ℚ) = ξ.convergent n := by - -- Porting note: `induction` got in trouble because of the irrelevant hypothesis `h` - clear h; have h := h'; clear h' induction v using Nat.strong_induction_on generalizing ξ u with | h v ih => ?_ rcases lt_trichotomy v 1 with (ht | rfl | ht) · replace h := h.2.2 diff --git a/Mathlib/NumberTheory/EulerProduct/Basic.lean b/Mathlib/NumberTheory/EulerProduct/Basic.lean index d3d44bd3e19d5..1b6b6f551ed15 100644 --- a/Mathlib/NumberTheory/EulerProduct/Basic.lean +++ b/Mathlib/NumberTheory/EulerProduct/Basic.lean @@ -64,14 +64,18 @@ In this section we consider multiplicative (on coprime arguments) functions `f : where `R` is a complete normed commutative ring. The main result is `EulerProduct.eulerProduct`. -/ -variable {R : Type*} [NormedCommRing R] [CompleteSpace R] {f : ℕ → R} -variable (hf₁ : f 1 = 1) (hmul : ∀ {m n}, Nat.Coprime m n → f (m * n) = f m * f n) +variable {R : Type*} [NormedCommRing R] {f : ℕ → R} -- local instance to speed up typeclass search @[local instance] private lemma instT0Space : T0Space R := MetricSpace.instT0Space +variable [CompleteSpace R] + namespace EulerProduct +variable (hf₁ : f 1 = 1) (hmul : ∀ {m n}, Nat.Coprime m n → f (m * n) = f m * f n) + +include hf₁ hmul in /-- We relate a finite product over primes in `s` to an infinite sum over `s`-factored numbers. -/ lemma summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum (hsum : ∀ {p : ℕ}, p.Prime → Summable (fun n : ℕ ↦ ‖f (p ^ n)‖)) (s : Finset ℕ) : @@ -100,6 +104,7 @@ lemma summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum apply summable_mul_of_summable_norm (hsum hpp) ih.1 · rwa [factoredNumbers_insert s hpp] +include hf₁ hmul in /-- A version of `EulerProduct.summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum` in terms of the value of the series. -/ lemma prod_filter_prime_tsum_eq_tsum_factoredNumbers (hsum : Summable (‖f ·‖)) (s : Finset ℕ) : @@ -125,6 +130,7 @@ lemma norm_tsum_factoredNumbers_sub_tsum_lt (hsum : Summable f) (hf₀ : f 0 = 0 -- Versions of the three lemmas above for `smoothNumbers N` +include hf₁ hmul in /-- We relate a finite product over primes to an infinite sum over smooth numbers. -/ lemma summable_and_hasSum_smoothNumbers_prod_primesBelow_tsum (hsum : ∀ {p : ℕ}, p.Prime → Summable (fun n : ℕ ↦ ‖f (p ^ n)‖)) (N : ℕ) : @@ -133,6 +139,7 @@ lemma summable_and_hasSum_smoothNumbers_prod_primesBelow_tsum rw [smoothNumbers_eq_factoredNumbers, primesBelow] exact summable_and_hasSum_factoredNumbers_prod_filter_prime_tsum hf₁ hmul hsum _ +include hf₁ hmul in /-- A version of `EulerProduct.summable_and_hasSum_smoothNumbers_prod_primesBelow_tsum` in terms of the value of the series. -/ lemma prod_primesBelow_tsum_eq_tsum_smoothNumbers (hsum : Summable (‖f ·‖)) (N : ℕ) : @@ -152,6 +159,7 @@ lemma norm_tsum_smoothNumbers_sub_tsum_lt (hsum : Summable f) (hf₀ : f 0 = 0) exact mem_range.mpr <| (lt_of_mem_primesBelow hp).trans_le hN +include hf₁ hmul in /-- The *Euler Product* for multiplicative (on coprime arguments) functions. If `f : ℕ → R`, where `R` is a complete normed commutative ring, `f 0 = 0`, `f 1 = 1`, `f` is @@ -173,6 +181,7 @@ theorem eulerProduct_hasProd (hsum : Summable (‖f ·‖)) (hf₀ : f 0 = 0) : norm_sub_rev] exact hN₀ s fun p hp ↦ hs <| mem_range.mpr <| lt_of_mem_primesBelow hp +include hf₁ hmul in /-- The *Euler Product* for multiplicative (on coprime arguments) functions. If `f : ℕ → R`, where `R` is a complete normed commutative ring, `f 0 = 0`, `f 1 = 1`, `f` i @@ -185,6 +194,7 @@ theorem eulerProduct_hasProd_mulIndicator (hsum : Summable (‖f ·‖)) (hf₀ exact eulerProduct_hasProd hf₁ hmul hsum hf₀ open Filter in +include hf₁ hmul in /-- The *Euler Product* for multiplicative (on coprime arguments) functions. If `f : ℕ → R`, where `R` is a complete normed commutative ring, `f 0 = 0`, `f 1 = 1`, `f` is @@ -200,6 +210,7 @@ theorem eulerProduct (hsum : Summable (‖f ·‖)) (hf₀ : f 0 = 0) : prod_mulIndicator_eq_prod_filter (range n) (fun _ ↦ F) (fun _ ↦ {p | Nat.Prime p}) id simpa only [H] +include hf₁ hmul in /-- The *Euler Product* for multiplicative (on coprime arguments) functions. If `f : ℕ → R`, where `R` is a complete normed commutative ring, `f 0 = 0`, `f 1 = 1`, `f` is diff --git a/Mathlib/NumberTheory/FLT/Three.lean b/Mathlib/NumberTheory/FLT/Three.lean index 09088e17b9fee..68fcc6a03f0ea 100644 --- a/Mathlib/NumberTheory/FLT/Three.lean +++ b/Mathlib/NumberTheory/FLT/Three.lean @@ -230,6 +230,7 @@ def Solution.multiplicity := S.toSolution'.multiplicity `S.c` is less or equal than the multiplicity in `S₁.c`. -/ def Solution.isMinimal : Prop := ∀ (S₁ : Solution hζ), S.multiplicity ≤ S₁.multiplicity +include S in /-- If there is a solution then there is a minimal one. -/ lemma Solution.exists_minimal : ∃ (S₁ : Solution hζ), S₁.isMinimal := by classical diff --git a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean index 51f004593f0b5..50aeb1c4b74dd 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/AddCharacter.lean @@ -293,7 +293,7 @@ end Ring section Field -variable (F : Type*) [Field F] [Finite F] [DecidableEq F] +variable (F : Type*) [Field F] [Finite F] private lemma ringChar_ne : ringChar ℂ ≠ ringChar F := by simpa only [ringChar.eq_zero] using (CharP.ringChar_ne_zero_of_finite F).symm diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean index 003aa76d38a34..3c7aab6f98055 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean @@ -52,15 +52,13 @@ for the Jacobi symbol. namespace legendreSym -variable (hp : p ≠ 2) - /-- `legendreSym p 2` is given by `χ₈ p`. -/ -theorem at_two : legendreSym p 2 = χ₈ p := by +theorem at_two (hp : p ≠ 2) : legendreSym p 2 = χ₈ p := by have : (2 : ZMod p) = (2 : ℤ) := by norm_cast rw [legendreSym, ← this, quadraticChar_two ((ringChar_zmod_n p).substr hp), card p] /-- `legendreSym p (-2)` is given by `χ₈' p`. -/ -theorem at_neg_two : legendreSym p (-2) = χ₈' p := by +theorem at_neg_two (hp : p ≠ 2) : legendreSym p (-2) = χ₈' p := by have : (-2 : ZMod p) = (-2 : ℤ) := by norm_cast rw [legendreSym, ← this, quadraticChar_neg_two ((ringChar_zmod_n p).substr hp), card p] @@ -68,10 +66,8 @@ end legendreSym namespace ZMod -variable (hp : p ≠ 2) - /-- `2` is a square modulo an odd prime `p` iff `p` is congruent to `1` or `7` mod `8`. -/ -theorem exists_sq_eq_two_iff : IsSquare (2 : ZMod p) ↔ p % 8 = 1 ∨ p % 8 = 7 := by +theorem exists_sq_eq_two_iff (hp : p ≠ 2) : IsSquare (2 : ZMod p) ↔ p % 8 = 1 ∨ p % 8 = 7 := by rw [FiniteField.isSquare_two_iff, card p] have h₁ := Prime.mod_two_eq_one_iff_ne_two.mpr hp rw [← mod_mod_of_dvd p (by decide : 2 ∣ 8)] at h₁ @@ -81,7 +77,7 @@ theorem exists_sq_eq_two_iff : IsSquare (2 : ZMod p) ↔ p % 8 = 1 ∨ p % 8 = 7 intros; interval_cases m <;> simp_all -- Porting note (#11043): was `decide!` /-- `-2` is a square modulo an odd prime `p` iff `p` is congruent to `1` or `3` mod `8`. -/ -theorem exists_sq_eq_neg_two_iff : IsSquare (-2 : ZMod p) ↔ p % 8 = 1 ∨ p % 8 = 3 := by +theorem exists_sq_eq_neg_two_iff (hp : p ≠ 2) : IsSquare (-2 : ZMod p) ↔ p % 8 = 1 ∨ p % 8 = 3 := by rw [FiniteField.isSquare_neg_two_iff, card p] have h₁ := Prime.mod_two_eq_one_iff_ne_two.mpr hp rw [← mod_mod_of_dvd p (by decide : 2 ∣ 8)] at h₁ diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index 28f0d4c866ce7..15ef0f0e6656e 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -150,6 +150,7 @@ theorem pow_sub_pow_of_prime {p : R} (hp : Prime p) {x y : R} (hxy : p ∣ x - y zero_add] variable (hp : Prime (p : R)) (hp1 : Odd p) (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) +include hp hp1 hxy hx theorem geom_sum₂_eq_one : multiplicity (↑p) (∑ i ∈ range p, x ^ i * y ^ (p - 1 - i)) = 1 := by rw [← Nat.cast_one] @@ -182,6 +183,7 @@ end IntegralDomain section LiftingTheExponent variable (hp : Nat.Prime p) (hp1 : Odd p) +include hp hp1 /-- **Lifting the exponent lemma** for odd primes. -/ theorem Int.pow_sub_pow {x y : ℤ} (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) (n : ℕ) : @@ -367,6 +369,7 @@ theorem pow_two_sub_pow (hyx : y < x) (hxy : 2 ∣ x - y) (hx : ¬2 ∣ x) {n : · simp only [tsub_pos_iff_lt, pow_lt_pow_left hyx zero_le' hn] variable {p : ℕ} [hp : Fact p.Prime] (hp1 : Odd p) +include hp hp1 theorem pow_sub_pow (hyx : y < x) (hxy : p ∣ x - y) (hx : ¬p ∣ x) {n : ℕ} (hn : n ≠ 0) : padicValNat p (x ^ n - y ^ n) = padicValNat p (x - y) + padicValNat p n := by diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index 9bcddc91ed612..a1d88ada36d6b 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -269,6 +269,7 @@ noncomputable abbrev boundOfDiscBdd : ℝ≥0 := sqrt N * (2 : ℝ≥0) ^ rankOf variable {N} (hK : |discr K| ≤ N) +include hK in /-- If `|discr K| ≤ N` then the degree of `K` is at most `rankOfDiscrBdd`. -/ theorem rank_le_rankOfDiscrBdd : finrank ℚ K ≤ rankOfDiscrBdd N := by @@ -295,6 +296,7 @@ theorem rank_le_rankOfDiscrBdd : exact le_trans (by norm_num) (Nat.one_le_cast.mpr (Nat.one_le_iff_ne_zero.mpr h_nz)) · exact le_max_of_le_left h +include hK in /-- If `|discr K| ≤ N` then the Minkowski bound of `K` is less than `boundOfDiscrBdd`. -/ theorem minkowskiBound_lt_boundOfDiscBdd : minkowskiBound K ↑1 < boundOfDiscBdd N := by have : boundOfDiscBdd N - 1 < boundOfDiscBdd N := by @@ -311,6 +313,7 @@ theorem minkowskiBound_lt_boundOfDiscBdd : minkowskiBound K ↑1 < boundOfDiscBd · exact one_le_two · exact rank_le_rankOfDiscrBdd hK +include hK in theorem natDegree_le_rankOfDiscrBdd (a : 𝓞 K) (h : ℚ⟮(a : K)⟯ = ⊤) : natDegree (minpoly ℤ (a : K)) ≤ rankOfDiscrBdd N := by rw [Field.primitive_element_iff_minpoly_natDegree_eq, diff --git a/Mathlib/NumberTheory/NumberField/House.lean b/Mathlib/NumberTheory/NumberField/House.lean index 87d342f35542d..7cceb1f1e7e7c 100644 --- a/Mathlib/NumberTheory/NumberField/House.lean +++ b/Mathlib/NumberTheory/NumberField/House.lean @@ -128,6 +128,7 @@ private def asiegel : Matrix (α × (K →+* ℂ)) (β × (K →+* ℂ)) ℤ := variable (ha : a ≠ 0) +include ha in private theorem asiegel_ne_0 : asiegel K a ≠ 0 := by simp (config := { unfoldPartialApp := true }) only [asiegel, a'] simp only [ne_eq] @@ -151,6 +152,7 @@ variable {p q : ℕ} (h0p : 0 < p) (hpq : p < q) (x : β × (K →+* ℂ) → /-- `ξ` is the the product of `x (l, r)` and the `r`-th basis element of the newBasis of `K`. -/ private def ξ : β → 𝓞 K := fun l => ∑ r : K →+* ℂ, x (l, r) * (newBasis K r) +include hxl in private theorem ξ_ne_0 : ξ K x ≠ 0 := by intro H apply hxl @@ -166,6 +168,7 @@ private theorem lin_1 (l k r) : a k l * (newBasis K) r = variable [Fintype β] (cardβ : Fintype.card β = q) (hmulvec0 : asiegel K a *ᵥ x = 0) +include hxl hmulvec0 in private theorem ξ_mulVec_eq_0 : a *ᵥ ξ K x = 0 := by funext k; simp only [Pi.zero_apply]; rw [eq_comm] @@ -209,6 +212,7 @@ private theorem c₂_nonneg : 0 ≤ c₂ K := variable [Fintype α] (cardα : Fintype.card α = p) (Apos : 0 ≤ A) (hxbound : ‖x‖ ≤ (q * finrank ℚ K * ‖asiegel K a‖) ^ ((p : ℝ) / (q - p))) +include habs Apos in private theorem asiegel_remark : ‖asiegel K a‖ ≤ c₂ K * A := by rw [Matrix.norm_le_iff] · intro kr lu @@ -244,6 +248,7 @@ private theorem asiegel_remark : ‖asiegel K a‖ ≤ c₂ K * A := by /-- `c₁ K` is the product of `finrank ℚ K` and `c₂ K` and depends on `K`. -/ private def c₁ := finrank ℚ K * c₂ K +include habs Apos hxbound hpq in private theorem house_le_bound : ∀ l, house (ξ K x l).1 ≤ (c₁ K) * ((c₁ K * q * A)^((p : ℝ) / (q - p))) := by let h := finrank ℚ K @@ -283,6 +288,7 @@ private theorem house_le_bound : ∀ l, house (ξ K x l).1 ≤ (c₁ K) * (supOfBasis_nonneg _)) · rw [mul_comm (q : ℝ) (c₁ K)]; rfl +include hpq h0p cardα cardβ ha habs in /-- There exists a "small" non-zero algebraic integral solution of an non-trivial underdetermined system of linear equations with algebraic integer coefficients.-/ theorem exists_ne_zero_int_vec_house_le : diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index db8688cc1ec6d..33f35ed619614 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -204,6 +204,7 @@ open NumberField.mixedEmbedding NNReal variable (w₁ : InfinitePlace K) {B : ℕ} (hB : minkowskiBound K 1 < (convexBodyLTFactor K) * B) +include hB in /-- This result shows that there always exists a next term in the sequence. -/ theorem seq_next {x : 𝓞 K} (hx : x ≠ 0) : ∃ y : 𝓞 K, y ≠ 0 ∧ diff --git a/Mathlib/NumberTheory/Ostrowski.lean b/Mathlib/NumberTheory/Ostrowski.lean index bfdd99846ba8c..c41547159ecef 100644 --- a/Mathlib/NumberTheory/Ostrowski.lean +++ b/Mathlib/NumberTheory/Ostrowski.lean @@ -84,6 +84,7 @@ def mulRingNorm_padic (p : ℕ) [Fact p.Prime] : MulRingNorm ℚ := variable (hf_nontriv : f ≠ 1) (bdd : ∀ n : ℕ, f n ≤ 1) +include hf_nontriv bdd in /-- There exists a minimal positive integer with absolute value smaller than 1. -/ lemma exists_minimal_nat_zero_lt_mulRingNorm_lt_one : ∃ p : ℕ, (0 < f p ∧ f p < 1) ∧ ∀ m : ℕ, 0 < f m ∧ f m < 1 → p ≤ m := by @@ -104,6 +105,7 @@ lemma exists_minimal_nat_zero_lt_mulRingNorm_lt_one : ∃ p : ℕ, (0 < f p ∧ variable {p : ℕ} (hp0 : 0 < f p) (hp1 : f p < 1) (hmin : ∀ m : ℕ, 0 < f m ∧ f m < 1 → p ≤ m) +include hp0 hp1 hmin in /-- The minimal positive integer with absolute value smaller than 1 is a prime number.-/ lemma is_prime_of_minimal_nat_zero_lt_mulRingNorm_lt_one : p.Prime := by rw [← Nat.irreducible_iff_nat_prime] @@ -131,6 +133,7 @@ lemma is_prime_of_minimal_nat_zero_lt_mulRingNorm_lt_one : p.Prime := by open Real +include hp0 hp1 hmin bdd in /-- A natural number not divible by `p` has absolute value 1. -/ lemma mulRingNorm_eq_one_of_not_dvd {m : ℕ} (hpm : ¬ p ∣ m) : f m = 1 := by apply le_antisymm (bdd m) @@ -173,6 +176,7 @@ lemma mulRingNorm_eq_one_of_not_dvd {m : ℕ} (hpm : ¬ p ∣ m) : f m = 1 := by -- ## Step 4: f p = p ^ (- t) for some positive real t +include hp0 hp1 hmin in /-- The absolute value of `p` is `p ^ (-t)` for some positive real number `t`. -/ lemma exists_pos_mulRingNorm_eq_pow_neg : ∃ t : ℝ, 0 < t ∧ f p = p ^ (-t) := by have pprime := is_prime_of_minimal_nat_zero_lt_mulRingNorm_lt_one hp0 hp1 hmin @@ -183,6 +187,7 @@ lemma exists_pos_mulRingNorm_eq_pow_neg : ∃ t : ℝ, 0 < t ∧ f p = p ^ (-t) -- ## Non-archimedean case: end goal +include hf_nontriv bdd in /-- If `f` is bounded and not trivial, then it is equivalent to a p-adic absolute value. -/ theorem mulRingNorm_equiv_padic_of_bounded : ∃! p, ∃ (hp : Fact (p.Prime)), MulRingNorm.equiv f (mulRingNorm_padic p) := by diff --git a/Mathlib/NumberTheory/Padics/Hensel.lean b/Mathlib/NumberTheory/Padics/Hensel.lean index 80a8898200c2a..d32fe3c1f17a2 100644 --- a/Mathlib/NumberTheory/Padics/Hensel.lean +++ b/Mathlib/NumberTheory/Padics/Hensel.lean @@ -56,15 +56,16 @@ section -- Porting note: replaced `parameter` with `variable` variable {p : ℕ} [Fact p.Prime] {ncs : CauSeq ℤ_[p] norm} {F : Polynomial ℤ_[p]} {a : ℤ_[p]} (ncs_der_val : ∀ n, ‖F.derivative.eval (ncs n)‖ = ‖F.derivative.eval a‖) +private theorem ncs_tendsto_lim : + Tendsto (fun i => ‖F.derivative.eval (ncs i)‖) atTop (𝓝 ‖F.derivative.eval ncs.lim‖) := + Tendsto.comp (continuous_iff_continuousAt.1 continuous_norm _) (comp_tendsto_lim _) + +include ncs_der_val private theorem ncs_tendsto_const : Tendsto (fun i => ‖F.derivative.eval (ncs i)‖) atTop (𝓝 ‖F.derivative.eval a‖) := by convert @tendsto_const_nhds ℝ ℕ _ _ _; rw [ncs_der_val] -private theorem ncs_tendsto_lim : - Tendsto (fun i => ‖F.derivative.eval (ncs i)‖) atTop (𝓝 ‖F.derivative.eval ncs.lim‖) := - Tendsto.comp (continuous_iff_continuousAt.1 continuous_norm _) (comp_tendsto_lim _) - private theorem norm_deriv_eq : ‖F.derivative.eval ncs.lim‖ = ‖F.derivative.eval a‖ := tendsto_nhds_unique ncs_tendsto_lim (ncs_tendsto_const ncs_der_val) @@ -75,6 +76,7 @@ section -- Porting note: replaced `parameter` with `variable` variable {p : ℕ} [Fact p.Prime] {ncs : CauSeq ℤ_[p] norm} {F : Polynomial ℤ_[p]} (hnorm : Tendsto (fun i => ‖F.eval (ncs i)‖) atTop (𝓝 0)) +include hnorm private theorem tendsto_zero_of_norm_tendsto_zero : Tendsto (fun i => F.eval (ncs i)) atTop (𝓝 0) := tendsto_iff_norm_sub_tendsto_zero.2 (by simpa using hnorm) @@ -90,7 +92,6 @@ open Nat -- Porting note: replaced `parameter` with `variable` variable (p : ℕ) [Fact p.Prime] (F : Polynomial ℤ_[p]) (a : ℤ_[p]) - (hnorm : ‖F.eval a‖ < ‖F.derivative.eval a‖ ^ 2) (hnsol : F.eval a ≠ 0) -- Porting note: renamed this `def` and used a local notation to provide arguments automatically /-- `T` is an auxiliary value that is used to control the behavior of the polynomial `F`. -/ @@ -100,6 +101,16 @@ local notation "T" => @T_gen p _ F a variable {p F a} +private theorem T_def : T = ‖F.eval a‖ / ‖F.derivative.eval a‖ ^ 2 := by + simp [T_gen, ← PadicInt.norm_def] + +private theorem T_nonneg : 0 ≤ T := norm_nonneg _ + +private theorem T_pow_nonneg (n : ℕ) : 0 ≤ T ^ n := pow_nonneg T_nonneg _ + +variable (hnorm : ‖F.eval a‖ < ‖F.derivative.eval a‖ ^ 2) +include hnorm + private theorem deriv_sq_norm_pos : 0 < ‖F.derivative.eval a‖ ^ 2 := lt_of_le_of_lt (norm_nonneg _) hnorm @@ -115,17 +126,11 @@ private theorem deriv_norm_pos : 0 < ‖F.derivative.eval a‖ := private theorem deriv_ne_zero : F.derivative.eval a ≠ 0 := mt norm_eq_zero.2 (deriv_norm_ne_zero hnorm) -private theorem T_def : T = ‖F.eval a‖ / ‖F.derivative.eval a‖ ^ 2 := by - simp [T_gen, ← PadicInt.norm_def] private theorem T_lt_one : T < 1 := by have h := (div_lt_one (deriv_sq_norm_pos hnorm)).2 hnorm rw [T_def]; exact h -private theorem T_nonneg : 0 ≤ T := norm_nonneg _ - -private theorem T_pow_nonneg (n : ℕ) : 0 ≤ T ^ n := pow_nonneg T_nonneg _ - private theorem T_pow {n : ℕ} (hn : n ≠ 0) : T ^ n < 1 := pow_lt_one T_nonneg (T_lt_one hnorm) hn private theorem T_pow' (n : ℕ) : T ^ 2 ^ n < 1 := T_pow hnorm (pow_ne_zero _ two_ne_zero) @@ -190,7 +195,6 @@ private def calc_eval_z' {z z' z1 : ℤ_[p]} (hz' : z' = z - z1) {n} (hz : ih n exact ⟨q, by simpa only [sub_eq_add_neg, this, hz', add_neg_cancel, neg_sq, zero_add] using hq⟩ - private def calc_eval_z'_norm {z z' z1 : ℤ_[p]} {n} (hz : ih n z) {q} (heq : F.eval z' = q * z1 ^ 2) (h1 : ‖(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)‖ ≤ 1) (hzeq : z1 = ⟨_, h1⟩) : ‖F.eval z'‖ ≤ ‖F.derivative.eval a‖ ^ 2 * T ^ 2 ^ (n + 1) := by @@ -266,30 +270,6 @@ private theorem newton_seq_succ_dist (n : ℕ) : ((div_le_div_right (deriv_norm_pos hnorm)).2 (newton_seq_norm_le hnorm _)) _ = ‖F.derivative.eval a‖ * T ^ 2 ^ n := div_sq_cancel _ _ -private theorem T_pos : T > 0 := by - rw [T_def] - exact div_pos (norm_pos_iff.2 hnsol) (deriv_sq_norm_pos hnorm) - -private theorem newton_seq_succ_dist_weak (n : ℕ) : - ‖newton_seq (n + 2) - newton_seq (n + 1)‖ < ‖F.eval a‖ / ‖F.derivative.eval a‖ := - have : 2 ≤ 2 ^ (n + 1) := by - have := pow_le_pow_right (by norm_num : 1 ≤ 2) (Nat.le_add_left _ _ : 1 ≤ n + 1) - simpa using this - calc - ‖newton_seq (n + 2) - newton_seq (n + 1)‖ ≤ ‖F.derivative.eval a‖ * T ^ 2 ^ (n + 1) := - newton_seq_succ_dist hnorm _ - _ ≤ ‖F.derivative.eval a‖ * T ^ 2 := - (mul_le_mul_of_nonneg_left (pow_le_pow_of_le_one (norm_nonneg _) - (le_of_lt (T_lt_one hnorm)) this) (norm_nonneg _)) - _ < ‖F.derivative.eval a‖ * T ^ 1 := - (mul_lt_mul_of_pos_left (pow_lt_pow_right_of_lt_one (T_pos hnorm hnsol) - (T_lt_one hnorm) (by norm_num)) (deriv_norm_pos hnorm)) - _ = ‖F.eval a‖ / ‖F.derivative.eval a‖ := by - rw [T_gen, sq, pow_one, norm_div, ← mul_div_assoc, PadicInt.padic_norm_e_of_padicInt, - PadicInt.coe_mul, padicNormE.mul] - apply mul_div_mul_left - apply deriv_norm_ne_zero; assumption - private theorem newton_seq_dist_aux (n : ℕ) : ∀ k : ℕ, ‖newton_seq (n + k) - newton_seq n‖ ≤ ‖F.derivative.eval a‖ * T ^ 2 ^ n | 0 => by simp [T_pow_nonneg, mul_nonneg] @@ -312,31 +292,12 @@ private theorem newton_seq_dist_aux (n : ℕ) : mul_le_mul_of_nonneg_left (pow_le_pow_of_le_one (norm_nonneg _) (le_of_lt (T_lt_one hnorm)) this) (norm_nonneg _) - private theorem newton_seq_dist {n k : ℕ} (hnk : n ≤ k) : ‖newton_seq k - newton_seq n‖ ≤ ‖F.derivative.eval a‖ * T ^ 2 ^ n := by have hex : ∃ m, k = n + m := Nat.exists_eq_add_of_le hnk let ⟨_, hex'⟩ := hex rw [hex']; apply newton_seq_dist_aux -private theorem newton_seq_dist_to_a : - ∀ n : ℕ, 0 < n → ‖newton_seq n - a‖ = ‖F.eval a‖ / ‖F.derivative.eval a‖ - | 1, _h => by simp [sub_eq_add_neg, add_assoc, newton_seq_gen, newton_seq_aux, ih_n] - | k + 2, _h => - have hlt : ‖newton_seq (k + 2) - newton_seq (k + 1)‖ < ‖newton_seq (k + 1) - a‖ := by - rw [newton_seq_dist_to_a (k + 1) (succ_pos _)]; apply newton_seq_succ_dist_weak - assumption - have hne' : ‖newton_seq (k + 2) - newton_seq (k + 1)‖ ≠ ‖newton_seq (k + 1) - a‖ := ne_of_lt hlt - calc - ‖newton_seq (k + 2) - a‖ = - ‖newton_seq (k + 2) - newton_seq (k + 1) + (newton_seq (k + 1) - a)‖ := by - rw [← sub_add_sub_cancel] - _ = max ‖newton_seq (k + 2) - newton_seq (k + 1)‖ ‖newton_seq (k + 1) - a‖ := - (PadicInt.norm_add_eq_max_of_ne hne') - _ = ‖newton_seq (k + 1) - a‖ := max_eq_right_of_lt hlt - _ = ‖Polynomial.eval a F‖ / ‖Polynomial.eval a (Polynomial.derivative F)‖ := - newton_seq_dist_to_a (k + 1) (succ_pos _) - private theorem bound' : Tendsto (fun n : ℕ => ‖F.derivative.eval a‖ * T ^ 2 ^ n) atTop (𝓝 0) := by rw [← mul_zero ‖F.derivative.eval a‖] exact @@ -378,25 +339,70 @@ private theorem newton_seq_norm_tendsto_zero : Tendsto (fun i => ‖F.eval (newton_cau_seq hnorm i)‖) atTop (𝓝 0) := squeeze_zero (fun _ => norm_nonneg _) (newton_seq_norm_le hnorm) (bound'_sq hnorm) +private theorem newton_seq_dist_tendsto' : + Tendsto (fun n => ‖newton_cau_seq hnorm n - a‖) atTop (𝓝 ‖soln - a‖) := + (continuous_norm.tendsto _).comp ((newton_cau_seq hnorm).tendsto_limit.sub tendsto_const_nhds) + +private theorem eval_soln : F.eval soln = 0 := + limit_zero_of_norm_tendsto_zero (newton_seq_norm_tendsto_zero hnorm) + +variable (hnsol : F.eval a ≠ 0) +include hnsol + +private theorem T_pos : T > 0 := by + rw [T_def] + exact div_pos (norm_pos_iff.2 hnsol) (deriv_sq_norm_pos hnorm) + +private theorem newton_seq_succ_dist_weak (n : ℕ) : + ‖newton_seq (n + 2) - newton_seq (n + 1)‖ < ‖F.eval a‖ / ‖F.derivative.eval a‖ := + have : 2 ≤ 2 ^ (n + 1) := by + have := pow_le_pow_right (by norm_num : 1 ≤ 2) (Nat.le_add_left _ _ : 1 ≤ n + 1) + simpa using this + calc + ‖newton_seq (n + 2) - newton_seq (n + 1)‖ ≤ ‖F.derivative.eval a‖ * T ^ 2 ^ (n + 1) := + newton_seq_succ_dist hnorm _ + _ ≤ ‖F.derivative.eval a‖ * T ^ 2 := + (mul_le_mul_of_nonneg_left (pow_le_pow_of_le_one (norm_nonneg _) + (le_of_lt (T_lt_one hnorm)) this) (norm_nonneg _)) + _ < ‖F.derivative.eval a‖ * T ^ 1 := + (mul_lt_mul_of_pos_left (pow_lt_pow_right_of_lt_one (T_pos hnorm hnsol) + (T_lt_one hnorm) (by norm_num)) (deriv_norm_pos hnorm)) + _ = ‖F.eval a‖ / ‖F.derivative.eval a‖ := by + rw [T_gen, sq, pow_one, norm_div, ← mul_div_assoc, PadicInt.padic_norm_e_of_padicInt, + PadicInt.coe_mul, padicNormE.mul] + apply mul_div_mul_left + apply deriv_norm_ne_zero; assumption + +private theorem newton_seq_dist_to_a : + ∀ n : ℕ, 0 < n → ‖newton_seq n - a‖ = ‖F.eval a‖ / ‖F.derivative.eval a‖ + | 1, _h => by simp [sub_eq_add_neg, add_assoc, newton_seq_gen, newton_seq_aux, ih_n] + | k + 2, _h => + have hlt : ‖newton_seq (k + 2) - newton_seq (k + 1)‖ < ‖newton_seq (k + 1) - a‖ := by + rw [newton_seq_dist_to_a (k + 1) (succ_pos _)]; apply newton_seq_succ_dist_weak + assumption + have hne' : ‖newton_seq (k + 2) - newton_seq (k + 1)‖ ≠ ‖newton_seq (k + 1) - a‖ := ne_of_lt hlt + calc + ‖newton_seq (k + 2) - a‖ = + ‖newton_seq (k + 2) - newton_seq (k + 1) + (newton_seq (k + 1) - a)‖ := by + rw [← sub_add_sub_cancel] + _ = max ‖newton_seq (k + 2) - newton_seq (k + 1)‖ ‖newton_seq (k + 1) - a‖ := + (PadicInt.norm_add_eq_max_of_ne hne') + _ = ‖newton_seq (k + 1) - a‖ := max_eq_right_of_lt hlt + _ = ‖Polynomial.eval a F‖ / ‖Polynomial.eval a (Polynomial.derivative F)‖ := + newton_seq_dist_to_a (k + 1) (succ_pos _) + private theorem newton_seq_dist_tendsto : Tendsto (fun n => ‖newton_cau_seq hnorm n - a‖) atTop (𝓝 (‖F.eval a‖ / ‖F.derivative.eval a‖)) := tendsto_const_nhds.congr' (eventually_atTop.2 ⟨1, fun _ hx => (newton_seq_dist_to_a hnorm hnsol _ hx).symm⟩) -private theorem newton_seq_dist_tendsto' : - Tendsto (fun n => ‖newton_cau_seq hnorm n - a‖) atTop (𝓝 ‖soln - a‖) := - (continuous_norm.tendsto _).comp ((newton_cau_seq hnorm).tendsto_limit.sub tendsto_const_nhds) - private theorem soln_dist_to_a : ‖soln - a‖ = ‖F.eval a‖ / ‖F.derivative.eval a‖ := tendsto_nhds_unique (newton_seq_dist_tendsto' hnorm) (newton_seq_dist_tendsto hnorm hnsol) private theorem soln_dist_to_a_lt_deriv : ‖soln - a‖ < ‖F.derivative.eval a‖ := by rw [soln_dist_to_a, div_lt_iff (deriv_norm_pos _), ← sq] <;> assumption -private theorem eval_soln : F.eval soln = 0 := - limit_zero_of_norm_tendsto_zero (newton_seq_norm_tendsto_zero hnorm) - private theorem soln_unique (z : ℤ_[p]) (hev : F.eval z = 0) (hnlt : ‖z - a‖ < ‖F.derivative.eval a‖) : z = soln := have soln_dist : ‖z - soln‖ < ‖F.derivative.eval a‖ := @@ -460,6 +466,7 @@ private theorem a_soln_is_unique (ha : F.eval a = 0) (z' : ℤ_[p]) (hz' : F.eva eq_of_sub_eq_zero (by rw [← this]) variable (hnorm : ‖F.eval a‖ < ‖F.derivative.eval a‖ ^ 2) +include hnorm private theorem a_is_soln (ha : F.eval a = 0) : F.eval a = 0 ∧ diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index 841b79aa5d791..31ba8f2ff384a 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -451,8 +451,8 @@ section lift open CauSeq PadicSeq -variable {R : Type*} [NonAssocSemiring R] (f : ∀ k : ℕ, R →+* ZMod (p ^ k)) - (f_compat : ∀ (k1 k2) (hk : k1 ≤ k2), (ZMod.castHom (pow_dvd_pow p hk) _).comp (f k2) = f k1) +variable {R : Type*} [NonAssocSemiring R] {p : Nat} (f : ∀ k : ℕ, R →+* ZMod (p ^ k)) + /-- Given a family of ring homs `f : Π n : ℕ, R →+* ZMod (p ^ n)`, `nthHom f r` is an integer-valued sequence @@ -467,6 +467,12 @@ theorem nthHom_zero : nthHom f 0 = 0 := by rfl variable {f} +variable [hp_prime : Fact p.Prime] + +section +variable + (f_compat : ∀ (k1 k2) (hk : k1 ≤ k2), (ZMod.castHom (pow_dvd_pow p hk) _).comp (f k2) = f k1) +include f_compat theorem pow_dvd_nthHom_sub (r : R) (i j : ℕ) (h : i ≤ j) : (p : ℤ) ^ i ∣ nthHom f r j - nthHom f r i := by @@ -617,10 +623,12 @@ theorem lift_unique (g : R →+* ℤ_[p]) (hg : ∀ n, (toZModPow n).comp g = f rw [dist_eq_norm, norm_le_pow_iff_mem_span_pow, ← ker_toZModPow, RingHom.mem_ker, RingHom.map_sub, ← RingHom.comp_apply, ← RingHom.comp_apply, lift_spec, hg, sub_self] +end + @[simp] -theorem lift_self (z : ℤ_[p]) : @lift p _ ℤ_[p] _ toZModPow zmod_cast_comp_toZModPow z = z := by +theorem lift_self (z : ℤ_[p]) : lift zmod_cast_comp_toZModPow z = z := by show _ = RingHom.id _ z - rw [@lift_unique p _ ℤ_[p] _ _ zmod_cast_comp_toZModPow (RingHom.id ℤ_[p])] + rw [lift_unique zmod_cast_comp_toZModPow (RingHom.id ℤ_[p])] intro; rw [RingHom.comp_id] end lift diff --git a/Mathlib/NumberTheory/PellMatiyasevic.lean b/Mathlib/NumberTheory/PellMatiyasevic.lean index 806449a094861..bd603798cbd3d 100644 --- a/Mathlib/NumberTheory/PellMatiyasevic.lean +++ b/Mathlib/NumberTheory/PellMatiyasevic.lean @@ -151,6 +151,7 @@ def az (a : ℕ) : ℤ := end +include a1 in theorem asq_pos : 0 < a * a := le_trans (le_of_lt a1) (by have := @Nat.mul_le_mul_left 1 a a (le_of_lt a1); rwa [mul_one] at this) @@ -228,6 +229,7 @@ theorem xn_ge_a_pow : ∀ n : ℕ, a ^ n ≤ xn a1 n simp only [_root_.pow_succ, xn_succ] exact le_trans (Nat.mul_le_mul_right _ (xn_ge_a_pow n)) (Nat.le_add_right _ _) +include a1 in theorem n_lt_a_pow : ∀ n : ℕ, n < a ^ n | 0 => Nat.le_refl 1 | n + 1 => by diff --git a/Mathlib/NumberTheory/PythagoreanTriples.lean b/Mathlib/NumberTheory/PythagoreanTriples.lean index 79011cee3267e..a9e537bc9d8e7 100644 --- a/Mathlib/NumberTheory/PythagoreanTriples.lean +++ b/Mathlib/NumberTheory/PythagoreanTriples.lean @@ -54,17 +54,18 @@ theorem PythagoreanTriple.zero : PythagoreanTriple 0 0 0 := by namespace PythagoreanTriple -variable {x y z : ℤ} (h : PythagoreanTriple x y z) +variable {x y z : ℤ} -theorem eq : x * x + y * y = z * z := +theorem eq (h : PythagoreanTriple x y z) : x * x + y * y = z * z := h @[symm] -theorem symm : PythagoreanTriple y x z := by rwa [pythagoreanTriple_comm] +theorem symm (h : PythagoreanTriple x y z) : PythagoreanTriple y x z := by + rwa [pythagoreanTriple_comm] /-- A triple is still a triple if you multiply `x`, `y` and `z` by a constant `k`. -/ -theorem mul (k : ℤ) : PythagoreanTriple (k * x) (k * y) (k * z) := +theorem mul (h : PythagoreanTriple x y z) (k : ℤ) : PythagoreanTriple (k * x) (k * y) (k * z) := calc k * x * (k * x) + k * y * (k * y) = k ^ 2 * (x * x + y * y) := by ring _ = k ^ 2 * (z * z) := by rw [h.eq] @@ -102,6 +103,9 @@ def IsPrimitiveClassified (_ : PythagoreanTriple x y z) := (x = m ^ 2 - n ^ 2 ∧ y = 2 * m * n ∨ x = 2 * m * n ∧ y = m ^ 2 - n ^ 2) ∧ Int.gcd m n = 1 ∧ (m % 2 = 0 ∧ n % 2 = 1 ∨ m % 2 = 1 ∧ n % 2 = 0) +variable (h : PythagoreanTriple x y z) +include h + theorem mul_isClassified (k : ℤ) (hc : h.IsClassified) : (h.mul k).IsClassified := by obtain ⟨l, m, n, ⟨⟨rfl, rfl⟩ | ⟨rfl, rfl⟩, co⟩⟩ := hc · use k * l, m, n @@ -569,7 +573,6 @@ theorem coprime_classification : (x = m ^ 2 - n ^ 2 ∧ y = 2 * m * n ∨ x = 2 * m * n ∧ y = m ^ 2 - n ^ 2) ∧ (z = m ^ 2 + n ^ 2 ∨ z = -(m ^ 2 + n ^ 2)) ∧ Int.gcd m n = 1 ∧ (m % 2 = 0 ∧ n % 2 = 1 ∨ m % 2 = 1 ∧ n % 2 = 0) := by - clear h -- Porting note: don't want this variable, but can't use `include` / `omit` constructor · intro h obtain ⟨m, n, H⟩ := h.left.isPrimitiveClassified_of_coprime h.right @@ -650,7 +653,6 @@ theorem classification : (x = k * (m ^ 2 - n ^ 2) ∧ y = k * (2 * m * n) ∨ x = k * (2 * m * n) ∧ y = k * (m ^ 2 - n ^ 2)) ∧ (z = k * (m ^ 2 + n ^ 2) ∨ z = -k * (m ^ 2 + n ^ 2)) := by - clear h constructor · intro h obtain ⟨k, m, n, H⟩ := h.classified diff --git a/Mathlib/Order/Atoms.lean b/Mathlib/Order/Atoms.lean index 3595346bda833..bcafc0ff30746 100644 --- a/Mathlib/Order/Atoms.lean +++ b/Mathlib/Order/Atoms.lean @@ -636,6 +636,7 @@ namespace IsSimpleOrder section Preorder variable [Preorder α] [BoundedOrder α] [IsSimpleOrder α] {a b : α} (h : a < b) +include h theorem eq_bot_of_lt : a = ⊥ := (IsSimpleOrder.eq_bot_or_eq_top _).resolve_right h.ne_top @@ -971,6 +972,7 @@ variable [Lattice α] [BoundedOrder α] [IsModularLattice α] namespace IsCompl variable {a b : α} (hc : IsCompl a b) +include hc theorem isAtom_iff_isCoatom : IsAtom a ↔ IsCoatom b := Set.isSimpleOrder_Iic_iff_isAtom.symm.trans <| diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index 66094dcfca16e..e1b7a921c750f 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -1058,13 +1058,15 @@ namespace Monotone variable [Preorder α] [Preorder β] {f : α → β} (Hf : Monotone f) {a : α} {s : Set α} +include Hf + theorem mem_upperBounds_image (Ha : a ∈ upperBounds s) : f a ∈ upperBounds (f '' s) := forall_mem_image.2 fun _ H => Hf (Ha H) theorem mem_lowerBounds_image (Ha : a ∈ lowerBounds s) : f a ∈ lowerBounds (f '' s) := forall_mem_image.2 fun _ H => Hf (Ha H) -theorem image_upperBounds_subset_upperBounds_image (Hf : Monotone f) : +theorem image_upperBounds_subset_upperBounds_image : f '' upperBounds s ⊆ upperBounds (f '' s) := by rintro _ ⟨a, ha, rfl⟩ exact Hf.mem_upperBounds_image ha @@ -1096,6 +1098,8 @@ namespace Antitone variable [Preorder α] [Preorder β] {f : α → β} (hf : Antitone f) {a : α} {s : Set α} +include hf + theorem mem_upperBounds_image : a ∈ lowerBounds s → f a ∈ upperBounds (f '' s) := hf.dual_right.mem_lowerBounds_image @@ -1135,6 +1139,8 @@ section MonotoneMonotone variable (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Monotone (f a)) +include h₀ h₁ + theorem mem_upperBounds_image2 (ha : a ∈ upperBounds s) (hb : b ∈ upperBounds t) : f a b ∈ upperBounds (image2 f s t) := forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy @@ -1152,13 +1158,13 @@ theorem image2_lowerBounds_lowerBounds_subset : image2_subset_iff.2 fun _ ha _ hb ↦ mem_lowerBounds_image2 h₀ h₁ ha hb /-- See also `Monotone.map_bddAbove`. -/ -protected theorem BddAbove.image2 (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Monotone (f a)) : +protected theorem BddAbove.image2 : BddAbove s → BddAbove t → BddAbove (image2 f s t) := by rintro ⟨a, ha⟩ ⟨b, hb⟩ exact ⟨f a b, mem_upperBounds_image2 h₀ h₁ ha hb⟩ /-- See also `Monotone.map_bddBelow`. -/ -protected theorem BddBelow.image2 (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Monotone (f a)) : +protected theorem BddBelow.image2 : BddBelow s → BddBelow t → BddBelow (image2 f s t) := by rintro ⟨a, ha⟩ ⟨b, hb⟩ exact ⟨f a b, mem_lowerBounds_image2 h₀ h₁ ha hb⟩ @@ -1177,6 +1183,8 @@ section MonotoneAntitone variable (h₀ : ∀ b, Monotone (swap f b)) (h₁ : ∀ a, Antitone (f a)) +include h₀ h₁ + theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_lowerBounds (ha : a ∈ upperBounds s) (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy @@ -1221,6 +1229,8 @@ section AntitoneAntitone variable (h₀ : ∀ b, Antitone (swap f b)) (h₁ : ∀ a, Antitone (f a)) +include h₀ h₁ + theorem mem_upperBounds_image2_of_mem_lowerBounds (ha : a ∈ lowerBounds s) (hb : b ∈ lowerBounds t) : f a b ∈ upperBounds (image2 f s t) := forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy @@ -1261,6 +1271,8 @@ section AntitoneMonotone variable (h₀ : ∀ b, Antitone (swap f b)) (h₁ : ∀ a, Monotone (f a)) +include h₀ h₁ + theorem mem_upperBounds_image2_of_mem_upperBounds_of_mem_upperBounds (ha : a ∈ lowerBounds s) (hb : b ∈ upperBounds t) : f a b ∈ upperBounds (image2 f s t) := forall_image2_iff.2 fun _ hx _ hy => (h₀ _ <| ha hx).trans <| h₁ _ <| hb hy diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index 9ab2e4ae32457..783720c3aa7e9 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -1356,6 +1356,7 @@ end WithTop namespace Monotone variable [Preorder α] [ConditionallyCompleteLattice β] {f : α → β} (h_mono : Monotone f) +include h_mono /-! A monotone function into a conditionally complete lattice preserves the ordering properties of `sSup` and `sInf`. -/ diff --git a/Mathlib/Order/Extension/Well.lean b/Mathlib/Order/Extension/Well.lean index 54045b44754b4..6c7088100dd59 100644 --- a/Mathlib/Order/Extension/Well.lean +++ b/Mathlib/Order/Extension/Well.lean @@ -60,6 +60,7 @@ instance wellOrderExtension.isWellFounded_lt : IsWellFounded α hwf.wellOrderExt ⟨InvImage.wf (fun a : α => (hwf.rank a, embeddingToCardinal a)) <| Ordinal.lt_wf.prod_lex Cardinal.lt_wf⟩ +include hwf in /-- Any well-founded relation can be extended to a well-ordering on that type. -/ theorem exists_well_order_ge : ∃ s, r ≤ s ∧ IsWellOrder α s := ⟨hwf.wellOrderExtension.lt, fun _ _ h => Prod.Lex.left _ _ (hwf.rank_lt_of_rel h), ⟨⟩⟩ diff --git a/Mathlib/Order/GaloisConnection.lean b/Mathlib/Order/GaloisConnection.lean index 5c28a2b3b84da..95ba690deb3fb 100644 --- a/Mathlib/Order/GaloisConnection.lean +++ b/Mathlib/Order/GaloisConnection.lean @@ -67,7 +67,7 @@ namespace GaloisConnection section -variable [Preorder α] [Preorder β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) +variable [Preorder α] [Preorder β] {l : α → β} {u : β → α} theorem monotone_intro (hu : Monotone u) (hl : Monotone l) (hul : ∀ a, a ≤ u (l a)) (hlu : ∀ a, l (u a) ≤ a) : GaloisConnection l u := fun _ _ => @@ -78,6 +78,9 @@ protected theorem dual {l : α → β} {u : β → α} (gc : GaloisConnection l (OrderDual.toDual ∘ l ∘ OrderDual.ofDual) := fun a b => (gc b a).symm +variable (gc : GaloisConnection l u) +include gc + theorem le_iff_le {a : α} {b : β} : l a ≤ b ↔ a ≤ u b := gc _ _ @@ -98,11 +101,11 @@ theorem monotone_u : Monotone u := fun a _ H => gc.le_u ((gc.l_u_le a).trans H) theorem monotone_l : Monotone l := gc.dual.monotone_u.dual -theorem upperBounds_l_image (gc : GaloisConnection l u) (s : Set α) : +theorem upperBounds_l_image (s : Set α) : upperBounds (l '' s) = u ⁻¹' upperBounds s := Set.ext fun b => by simp [upperBounds, gc _ _] -theorem lowerBounds_u_image (gc : GaloisConnection l u) (s : Set β) : +theorem lowerBounds_u_image (s : Set β) : lowerBounds (u '' s) = l ⁻¹' lowerBounds s := gc.dual.upperBounds_l_image s @@ -146,6 +149,7 @@ end section PartialOrder variable [PartialOrder α] [Preorder β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) +include gc theorem u_l_u_eq_u (b : β) : u (l (u b)) = u b := (gc.monotone_u (gc.l_u_le _)).antisymm (gc.le_u_l _) @@ -161,7 +165,7 @@ theorem u_unique {l' : α → β} {u' : β → α} (gc' : GaloisConnection l' u' theorem exists_eq_u (a : α) : (∃ b : β, a = u b) ↔ a = u (l a) := ⟨fun ⟨_, hS⟩ => hS.symm ▸ (gc.u_l_u_eq_u _).symm, fun HI => ⟨_, HI⟩⟩ -theorem u_eq (gc : GaloisConnection l u) {z : α} {y : β} : u y = z ↔ ∀ x, x ≤ z ↔ l x ≤ y := by +theorem u_eq {z : α} {y : β} : u y = z ↔ ∀ x, x ≤ z ↔ l x ≤ y := by constructor · rintro rfl x exact (gc x y).symm @@ -173,6 +177,7 @@ end PartialOrder section PartialOrder variable [Preorder α] [PartialOrder β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) +include gc theorem l_u_l_eq_l (a : α) : l (u (l a)) = l a := gc.dual.u_l_u_eq_u _ @@ -215,18 +220,18 @@ end OrderBot section SemilatticeSup -variable [SemilatticeSup α] [SemilatticeSup β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) +variable [SemilatticeSup α] [SemilatticeSup β] {l : α → β} {u : β → α} -theorem l_sup : l (a₁ ⊔ a₂) = l a₁ ⊔ l a₂ := +theorem l_sup (gc : GaloisConnection l u) : l (a₁ ⊔ a₂) = l a₁ ⊔ l a₂ := (gc.isLUB_l_image isLUB_pair).unique <| by simp only [image_pair, isLUB_pair] end SemilatticeSup section SemilatticeInf -variable [SemilatticeInf α] [SemilatticeInf β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) +variable [SemilatticeInf α] [SemilatticeInf β] {l : α → β} {u : β → α} -theorem u_inf : u (b₁ ⊓ b₂) = u b₁ ⊓ u b₂ := gc.dual.l_sup +theorem u_inf (gc : GaloisConnection l u) : u (b₁ ⊓ b₂) = u b₁ ⊓ u b₂ := gc.dual.l_sup end SemilatticeInf @@ -245,6 +250,7 @@ theorem l_iSup₂ (gc : GaloisConnection l u) {f : ∀ i, κ i → α} : simp_rw [gc.l_iSup] variable (gc : GaloisConnection l u) +include gc theorem u_iInf {f : ι → β} : u (iInf f) = ⨅ i, u (f i) := gc.dual.l_iSup @@ -252,7 +258,7 @@ theorem u_iInf {f : ι → β} : u (iInf f) = ⨅ i, u (f i) := theorem u_iInf₂ {f : ∀ i, κ i → β} : u (⨅ (i) (j), f i j) = ⨅ (i) (j), u (f i j) := gc.dual.l_iSup₂ -theorem l_sSup (gc : GaloisConnection l u) {s : Set α} : l (sSup s) = ⨆ a ∈ s, l a := by +theorem l_sSup {s : Set α} : l (sSup s) = ⨆ a ∈ s, l a := by simp only [sSup_eq_iSup, gc.l_iSup] theorem u_sInf {s : Set β} : u (sInf s) = ⨅ a ∈ s, u a := @@ -262,9 +268,9 @@ end CompleteLattice section LinearOrder -variable [LinearOrder α] [LinearOrder β] {l : α → β} {u : β → α} (gc : GaloisConnection l u) +variable [LinearOrder α] [LinearOrder β] {l : α → β} {u : β → α} -theorem lt_iff_lt {a : α} {b : β} : b < l a ↔ u b < a := +theorem lt_iff_lt (gc : GaloisConnection l u) {a : α} {b : β} : b < l a ↔ u b < a := lt_iff_lt_of_le_iff_le (gc a b) end LinearOrder diff --git a/Mathlib/Order/Hom/Lattice.lean b/Mathlib/Order/Hom/Lattice.lean index a2721e9126d8e..a6513d866da90 100644 --- a/Mathlib/Order/Hom/Lattice.lean +++ b/Mathlib/Order/Hom/Lattice.lean @@ -1017,7 +1017,8 @@ end LatticeHom namespace OrderHomClass -variable (α β) [LinearOrder α] [Lattice β] [FunLike F α β] [OrderHomClass F α β] +variable (α β) +variable [LinearOrder α] [Lattice β] [OrderHomClass F α β] /-- An order homomorphism from a linear order is a lattice homomorphism. -/ -- Porting note: made it an `instance` because we're no longer afraid of loops diff --git a/Mathlib/Order/Interval/Finset/Fin.lean b/Mathlib/Order/Interval/Finset/Fin.lean index 96e602971b461..f0a6f5f7b6aef 100644 --- a/Mathlib/Order/Interval/Finset/Fin.lean +++ b/Mathlib/Order/Interval/Finset/Fin.lean @@ -143,9 +143,6 @@ theorem Iio_eq_finset_subtype : Iio b = (Iio (b : ℕ)).fin n := @[simp] theorem map_valEmbedding_Ici : (Ici a).map Fin.valEmbedding = Icc ↑a (n - 1) := by --- Porting note: without `clear b` Lean includes `b` in the statement (because the `rfl`) in the --- `rintro` below acts on it. - clear b ext x simp only [exists_prop, Embedding.coe_subtype, mem_Ici, mem_map, mem_Icc] constructor @@ -157,9 +154,6 @@ theorem map_valEmbedding_Ici : (Ici a).map Fin.valEmbedding = Icc ↑a (n - 1) : @[simp] theorem map_valEmbedding_Ioi : (Ioi a).map Fin.valEmbedding = Ioc ↑a (n - 1) := by --- Porting note: without `clear b` Lean includes `b` in the statement (because the `rfl`) in the --- `rintro` below acts on it. - clear b ext x simp only [exists_prop, Embedding.coe_subtype, mem_Ioi, mem_map, mem_Ioc] constructor @@ -179,8 +173,6 @@ theorem map_valEmbedding_Iio : (Iio b).map Fin.valEmbedding = Iio ↑b := by @[simp] theorem card_Ici : (Ici a).card = n - a := by --- Porting note: without `clear b` Lean includes `b` in the statement. - clear b cases n with | zero => exact Fin.elim0 a | succ => diff --git a/Mathlib/Order/Interval/Set/Infinite.lean b/Mathlib/Order/Interval/Set/Infinite.lean index 5138978cb5210..4390b048091c8 100644 --- a/Mathlib/Order/Interval/Set/Infinite.lean +++ b/Mathlib/Order/Interval/Set/Infinite.lean @@ -32,6 +32,7 @@ namespace Set section DenselyOrdered variable [DenselyOrdered α] {a b : α} (h : a < b) +include h theorem Ioo.infinite : Infinite (Ioo a b) := @NoMaxOrder.infinite _ _ (nonempty_Ioo_subtype h) _ diff --git a/Mathlib/Order/Interval/Set/OrdConnectedComponent.lean b/Mathlib/Order/Interval/Set/OrdConnectedComponent.lean index f23da9a3652e0..54ddf55dae58b 100644 --- a/Mathlib/Order/Interval/Set/OrdConnectedComponent.lean +++ b/Mathlib/Order/Interval/Set/OrdConnectedComponent.lean @@ -167,7 +167,6 @@ def ordT5Nhd (s t : Set α) : Set α := ⋃ x ∈ s, ordConnectedComponent (tᶜ ∩ (ordConnectedSection <| ordSeparatingSet s t)ᶜ) x theorem disjoint_ordT5Nhd : Disjoint (ordT5Nhd s t) (ordT5Nhd t s) := by - clear x y z rw [disjoint_iff_inf_le] rintro x ⟨hx₁, hx₂⟩ rcases mem_iUnion₂.1 hx₁ with ⟨a, has, ha⟩ diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index 787fe8c47cab0..e3a967e7b0c55 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -274,9 +274,11 @@ variable {s : Set α} (hs : SetIndependent s) theorem setIndependent_empty : SetIndependent (∅ : Set α) := fun x hx => (Set.not_mem_empty x hx).elim +include hs in theorem SetIndependent.mono {t : Set α} (hst : t ⊆ s) : SetIndependent t := fun _ ha => (hs (hst ha)).mono_right (sSup_le_sSup (diff_subset_diff_left hst)) +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 => disjoint_sSup_right (hs hx) ((mem_diff y).mpr ⟨hy, h.symm⟩) @@ -295,6 +297,7 @@ theorem setIndependent_pair {a b : α} (hab : a ≠ b) : · convert h.symm using 1 simp [hab, sSup_singleton] +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) : @@ -343,6 +346,7 @@ theorem independent_empty (t : Empty → α) : Independent t := theorem independent_pempty (t : PEmpty → α) : Independent t := nofun +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 => disjoint_sSup_right (ht x) ⟨y, iSup_pos h.symm⟩ diff --git a/Mathlib/Order/WellFounded.lean b/Mathlib/Order/WellFounded.lean index f49074d382109..107a3a70cbacc 100644 --- a/Mathlib/Order/WellFounded.lean +++ b/Mathlib/Order/WellFounded.lean @@ -140,8 +140,8 @@ private theorem eq_strictMono_iff_eq_range_aux {f g : β → γ} (hf : StrictMon · rw [← hc] exact hf.monotone hbc -theorem eq_strictMono_iff_eq_range (h : WellFounded ((· < ·) : β → β → Prop)) {f g : β → γ} - (hf : StrictMono f) (hg : StrictMono g) : +theorem eq_strictMono_iff_eq_range (h : WellFounded ((· < ·) : β → β → Prop)) + {f g : β → γ} (hf : StrictMono f) (hg : StrictMono g) : Set.range f = Set.range g ↔ f = g := ⟨fun hfg => by funext a @@ -151,8 +151,8 @@ theorem eq_strictMono_iff_eq_range (h : WellFounded ((· < ·) : β → β → P (eq_strictMono_iff_eq_range_aux hg hf hfg.symm fun a hab => (H a hab).symm), congr_arg _⟩ -theorem self_le_of_strictMono (h : WellFounded ((· < ·) : β → β → Prop)) {f : β → β} - (hf : StrictMono f) : ∀ n, n ≤ f n := by +theorem self_le_of_strictMono (h : WellFounded ((· < ·) : β → β → Prop)) + {f : β → β} (hf : StrictMono f) : ∀ n, n ≤ f n := by by_contra! h₁ have h₂ := h.min_mem _ h₁ exact h.not_lt_min _ h₁ (hf h₂) h₂ diff --git a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean index 54e2234837003..fce06642b56a2 100644 --- a/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean +++ b/Mathlib/Probability/Kernel/Disintegration/MeasurableStieltjes.lean @@ -48,7 +48,7 @@ open scoped NNReal ENNReal MeasureTheory Topology namespace ProbabilityTheory -variable {α β ι : Type*} [MeasurableSpace α] +variable {α β ι : Type*} section IsMeasurableRatCDF @@ -68,7 +68,7 @@ lemma isRatStieltjesPoint_unit_prod_iff (f : α → ℚ → ℝ) (a : α) : constructor <;> exact fun h ↦ ⟨h.mono, h.tendsto_atTop_one, h.tendsto_atBot_zero, h.iInf_rat_gt_eq⟩ -lemma measurableSet_isRatStieltjesPoint (hf : Measurable f) : +lemma measurableSet_isRatStieltjesPoint [MeasurableSpace α] (hf : Measurable f) : MeasurableSet {a | IsRatStieltjesPoint f a} := by have h1 : MeasurableSet {a | Monotone (f a)} := by change MeasurableSet {a | ∀ q r (_ : q ≤ r), f a q ≤ f a r} @@ -106,6 +106,7 @@ lemma IsRatStieltjesPoint.ite {f g : α → ℚ → ℝ} {a : α} (p : α → Pr split_ifs with h; exacts [(hf h).tendsto_atBot_zero, (hg h).tendsto_atBot_zero] iInf_rat_gt_eq := by split_ifs with h; exacts [(hf h).iInf_rat_gt_eq, (hg h).iInf_rat_gt_eq] +variable [MeasurableSpace α] /-- A function `f : α → ℚ → ℝ` is a (kernel) rational cumulative distribution function if it is measurable in the first argument and if `f a` satisfies a list of properties for all `a : α`: @@ -232,6 +233,13 @@ lemma toRatCDF_of_isRatStieltjesPoint {a : α} (h : IsRatStieltjesPoint f a) (q toRatCDF f a q = f a q := by rw [toRatCDF, if_pos h] +lemma toRatCDF_unit_prod (a : α) : + toRatCDF (fun (p : Unit × α) ↦ f p.2) ((), a) = toRatCDF f a := by + unfold toRatCDF + rw [isRatStieltjesPoint_unit_prod_iff] + +variable [MeasurableSpace α] + lemma measurable_toRatCDF (hf : Measurable f) : Measurable (toRatCDF f) := Measurable.ite (measurableSet_isRatStieltjesPoint hf) hf measurable_const @@ -243,11 +251,6 @@ lemma isMeasurableRatCDF_toRatCDF (hf : Measurable f) : (fun _ ↦ isRatStieltjesPoint_defaultRatCDF a) measurable := measurable_toRatCDF hf -lemma toRatCDF_unit_prod (a : α) : - toRatCDF (fun (p : Unit × α) ↦ f p.2) ((), a) = toRatCDF f a := by - unfold toRatCDF - rw [isRatStieltjesPoint_unit_prod_iff] - end ToRatCDF section IsMeasurableRatCDF.stieltjesFunction @@ -268,7 +271,8 @@ lemma IsMeasurableRatCDF.stieltjesFunctionAux_unit_prod {f : α → ℚ → ℝ} = IsMeasurableRatCDF.stieltjesFunctionAux f a := by simp_rw [IsMeasurableRatCDF.stieltjesFunctionAux_def'] -variable {f : α → ℚ → ℝ} (hf : IsMeasurableRatCDF f) +variable {f : α → ℚ → ℝ} [MeasurableSpace α] (hf : IsMeasurableRatCDF f) +include hf lemma IsMeasurableRatCDF.stieltjesFunctionAux_eq (a : α) (r : ℚ) : IsMeasurableRatCDF.stieltjesFunctionAux f a r = f a r := by @@ -435,7 +439,7 @@ end IsMeasurableRatCDF.stieltjesFunction section stieltjesOfMeasurableRat -variable {f : α → ℚ → ℝ} +variable {f : α → ℚ → ℝ} [MeasurableSpace α] /-- Turn a measurable function `f : α → ℚ → ℝ` into a measurable function `α → StieltjesFunction`. Composition of `toRatCDF` and `IsMeasurableRatCDF.stieltjesFunction`. -/ diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 04d5c444a9997..2f8ff3eb320c1 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -371,6 +371,7 @@ variable (X : ℕ → Ω → ℝ) (hint : Integrable (X 0)) (hindep : Pairwise fun i j => IndepFun (X i) (X j)) (hident : ∀ i, IdentDistrib (X i) (X 0)) (hnonneg : ∀ i ω, 0 ≤ X i ω) +include hint hindep hident hnonneg in /-- The truncation of `Xᵢ` up to `i` satisfies the strong law of large numbers (with respect to the truncated expectation) along the sequence `c^n`, for any `c > 1`, up to a given `ε > 0`. This follows from a variance control. -/ @@ -476,6 +477,7 @@ theorem strong_law_aux1 {c : ℝ} (c_one : 1 < c) {ε : ℝ} (εpos : 0 < ε) : simp_rw [S, not_le, mul_comm, sum_apply] at hω convert hω; simp only [sum_apply] +include hint hindep hident hnonneg in /- The truncation of `Xᵢ` up to `i` satisfies the strong law of large numbers (with respect to the truncated expectation) along the sequence `c^n`, for any `c > 1`. This follows from `strong_law_aux1` by varying `ε`. -/ @@ -493,6 +495,7 @@ theorem strong_law_aux2 {c : ℝ} (c_one : 1 < c) : simp only [Real.norm_eq_abs, abs_abs, Nat.abs_cast] exact hn.le.trans (mul_le_mul_of_nonneg_right hi.le (Nat.cast_nonneg _)) +include hint hident in /-- The expectation of the truncated version of `Xᵢ` behaves asymptotically like the whole expectation. This follows from convergence and Cesàro averaging. -/ theorem strong_law_aux3 : @@ -507,6 +510,7 @@ theorem strong_law_aux3 : rw [integral_finset_sum _ fun i _ => ?_] exact ((hident i).symm.integrable_snd hint).1.integrable_truncation +include hint hindep hident hnonneg in /- The truncation of `Xᵢ` up to `i` satisfies the strong law of large numbers (with respect to the original expectation) along the sequence `c^n`, for any `c > 1`. This follows from the version from the truncated expectation, and the @@ -521,6 +525,7 @@ theorem strong_law_aux4 {c : ℝ} (c_one : 1 < c) : ext1 n simp +include hint hident hnonneg in /-- The truncated and non-truncated versions of `Xᵢ` have the same asymptotic behavior, as they almost surely coincide at all but finitely many steps. This follows from a probability computation and Borel-Cantelli. -/ @@ -548,6 +553,7 @@ theorem strong_law_aux5 : ext n rw [sum_sub_distrib] +include hint hindep hident hnonneg in /- `Xᵢ` satisfies the strong law of large numbers along the sequence `c^n`, for any `c > 1`. This follows from the version for the truncated `Xᵢ`, and the fact that `Xᵢ` and its truncated version have the same asymptotic behavior. -/ @@ -570,6 +576,7 @@ theorem strong_law_aux6 {c : ℝ} (c_one : 1 < c) : convert L.mul_isBigO (isBigO_refl (fun n : ℕ => (⌊c ^ n⌋₊ : ℝ)⁻¹) atTop) using 1 <;> (ext1 n; field_simp [(H n).ne']) +include hint hindep hident hnonneg in /-- `Xᵢ` satisfies the strong law of large numbers along all integers. This follows from the corresponding fact along the sequences `c^n`, and the fact that any integer can be sandwiched between `c^n` and `c^(n+1)` with comparably small error if `c` is close enough to `1` @@ -621,7 +628,7 @@ section StrongLawVectorSpace variable {Ω : Type*} [MeasureSpace Ω] [IsProbabilityMeasure (ℙ : Measure Ω)] {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [CompleteSpace E] - [MeasurableSpace E] [BorelSpace E] + [MeasurableSpace E] open Set TopologicalSpace @@ -672,6 +679,8 @@ lemma strong_law_ae_simpleFunc_comp (X : ℕ → Ω → E) (h' : Measurable (X 0 · exact (φ.comp (X 0) h').integrable_of_isFiniteMeasure · exact (ψ.comp (X 0) h').integrable_of_isFiniteMeasure +variable [BorelSpace E] + /-- Preliminary lemma for the strong law of large numbers for vector-valued random variables, assuming measurability in addition to integrability. This is weakened to ae measurability in the full version `ProbabilityTheory.strong_law_ae`. -/ diff --git a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean index 7d07464cc61ca..afc24e101812b 100644 --- a/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean +++ b/Mathlib/RingTheory/AdicCompletion/AsTensorProduct.lean @@ -170,7 +170,6 @@ section Noetherian variable {R : Type u} [CommRing R] (I : Ideal R) variable (M : Type u) [AddCommGroup M] [Module R M] -variable [IsNoetherianRing R] /-! @@ -203,7 +202,7 @@ open CategoryTheory section -variable {ι : Type} [Fintype ι] [DecidableEq ι] (f : (ι → R) →ₗ[R] M) (hf : Function.Surjective f) +variable {ι : Type} (f : (ι → R) →ₗ[R] M) /- The first horizontal arrow in the top row. -/ private @@ -216,21 +215,26 @@ private def lTensorf : AdicCompletion I R ⊗[R] (ι → R) →ₗ[AdicCompletion I R] AdicCompletion I R ⊗[R] M := AlgebraTensorModule.map LinearMap.id f -private lemma if_exact : Function.Exact (LinearMap.ker f).subtype f := fun x ↦ - ⟨fun hx ↦ ⟨⟨x, hx⟩, rfl⟩, by rintro ⟨x, rfl⟩; exact x.property⟩ +variable (hf : Function.Surjective f) + +section +include hf private lemma tens_exact : Function.Exact (lTensorKerIncl I M f) (lTensorf I M f) := - lTensor_exact (AdicCompletion I R) (if_exact M f) hf + lTensor_exact (AdicCompletion I R) (f.exact_subtype_ker_map) hf private lemma tens_surj : Function.Surjective (lTensorf I M f) := LinearMap.lTensor_surjective (AdicCompletion I R) hf -private lemma adic_exact : Function.Exact (map I (LinearMap.ker f).subtype) (map I f) := - map_exact (Submodule.injective_subtype _) (if_exact M f) hf +private lemma adic_exact [IsNoetherianRing R] [Fintype ι] : + Function.Exact (map I (LinearMap.ker f).subtype) (map I f) := + map_exact (Submodule.injective_subtype _) (f.exact_subtype_ker_map) hf private lemma adic_surj : Function.Surjective (map I f) := map_surjective I hf +end + /- Instance to speed up instance inference. -/ private instance : AddCommGroup (AdicCompletion I R ⊗[R] (LinearMap.ker f)) := inferInstance @@ -249,6 +253,8 @@ private def secondRow : ComposableArrows (ModuleCat (AdicCompletion I R)) 4 := (ModuleCat.ofHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) (ModuleCat.ofHom (0 : _ →ₗ[AdicCompletion I R] PUnit)) +include hf + private lemma firstRow_exact : (firstRow I M f).Exact where zero k _ := match k with | 0 => (tens_exact I M f hf).linearMap_comp_eq_zero @@ -261,7 +267,7 @@ private lemma firstRow_exact : (firstRow I M f).Exact where | 1 => intro x _; exact (tens_surj I M f hf) x | 2 => intro _ _; exact ⟨0, rfl⟩ -private lemma secondRow_exact : (secondRow I M f).Exact where +private lemma secondRow_exact [Fintype ι] [IsNoetherianRing R] : (secondRow I M f).Exact where zero k _ := match k with | 0 => (adic_exact I M f hf).linearMap_comp_eq_zero | 1 => LinearMap.zero_comp (map I f) @@ -286,7 +292,8 @@ private def firstRowToSecondRow : firstRow I M f ⟶ secondRow I M f := rfl rfl -private lemma ofTensorProduct_iso : IsIso (ModuleCat.ofHom (ofTensorProduct I M)) := by +private lemma ofTensorProduct_iso [Fintype ι] [IsNoetherianRing R] : + IsIso (ModuleCat.ofHom (ofTensorProduct I M)) := by refine Abelian.isIso_of_epi_of_isIso_of_isIso_of_mono (firstRow_exact I M f hf) (secondRow_exact I M f hf) (firstRowToSecondRow I M f) ?_ ?_ ?_ ?_ · apply ConcreteCategory.epi_of_surjective @@ -301,16 +308,20 @@ private lemma ofTensorProduct_iso : IsIso (ModuleCat.ofHom (ofTensorProduct I M) rfl private -lemma ofTensorProduct_bijective_of_map_from_fin : Function.Bijective (ofTensorProduct I M) := by +lemma ofTensorProduct_bijective_of_map_from_fin [Fintype ι] [IsNoetherianRing R] : + Function.Bijective (ofTensorProduct I M) := by have : IsIso (ModuleCat.ofHom (ofTensorProduct I M)) := ofTensorProduct_iso I M f hf exact ConcreteCategory.bijective_of_isIso (ModuleCat.ofHom (ofTensorProduct I M)) end +variable [IsNoetherianRing R] + /-- If `R` is a Noetherian ring and `M` is a finite `R`-module, then the natural map given by `AdicCompletion.ofTensorProduct` is an isomorphism. -/ -theorem ofTensorProduct_bijective_of_finite_of_isNoetherian [Module.Finite R M] : +theorem ofTensorProduct_bijective_of_finite_of_isNoetherian + [Module.Finite R M] : Function.Bijective (ofTensorProduct I M) := by obtain ⟨n, f, hf⟩ := Module.Finite.exists_fin' R M exact ofTensorProduct_bijective_of_map_from_fin I M f hf @@ -329,7 +340,8 @@ lemma ofTensorProductEquivOfFiniteNoetherian_apply [Module.Finite R M] rfl @[simp] -lemma ofTensorProductEquivOfFiniteNoetherian_symm_of [Module.Finite R M] (x : M) : +lemma ofTensorProductEquivOfFiniteNoetherian_symm_of + [Module.Finite R M] (x : M) : (ofTensorProductEquivOfFiniteNoetherian I M).symm ((of I M) x) = 1 ⊗ₜ x := by have h : (of I M) x = ofTensorProductEquivOfFiniteNoetherian I M (1 ⊗ₜ x) := by simp diff --git a/Mathlib/RingTheory/AdicCompletion/Exactness.lean b/Mathlib/RingTheory/AdicCompletion/Exactness.lean index 1012a6c0fadf5..347ac711c8b05 100644 --- a/Mathlib/RingTheory/AdicCompletion/Exactness.lean +++ b/Mathlib/RingTheory/AdicCompletion/Exactness.lean @@ -40,6 +40,7 @@ variable {M : Type v} [AddCommGroup M] [Module R M] variable {N : Type w} [AddCommGroup N] [Module R N] variable {f : M →ₗ[R] N} (hf : Function.Surjective f) +include hf /- In each step, a preimage is constructed from the preimage of the previous step by subtracting this delta. -/ @@ -147,6 +148,7 @@ private noncomputable def mapExactAuxDelta {n : ℕ} {d : N} open Submodule +include hfg in /- Inductively construct preimage of cauchy sequence in kernel of `g.adicCompletion I`. -/ private noncomputable def mapExactAux : (n : ℕ) → { a : M | f a - x (k + n) ∈ (I ^ (k + n) • ⊤ : Submodule R N) } @@ -177,6 +179,7 @@ private noncomputable def mapExactAux : end +include hf hfg hg in /-- `AdicCompletion` over a Noetherian ring is exact on finitely generated modules. -/ theorem map_exact : Function.Exact (map I f) (map I g) := by refine LinearMap.exact_of_comp_eq_zero_of_ker_le_range ?_ (fun y ↦ ?_) diff --git a/Mathlib/RingTheory/Adjoin/PowerBasis.lean b/Mathlib/RingTheory/Adjoin/PowerBasis.lean index ee1dfa344a2bf..115e9eae03f82 100644 --- a/Mathlib/RingTheory/Adjoin/PowerBasis.lean +++ b/Mathlib/RingTheory/Adjoin/PowerBasis.lean @@ -83,13 +83,13 @@ open Polynomial variable {R : Type*} [CommRing R] [Algebra R S] [Algebra R K] [IsScalarTower R K S] variable {A : Type*} [CommRing A] [Algebra R A] [Algebra S A] -variable [IsScalarTower R S A] {B : PowerBasis S A} (hB : IsIntegral R B.gen) +variable [IsScalarTower R S A] {B : PowerBasis S A} /-- If `B : PowerBasis S A` is such that `IsIntegral R B.gen`, then `IsIntegral R (B.basis.repr (B.gen ^ n) i)` for all `i` if `minpoly S B.gen = (minpoly R B.gen).map (algebraMap R S)`. This is the case if `R` is a GCD domain and `S` is its fraction ring. -/ -theorem repr_gen_pow_isIntegral [IsDomain S] +theorem repr_gen_pow_isIntegral (hB : IsIntegral R B.gen) [IsDomain S] (hmin : minpoly S B.gen = (minpoly R B.gen).map (algebraMap R S)) (n : ℕ) : ∀ i, IsIntegral R (B.basis.repr (B.gen ^ n) i) := by intro i @@ -121,8 +121,8 @@ theorem repr_gen_pow_isIntegral [IsDomain S] integral coordinates in the base `B.basis`. Then `IsIntegral R ((B.basis.repr (x * y) i)` for all `i` if `minpoly S B.gen = (minpoly R B.gen).map (algebraMap R S)`. This is the case if `R` is a GCD domain and `S` is its fraction ring. -/ -theorem repr_mul_isIntegral [IsDomain S] {x y : A} (hx : ∀ i, IsIntegral R (B.basis.repr x i)) - (hy : ∀ i, IsIntegral R (B.basis.repr y i)) +theorem repr_mul_isIntegral (hB : IsIntegral R B.gen) [IsDomain S] {x y : A} + (hx : ∀ i, IsIntegral R (B.basis.repr x i)) (hy : ∀ i, IsIntegral R (B.basis.repr y i)) (hmin : minpoly S B.gen = (minpoly R B.gen).map (algebraMap R S)) : ∀ i, IsIntegral R (B.basis.repr (x * y) i) := by intro i @@ -139,7 +139,8 @@ theorem repr_mul_isIntegral [IsDomain S] {x y : A} (hx : ∀ i, IsIntegral R (B. with integral coordinates in the base `B.basis`. Then `IsIntegral R ((B.basis.repr (x ^ n) i)` for all `i` and all `n` if `minpoly S B.gen = (minpoly R B.gen).map (algebraMap R S)`. This is the case if `R` is a GCD domain and `S` is its fraction ring. -/ -theorem repr_pow_isIntegral [IsDomain S] {x : A} (hx : ∀ i, IsIntegral R (B.basis.repr x i)) +theorem repr_pow_isIntegral [IsDomain S] (hB : IsIntegral R B.gen) {x : A} + (hx : ∀ i, IsIntegral R (B.basis.repr x i)) (hmin : minpoly S B.gen = (minpoly R B.gen).map (algebraMap R S)) (n : ℕ) : ∀ i, IsIntegral R (B.basis.repr (x ^ n) i) := by nontriviality A using Subsingleton.elim (x ^ n) 0, isIntegral_zero diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean index 8898f192deb78..6a5c5db0549e9 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent.lean @@ -90,7 +90,14 @@ theorem algebraicIndependent_empty_type_iff [IsEmpty ι] : namespace AlgebraicIndependent +theorem of_comp (f : A →ₐ[R] A') (hfv : AlgebraicIndependent R (f ∘ x)) : + AlgebraicIndependent R x := by + have : aeval (f ∘ x) = f.comp (aeval x) := by ext; simp + rw [AlgebraicIndependent, this, AlgHom.coe_comp] at hfv + exact hfv.of_comp + variable (hx : AlgebraicIndependent R x) +include hx theorem algebraMap_injective : Injective (algebraMap R A) := by simpa [Function.comp] using @@ -137,12 +144,6 @@ theorem map {f : A →ₐ[R] A'} (hf_inj : Set.InjOn f (adjoin R (range x))) : theorem map' {f : A →ₐ[R] A'} (hf_inj : Injective f) : AlgebraicIndependent R (f ∘ x) := hx.map hf_inj.injOn -theorem of_comp (f : A →ₐ[R] A') (hfv : AlgebraicIndependent R (f ∘ x)) : - AlgebraicIndependent R x := by - have : aeval (f ∘ x) = f.comp (aeval x) := by ext; simp - rw [AlgebraicIndependent, this, AlgHom.coe_comp] at hfv - exact hfv.of_comp - end AlgebraicIndependent open AlgebraicIndependent diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian.lean index 916bb9144d3c4..65ecbf2b5c896 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian.lean @@ -446,6 +446,7 @@ theorem isNilpotent_jacobson_bot : IsNilpotent (Ideal.jacobson (⊥ : Ideal R)) section Localization variable (S : Submonoid R) (L : Type*) [CommRing L] [Algebra R L] [IsLocalization S L] +include S /-- Localizing an artinian ring can only reduce the amount of elements. -/ theorem localization_surjective : Function.Surjective (algebraMap R L) := by diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 5cdfa4e768306..6c656b177371a 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -261,6 +261,7 @@ theorem FractionalIdeal.adjoinIntegral_eq_one_of_isUnit [Algebra A K] [IsFractio namespace IsDedekindDomainInv variable [Algebra A K] [IsFractionRing A K] (h : IsDedekindDomainInv A) +include h theorem mul_inv_eq_one {I : FractionalIdeal A⁰ K} (hI : I ≠ 0) : I * I⁻¹ = 1 := isDedekindDomainInv_iff.mp h I hI @@ -1068,8 +1069,6 @@ theorem idealFactorsEquivOfQuotEquiv_is_dvd_iso {L M : Ideal R} (hL : L ∣ I) ( open UniqueFactorizationMonoid -variable [DecidableEq (Ideal R)] [DecidableEq (Ideal A)] - theorem idealFactorsEquivOfQuotEquiv_mem_normalizedFactors_of_mem_normalizedFactors (hJ : J ≠ ⊥) {L : Ideal R} (hL : L ∈ normalizedFactors I) : ↑(idealFactorsEquivOfQuotEquiv f ⟨L, dvd_of_mem_normalizedFactors hL⟩) @@ -1375,7 +1374,8 @@ theorem multiplicity_eq_multiplicity_span [DecidableRel ((· ∣ ·) : R → R rw [Ideal.span_singleton_pow, span_singleton_dvd_span_singleton_iff_dvd] exact not_finite_iff_forall.mp h n -variable [DecidableEq R] [DecidableEq (Ideal R)] [NormalizationMonoid R] +section NormalizationMonoid +variable [NormalizationMonoid R] /-- The bijection between the (normalized) prime factors of `r` and the (normalized) prime factors of `span {r}` -/ @@ -1428,6 +1428,8 @@ theorem multiplicity_normalizedFactorsEquivSpanNormalizedFactors_symm_eq_multipl rw [hx.symm, Equiv.symm_apply_apply, Subtype.coe_mk, multiplicity_normalizedFactorsEquivSpanNormalizedFactors_eq_multiplicity hr ha] +variable [DecidableEq R] [DecidableEq (Ideal R)] + /-- The bijection between the set of prime factors of the ideal `⟨r⟩` and the set of prime factors of `r` preserves `count` of the corresponding multisets. See `multiplicity_normalizedFactorsEquivSpanNormalizedFactors_eq_multiplicity` for the version @@ -1449,6 +1451,10 @@ theorem count_span_normalizedFactors_eq_of_normUnit {r X : R} Multiset.count X (normalizedFactors r) := by simpa [hX₁] using count_span_normalizedFactors_eq hr hX +end NormalizationMonoid + +variable [DecidableEq (Ideal R)] + /-- The number of times an ideal `I` occurs as normalized factor of another ideal `J` is stable when regarding at these ideals as associated elements of the monoid of ideals.-/ theorem count_associates_factors_eq [DecidableEq <| Associates (Ideal R)] diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index a570a392aa683..8f37a34bfe33b 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -55,6 +55,7 @@ variable [Algebra A K] [IsFractionRing A K] variable (L : Type*) [Field L] (C : Type*) [CommRing C] variable [Algebra K L] [Algebra A L] [IsScalarTower A K L] variable [Algebra C L] [IsIntegralClosure C A L] [Algebra A C] [IsScalarTower A C L] +include K L /-- If `L` is an algebraic extension of `K = Frac(A)` and `L` has no zero smul divisors by `A`, then `L` is the localization of the integral closure `C` of `A` in `L` at `A⁰`. -/ diff --git a/Mathlib/RingTheory/DedekindDomain/PID.lean b/Mathlib/RingTheory/DedekindDomain/PID.lean index b80cae3707d99..881352fd3f2ce 100644 --- a/Mathlib/RingTheory/DedekindDomain/PID.lean +++ b/Mathlib/RingTheory/DedekindDomain/PID.lean @@ -189,6 +189,7 @@ variable (p : Ideal R) (hp0 : p ≠ ⊥) [IsPrime p] variable {Sₚ : Type*} [CommRing Sₚ] [Algebra S Sₚ] variable [IsLocalization (Algebra.algebraMapSubmonoid S p.primeCompl) Sₚ] variable [Algebra R Sₚ] [IsScalarTower R S Sₚ] +include hp0 /- The first hypothesis below follows from properties of the localization but is needed for the second, so we leave it to the user to provide (automatically). -/ diff --git a/Mathlib/RingTheory/Etale/Basic.lean b/Mathlib/RingTheory/Etale/Basic.lean index 319a387c717e6..0a5d43bc3a05a 100644 --- a/Mathlib/RingTheory/Etale/Basic.lean +++ b/Mathlib/RingTheory/Etale/Basic.lean @@ -139,6 +139,7 @@ variable [Algebra R S] [Algebra R Sₘ] [Algebra S Sₘ] [Algebra R Rₘ] [Algeb variable [IsScalarTower R Rₘ Sₘ] [IsScalarTower R S Sₘ] /-! and that Rₘ and Sₘ are localizations of R and S at M. -/ variable [IsLocalization M Rₘ] [IsLocalization (M.map (algebraMap R S)) Sₘ] +include M -- Porting note: no longer supported -- attribute [local elab_as_elim] Ideal.IsNilpotent.induction_on diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index b0846e78aa5bb..04c32e6e803d6 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -25,7 +25,7 @@ variable {R : Type u} {S : Type v} section Semiring variable {F : Type*} [Semiring R] [Semiring S] -variable [FunLike F R S] [RingHomClass F R S] +variable [FunLike F R S] variable (f : F) variable {I J : Ideal R} {K L : Ideal S} @@ -35,7 +35,7 @@ def map (I : Ideal R) : Ideal S := span (f '' I) /-- `I.comap f` is the preimage of `I` under `f`. -/ -def comap (I : Ideal S) : Ideal R where +def comap [RingHomClass F R S] (I : Ideal S) : Ideal R where carrier := f ⁻¹' I add_mem' {x y} hx hy := by simp only [Set.mem_preimage, SetLike.mem_coe, map_add f] at hx hy ⊢ @@ -46,7 +46,7 @@ def comap (I : Ideal S) : Ideal R where exact mul_mem_left I _ hx @[simp] -theorem coe_comap (I : Ideal S) : (comap f I : Set R) = f ⁻¹' I := rfl +theorem coe_comap [RingHomClass F R S] (I : Ideal S) : (comap f I : Set R) = f ⁻¹' I := rfl variable {f} @@ -59,19 +59,19 @@ theorem mem_map_of_mem (f : F) {I : Ideal R} {x : R} (h : x ∈ I) : f x ∈ map theorem apply_coe_mem_map (f : F) (I : Ideal R) (x : I) : f x ∈ I.map f := mem_map_of_mem f x.2 -theorem map_le_iff_le_comap : map f I ≤ K ↔ I ≤ comap f K := +theorem map_le_iff_le_comap [RingHomClass F R S] : map f I ≤ K ↔ I ≤ comap f K := span_le.trans Set.image_subset_iff @[simp] -theorem mem_comap {x} : x ∈ comap f K ↔ f x ∈ K := +theorem mem_comap [RingHomClass F R S] {x} : x ∈ comap f K ↔ f x ∈ K := Iff.rfl -theorem comap_mono (h : K ≤ L) : comap f K ≤ comap f L := +theorem comap_mono [RingHomClass F R S] (h : K ≤ L) : comap f K ≤ comap f L := Set.preimage_mono fun _ hx => h hx variable (f) -theorem comap_ne_top (hK : K ≠ ⊤) : comap f K ≠ ⊤ := +theorem comap_ne_top [RingHomClass F R S] (hK : K ≠ ⊤) : comap f K ≠ ⊤ := (ne_top_iff_one _).2 <| by rw [mem_comap, map_one]; exact (ne_top_iff_one _).1 hK variable {G : Type*} [FunLike G S R] @@ -84,7 +84,7 @@ theorem map_le_comap_of_inv_on [RingHomClass G S R] (g : G) (I : Ideal R) rw [SetLike.mem_coe, mem_comap, hf hx] exact hx -theorem comap_le_map_of_inv_on (g : G) (I : Ideal S) +theorem comap_le_map_of_inv_on [RingHomClass F R S] (g : G) (I : Ideal S) (hf : Set.LeftInvOn g f (f ⁻¹' I)) : I.comap f ≤ I.map g := fun x (hx : f x ∈ I) => hf hx ▸ Ideal.mem_map_of_mem g hx @@ -95,6 +95,8 @@ theorem map_le_comap_of_inverse [RingHomClass G S R] (g : G) (I : Ideal R) I.map f ≤ I.comap g := map_le_comap_of_inv_on _ _ _ <| h.leftInvOn _ +variable [RingHomClass F R S] + /-- The `Ideal` version of `Set.preimage_subset_image_of_inverse`. -/ theorem comap_le_map_of_inverse (g : G) (I : Ideal S) (h : Function.LeftInverse g f) : I.comap f ≤ I.map g := @@ -244,7 +246,8 @@ theorem restrictScalars_mul {R S : Type*} [CommSemiring R] [CommSemiring S] [Alg section Surjective -variable (hf : Function.Surjective f) +section variable (hf : Function.Surjective f) +include hf open Function @@ -292,6 +295,8 @@ theorem mem_map_iff_of_surjective {I : Ideal R} {y} : y ∈ map f I ↔ ∃ x, x theorem le_map_of_comap_le_of_surjective : comap f K ≤ I → K ≤ map f I := fun h => map_comap_of_surjective f hf K ▸ map_mono h +end + theorem map_eq_submodule_map (f : R →+* S) [h : RingHomSurjective f] (I : Ideal R) : I.map f = Submodule.map f.toSemilinearMap I := Submodule.ext fun _ => mem_map_iff_of_surjective f h.1 diff --git a/Mathlib/RingTheory/Ideal/MinimalPrime.lean b/Mathlib/RingTheory/Ideal/MinimalPrime.lean index b3b7e15843969..50f2688455828 100644 --- a/Mathlib/RingTheory/Ideal/MinimalPrime.lean +++ b/Mathlib/RingTheory/Ideal/MinimalPrime.lean @@ -105,7 +105,7 @@ theorem Ideal.exists_comap_eq_of_mem_minimalPrimes_of_injective {f : R →+* S} obtain ⟨M, hM⟩ := Ideal.exists_maximal (Localization (Submonoid.map f p.primeCompl)) refine ⟨M.comap (algebraMap S <| Localization (Submonoid.map f p.primeCompl)), inferInstance, ?_⟩ rw [Ideal.comap_comap, ← @IsLocalization.map_comp _ _ _ _ _ _ _ _ Localization.isLocalization - _ _ _ _ p.primeCompl.le_comap_map _ Localization.isLocalization, + _ _ _ _ _ Localization.isLocalization p.primeCompl.le_comap_map, ← Ideal.comap_comap] suffices _ ≤ p by exact this.antisymm (H.2 ⟨inferInstance, bot_le⟩ this) intro x hx @@ -210,6 +210,7 @@ end namespace Localization.AtPrime variable {R : Type*} [CommSemiring R] {I : Ideal R} [hI : I.IsPrime] (hMin : I ∈ minimalPrimes R) +include hMin theorem _root_.IsLocalization.AtPrime.prime_unique_of_minimal {S} [CommSemiring S] [Algebra R S] [IsLocalization.AtPrime S I] {J K : Ideal S} [J.IsPrime] [K.IsPrime] : J = K := diff --git a/Mathlib/RingTheory/Ideal/Over.lean b/Mathlib/RingTheory/Ideal/Over.lean index ec51acc94cca5..f8586e0f2559a 100644 --- a/Mathlib/RingTheory/Ideal/Over.lean +++ b/Mathlib/RingTheory/Ideal/Over.lean @@ -274,6 +274,7 @@ section IsIntegralClosure variable (S) {A : Type*} [CommRing A] variable [Algebra R A] [Algebra A S] [IsScalarTower R A S] [IsIntegralClosure A R S] +include S theorem IsIntegralClosure.comap_lt_comap {I J : Ideal A} [I.IsPrime] (I_lt_J : I < J) : I.comap (algebraMap R A) < J.comap (algebraMap R A) := @@ -340,9 +341,9 @@ theorem exists_ideal_over_prime_of_isIntegral_of_isDomain [Algebra.IsIntegral R refine ⟨comap (algebraMap S Sₚ) Qₚ, ⟨comap_isPrime _ Qₚ, ?_⟩⟩ convert Localization.AtPrime.comap_maximalIdeal (I := P) rw [comap_comap, ← LocalRing.eq_maximalIdeal Qₚ_max, - ← @IsLocalization.map_comp (P := S) (Q := Sₚ) (g := algebraMap R S) - (M := P.primeCompl) (T := Algebra.algebraMapSubmonoid S P.primeCompl) (S := Rₚ) _ - _ _ _ _ _ (fun p hp => Algebra.mem_algebraMapSubmonoid_of_mem ⟨p, hp⟩) _ _] + ← IsLocalization.map_comp (P := S) (Q := Sₚ) (g := algebraMap R S) + (M := P.primeCompl) (T := Algebra.algebraMapSubmonoid S P.primeCompl) (S := Rₚ) + (fun p hp => Algebra.mem_algebraMapSubmonoid_of_mem ⟨p, hp⟩) ] rfl end @@ -407,7 +408,7 @@ theorem exists_ideal_over_maximal_of_isIntegral [Algebra.IsIntegral R S] obtain ⟨Q, -, Q_prime, hQ⟩ := exists_ideal_over_prime_of_isIntegral P ⊥ hP exact ⟨Q, isMaximal_of_isIntegral_of_isMaximal_comap _ (hQ.symm ▸ P_max), hQ⟩ -lemma map_eq_top_iff_of_ker_le +lemma map_eq_top_iff_of_ker_le {R S} [CommRing R] [CommRing S] (f : R →+* S) {I : Ideal R} (hf₁ : RingHom.ker f ≤ I) (hf₂ : f.IsIntegral) : I.map f = ⊤ ↔ I = ⊤ := by constructor; swap @@ -421,7 +422,7 @@ lemma map_eq_top_iff_of_ker_le rw [← map_le_iff_le_comap] at hm exact (hm.trans_lt (lt_top_iff_ne_top.mpr (IsMaximal.ne_top ‹_›))).ne -lemma map_eq_top_iff +lemma map_eq_top_iff {R S} [CommRing R] [CommRing S] (f : R →+* S) {I : Ideal R} (hf₁ : Function.Injective f) (hf₂ : f.IsIntegral) : I.map f = ⊤ ↔ I = ⊤ := map_eq_top_iff_of_ker_le f (by simp [f.injective_iff_ker_eq_bot.mp hf₁]) hf₂ diff --git a/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean b/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean index e65a17d8ac054..f0f556d8c02f6 100644 --- a/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean +++ b/Mathlib/RingTheory/IntegralClosure/IntegralRestrict.lean @@ -227,6 +227,7 @@ open nonZeroDivisors variable [IsDomain Aₘ] [IsIntegrallyClosed Aₘ] [IsDomain Bₘ] [IsIntegrallyClosed Bₘ] variable [NoZeroSMulDivisors Aₘ Bₘ] [Module.Finite Aₘ Bₘ] +include M in lemma Algebra.intTrace_eq_of_isLocalization (x : B) : algebraMap A Aₘ (Algebra.intTrace A B x) = Algebra.intTrace Aₘ Bₘ (algebraMap B Bₘ x) := by @@ -383,6 +384,7 @@ variable [IsDomain Aₘ] [IsIntegrallyClosed Aₘ] [IsDomain Bₘ] [IsIntegrally variable [NoZeroSMulDivisors Aₘ Bₘ] [Module.Finite Aₘ Bₘ] variable [Algebra.IsSeparable (FractionRing Aₘ) (FractionRing Bₘ)] +include M in lemma Algebra.intNorm_eq_of_isLocalization (x : B) : algebraMap A Aₘ (Algebra.intNorm A B x) = Algebra.intNorm Aₘ Bₘ (algebraMap B Bₘ x) := by by_cases hM : 0 ∈ M diff --git a/Mathlib/RingTheory/IsAdjoinRoot.lean b/Mathlib/RingTheory/IsAdjoinRoot.lean index a1951348cbd0f..f48ab8430a1bd 100644 --- a/Mathlib/RingTheory/IsAdjoinRoot.lean +++ b/Mathlib/RingTheory/IsAdjoinRoot.lean @@ -177,7 +177,11 @@ theorem ext (h h' : IsAdjoinRoot S f) (eq : h.root = h'.root) : h = h' := section lift -variable {T : Type*} [CommRing T] {i : R →+* T} {x : T} (hx : f.eval₂ i x = 0) +variable {T : Type*} [CommRing T] {i : R →+* T} {x : T} + +section +variable (hx : f.eval₂ i x = 0) +include hx /-- Auxiliary lemma for `IsAdjoinRoot.lift` -/ theorem eval₂_repr_eq_eval₂_of_map_eq (h : IsAdjoinRoot S f) (z : S) (w : R[X]) @@ -236,6 +240,8 @@ theorem eq_lift (h : IsAdjoinRoot S f) (g : S →+* T) (hmap : ∀ a, g (algebra (hroot : g h.root = x) : g = h.lift i x hx := RingHom.ext (h.apply_eq_lift hx g hmap hroot) +end + variable [Algebra R T] (hx' : aeval x f = 0) variable (x) diff --git a/Mathlib/RingTheory/IsTensorProduct.lean b/Mathlib/RingTheory/IsTensorProduct.lean index f3699b8d59848..07110fac07464 100644 --- a/Mathlib/RingTheory/IsTensorProduct.lean +++ b/Mathlib/RingTheory/IsTensorProduct.lean @@ -167,6 +167,9 @@ theorem IsBaseChange.lift_comp (g : M →ₗ[R] Q) : ((h.lift g).restrictScalars end +section +include h + @[elab_as_elim] nonrec theorem IsBaseChange.inductionOn (x : N) (P : N → Prop) (h₁ : P 0) (h₂ : ∀ m : M, P (f m)) (h₃ : ∀ (s : S) (n), P n → P (s • n)) (h₄ : ∀ n₁ n₂, P n₁ → P n₂ → P (n₁ + n₂)) : P x := @@ -186,6 +189,8 @@ theorem IsBaseChange.algHom_ext' [Module R Q] [IsScalarTower R S Q] (g₁ g₂ : (e : (g₁.restrictScalars R).comp f = (g₂.restrictScalars R).comp f) : g₁ = g₂ := h.algHom_ext g₁ g₂ (LinearMap.congr_fun e) +end + variable (R M N S) theorem TensorProduct.isBaseChange : IsBaseChange S (TensorProduct.mk R S M 1) := by diff --git a/Mathlib/RingTheory/Jacobson.lean b/Mathlib/RingTheory/Jacobson.lean index f74c1ecaaeafa..f9ca45cdd5b0b 100644 --- a/Mathlib/RingTheory/Jacobson.lean +++ b/Mathlib/RingTheory/Jacobson.lean @@ -207,6 +207,7 @@ def orderIsoOfMaximal [IsJacobson R] : map_comap (powers y) S I.val ▸ map_comap (powers y) S I'.val ▸ Ideal.map_mono h, fun h _ hx => h hx⟩ +include y in /-- If `S` is the localization of the Jacobson ring `R` at the submonoid generated by `y : R`, then `S` is Jacobson. -/ theorem isJacobson_localization [H : IsJacobson R] : IsJacobson S := by @@ -441,10 +442,11 @@ end CommRing section -variable {R : Type*} [CommRing R] [IsJacobson R] +variable {R : Type*} [CommRing R] variable (P : Ideal R[X]) [hP : P.IsMaximal] -theorem isMaximal_comap_C_of_isMaximal [Nontrivial R] (hP' : ∀ x : R, C x ∈ P → x = 0) : +theorem isMaximal_comap_C_of_isMaximal [IsJacobson R] [Nontrivial R] + (hP' : ∀ x : R, C x ∈ P → x = 0) : IsMaximal (comap (C : R →+* R[X]) P : Ideal R) := by let P' := comap (C : R →+* R[X]) P haveI hP'_prime : P'.IsPrime := comap_isPrime C P @@ -522,6 +524,8 @@ private theorem quotient_mk_comp_C_isIntegral_of_jacobson' [Nontrivial R] (hR : (Localization M) (Localization M') _ _ P pX hpX _ _ _ isloc rw [IsLocalization.map_comp M.le_comap_map] +variable [IsJacobson R] + /-- If `R` is a Jacobson ring, and `P` is a maximal ideal of `R[X]`, then `R → R[X]/P` is an integral map. -/ theorem quotient_mk_comp_C_isIntegral_of_jacobson : diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index ba35a7d6cb2be..b9d1f5ee13512 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -328,13 +328,11 @@ theorem ne_zero_of_mk'_ne_zero {x : R} {y : M} (hxy : IsLocalization.mk' S x y section Ext -variable [Algebra R P] [IsLocalization M P] - -theorem eq_iff_eq {x y} : +theorem eq_iff_eq [Algebra R P] [IsLocalization M P] {x y} : algebraMap R S x = algebraMap R S y ↔ algebraMap R P x = algebraMap R P y := (toLocalizationMap M S).eq_iff_eq (toLocalizationMap M P) -theorem mk'_eq_iff_mk'_eq {x₁ x₂} {y₁ y₂ : M} : +theorem mk'_eq_iff_mk'_eq [Algebra R P] [IsLocalization M P] {x₁ x₂} {y₁ y₂ : M} : mk' S x₁ y₁ = mk' S x₂ y₂ ↔ mk' P x₁ y₁ = mk' P x₂ y₂ := (toLocalizationMap M S).mk'_eq_iff_mk'_eq (toLocalizationMap M P) @@ -492,6 +490,9 @@ theorem lift_of_comp (j : S →+* P) : lift (isUnit_comp M j) = j := variable (M) +section +include M + /-- See note [partially-applied ext lemmas] -/ theorem monoidHom_ext ⦃j k : S →* P⦄ (h : j.comp (algebraMap R S : R →* S) = k.comp (algebraMap R S)) : j = k := @@ -520,6 +521,7 @@ protected theorem ext (j k : S → P) (hj1 : j 1 = 1) (hk1 : k 1 = 1) { toFun := k, map_one' := hk1, map_mul' := hkm } have : j' = k' := monoidHom_ext M (MonoidHom.ext h) show j'.toFun = k'.toFun by rw [this] +end variable {M} @@ -543,7 +545,7 @@ theorem lift_injective_iff : section Map -variable {T : Submonoid P} {Q : Type*} [CommSemiring Q] (hy : M ≤ T.comap g) +variable {T : Submonoid P} {Q : Type*} [CommSemiring Q] variable [Algebra P Q] [IsLocalization T Q] section @@ -561,6 +563,10 @@ noncomputable def map (g : R →+* P) (hy : M ≤ T.comap g) : S →+* Q := end +section +variable (hy : M ≤ T.comap g) +include hy + -- Porting note: added `simp` attribute, since it proves very similar lemmas marked `simp` @[simp] theorem map_eq (x) : map Q g hy ((algebraMap R S) x) = algebraMap P Q (g x) := @@ -574,16 +580,6 @@ theorem map_mk' (x) (y : M) : map Q g hy (mk' S x y) = mk' Q (g x) ⟨g y, hy y. Submonoid.LocalizationMap.map_mk' (toLocalizationMap M S) (g := g.toMonoidHom) (fun y => hy y.2) (k := toLocalizationMap T Q) .. -@[simp] -theorem map_id_mk' {Q : Type*} [CommSemiring Q] [Algebra R Q] [IsLocalization M Q] (x) (y : M) : - map Q (RingHom.id R) (le_refl M) (mk' S x y) = mk' Q x y := - map_mk' .. - -@[simp] -theorem map_id (z : S) (h : M ≤ M.comap (RingHom.id R) := le_refl M) : - map S (RingHom.id _) h z = z := - lift_id _ - theorem map_unique (j : S →+* Q) (hj : ∀ x : R, j (algebraMap R S x) = algebraMap P Q (g x)) : map Q g hy = j := lift_unique (fun y => map_units _ ⟨g y, hy y.2⟩) hj @@ -607,6 +603,18 @@ theorem map_map {A : Type*} [CommSemiring A] {U : Submonoid A} {W} [CommSemiring theorem map_smul (x : S) (z : R) : map Q g hy (z • x : S) = g z • map Q g hy x := by rw [Algebra.smul_def, Algebra.smul_def, RingHom.map_mul, map_eq] +end + +@[simp] +theorem map_id_mk' {Q : Type*} [CommSemiring Q] [Algebra R Q] [IsLocalization M Q] (x) (y : M) : + map Q (RingHom.id R) (le_refl M) (mk' S x y) = mk' Q x y := + map_mk' .. + +@[simp] +theorem map_id (z : S) (h : M ≤ M.comap (RingHom.id R) := le_refl M) : + map S (RingHom.id _) h z = z := + lift_id _ + section variable (S Q) @@ -675,6 +683,7 @@ theorem algEquiv_mk' (x : R) (y : M) : algEquiv M S Q (mk' S x y) = mk' Q x y := theorem algEquiv_symm_mk' (x : R) (y : M) : (algEquiv M S Q).symm (mk' Q x y) = mk' S x y := by simp variable (M) in +include M in protected lemma bijective (f : S →+* Q) (hf : f.comp (algebraMap R S) = algebraMap R Q) : Function.Bijective f := (show f = IsLocalization.algEquiv M S Q by @@ -684,8 +693,7 @@ protected lemma bijective (f : S →+* Q) (hf : f.comp (algebraMap R S) = algebr end AlgEquiv section at_units - -lemma at_units {R : Type*} [CommSemiring R] (S : Submonoid R) +lemma at_units (S : Submonoid R) (hS : S ≤ IsUnit.submonoid R) : IsLocalization S R where map_units' y := hS y.prop surj' := fun s ↦ ⟨⟨s, 1⟩, by simp⟩ @@ -1063,7 +1071,7 @@ section Algebra variable {S} {Rₘ Sₘ : Type*} [CommRing Rₘ] [CommRing Sₘ] variable [Algebra R Rₘ] [IsLocalization M Rₘ] variable [Algebra S Sₘ] [i : IsLocalization (Algebra.algebraMapSubmonoid S M) Sₘ] - +include S section variable (S M) @@ -1091,7 +1099,7 @@ variable [Algebra Rₘ Sₘ] [Algebra R Sₘ] [IsScalarTower R Rₘ Sₘ] [IsSca variable (S Rₘ Sₘ) theorem IsLocalization.map_units_map_submonoid - [IsLocalization (Algebra.algebraMapSubmonoid S M) Sₘ] (y : M) : + (y : M) : IsUnit (algebraMap R Sₘ y) := by rw [IsScalarTower.algebraMap_apply _ S] exact IsLocalization.map_units Sₘ ⟨algebraMap R S y, Algebra.mem_algebraMapSubmonoid_of_mem y⟩ diff --git a/Mathlib/RingTheory/Localization/Finiteness.lean b/Mathlib/RingTheory/Localization/Finiteness.lean index 0f85d35308b26..10abd059206ce 100644 --- a/Mathlib/RingTheory/Localization/Finiteness.lean +++ b/Mathlib/RingTheory/Localization/Finiteness.lean @@ -37,6 +37,7 @@ variable {M : Type w} [AddCommMonoid M] [Module R M] variable {Mₚ : Type t} [AddCommMonoid Mₚ] [Module R Mₚ] [Module Rₚ Mₚ] [IsScalarTower R Rₚ Mₚ] variable (f : M →ₗ[R] Mₚ) [IsLocalizedModule S f] +include S f in lemma of_isLocalizedModule [Module.Finite R M] : Module.Finite Rₚ Mₚ := by classical obtain ⟨T, hT⟩ := ‹Module.Finite R M› diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index 64199851be828..1960bbcd7e01d 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -93,6 +93,7 @@ protected theorem to_map_ne_zero_of_mem_nonZeroDivisors [Nontrivial R] {x : R} variable (A) +include A in /-- A `CommRing` `K` which is the localization of an integral domain `R` at `R - {0}` is an integral domain. -/ protected theorem isDomain : IsDomain K := diff --git a/Mathlib/RingTheory/Localization/Ideal.lean b/Mathlib/RingTheory/Localization/Ideal.lean index fc8d3b40ea966..3882ffbd9f01a 100644 --- a/Mathlib/RingTheory/Localization/Ideal.lean +++ b/Mathlib/RingTheory/Localization/Ideal.lean @@ -68,7 +68,8 @@ lemma mk'_mem_map_algebraMap_iff (I : Ideal R) (x : R) (s : M) : exact ⟨fun ⟨⟨y, t⟩, c, h⟩ ↦ ⟨_, (c * t).2, h ▸ I.mul_mem_left c.1 y.2⟩, fun ⟨s, hs, h⟩ ↦ ⟨⟨⟨_, h⟩, ⟨s, hs⟩⟩, 1, by simp⟩⟩ -theorem map_comap [IsLocalization M S] (J : Ideal S) : +include M in +theorem map_comap (J : Ideal S) : Ideal.map (algebraMap R S) (Ideal.comap (algebraMap R S) J) = J := le_antisymm (Ideal.map_le_iff_le_comap.2 le_rfl) fun x hJ => by obtain ⟨r, s, hx⟩ := mk'_surjective M x @@ -164,6 +165,7 @@ section CommRing variable {R : Type*} [CommRing R] (M : Submonoid R) (S : Type*) [CommRing S] variable [Algebra R S] [IsLocalization M S] +include M in /-- `quotientMap` applied to maximal ideals of a localization is `surjective`. The quotient by a maximal ideal is a field, so inverses to elements already exist, and the localization necessarily maps the equivalence class of the inverse in the localization -/ diff --git a/Mathlib/RingTheory/Localization/LocalizationLocalization.lean b/Mathlib/RingTheory/Localization/LocalizationLocalization.lean index a74d15537efdc..bc2b5cc8d572a 100644 --- a/Mathlib/RingTheory/Localization/LocalizationLocalization.lean +++ b/Mathlib/RingTheory/Localization/LocalizationLocalization.lean @@ -111,6 +111,7 @@ theorem localization_localization_isLocalization [IsLocalization N T] : surj' := localization_localization_surj M N T exists_of_eq := localization_localization_exists_of_eq M N T _ _ } +include M in /-- Given submodules `M ⊆ R` and `N ⊆ S = M⁻¹R`, with `f : R →+* S` the localization map, if `N` contains all the units of `S`, then `N ⁻¹ S = T = (f⁻¹ N) ⁻¹ R`. I.e., the localization of a localization is a localization. @@ -125,6 +126,7 @@ theorem localization_localization_isLocalization_of_has_all_units [IsLocalizatio rintro _ ⟨x, hx, rfl⟩ exact H _ (IsLocalization.map_units _ ⟨x, hx⟩) +include M in /-- Given a submodule `M ⊆ R` and a prime ideal `p` of `S = M⁻¹R`, with `f : R →+* S` the localization map, then `T = Sₚ` is the localization of `R` at `f⁻¹(p)`. diff --git a/Mathlib/RingTheory/Localization/Module.lean b/Mathlib/RingTheory/Localization/Module.lean index 1c96e3fd8daf3..92f8b221b0cbb 100644 --- a/Mathlib/RingTheory/Localization/Module.lean +++ b/Mathlib/RingTheory/Localization/Module.lean @@ -38,10 +38,12 @@ section AddCommMonoid open Submodule variable [CommSemiring Rₛ] [Algebra R Rₛ] [hT : IsLocalization S Rₛ] -variable {M M' : Type*} [AddCommMonoid M] [Module R M] [Module Rₛ M] [IsScalarTower R Rₛ M] +variable {M M' : Type*} [AddCommMonoid M] [Module R M] [AddCommMonoid M'] [Module R M'] [Module Rₛ M'] [IsScalarTower R Rₛ M'] (f : M →ₗ[R] M') [IsLocalizedModule S f] +include S + theorem span_eq_top_of_isLocalizedModule {v : Set M} (hv : span R v = ⊤) : span Rₛ (f '' v) = ⊤ := top_unique fun x _ ↦ by obtain ⟨⟨m, s⟩, h⟩ := IsLocalizedModule.surj S f x @@ -68,6 +70,8 @@ theorem LinearIndependent.of_isLocalizedModule {ι : Type*} {v : ι → M} rw [← (IsLocalization.map_units Rₛ a).mul_right_eq_zero, ← Algebra.smul_def, ← hg' i hi] exact (IsLocalization.map_eq_zero_iff S _ _).2 ⟨s, hv⟩ +variable [Module Rₛ M] [IsScalarTower R Rₛ M] + theorem LinearIndependent.localization {ι : Type*} {b : ι → M} (hli : LinearIndependent R b) : LinearIndependent Rₛ b := by have := isLocalizedModule_id S M Rₛ @@ -81,7 +85,7 @@ variable [CommRing Rₛ] [Algebra R Rₛ] [hT : IsLocalization S Rₛ] open Submodule -variable {M Mₛ : Type*} [AddCommGroup M] [AddCommGroup Mₛ] [Module R M] [Module R Mₛ] [Module R Mₛ] +variable {M Mₛ : Type*} [AddCommGroup M] [AddCommGroup Mₛ] [Module R M] [Module R Mₛ] [Module Rₛ Mₛ] (f : M →ₗ[R] Mₛ) [IsLocalizedModule S f] [IsScalarTower R Rₛ Mₛ] {ι : Type*} (b : Basis ι R M) @@ -119,8 +123,8 @@ end IsLocalizedModule section LocalizationLocalization -variable {R : Type*} (Rₛ : Type*) [CommSemiring R] [CommRing Rₛ] [Algebra R Rₛ] -variable (S : Submonoid R) [hT : IsLocalization S Rₛ] +variable [CommRing Rₛ] [Algebra R Rₛ] +variable [hT : IsLocalization S Rₛ] variable {A : Type*} [CommRing A] [Algebra R A] variable (Aₛ : Type*) [CommRing Aₛ] [Algebra A Aₛ] variable [Algebra Rₛ Aₛ] [Algebra R Aₛ] [IsScalarTower R Rₛ Aₛ] [IsScalarTower R A Aₛ] @@ -128,6 +132,8 @@ variable [hA : IsLocalization (Algebra.algebraMapSubmonoid A S) Aₛ] open Submodule +include S + theorem LinearIndependent.localization_localization {ι : Type*} {v : ι → A} (hv : LinearIndependent R v) : LinearIndependent Rₛ ((algebraMap A Aₛ) ∘ v) := hv.of_isLocalizedModule Rₛ S (IsScalarTower.toAlgHom R A Aₛ).toLinearMap diff --git a/Mathlib/RingTheory/Localization/NormTrace.lean b/Mathlib/RingTheory/Localization/NormTrace.lean index 832b1621e7696..1154b1242ff4a 100644 --- a/Mathlib/RingTheory/Localization/NormTrace.lean +++ b/Mathlib/RingTheory/Localization/NormTrace.lean @@ -42,7 +42,7 @@ variable {Rₘ Sₘ : Type*} [CommRing Rₘ] [Algebra R Rₘ] [CommRing Sₘ] [A variable (M : Submonoid R) variable [IsLocalization M Rₘ] [IsLocalization (Algebra.algebraMapSubmonoid S M) Sₘ] variable [Algebra Rₘ Sₘ] [Algebra R Sₘ] [IsScalarTower R Rₘ Sₘ] [IsScalarTower R S Sₘ] - +include M open Algebra theorem Algebra.map_leftMulMatrix_localization {ι : Type*} [Fintype ι] [DecidableEq ι] diff --git a/Mathlib/RingTheory/Localization/Submodule.lean b/Mathlib/RingTheory/Localization/Submodule.lean index eecf67203492a..b00dfdb755232 100644 --- a/Mathlib/RingTheory/Localization/Submodule.lean +++ b/Mathlib/RingTheory/Localization/Submodule.lean @@ -77,14 +77,11 @@ variable {T : Submonoid P} (hy : M ≤ T.comap g) {Q : Type*} [CommRing Q] variable [Algebra P Q] [IsLocalization T Q] variable [IsLocalization M S] -section - +include M in theorem isNoetherianRing (h : IsNoetherianRing R) : IsNoetherianRing S := by rw [isNoetherianRing_iff, isNoetherian_iff_wellFounded] at h ⊢ exact OrderEmbedding.wellFounded (IsLocalization.orderEmbedding M S).dual h -end - variable {S M} @[mono] diff --git a/Mathlib/RingTheory/Perfection.lean b/Mathlib/RingTheory/Perfection.lean index 685ba13d6aae3..341d08d64c115 100644 --- a/Mathlib/RingTheory/Perfection.lean +++ b/Mathlib/RingTheory/Perfection.lean @@ -333,17 +333,15 @@ def ModP (K : Type u₁) [Field K] (v : Valuation K ℝ≥0) (O : Type u₂) [Co (_ : v.Integers O) (p : ℕ) := O ⧸ (Ideal.span {(p : O)} : Ideal O) -variable [hp : Fact p.Prime] [hvp : Fact (v p ≠ 1)] - namespace ModP instance commRing : CommRing (ModP K v O hv p) := Ideal.Quotient.commRing (Ideal.span {(p : O)} : Ideal O) -instance charP : CharP (ModP K v O hv p) p := +instance charP [Fact p.Prime] [hvp : Fact (v p ≠ 1)] : CharP (ModP K v O hv p) p := CharP.quotient O p <| mt hv.one_of_isUnit <| (map_natCast (algebraMap O K) p).symm ▸ hvp.1 -instance : Nontrivial (ModP K v O hv p) := +instance [hp : Fact p.Prime] [Fact (v p ≠ 1)] : Nontrivial (ModP K v O hv p) := CharP.nontrivial_of_char_ne_one hp.1.ne_one section Classical @@ -406,6 +404,7 @@ theorem preVal_eq_zero {x : ModP K v O hv p} : preVal K v O hv p x = 0 ↔ x = 0 fun hx => hx.symm ▸ preVal_zero⟩ variable (hv) -- Porting note: Originally `(hv hvp)`. Removed `(hvp)` because it caused an error. +include hv theorem v_p_lt_val {x : O} : v p < v (algebraMap O K x) ↔ (Ideal.Quotient.mk _ x : ModP K v O hv p) ≠ 0 := by @@ -415,6 +414,7 @@ theorem v_p_lt_val {x : O} : open NNReal variable {hv} -- Porting note: Originally `{hv} (hvp)`. Removed `(hvp)` because it caused an error. +variable [hp : Fact p.Prime] theorem mul_ne_zero_of_pow_p_ne_zero {x y : ModP K v O hv p} (hx : x ^ p ≠ 0) (hy : y ^ p ≠ 0) : x * y ≠ 0 := by @@ -446,6 +446,8 @@ def PreTilt := namespace PreTilt +variable [Fact p.Prime] [Fact (v p ≠ 1)] + instance : CommRing (PreTilt K v O hv p) := Perfection.commRing p _ @@ -561,7 +563,7 @@ theorem map_eq_zero {f : PreTilt K v O hv p} : val K v O hv p f = 0 ↔ f = 0 := end Classical -instance : IsDomain (PreTilt K v O hv p) := by +instance [hp : Fact p.Prime] : IsDomain (PreTilt K v O hv p) := by haveI : Nontrivial (PreTilt K v O hv p) := ⟨(CharP.nontrivial_of_char_ne_one hp.1.ne_one).1⟩ haveI : NoZeroDivisors (PreTilt K v O hv p) := ⟨fun hfg => by @@ -575,12 +577,12 @@ end PreTilt [scholze2011perfectoid]. Given a field `K` with valuation `K → ℝ≥0` and ring of integers `O`, this is implemented as the fraction field of the perfection of `O/(p)`. -/ -- @[nolint has_nonempty_instance] -- Porting note(#5171): This linter does not exist yet. -def Tilt := +def Tilt [Fact p.Prime] [Fact (v p ≠ 1)] := FractionRing (PreTilt K v O hv p) namespace Tilt -noncomputable instance : Field (Tilt K v O hv p) := +noncomputable instance [Fact p.Prime] [Fact (v p ≠ 1)] : Field (Tilt K v O hv p) := FractionRing.field _ end Tilt diff --git a/Mathlib/RingTheory/Polynomial/GaussLemma.lean b/Mathlib/RingTheory/Polynomial/GaussLemma.lean index ea918e9caf17c..ca29983d56934 100644 --- a/Mathlib/RingTheory/Polynomial/GaussLemma.lean +++ b/Mathlib/RingTheory/Polynomial/GaussLemma.lean @@ -67,7 +67,7 @@ theorem integralClosure.mem_lifts_of_monic_of_dvd_map {f : R[X]} (hf : f.Monic) refine Multiset.mem_of_le (roots.le_of_dvd ((hf.map _).map _).ne_zero ?_) ha exact map_dvd (algebraMap K g.SplittingField) hd -variable [IsDomain R] [IsFractionRing R K] +variable [IsFractionRing R K] /-- If `K = Frac(R)` and `g : K[X]` divides a monic polynomial with coefficients in `R`, then `g * (C g.leadingCoeff⁻¹)` has coefficients in `R` -/ @@ -106,6 +106,7 @@ section variable {S : Type*} [CommRing S] [IsDomain S] variable {φ : R →+* S} (hinj : Function.Injective φ) {f : R[X]} (hf : f.IsPrimitive) +include hinj hf theorem IsPrimitive.isUnit_iff_isUnit_map_of_injective : IsUnit f ↔ IsUnit (map φ f) := by refine ⟨(mapRingHom φ).isUnit_map, fun h => ?_⟩ @@ -133,8 +134,6 @@ theorem IsPrimitive.isUnit_iff_isUnit_map {p : R[X]} (hp : p.IsPrimitive) : IsUnit p ↔ IsUnit (p.map (algebraMap R K)) := hp.isUnit_iff_isUnit_map_of_injective (IsFractionRing.injective _ _) -variable [IsDomain R] - section IsIntegrallyClosed open IsIntegrallyClosed @@ -177,7 +176,7 @@ theorem Monic.irreducible_iff_irreducible_map_fraction_map [IsIntegrallyClosed R /-- Integrally closed domains are precisely the domains for in which Gauss's lemma holds for monic polynomials -/ -theorem isIntegrallyClosed_iff' : +theorem isIntegrallyClosed_iff' [IsDomain R] : IsIntegrallyClosed R ↔ ∀ p : R[X], p.Monic → (Irreducible p ↔ Irreducible (p.map <| algebraMap R K)) := by constructor @@ -212,7 +211,7 @@ open IsLocalization section NormalizedGCDMonoid -variable [NormalizedGCDMonoid R] +variable [IsDomain R] [NormalizedGCDMonoid R] theorem isUnit_or_eq_zero_of_isUnit_integerNormalization_primPart {p : K[X]} (h0 : p ≠ 0) (h : IsUnit (integerNormalization R⁰ p).primPart) : IsUnit p := by diff --git a/Mathlib/RingTheory/Polynomial/SeparableDegree.lean b/Mathlib/RingTheory/Polynomial/SeparableDegree.lean index 466a2146a04c1..6d5e04ba59670 100644 --- a/Mathlib/RingTheory/Polynomial/SeparableDegree.lean +++ b/Mathlib/RingTheory/Polynomial/SeparableDegree.lean @@ -76,7 +76,7 @@ theorem IsSeparableContraction.dvd_degree' {g} (hf : IsSeparableContraction q f rw [natDegree_expand] theorem HasSeparableContraction.dvd_degree' : ∃ m : ℕ, hf.degree * q ^ m = f.natDegree := - (Classical.choose_spec hf).dvd_degree' hf + (Classical.choose_spec hf).dvd_degree' /-- The separable degree divides the degree. -/ theorem HasSeparableContraction.dvd_degree : hf.degree ∣ f.natDegree := @@ -111,7 +111,7 @@ theorem contraction_degree_eq_or_insep [hq : NeZero q] [CharP F q] (g g' : F[X]) (h_expand : expand F (q ^ m) g = expand F (q ^ m') g') (hg : g.Separable) (hg' : g'.Separable) : g.natDegree = g'.natDegree := by wlog hm : m ≤ m' - · exact (this q hf g' g m' m h_expand.symm hg' hg (le_of_not_le hm)).symm + · exact (this q g' g m' m h_expand.symm hg' hg (le_of_not_le hm)).symm obtain ⟨s, rfl⟩ := exists_add_of_le hm rw [pow_add, expand_mul, expand_inj (pow_pos (NeZero.pos q) m)] at h_expand subst h_expand diff --git a/Mathlib/RingTheory/QuotientNilpotent.lean b/Mathlib/RingTheory/QuotientNilpotent.lean index ef0ee6faba2e0..79300ab5bd018 100644 --- a/Mathlib/RingTheory/QuotientNilpotent.lean +++ b/Mathlib/RingTheory/QuotientNilpotent.lean @@ -17,7 +17,6 @@ theorem Ideal.isRadical_iff_quotient_reduced {R : Type*} [CommRing R] (I : Ideal variable {R S : Type*} [CommSemiring R] [CommRing S] [Algebra R S] (I : Ideal S) - /-- Let `P` be a property on ideals. If `P` holds for square-zero ideals, and if `P I → P (J ⧸ I) → P J`, then `P` holds for all nilpotent ideals. -/ theorem Ideal.IsNilpotent.induction_on (hI : IsNilpotent I) @@ -51,7 +50,7 @@ theorem IsNilpotent.isUnit_quotient_mk_iff {R : Type*} [CommRing R] {I : Ideal R (hI : IsNilpotent I) {x : R} : IsUnit (Ideal.Quotient.mk I x) ↔ IsUnit x := by refine ⟨?_, fun h => h.map <| Ideal.Quotient.mk I⟩ revert x - apply Ideal.IsNilpotent.induction_on (R := R) (S := R) I hI <;> clear hI I + apply Ideal.IsNilpotent.induction_on (S := R) I hI <;> clear hI I swap · introv e h₁ h₂ h₃ apply h₁ diff --git a/Mathlib/RingTheory/Regular/IsSMulRegular.lean b/Mathlib/RingTheory/Regular/IsSMulRegular.lean index f2bab0d7fab5e..34012a4084c2b 100644 --- a/Mathlib/RingTheory/Regular/IsSMulRegular.lean +++ b/Mathlib/RingTheory/Regular/IsSMulRegular.lean @@ -52,6 +52,7 @@ open scoped TensorProduct variable (M) [CommRing R] [AddCommGroup M] [AddCommGroup M'] [Module R M] [Module R M'] [Module.Flat R M] {r : R} (h : IsSMulRegular M' r) +include h lemma IsSMulRegular.lTensor : IsSMulRegular (M ⊗[R] M') r := have h1 := congrArg DFunLike.coe (LinearMap.lTensor_smul_action M M' r) diff --git a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean index dc2c6b41f0733..18723ca9b73e4 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Minpoly.lean @@ -30,6 +30,7 @@ namespace IsPrimitiveRoot section CommRing variable {n : ℕ} {K : Type*} [CommRing K] {μ : K} (h : IsPrimitiveRoot μ n) +include h /-- `μ` is integral over `ℤ`. -/ -- Porting note: `hpos` was in the `variable` line, with an `omit` in mathlib3 just after this diff --git a/Mathlib/RingTheory/Smooth/Basic.lean b/Mathlib/RingTheory/Smooth/Basic.lean index 50bb4c25e4046..d61b49519ff71 100644 --- a/Mathlib/RingTheory/Smooth/Basic.lean +++ b/Mathlib/RingTheory/Smooth/Basic.lean @@ -68,7 +68,7 @@ theorem exists_lift {B : Type u} [CommRing B] [_RB : Algebra R B] revert g change Function.Surjective (Ideal.Quotient.mkₐ R I).comp revert _RB - apply Ideal.IsNilpotent.induction_on (R := B) I hI + apply Ideal.IsNilpotent.induction_on (S := B) I hI · intro B _ I hI _; exact FormallySmooth.comp_surjective I hI · intro B _ I J hIJ h₁ h₂ _ g let this : ((B ⧸ I) ⧸ J.map (Ideal.Quotient.mk I)) ≃ₐ[R] B ⧸ J := @@ -189,7 +189,7 @@ section OfSurjective variable {R S : Type u} [CommRing R] [CommSemiring S] variable {P A : Type u} [CommRing A] [Algebra R A] [CommRing P] [Algebra R P] -variable (I : Ideal P) (f : P →ₐ[R] A) (hf : Function.Surjective f) +variable (I : Ideal P) (f : P →ₐ[R] A) theorem of_split [FormallySmooth R P] (g : A →ₐ[R] P ⧸ (RingHom.ker f.toRingHom) ^ 2) (hg : f.kerSquareLift.comp g = AlgHom.id R A) : FormallySmooth R A := by @@ -211,6 +211,9 @@ theorem of_split [FormallySmooth R P] (g : A →ₐ[R] P ⧸ (RingHom.ker f.toRi exact (FormallySmooth.mk_lift I ⟨2, hI⟩ (i.comp f) x).symm exact ⟨l.comp g, by rw [← AlgHom.comp_assoc, ← this, AlgHom.comp_assoc, hg, AlgHom.comp_id]⟩ +variable (hf : Function.Surjective f) +include hf + /-- Let `P →ₐ[R] A` be a surjection with kernel `J`, and `P` a formally smooth `R`-algebra, then `A` is formally smooth over `R` iff the surjection `P ⧸ J ^ 2 →ₐ[R] A` has a section. @@ -281,6 +284,7 @@ variable (M : Submonoid R) variable [Algebra R S] [Algebra R Sₘ] [Algebra S Sₘ] [Algebra R Rₘ] [Algebra Rₘ Sₘ] variable [IsScalarTower R Rₘ Sₘ] [IsScalarTower R S Sₘ] variable [IsLocalization M Rₘ] [IsLocalization (M.map (algebraMap R S)) Sₘ] +include M -- Porting note: no longer supported -- attribute [local elab_as_elim] Ideal.IsNilpotent.induction_on diff --git a/Mathlib/RingTheory/Smooth/Kaehler.lean b/Mathlib/RingTheory/Smooth/Kaehler.lean index 1a649e6bd938c..a5825c5318f5f 100644 --- a/Mathlib/RingTheory/Smooth/Kaehler.lean +++ b/Mathlib/RingTheory/Smooth/Kaehler.lean @@ -30,13 +30,13 @@ open TensorProduct KaehlerDifferential open Function (Surjective) variable {R P S : Type u} [CommRing R] [CommRing P] [CommRing S] -variable [Algebra R S] [Algebra R P] [Algebra P S] [IsScalarTower R P S] -variable (hf : Surjective (algebraMap P S)) (hf' : (RingHom.ker (algebraMap P S)) ^ 2 = ⊥) +variable [Algebra R P] [Algebra P S] section ofSection +variable [Algebra R S] [IsScalarTower R P S] -- Suppose we have a section (as alghom) of `P →ₐ[R] S`. -variable (g : S →ₐ[R] P) (hg : (IsScalarTower.toAlgHom R P S).comp g = AlgHom.id R S) +variable (g : S →ₐ[R] P) /-- Given a surjective algebra homomorphism `f : P →ₐ[R] S` with square-zero kernel `I`, @@ -69,6 +69,10 @@ def derivationOfSectionOfKerSqZero (f : P →ₐ[R] S) (hf' : (RingHom.ker f) ^ mul_sub, AddSubmonoid.mk_add_mk, sub_mul, neg_sub] ring +variable (hf' : (RingHom.ker (algebraMap P S)) ^ 2 = ⊥) + (hg : (IsScalarTower.toAlgHom R P S).comp g = AlgHom.id R S) +include hf' hg + lemma isScalarTower_of_section_of_ker_sqZero : letI := g.toRingHom.toAlgebra; IsScalarTower P S (RingHom.ker (algebraMap P S)) := by letI := g.toRingHom.toAlgebra @@ -90,24 +94,24 @@ we get a retraction of the injection `I → S ⊗[P] Ω[P/R]`. noncomputable def retractionOfSectionOfKerSqZero : S ⊗[P] Ω[P⁄R] →ₗ[P] RingHom.ker (algebraMap P S) := letI := g.toRingHom.toAlgebra - haveI := isScalarTower_of_section_of_ker_sqZero hf' g hg + haveI := isScalarTower_of_section_of_ker_sqZero g hf' hg letI f : _ →ₗ[P] RingHom.ker (algebraMap P S) := (derivationOfSectionOfKerSqZero (IsScalarTower.toAlgHom R P S) hf' g hg).liftKaehlerDifferential (f.liftBaseChange S).restrictScalars P @[simp] lemma retractionOfSectionOfKerSqZero_tmul_D (s : S) (t : P) : - retractionOfSectionOfKerSqZero hf' g hg (s ⊗ₜ .D _ _ t) = + retractionOfSectionOfKerSqZero g hf' hg (s ⊗ₜ .D _ _ t) = g s * t - g s * g (algebraMap _ _ t) := by letI := g.toRingHom.toAlgebra - haveI := isScalarTower_of_section_of_ker_sqZero hf' g hg + haveI := isScalarTower_of_section_of_ker_sqZero g hf' hg simp only [retractionOfSectionOfKerSqZero, AlgHom.toRingHom_eq_coe, LinearMap.coe_restrictScalars, LinearMap.liftBaseChange_tmul, SetLike.val_smul_of_tower] erw [Derivation.liftKaehlerDifferential_comp_D] exact mul_sub (g s) t (g (algebraMap P S t)) lemma retractionOfSectionOfKerSqZero_comp_kerToTensor : - (retractionOfSectionOfKerSqZero hf' g hg).comp (kerToTensor R P S) = LinearMap.id := by + (retractionOfSectionOfKerSqZero g hf' hg).comp (kerToTensor R P S) = LinearMap.id := by ext x; simp [(RingHom.mem_ker _).mp x.2] end ofSection @@ -116,6 +120,7 @@ section ofRetraction variable (l : S ⊗[P] Ω[P⁄R] →ₗ[P] RingHom.ker (algebraMap P S)) variable (hl : l.comp (kerToTensor R P S) = LinearMap.id) +include hl -- suppose we have a (set-theoretic) section variable (σ : S → P) (hσ : ∀ x, algebraMap P S (σ x) = x) @@ -126,6 +131,10 @@ lemma sectionOfRetractionKerToTensorAux_prop (x y) (h : algebraMap P S x = algeb ← map_sub, ← tmul_sub, ← map_sub] exact congr_arg Subtype.val (LinearMap.congr_fun hl.symm ⟨x - y, by simp [RingHom.mem_ker, h]⟩) +variable [Algebra R S] [IsScalarTower R P S] +variable (hf' : (RingHom.ker (algebraMap P S)) ^ 2 = ⊥) +include hf' + /-- Given a surjective algebra homomorphism `f : P →ₐ[R] S` with square-zero kernel `I`. Let `σ` be an arbitrary (set-theoretic) section of `f`. @@ -152,12 +161,15 @@ def sectionOfRetractionKerToTensorAux : S →ₐ[R] P where (σ (algebraMap R S r)) (algebraMap R P r) (by simp [hσ, ← IsScalarTower.algebraMap_apply])] lemma sectionOfRetractionKerToTensorAux_algebraMap (x : P) : - sectionOfRetractionKerToTensorAux hf' l hl σ hσ (algebraMap P S x) = x - l (1 ⊗ₜ .D _ _ x) := + sectionOfRetractionKerToTensorAux l hl σ hσ hf' (algebraMap P S x) = x - l (1 ⊗ₜ .D _ _ x) := sectionOfRetractionKerToTensorAux_prop l hl _ x (by simp [hσ]) +variable (hf : Surjective (algebraMap P S)) +include hf + lemma toAlgHom_comp_sectionOfRetractionKerToTensorAux : (IsScalarTower.toAlgHom R P S).comp - (sectionOfRetractionKerToTensorAux hf' l hl σ hσ) = AlgHom.id _ _ := by + (sectionOfRetractionKerToTensorAux l hl σ hσ hf') = AlgHom.id _ _ := by ext x obtain ⟨x, rfl⟩ := hf x simp [sectionOfRetractionKerToTensorAux_algebraMap, (RingHom.mem_ker _).mp] @@ -169,21 +181,25 @@ Suppose we have a retraction `l` of the injection `I →ₗ[P] S ⊗[P] Ω[P/R]` where `σ` is an arbitrary (set-theoretic) section of `f` -/ noncomputable def sectionOfRetractionKerToTensor : S →ₐ[R] P := - sectionOfRetractionKerToTensorAux hf' l hl _ (fun x ↦ (hf x).choose_spec) + sectionOfRetractionKerToTensorAux l hl _ (fun x ↦ (hf x).choose_spec) hf' @[simp] lemma sectionOfRetractionKerToTensor_algebraMap (x : P) : - sectionOfRetractionKerToTensor hf hf' l hl (algebraMap P S x) = x - l (1 ⊗ₜ .D _ _ x) := - sectionOfRetractionKerToTensorAux_algebraMap hf' l hl _ _ x + sectionOfRetractionKerToTensor l hl hf' hf (algebraMap P S x) = x - l (1 ⊗ₜ .D _ _ x) := + sectionOfRetractionKerToTensorAux_algebraMap l hl _ _ hf' x @[simp] lemma toAlgHom_comp_sectionOfRetractionKerToTensor : (IsScalarTower.toAlgHom R P S).comp - (sectionOfRetractionKerToTensor hf hf' l hl) = AlgHom.id _ _ := - toAlgHom_comp_sectionOfRetractionKerToTensorAux hf _ _ _ _ _ + (sectionOfRetractionKerToTensor l hl hf' hf) = AlgHom.id _ _ := + toAlgHom_comp_sectionOfRetractionKerToTensorAux (hf := hf) .. end ofRetraction +variable [Algebra R S] [IsScalarTower R P S] +variable (hf' : (RingHom.ker (algebraMap P S)) ^ 2 = ⊥) (hf : Surjective (algebraMap P S)) +include hf' hf + /-- Given a surjective algebra homomorphism `f : P →ₐ[R] S` with square-zero kernel `I`, there is a one-to-one correspondence between `P`-linear retractions of `I →ₗ[P] S ⊗[P] Ω[P/R]` @@ -193,8 +209,8 @@ noncomputable def retractionKerToTensorEquivSection : { l // l ∘ₗ (kerToTensor R P S) = LinearMap.id } ≃ { g // (IsScalarTower.toAlgHom R P S).comp g = AlgHom.id R S } where - toFun l := ⟨_, toAlgHom_comp_sectionOfRetractionKerToTensor hf hf' _ l.2⟩ - invFun g := ⟨_, retractionOfSectionOfKerSqZero_comp_kerToTensor hf' _ g.2⟩ + toFun l := ⟨_, toAlgHom_comp_sectionOfRetractionKerToTensor _ l.2 hf' hf⟩ + invFun g := ⟨_, retractionOfSectionOfKerSqZero_comp_kerToTensor _ hf' g.2⟩ left_inv l := by ext s p obtain ⟨s, rfl⟩ := hf s diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index fd3b1c5921f22..395b39b238df8 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -293,6 +293,7 @@ section ExistsPrimeFactors variable [CancelCommMonoidWithZero α] variable (pf : ∀ a : α, a ≠ 0 → ∃ f : Multiset α, (∀ b ∈ f, Prime b) ∧ f.prod ~ᵤ a) +include pf theorem WfDvdMonoid.of_exists_prime_factors : WfDvdMonoid α := ⟨by @@ -914,10 +915,12 @@ theorem pow_eq_pow_iff {a : R} (ha0 : a ≠ 0) (ha1 : ¬IsUnit a) {i j : ℕ} : section multiplicity variable [NormalizationMonoid R] -variable [DecidableRel (Dvd.dvd : R → R → Prop)] open multiplicity Multiset +section +variable [DecidableRel (Dvd.dvd : R → R → Prop)] + theorem le_multiplicity_iff_replicate_le_normalizedFactors {a b : R} {n : ℕ} (ha : Irreducible a) (hb : b ≠ 0) : ↑n ≤ multiplicity a b ↔ replicate n (normalize a) ≤ normalizedFactors b := by @@ -951,6 +954,7 @@ theorem multiplicity_eq_count_normalizedFactors [DecidableEq R] {a b : R} (ha : simp rw [le_multiplicity_iff_replicate_le_normalizedFactors ha hb, ← le_count_iff_replicate_le] +end /-- The number of times an irreducible factor `p` appears in `normalizedFactors x` is defined by the number of times it divides `x`. @@ -984,13 +988,13 @@ theorem count_normalizedFactors_eq' [DecidableEq R] {p x : R} (hp : p = 0 ∨ Ir exact absurd hle hlt · exact count_normalizedFactors_eq hp hnorm hle hlt +end multiplicity + /-- Deprecated. Use `WfDvdMonoid.max_power_factor` instead. -/ @[deprecated WfDvdMonoid.max_power_factor (since := "2024-03-01")] theorem max_power_factor {a₀ x : R} (h : a₀ ≠ 0) (hx : Irreducible x) : ∃ n : ℕ, ∃ a : R, ¬x ∣ a ∧ a₀ = x ^ n * a := WfDvdMonoid.max_power_factor h hx -end multiplicity - section Multiplicative variable [CancelCommMonoidWithZero α] [UniqueFactorizationMonoid α] diff --git a/Mathlib/RingTheory/Unramified/Basic.lean b/Mathlib/RingTheory/Unramified/Basic.lean index 59ee1bcd9e292..af45772e8e3eb 100644 --- a/Mathlib/RingTheory/Unramified/Basic.lean +++ b/Mathlib/RingTheory/Unramified/Basic.lean @@ -70,7 +70,7 @@ theorem lift_unique {B : Type u} [CommRing B] [_RB : Algebra R B] revert g₁ g₂ change Function.Injective (Ideal.Quotient.mkₐ R I).comp revert _RB - apply Ideal.IsNilpotent.induction_on (R := B) I hI + apply Ideal.IsNilpotent.induction_on (S := B) I hI · intro B _ I hI _; exact FormallyUnramified.comp_injective I hI · intro B _ I J hIJ h₁ h₂ _ g₁ g₂ e apply h₁ @@ -181,13 +181,14 @@ variable {R S Rₘ Sₘ : Type u} [CommRing R] [CommRing S] [CommRing Rₘ] [Com variable (M : Submonoid R) variable [Algebra R S] [Algebra R Sₘ] [Algebra S Sₘ] [Algebra R Rₘ] [Algebra Rₘ Sₘ] variable [IsScalarTower R Rₘ Sₘ] [IsScalarTower R S Sₘ] -variable [IsLocalization M Rₘ] [IsLocalization (M.map (algebraMap R S)) Sₘ] +variable [IsLocalization (M.map (algebraMap R S)) Sₘ] +include M -- Porting note: no longer supported -- attribute [local elab_as_elim] Ideal.IsNilpotent.induction_on /-- This holds in general for epimorphisms. -/ -theorem of_isLocalization : FormallyUnramified R Rₘ := by +theorem of_isLocalization [IsLocalization M Rₘ] : FormallyUnramified R Rₘ := by constructor intro Q _ _ I _ f₁ f₂ _ apply AlgHom.coe_ringHom_injective diff --git a/Mathlib/RingTheory/Valuation/Integral.lean b/Mathlib/RingTheory/Valuation/Integral.lean index 993239cc20485..35bd2ff83d00d 100644 --- a/Mathlib/RingTheory/Valuation/Integral.lean +++ b/Mathlib/RingTheory/Valuation/Integral.lean @@ -23,6 +23,7 @@ section CommRing variable {R : Type u} {Γ₀ : Type v} [CommRing R] [LinearOrderedCommGroupWithZero Γ₀] variable {v : Valuation R Γ₀} {O : Type w} [CommRing O] [Algebra O R] (hv : Integers v O) +include hv open Polynomial @@ -49,10 +50,11 @@ end CommRing section FractionField variable {K : Type u} {Γ₀ : Type v} [Field K] [LinearOrderedCommGroupWithZero Γ₀] -variable {v : Valuation K Γ₀} {O : Type w} [CommRing O] [IsDomain O] +variable {v : Valuation K Γ₀} {O : Type w} [CommRing O] variable [Algebra O K] [IsFractionRing O K] variable (hv : Integers v O) +include hv in theorem integrallyClosed : IsIntegrallyClosed O := (IsIntegrallyClosed.integralClosure_eq_bot_iff K).mp (Valuation.Integers.integralClosure hv) diff --git a/Mathlib/RingTheory/WittVector/Truncated.lean b/Mathlib/RingTheory/WittVector/Truncated.lean index ce513738df05e..ad0dab556c7ae 100644 --- a/Mathlib/RingTheory/WittVector/Truncated.lean +++ b/Mathlib/RingTheory/WittVector/Truncated.lean @@ -39,7 +39,7 @@ open Function (Injective Surjective) noncomputable section -variable {p : ℕ} [hp : Fact p.Prime] (n : ℕ) (R : Type*) +variable {p : ℕ} (n : ℕ) (R : Type*) local notation "𝕎" => WittVector p -- type as `\bbW` @@ -151,6 +151,7 @@ theorem truncateFun_out (x : TruncatedWittVector p n R) : x.out.truncateFun n = open WittVector variable (p n R) +variable [Fact p.Prime] instance : Zero (TruncatedWittVector p n R) := ⟨truncateFun n 0⟩ @@ -215,6 +216,8 @@ variable [CommRing R] theorem truncateFun_surjective : Surjective (@truncateFun p n R) := Function.RightInverse.surjective TruncatedWittVector.truncateFun_out +variable [Fact p.Prime] + @[simp] theorem truncateFun_zero : truncateFun n (0 : 𝕎 R) = 0 := rfl @@ -267,6 +270,7 @@ open WittVector variable (p n R) variable [CommRing R] +variable [Fact p.Prime] instance instCommRing : CommRing (TruncatedWittVector p n R) := (truncateFun_surjective p n R).commRing _ (truncateFun_zero p n R) (truncateFun_one p n R) @@ -282,6 +286,7 @@ open TruncatedWittVector variable (n) variable [CommRing R] +variable [Fact p.Prime] /-- `truncate n` is a ring homomorphism that truncates `x` to its first `n` entries to obtain a `TruncatedWittVector`, which has the same base `p` as `x`. -/ @@ -306,7 +311,7 @@ theorem coeff_truncate (x : 𝕎 R) (i : Fin n) : (truncate n x).coeff i = x.coe variable (n) theorem mem_ker_truncate (x : 𝕎 R) : - x ∈ RingHom.ker (@truncate p _ n R _) ↔ ∀ i < n, x.coeff i = 0 := by + x ∈ RingHom.ker (truncate (p := p) n) ↔ ∀ i < n, x.coeff i = 0 := by simp only [RingHom.mem_ker, truncate, truncateFun, RingHom.coe_mk, TruncatedWittVector.ext_iff, TruncatedWittVector.coeff_mk, coeff_zero] exact Fin.forall_iff @@ -325,6 +330,9 @@ namespace TruncatedWittVector variable [CommRing R] +section +variable [Fact p.Prime] + /-- A ring homomorphism that truncates a truncated Witt vector of length `m` to a truncated Witt vector of length `n`, for `n ≤ m`. -/ @@ -338,7 +346,7 @@ def truncate {m : ℕ} (hm : n ≤ m) : TruncatedWittVector p m R →+* Truncate @[simp] theorem truncate_comp_wittVector_truncate {m : ℕ} (hm : n ≤ m) : - (@truncate p _ n R _ m hm).comp (WittVector.truncate m) = WittVector.truncate n := + (truncate (p := p) (R := R) hm).comp (WittVector.truncate m) = WittVector.truncate n := RingHom.liftOfRightInverse_comp _ _ _ _ @[simp] @@ -350,17 +358,17 @@ theorem truncate_wittVector_truncate {m : ℕ} (hm : n ≤ m) (x : 𝕎 R) : theorem truncate_truncate {n₁ n₂ n₃ : ℕ} (h1 : n₁ ≤ n₂) (h2 : n₂ ≤ n₃) (x : TruncatedWittVector p n₃ R) : (truncate h1) (truncate h2 x) = truncate (h1.trans h2) x := by - obtain ⟨x, rfl⟩ := @WittVector.truncate_surjective p _ n₃ R _ x + obtain ⟨x, rfl⟩ := WittVector.truncate_surjective (p := p) n₃ R x simp only [truncate_wittVector_truncate] @[simp] theorem truncate_comp {n₁ n₂ n₃ : ℕ} (h1 : n₁ ≤ n₂) (h2 : n₂ ≤ n₃) : - (@truncate p _ _ R _ _ h1).comp (truncate h2) = truncate (h1.trans h2) := by + (truncate (p := p) (R := R) h1).comp (truncate h2) = truncate (h1.trans h2) := by ext1 x; simp only [truncate_truncate, Function.comp_apply, RingHom.coe_comp] -theorem truncate_surjective {m : ℕ} (hm : n ≤ m) : Surjective (@truncate p _ _ R _ _ hm) := by +theorem truncate_surjective {m : ℕ} (hm : n ≤ m) : Surjective (truncate (p := p) (R := R) hm) := by intro x - obtain ⟨x, rfl⟩ := @WittVector.truncate_surjective p _ _ R _ x + obtain ⟨x, rfl⟩ := WittVector.truncate_surjective (p := p) _ R x exact ⟨WittVector.truncate _ x, truncate_wittVector_truncate _ _⟩ @[simp] @@ -369,6 +377,8 @@ theorem coeff_truncate {m : ℕ} (hm : n ≤ m) (i : Fin n) (x : TruncatedWittVe obtain ⟨y, rfl⟩ := @WittVector.truncate_surjective p _ _ _ _ x simp only [truncate_wittVector_truncate, WittVector.coeff_truncate, Fin.coe_castLE] +end + section Fintype instance {R : Type*} [Fintype R] : Fintype (TruncatedWittVector p n R) := @@ -382,7 +392,9 @@ theorem card {R : Type*} [Fintype R] : end Fintype -theorem iInf_ker_truncate : ⨅ i : ℕ, RingHom.ker (@WittVector.truncate p _ i R _) = ⊥ := by +variable [Fact p.Prime] + +theorem iInf_ker_truncate : ⨅ i : ℕ, RingHom.ker (WittVector.truncate (p := p) (R := R) i) = ⊥ := by rw [Submodule.eq_bot_iff] intro x hx ext @@ -398,6 +410,7 @@ open TruncatedWittVector hiding truncate coeff section lift variable [CommRing R] +variable [Fact p.Prime] variable {S : Type*} [Semiring S] variable (f : ∀ k : ℕ, S →+* TruncatedWittVector p k R) variable @@ -413,6 +426,7 @@ def liftFun (s : S) : 𝕎 R := variable {f} +include f_compat in @[simp] theorem truncate_liftFun (s : S) : WittVector.truncate n (liftFun f s) = f n s := by ext i diff --git a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean index e7d18534a96a9..8b1c1ca19d66a 100644 --- a/Mathlib/Topology/Algebra/Module/FiniteDimension.lean +++ b/Mathlib/Topology/Algebra/Module/FiniteDimension.lean @@ -485,9 +485,10 @@ section UniformAddGroup variable (𝕜 E : Type*) [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] [AddCommGroup E] [UniformSpace E] [T2Space E] [UniformAddGroup E] - [Module 𝕜 E] [ContinuousSMul 𝕜 E] [FiniteDimensional 𝕜 E] + [Module 𝕜 E] [ContinuousSMul 𝕜 E] -theorem FiniteDimensional.complete : CompleteSpace E := by +include 𝕜 in +theorem FiniteDimensional.complete [FiniteDimensional 𝕜 E] : CompleteSpace E := by set e := ContinuousLinearEquiv.ofFinrankEq (@finrank_fin_fun 𝕜 _ _ (finrank 𝕜 E)).symm have : UniformEmbedding e.toLinearEquiv.toEquiv.symm := e.symm.uniformEmbedding exact (completeSpace_congr this).1 (by infer_instance) @@ -503,20 +504,21 @@ theorem Submodule.complete_of_finiteDimensional (s : Submodule 𝕜 E) [FiniteDi end UniformAddGroup variable {𝕜 E F : Type*} [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] - [AddCommGroup E] [TopologicalSpace E] [T2Space E] [TopologicalAddGroup E] [Module 𝕜 E] + [AddCommGroup E] [TopologicalSpace E] [TopologicalAddGroup E] [Module 𝕜 E] [ContinuousSMul 𝕜 E] [AddCommGroup F] [TopologicalSpace F] [T2Space F] [TopologicalAddGroup F] [Module 𝕜 F] [ContinuousSMul 𝕜 F] /-- A finite-dimensional subspace is closed. -/ -theorem Submodule.closed_of_finiteDimensional (s : Submodule 𝕜 E) [FiniteDimensional 𝕜 s] : +theorem Submodule.closed_of_finiteDimensional + [T2Space E] (s : Submodule 𝕜 E) [FiniteDimensional 𝕜 s] : IsClosed (s : Set E) := letI := TopologicalAddGroup.toUniformSpace E haveI : UniformAddGroup E := comm_topologicalAddGroup_is_uniform s.complete_of_finiteDimensional.isClosed /-- An injective linear map with finite-dimensional domain is a closed embedding. -/ -theorem LinearMap.closedEmbedding_of_injective [FiniteDimensional 𝕜 E] {f : E →ₗ[𝕜] F} +theorem LinearMap.closedEmbedding_of_injective [T2Space E] [FiniteDimensional 𝕜 E] {f : E →ₗ[𝕜] F} (hf : LinearMap.ker f = ⊥) : ClosedEmbedding f := let g := LinearEquiv.ofInjective f (LinearMap.ker_eq_bot.mp hf) { embedding_subtype_val.comp g.toContinuousLinearEquiv.toHomeomorph.embedding with @@ -524,11 +526,12 @@ theorem LinearMap.closedEmbedding_of_injective [FiniteDimensional 𝕜 E] {f : E haveI := f.finiteDimensional_range simpa [LinearMap.range_coe f] using f.range.closed_of_finiteDimensional } -theorem closedEmbedding_smul_left {c : E} (hc : c ≠ 0) : ClosedEmbedding fun x : 𝕜 => x • c := +theorem closedEmbedding_smul_left [T2Space E] {c : E} (hc : c ≠ 0) : + ClosedEmbedding fun x : 𝕜 => x • c := LinearMap.closedEmbedding_of_injective (LinearMap.ker_toSpanSingleton 𝕜 E hc) -- `smul` is a closed map in the first argument. -theorem isClosedMap_smul_left (c : E) : IsClosedMap fun x : 𝕜 => x • c := by +theorem isClosedMap_smul_left [T2Space E] (c : E) : IsClosedMap fun x : 𝕜 => x • c := by by_cases hc : c = 0 · simp_rw [hc, smul_zero] exact isClosedMap_const diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index a9d04af3526f0..dd2ee7af35ad6 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -148,6 +148,7 @@ section tendsto_nhds variable {𝕜 : Type*} [Preorder 𝕜] [Zero 𝕜] [Mul 𝕜] [TopologicalSpace 𝕜] [ContinuousMul 𝕜] {l : Filter α} {f : α → 𝕜} {b c : 𝕜} (hb : 0 < b) +include hb theorem Filter.TendstoNhdsWithinIoi.const_mul [PosMulStrictMono 𝕜] [PosMulReflectLT 𝕜] (h : Tendsto f l (𝓝[>] c)) : Tendsto (fun a => b * f a) l (𝓝[>] (b * c)) := diff --git a/Mathlib/Topology/Algebra/MulAction.lean b/Mathlib/Topology/Algebra/MulAction.lean index 230401683c420..30ec72337a896 100644 --- a/Mathlib/Topology/Algebra/MulAction.lean +++ b/Mathlib/Topology/Algebra/MulAction.lean @@ -269,6 +269,7 @@ section AddTorsor variable (G : Type*) (P : Type*) [AddGroup G] [AddTorsor G P] [TopologicalSpace G] variable [PreconnectedSpace G] [TopologicalSpace P] [ContinuousVAdd G P] +include G in /-- An `AddTorsor` for a connected space is a connected space. This is not an instance because it loops for a group as a torsor over itself. -/ protected theorem AddTorsor.connectedSpace : ConnectedSpace P := diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean b/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean index a74825af0df42..b1e78a20544fa 100644 --- a/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean +++ b/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean @@ -344,6 +344,7 @@ view definitionaly gives the same topology on `A`. -/ variable [TopologicalSpace R] {B : ι → Submodule R A} (hB : SubmodulesRingBasis B) (hsmul : ∀ (m : A) (i : ι), ∀ᶠ a : R in 𝓝 0, a • m ∈ B i) +include hB hsmul theorem SubmodulesRingBasis.toSubmodulesBasis : SubmodulesBasis B := { inter := hB.inter diff --git a/Mathlib/Topology/Algebra/UniformGroup.lean b/Mathlib/Topology/Algebra/UniformGroup.lean index bd3442028ba4a..f667159ebd3f8 100644 --- a/Mathlib/Topology/Algebra/UniformGroup.lean +++ b/Mathlib/Topology/Algebra/UniformGroup.lean @@ -594,10 +594,10 @@ variable [TopologicalSpace α] [Group α] [TopologicalGroup α] -- β is a dense subgroup of α, inclusion is denoted by e variable [TopologicalSpace β] [Group β] -variable [FunLike hom β α] [MonoidHomClass hom β α] {e : hom} (de : DenseInducing e) +variable [FunLike hom β α] [MonoidHomClass hom β α] {e : hom} @[to_additive] -theorem tendsto_div_comap_self (x₀ : α) : +theorem tendsto_div_comap_self (de : DenseInducing e) (x₀ : α) : Tendsto (fun t : β × β => t.2 / t.1) ((comap fun p : β × β => (e p.1, e p.2)) <| 𝓝 (x₀, x₀)) (𝓝 1) := by have comm : ((fun x : α × α => x.2 / x.1) ∘ fun t : β × β => (e t.1, e t.2)) = @@ -619,16 +619,18 @@ variable {G : Type*} -- β is a dense subgroup of α, inclusion is denoted by e -- δ is a dense subgroup of γ, inclusion is denoted by f variable [TopologicalSpace α] [AddCommGroup α] [TopologicalAddGroup α] -variable [TopologicalSpace β] [AddCommGroup β] [TopologicalAddGroup β] +variable [TopologicalSpace β] [AddCommGroup β] variable [TopologicalSpace γ] [AddCommGroup γ] [TopologicalAddGroup γ] -variable [TopologicalSpace δ] [AddCommGroup δ] [TopologicalAddGroup δ] -variable [UniformSpace G] [AddCommGroup G] [UniformAddGroup G] [T0Space G] [CompleteSpace G] +variable [TopologicalSpace δ] [AddCommGroup δ] +variable [UniformSpace G] [AddCommGroup G] variable {e : β →+ α} (de : DenseInducing e) variable {f : δ →+ γ} (df : DenseInducing f) variable {φ : β →+ δ →+ G} variable (hφ : Continuous (fun p : β × δ => φ p.1 p.2)) variable {W' : Set G} (W'_nhd : W' ∈ 𝓝 (0 : G)) +include de hφ +include W'_nhd in private theorem extend_Z_bilin_aux (x₀ : α) (y₁ : δ) : ∃ U₂ ∈ comap e (𝓝 x₀), ∀ x ∈ U₂, ∀ x' ∈ U₂, (fun p : β × δ => φ p.1 p.2) (x' - x, y₁) ∈ W' := by let Nx := 𝓝 x₀ @@ -646,6 +648,9 @@ private theorem extend_Z_bilin_aux (x₀ : α) (y₁ : δ) : ∃ U₂ ∈ comap simp_rw [forall_mem_comm] exact lim W' W'_nhd +variable [UniformAddGroup G] + +include df W'_nhd in private theorem extend_Z_bilin_key (x₀ : α) (y₀ : γ) : ∃ U ∈ comap e (𝓝 x₀), ∃ V ∈ comap f (𝓝 y₀), ∀ x ∈ U, ∀ x' ∈ U, ∀ (y) (_ : y ∈ V) (y') (_ : y' ∈ V), (fun p : β × δ => φ p.1 p.2) (x', y') - (fun p : β × δ => φ p.1 p.2) (x, y) ∈ W' := by @@ -695,6 +700,8 @@ private theorem extend_Z_bilin_key (x₀ : α) (y₀ : γ) : ∃ U ∈ comap e ( open DenseInducing +variable [T0Space G] [CompleteSpace G] + /-- Bourbaki GT III.6.5 Theorem I: ℤ-bilinear continuous maps from dense images into a complete Hausdorff group extend by continuity. Note: Bourbaki assumes that α and β are also complete Hausdorff, but this is not necessary. -/ diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index c54867416cfdc..94b7c9055f65a 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -49,7 +49,7 @@ namespace Profinite namespace NobelingProof -variable {I : Type u} [LinearOrder I] [IsWellOrder I (·<·)] (C : Set (I → Bool)) +variable {I : Type u} (C : Set (I → Bool)) open Profinite ContinuousMap CategoryTheory Limits Opposite Submodule @@ -291,6 +291,8 @@ def e (i : I) : LocallyConstant C ℤ where exact (continuous_of_discreteTopology (f := fun (a : Bool) ↦ (if a then (1 : ℤ) else 0))).comp ((continuous_apply i).comp continuous_subtype_val) +variable [LinearOrder I] + /-- `Products I` is the type of lists of decreasing elements of `I`, so a typical element is `[i₁, i₂, ...]` with `i₁ > i₂ > ...`. We order `Products I` lexicographically, so `[] < [i₁, ...]`, @@ -310,7 +312,7 @@ instance : LinearOrder (Products I) := theorem lt_iff_lex_lt (l m : Products I) : l < m ↔ List.Lex (·<·) l.val m.val := by cases l; cases m; rw [Subtype.mk_lt_mk]; exact Iff.rfl -instance : IsWellFounded (Products I) (·<·) := by +instance [IsWellOrder I (· < ·)] : IsWellFounded (Products I) (·<·) := by have : (· < · : Products I → _ → _) = (fun l m ↦ List.Lex (·<·) l.val m.val) := by ext; exact lt_iff_lex_lt _ _ rw [this] @@ -427,8 +429,9 @@ theorem prop_of_isGood {l : Products I} (J : I → Prop) [∀ j, Decidable (J j end Products /-- The good products span `LocallyConstant C ℤ` if and only all the products do. -/ -theorem GoodProducts.span_iff_products : ⊤ ≤ Submodule.span ℤ (Set.range (eval C)) ↔ - ⊤ ≤ Submodule.span ℤ (Set.range (Products.eval C)) := by +theorem GoodProducts.span_iff_products [IsWellOrder I (· < ·)] : + ⊤ ≤ Submodule.span ℤ (Set.range (eval C)) ↔ + ⊤ ≤ Submodule.span ℤ (Set.range (Products.eval C)) := by refine ⟨fun h ↦ le_trans h (span_mono (fun a ⟨b, hb⟩ ↦ ⟨b.val, hb⟩)), fun h ↦ le_trans h ?_⟩ rw [span_le] rintro f ⟨l, rfl⟩ @@ -449,6 +452,8 @@ theorem GoodProducts.span_iff_products : ⊤ ≤ Submodule.span ℤ (Set.range ( end Products +variable [LinearOrder I] + section Span /-! ## The good products span @@ -522,7 +527,7 @@ def factors (x : π C (· ∈ s)) : List (LocallyConstant (π C (· ∈ s)) ℤ) List.map (fun i ↦ if x.val i = true then e (π C (· ∈ s)) i else (1 - (e (π C (· ∈ s)) i))) (s.sort (·≥·)) -theorem list_prod_apply (x : C) (l : List (LocallyConstant C ℤ)) : +theorem list_prod_apply {I} (C : Set (I → Bool)) (x : C) (l : List (LocallyConstant C ℤ)) : l.prod x = (l.map (LocallyConstant.evalMonoidHom x)).prod := by rw [← map_list_prod (LocallyConstant.evalMonoidHom x) l] rfl @@ -597,7 +602,8 @@ theorem GoodProducts.finsupp_sum_mem_span_eval {a : I} {as : List I} simp only [Products.eval, List.map, List.prod_cons] /-- If `s` is a finite subset of `I`, then the good products span. -/ -theorem GoodProducts.spanFin : ⊤ ≤ Submodule.span ℤ (Set.range (eval (π C (· ∈ s)))) := by +theorem GoodProducts.spanFin [IsWellOrder I (· < ·)] : + ⊤ ≤ Submodule.span ℤ (Set.range (eval (π C (· ∈ s)))) := by rw [span_iff_products] refine le_trans (spanFinBasis.span C s) ?_ rw [Submodule.span_le] @@ -662,7 +668,7 @@ theorem fin_comap_jointlySurjective exact ⟨(Opposite.unop J), g, h⟩ /-- The good products span all of `LocallyConstant C ℤ` if `C` is closed. -/ -theorem GoodProducts.span (hC : IsClosed C) : +theorem GoodProducts.span [IsWellOrder I (· < ·)] (hC : IsClosed C) : ⊤ ≤ Submodule.span ℤ (Set.range (eval C)) := by rw [span_iff_products] intro f _ @@ -674,6 +680,8 @@ theorem GoodProducts.span (hC : IsClosed C) : end Span +variable [IsWellOrder I (· < ·)] + section Ordinal /-! @@ -765,13 +773,13 @@ instance : IsEmpty { l // Products.isGood (∅ : Set (I → Bool)) l } := rw [subsingleton_iff.mp inferInstance (Products.eval ∅ l) 0] exact Submodule.zero_mem _ -theorem GoodProducts.linearIndependentEmpty : +theorem GoodProducts.linearIndependentEmpty {I} [LinearOrder I] : LinearIndependent ℤ (eval (∅ : Set (I → Bool))) := linearIndependent_empty_type /-- The empty list as a `Products` -/ def Products.nil : Products I := ⟨[], by simp only [List.chain'_nil]⟩ -theorem Products.lt_nil_empty : { m : Products I | m < Products.nil } = ∅ := by +theorem Products.lt_nil_empty {I} [LinearOrder I] : { m : Products I | m < Products.nil } = ∅ := by ext ⟨m, hm⟩ refine ⟨fun h ↦ ?_, by tauto⟩ simp only [Set.mem_setOf_eq, lt_iff_lex_lt, nil, List.Lex.not_nil_right] at h @@ -779,12 +787,13 @@ theorem Products.lt_nil_empty : { m : Products I | m < Products.nil } = ∅ := b instance {α : Type*} [TopologicalSpace α] [Nonempty α] : Nontrivial (LocallyConstant α ℤ) := ⟨0, 1, ne_of_apply_ne DFunLike.coe <| (Function.const_injective (β := ℤ)).ne zero_ne_one⟩ -theorem Products.isGood_nil : Products.isGood ({fun _ ↦ false} : Set (I → Bool)) Products.nil := by +theorem Products.isGood_nil {I} [LinearOrder I] : + Products.isGood ({fun _ ↦ false} : Set (I → Bool)) Products.nil := by intro h simp only [Products.lt_nil_empty, Products.eval, List.map, List.prod_nil, Set.image_empty, Submodule.span_empty, Submodule.mem_bot, one_ne_zero] at h -theorem Products.span_nil_eq_top : +theorem Products.span_nil_eq_top {I} [LinearOrder I] : Submodule.span ℤ (eval ({fun _ ↦ false} : Set (I → Bool)) '' {nil}) = ⊤ := by rw [Set.image_singleton, eq_top_iff] intro f _ @@ -822,7 +831,7 @@ instance (α : Type*) [TopologicalSpace α] : NoZeroSMulDivisors ℤ (LocallyCon simp [LocallyConstant.ext_iff] at h simpa [LocallyConstant.ext_iff] using h x -theorem GoodProducts.linearIndependentSingleton : +theorem GoodProducts.linearIndependentSingleton {I} [LinearOrder I] : LinearIndependent ℤ (eval ({fun _ ↦ false} : Set (I → Bool))) := by refine linearIndependent_unique (eval ({fun _ ↦ false} : Set (I → Bool))) ?_ simp only [eval, Products.eval, List.map, List.prod_nil, ne_eq, one_ne_zero, not_false_eq_true] @@ -1039,7 +1048,8 @@ theorem smaller_mono {o₁ o₂ : Ordinal} (h : o₁ ≤ o₂) : smaller C o₁ end GoodProducts -variable {o : Ordinal} (ho : o.IsLimit) (hsC : contained C o) +variable {o : Ordinal} (ho : o.IsLimit) +include ho theorem Products.limitOrdinal (l : Products I) : l.isGood (π C (ord I · < o)) ↔ ∃ (o' : Ordinal), o' < o ∧ l.isGood (π C (ord I · < o')) := by @@ -1057,6 +1067,9 @@ theorem Products.limitOrdinal (l : Products I) : l.isGood (π C (ord I · < o)) rwa [eval_πs_image' C (le_of_lt hslt) hlt, ← eval_πs' C (le_of_lt hslt) hlt, Submodule.apply_mem_span_image_iff_mem_span (injective_πs' C _)] +variable (hsC : contained C o) +include hsC + theorem GoodProducts.union : range C = ⋃ (e : {o' // o' < o}), (smaller C e.val) := by ext p simp only [smaller, range, Set.mem_iUnion, Set.mem_image, Set.mem_range, Subtype.exists] @@ -1082,9 +1095,9 @@ theorem GoodProducts.range_equiv_factorization : (fun (p : ⋃ (e : {o' // o' < o}), (smaller C e.val)) ↦ p.1) ∘ (range_equiv C ho hsC).toFun = (fun (p : range C) ↦ (p.1 : LocallyConstant C ℤ)) := rfl -theorem GoodProducts.linearIndependent_iff_union_smaller {o : Ordinal} (ho : o.IsLimit) - (hsC : contained C o) : LinearIndependent ℤ (GoodProducts.eval C) ↔ - LinearIndependent ℤ (fun (p : ⋃ (e : {o' // o' < o}), (smaller C e.val)) ↦ p.1) := by +theorem GoodProducts.linearIndependent_iff_union_smaller : + LinearIndependent ℤ (GoodProducts.eval C) ↔ + LinearIndependent ℤ (fun (p : ⋃ (e : {o' // o' < o}), (smaller C e.val)) ↦ p.1) := by rw [GoodProducts.linearIndependent_iff_range, ← range_equiv_factorization C ho hsC] exact linearIndependent_equiv (range_equiv C ho hsC) @@ -1166,11 +1179,13 @@ def C0 := C ∩ {f | f (term I ho) = false} /-- The subset of `C` consisting of those elements whose `o`-th entry is `true`. -/ def C1 := C ∩ {f | f (term I ho) = true} +include hC in theorem isClosed_C0 : IsClosed (C0 C ho) := by refine hC.inter ?_ have h : Continuous (fun (f : I → Bool) ↦ f (term I ho)) := continuous_apply (term I ho) exact IsClosed.preimage h (t := {false}) (isClosed_discrete _) +include hC in theorem isClosed_C1 : IsClosed (C1 C ho) := by refine hC.inter ?_ have h : Continuous (fun (f : I → Bool) ↦ f (term I ho)) := continuous_apply (term I ho) @@ -1190,6 +1205,7 @@ this set. -/ def C' := C0 C ho ∩ π (C1 C ho) (ord I · < o) +include hC in theorem isClosed_C' : IsClosed (C' C ho) := IsClosed.inter (isClosed_C0 _ hC _) (isClosed_proj _ _ (isClosed_C1 _ hC _)) @@ -1213,6 +1229,7 @@ theorem continuous_swapTrue : variable {o} +include hsC in theorem swapTrue_mem_C1 (f : π (C1 C ho) (ord I · < o)) : SwapTrue o f.val ∈ C1 C ho := by obtain ⟨f, g, hg, rfl⟩ := f @@ -1269,6 +1286,7 @@ theorem CC_comp_zero : ∀ y, (Linear_CC' C hsC ho) ((πs C o) y) = 0 := by intro h₁ h₂ exact (h₁.ne h₂).elim +include hsC in theorem C0_projOrd {x : I → Bool} (hx : x ∈ C0 C ho) : Proj (ord I · < o) x = x := by ext i simp only [Proj, Set.mem_setOf, ite_eq_left_iff, not_lt] @@ -1283,6 +1301,7 @@ theorem C0_projOrd {x : I → Bool} (hx : x ∈ C0 C ho) : Proj (ord I · < o) x rw [eq_comm, ord_term ho] at hi rw [← hx.2, hi] +include hsC in theorem C1_projOrd {x : I → Bool} (hx : x ∈ C1 C ho) : SwapTrue o (Proj (ord I · < o) x) = x := by ext i dsimp [SwapTrue, Proj] @@ -1297,6 +1316,7 @@ theorem C1_projOrd {x : I → Bool} (hx : x ∈ C1 C ho) : SwapTrue o (Proj (ord simp only [not_lt, Bool.not_eq_true, Order.succ_le_iff] at hsC exact (hsC h').symm +include hC in open scoped Classical in theorem CC_exact {f : LocallyConstant C ℤ} (hf : Linear_CC' C hsC ho f = 0) : ∃ y, πs C o y = f := by @@ -1338,6 +1358,7 @@ theorem succ_mono : CategoryTheory.Mono (ModuleCat.ofHom (πs C o)) := by rw [ModuleCat.mono_iff_injective] exact injective_πs _ _ +include hC in theorem succ_exact : (ShortComplex.mk (ModuleCat.ofHom (πs C o)) (ModuleCat.ofHom (Linear_CC' C hsC ho)) (by ext; apply CC_comp_zero)).Exact := by @@ -1357,6 +1378,7 @@ The `GoodProducts` in `C` that contain `o` (they necessarily start with `o`, see -/ def MaxProducts : Set (Products I) := {l | l.isGood C ∧ term I ho ∈ l.val} +include hsC in theorem union_succ : GoodProducts C = GoodProducts (π C (ord I · < o)) ∪ MaxProducts C ho := by ext l simp only [GoodProducts, MaxProducts, Set.mem_union, Set.mem_setOf_eq] @@ -1439,12 +1461,14 @@ def SumEval : GoodProducts (π C (ord I · < o)) ⊕ MaxProducts C ho → LocallyConstant C ℤ := Sum.elim (fun l ↦ l.1.eval C) (fun l ↦ l.1.eval C) +include hsC in theorem linearIndependent_iff_sum : LinearIndependent ℤ (eval C) ↔ LinearIndependent ℤ (SumEval C ho) := by rw [← linearIndependent_equiv (sum_equiv C hsC ho), SumEval, ← sum_equiv_comp_eval_eq_elim C hsC ho] exact Iff.rfl +include hsC in theorem span_sum : Set.range (eval C) = Set.range (Sum.elim (fun (l : GoodProducts (π C (ord I · < o))) ↦ Products.eval C l.1) (fun (l : MaxProducts C ho) ↦ Products.eval C l.1)) := by @@ -1483,6 +1507,7 @@ theorem Products.max_eq_o_cons_tail' [Inhabited I] (l : Products I) (hl : l.val simp_rw [← max_eq_o_cons_tail ho l hl hlh] rfl +include hsC in theorem GoodProducts.head!_eq_o_of_maxProducts [Inhabited I] (l : ↑(MaxProducts C ho)) : l.val.val.head! = term I ho := by rw [eq_comm, ← ord_term ho] @@ -1496,13 +1521,14 @@ theorem GoodProducts.head!_eq_o_of_maxProducts [Inhabited I] (l : ↑(MaxProduct exact Products.rel_head!_of_mem hm rwa [ord_term_aux] at h +include hsC in theorem GoodProducts.max_eq_o_cons_tail (l : MaxProducts C ho) : l.val.val = (term I ho) :: l.val.Tail.val := have : Inhabited I := ⟨term I ho⟩ Products.max_eq_o_cons_tail ho l.val (List.ne_nil_of_mem l.prop.2) (head!_eq_o_of_maxProducts _ hsC ho l) -theorem Products.evalCons {l : List I} {a : I} +theorem Products.evalCons {I} [LinearOrder I] {C : Set (I → Bool)} {l : List I} {a : I} (hla : (a::l).Chain' (·>·)) : Products.eval C ⟨a::l,hla⟩ = (e C a) * Products.eval C ⟨l,List.Chain'.sublist hla (List.tail_sublist (a::l))⟩ := by simp only [eval.eq_1, List.map, List.prod_cons] @@ -1550,6 +1576,7 @@ theorem max_eq_eval_unapply : ext1 l exact max_eq_eval _ _ _ _ +include hsC in theorem chain'_cons_of_lt (l : MaxProducts C ho) (q : Products I) (hq : q < l.val.Tail) : List.Chain' (fun x x_1 ↦ x > x_1) (term I ho :: q.val) := by @@ -1566,6 +1593,7 @@ theorem chain'_cons_of_lt (l : MaxProducts C ho) rw [max_eq_o_cons_tail C hsC ho l, List.chain'_iff_pairwise] at this exact List.rel_of_pairwise_cons this (List.head!_mem_self hM) +include hsC in theorem good_lt_maxProducts (q : GoodProducts (π C (ord I · < o))) (l : MaxProducts C ho) : List.Lex (·<·) q.val.val l.val.val := by have : Inhabited I := ⟨term I ho⟩ @@ -1578,6 +1606,7 @@ theorem good_lt_maxProducts (q : GoodProducts (π C (ord I · < o))) simp only [term, Ordinal.typein_enum] exact Products.prop_of_isGood C _ q.prop q.val.val.head! (List.head!_mem_self h) +include hC hsC in /-- Removing the leading `o` from a term of `MaxProducts C` yields a list which `isGood` with respect to `C'`. @@ -1667,6 +1696,7 @@ theorem maxToGood_injective dsimp [MaxToGood] at h rw [max_eq_o_cons_tail C hsC ho m, max_eq_o_cons_tail C hsC ho n, h] +include hC in theorem linearIndependent_comp_of_eval (h₁ : ⊤ ≤ Submodule.span ℤ (Set.range (eval (π C (ord I · < o))))) : LinearIndependent ℤ (eval (C' C ho)) → @@ -1749,6 +1779,7 @@ def GoodProducts.Basis (hC : IsClosed C) : end Induction variable {S : Profinite} {ι : S → I → Bool} (hι : ClosedEmbedding ι) +include hι /-- Given a profinite set `S` and a closed embedding `S → (I → Bool)`, the `ℤ`-module diff --git a/Mathlib/Topology/Category/Stonean/Limits.lean b/Mathlib/Topology/Category/Stonean/Limits.lean index 49cd92be54a6e..9a344de319c22 100644 --- a/Mathlib/Topology/Category/Stonean/Limits.lean +++ b/Mathlib/Topology/Category/Stonean/Limits.lean @@ -25,6 +25,7 @@ instance : HasExplicitFiniteCoproducts.{w, u} (fun Y ↦ ExtremallyDisconnected hasProp _ := { hasProp := show ExtremallyDisconnected (Σ (_a : _), _) from inferInstance} variable {X Y Z : Stonean} {f : X ⟶ Z} (i : Y ⟶ Z) (hi : OpenEmbedding f) +include hi lemma extremallyDisconnected_preimage : ExtremallyDisconnected (i ⁻¹' (Set.range f)) where open_closure U hU := by diff --git a/Mathlib/Topology/Connected/Clopen.lean b/Mathlib/Topology/Connected/Clopen.lean index e6f92357c8096..e99f549d58af3 100644 --- a/Mathlib/Topology/Connected/Clopen.lean +++ b/Mathlib/Topology/Connected/Clopen.lean @@ -135,6 +135,7 @@ section disjoint_subsets variable [PreconnectedSpace α] {s : ι → Set α} (h_nonempty : ∀ i, (s i).Nonempty) (h_disj : Pairwise (Disjoint on s)) +include h_nonempty h_disj /-- In a preconnected space, any disjoint family of non-empty clopen subsets has at most one element. -/ diff --git a/Mathlib/Topology/Covering.lean b/Mathlib/Topology/Covering.lean index 7a1b2b120f702..97aae6bbbf86b 100644 --- a/Mathlib/Topology/Covering.lean +++ b/Mathlib/Topology/Covering.lean @@ -141,6 +141,7 @@ theorem mk (F : X → Type*) [∀ x, TopologicalSpace (F x)] [∀ x, DiscreteTop variable {f} variable (hf : IsCoveringMap f) +include hf protected theorem continuous : Continuous f := continuous_iff_continuousOn_univ.mpr hf.isCoveringMapOn.continuousOn diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index ec2206468421b..62a0b258ffe46 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -23,7 +23,7 @@ open TopologicalSpace Set Filter open Topology Filter variable {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] {f : α → β} -variable {s : Set β} {ι : Type*} {U : ι → Opens β} (hU : iSup U = ⊤) +variable {s : Set β} {ι : Type*} {U : ι → Opens β} theorem Set.restrictPreimage_inducing (s : Set β) (h : Inducing f) : Inducing (s.restrictPreimage f) := by @@ -78,6 +78,9 @@ theorem IsOpenMap.restrictPreimage (H : IsOpenMap f) (s : Set β) : theorem Set.restrictPreimage_isOpenMap (s : Set β) (H : IsOpenMap f) : IsOpenMap (s.restrictPreimage f) := H.restrictPreimage s +variable (hU : iSup U = ⊤) +include hU + theorem isOpen_iff_inter_of_iSup_eq_top (s : Set β) : IsOpen s ↔ ∀ i, IsOpen (s ∩ U i) := by constructor · exact fun H i => H.inter (U i).2 diff --git a/Mathlib/Topology/MetricSpace/Contracting.lean b/Mathlib/Topology/MetricSpace/Contracting.lean index 27b133d78aa62..5831e19b36935 100644 --- a/Mathlib/Topology/MetricSpace/Contracting.lean +++ b/Mathlib/Topology/MetricSpace/Contracting.lean @@ -40,7 +40,7 @@ def ContractingWith [EMetricSpace α] (K : ℝ≥0) (f : α → α) := namespace ContractingWith -variable [EMetricSpace α] [cs : CompleteSpace α] {K : ℝ≥0} {f : α → α} +variable [EMetricSpace α] {K : ℝ≥0} {f : α → α} open EMetric Set @@ -80,6 +80,9 @@ theorem restrict (hf : ContractingWith K f) {s : Set α} (hs : MapsTo f s s) : ContractingWith K (hs.restrict f s s) := ⟨hf.1, fun x y ↦ hf.2 x y⟩ +section +variable [CompleteSpace α] + /-- Banach fixed-point theorem, contraction mapping theorem, `EMetricSpace` version. A contracting map on a complete metric space has a fixed point. We include more conclusions in this theorem to avoid proving them again later. @@ -143,6 +146,8 @@ theorem efixedPoint_eq_of_edist_lt_top (hf : ContractingWith K f) {x : α} (hx : trans y exacts [lt_top_iff_ne_top.2 h, hf.edist_efixedPoint_lt_top hy] +end + /-- Banach fixed-point theorem for maps contracting on a complete subset. -/ theorem exists_fixedPoint' {s : Set α} (hsc : IsComplete s) (hsf : MapsTo f s s) (hf : ContractingWith K <| hsf.restrict f s s) {x : α} (hxs : x ∈ s) (hx : edist x (f x) ≠ ∞) : @@ -230,11 +235,15 @@ end ContractingWith namespace ContractingWith -variable [MetricSpace α] {K : ℝ≥0} {f : α → α} (hf : ContractingWith K f) +variable [MetricSpace α] {K : ℝ≥0} {f : α → α} theorem one_sub_K_pos (hf : ContractingWith K f) : (0 : ℝ) < 1 - K := sub_pos.2 hf.1 +section +variable (hf : ContractingWith K f) +include hf + theorem dist_le_mul (x y : α) : dist (f x) (f y) ≤ K * dist x y := hf.toLipschitzWith.dist_le_mul x y @@ -301,6 +310,10 @@ theorem fixedPoint_lipschitz_in_map {g : α → α} (hg : ContractingWith K g) { (hfg : ∀ z, dist (f z) (g z) ≤ C) : dist (fixedPoint f hf) (fixedPoint g hg) ≤ C / (1 - K) := hf.dist_fixedPoint_fixedPoint_of_dist_le' g hf.fixedPoint_isFixedPt hg.fixedPoint_isFixedPt hfg +end + +variable [Nonempty α] [CompleteSpace α] + /-- If a map `f` has a contracting iterate `f^[n]`, then the fixed point of `f^[n]` is also a fixed point of `f`. -/ theorem isFixedPt_fixedPoint_iterate {n : ℕ} (hf : ContractingWith K f^[n]) : diff --git a/Mathlib/Topology/MetricSpace/Isometry.lean b/Mathlib/Topology/MetricSpace/Isometry.lean index 9e0c828a91991..b2464d0e4b65a 100644 --- a/Mathlib/Topology/MetricSpace/Isometry.lean +++ b/Mathlib/Topology/MetricSpace/Isometry.lean @@ -498,6 +498,7 @@ theorem diam_image (s : Set α) : Metric.diam (h '' s) = Metric.diam s := theorem diam_preimage (s : Set β) : Metric.diam (h ⁻¹' s) = Metric.diam s := by rw [← image_symm, diam_image] +include h in theorem diam_univ : Metric.diam (univ : Set α) = Metric.diam (univ : Set β) := congr_arg ENNReal.toReal h.ediam_univ diff --git a/Mathlib/Topology/Order/LawsonTopology.lean b/Mathlib/Topology/Order/LawsonTopology.lean index 9651b391ce05a..4c730963736e1 100644 --- a/Mathlib/Topology/Order/LawsonTopology.lean +++ b/Mathlib/Topology/Order/LawsonTopology.lean @@ -203,10 +203,11 @@ lemma lawsonClosed_iff_scottClosed_of_isLowerSet (s : Set α) (h : IsLowerSet s) rw [← @isOpen_compl_iff, ← isOpen_compl_iff, (lawsonOpen_iff_scottOpen_of_isUpperSet' L S _ (isUpperSet_compl.mpr h))] +include S in /-- A lower set is Lawson closed if and only if it is closed under sups of directed sets -/ lemma lawsonClosed_iff_dirSupClosed_of_isLowerSet (s : Set α) (h : IsLowerSet s) : IsClosed[L] s ↔ DirSupClosed s := by - rw [(lawsonClosed_iff_scottClosed_of_isLowerSet L S _ h), + rw [lawsonClosed_iff_scottClosed_of_isLowerSet L S _ h, @IsScott.isClosed_iff_isLowerSet_and_dirSupClosed] aesop diff --git a/Mathlib/Topology/Order/LeftRightLim.lean b/Mathlib/Topology/Order/LeftRightLim.lean index b7a3ee1b41c06..09a573dd0a4b5 100644 --- a/Mathlib/Topology/Order/LeftRightLim.lean +++ b/Mathlib/Topology/Order/LeftRightLim.lean @@ -89,6 +89,7 @@ namespace Monotone variable {α β : Type*} [LinearOrder α] [ConditionallyCompleteLinearOrder β] [TopologicalSpace β] [OrderTopology β] {f : α → β} (hf : Monotone f) {x y : α} +include hf theorem leftLim_eq_sSup [TopologicalSpace α] [OrderTopology α] (h : 𝓝[<] x ≠ ⊥) : leftLim f x = sSup (f '' Iio x) := @@ -253,6 +254,7 @@ namespace Antitone variable {α β : Type*} [LinearOrder α] [ConditionallyCompleteLinearOrder β] [TopologicalSpace β] [OrderTopology β] {f : α → β} (hf : Antitone f) {x y : α} +include hf theorem le_leftLim (h : x ≤ y) : f y ≤ leftLim f x := hf.dual_right.leftLim_le h diff --git a/Mathlib/Topology/SeparatedMap.lean b/Mathlib/Topology/SeparatedMap.lean index c43a9904d878f..67633a9c6069d 100644 --- a/Mathlib/Topology/SeparatedMap.lean +++ b/Mathlib/Topology/SeparatedMap.lean @@ -105,7 +105,7 @@ theorem IsSeparatedMap.pullback {f : X → Y} (sep : IsSeparatedMap f) (g : A refine sep.preimage (Continuous.mapPullback ?_ ?_) <;> apply_rules [continuous_fst, continuous_subtype_val, Continuous.comp] -theorem IsSeparatedMap.comp_left {f : X → Y} (sep : IsSeparatedMap f) {g : Y → A} +theorem IsSeparatedMap.comp_left {A} {f : X → Y} (sep : IsSeparatedMap f) {g : Y → A} (inj : g.Injective) : IsSeparatedMap (g ∘ f) := fun x₁ x₂ he ↦ sep x₁ x₂ (inj he) theorem IsSeparatedMap.comp_right {f : X → Y} (sep : IsSeparatedMap f) {g : A → X} @@ -159,7 +159,7 @@ theorem discreteTopology_iff_locallyInjective (y : Y) : convert hU; ext x'; refine ⟨?_, fun h ↦ inj h (mem_of_mem_nhds hU) rfl⟩ rintro rfl; exact mem_of_mem_nhds hU -theorem IsLocallyInjective.comp_left {f : X → Y} (hf : IsLocallyInjective f) {g : Y → A} +theorem IsLocallyInjective.comp_left {A} {f : X → Y} (hf : IsLocallyInjective f) {g : Y → A} (hg : g.Injective) : IsLocallyInjective (g ∘ f) := fun x ↦ let ⟨U, hU, hx, inj⟩ := hf x; ⟨U, hU, hx, hg.comp_injOn inj⟩ @@ -171,34 +171,37 @@ theorem IsLocallyInjective.comp_right {f : X → Y} (hf : IsLocallyInjective f) section eqLocus -variable {f : X → Y} (sep : IsSeparatedMap f) (inj : IsLocallyInjective f) - {g₁ g₂ : A → X} (h₁ : Continuous g₁) (h₂ : Continuous g₂) +variable {f : X → Y} {g₁ g₂ : A → X} (h₁ : Continuous g₁) (h₂ : Continuous g₂) +include h₁ h₂ -theorem IsSeparatedMap.isClosed_eqLocus (he : f ∘ g₁ = f ∘ g₂) : IsClosed {a | g₁ a = g₂ a} := +theorem IsSeparatedMap.isClosed_eqLocus (sep : IsSeparatedMap f) (he : f ∘ g₁ = f ∘ g₂) : + IsClosed {a | g₁ a = g₂ a} := let g : A → f.Pullback f := fun a ↦ ⟨⟨g₁ a, g₂ a⟩, congr_fun he a⟩ (isSeparatedMap_iff_isClosed_diagonal.mp sep).preimage (by fun_prop : Continuous g) -theorem IsLocallyInjective.isOpen_eqLocus (he : f ∘ g₁ = f ∘ g₂) : IsOpen {a | g₁ a = g₂ a} := +theorem IsLocallyInjective.isOpen_eqLocus (inj : IsLocallyInjective f) (he : f ∘ g₁ = f ∘ g₂) : + IsOpen {a | g₁ a = g₂ a} := let g : A → f.Pullback f := fun a ↦ ⟨⟨g₁ a, g₂ a⟩, congr_fun he a⟩ (isLocallyInjective_iff_isOpen_diagonal.mp inj).preimage (by fun_prop : Continuous g) end eqLocus -variable {E A : Type*} [TopologicalSpace E] [TopologicalSpace A] {p : E → X} +variable {X E A : Type*} [TopologicalSpace E] [TopologicalSpace A] {p : E → X} namespace IsSeparatedMap -variable (sep : IsSeparatedMap p) (inj : IsLocallyInjective p) {s : Set A} (hs : IsPreconnected s) - {g g₁ g₂ : A → E} +variable {s : Set A} {g g₁ g₂ : A → E} (sep : IsSeparatedMap p) (inj : IsLocallyInjective p) +include sep inj /-- If `p` is a locally injective separated map, and `A` is a connected space, then two lifts `g₁, g₂ : A → E` of a map `f : A → X` are equal if they agree at one point. -/ -theorem eq_of_comp_eq [PreconnectedSpace A] (h₁ : Continuous g₁) (h₂ : Continuous g₂) +theorem eq_of_comp_eq + [PreconnectedSpace A] (h₁ : Continuous g₁) (h₂ : Continuous g₂) (he : p ∘ g₁ = p ∘ g₂) (a : A) (ha : g₁ a = g₂ a) : g₁ = g₂ := funext fun a' ↦ by apply (IsClopen.eq_univ ⟨sep.isClosed_eqLocus h₁ h₂ he, inj.isOpen_eqLocus h₁ h₂ he⟩ ⟨a, ha⟩).symm ▸ Set.mem_univ a' -theorem eqOn_of_comp_eqOn (h₁ : ContinuousOn g₁ s) (h₂ : ContinuousOn g₂ s) +theorem eqOn_of_comp_eqOn (hs : IsPreconnected s) (h₁ : ContinuousOn g₁ s) (h₂ : ContinuousOn g₂ s) (he : s.EqOn (p ∘ g₁) (p ∘ g₂)) {a : A} (has : a ∈ s) (ha : g₁ a = g₂ a) : s.EqOn g₁ g₂ := by rw [← Set.restrict_eq_restrict_iff] at he ⊢ rw [continuousOn_iff_continuous_restrict] at h₁ h₂ @@ -209,7 +212,7 @@ theorem const_of_comp [PreconnectedSpace A] (cont : Continuous g) (he : ∀ a a', p (g a) = p (g a')) (a a') : g a = g a' := congr_fun (sep.eq_of_comp_eq inj cont continuous_const (funext fun a ↦ he a a') a' rfl) a -theorem constOn_of_comp (cont : ContinuousOn g s) +theorem constOn_of_comp (hs : IsPreconnected s) (cont : ContinuousOn g s) (he : ∀ a ∈ s, ∀ a' ∈ s, p (g a) = p (g a')) {a a'} (ha : a ∈ s) (ha' : a' ∈ s) : g a = g a' := sep.eqOn_of_comp_eqOn inj hs cont continuous_const.continuousOn diff --git a/Mathlib/Topology/Separation.lean b/Mathlib/Topology/Separation.lean index af4b880e2169a..949fd616fcc6b 100644 --- a/Mathlib/Topology/Separation.lean +++ b/Mathlib/Topology/Separation.lean @@ -337,7 +337,6 @@ instance SeparationQuotient.instT0Space : T0Space (SeparationQuotient X) := theorem minimal_nonempty_closed_subsingleton [T0Space X] {s : Set X} (hs : IsClosed s) (hmin : ∀ t, t ⊆ s → t.Nonempty → IsClosed t → t = s) : s.Subsingleton := by - clear Y -- Porting note: added refine fun x hx y hy => of_not_not fun hxy => ?_ rcases exists_isOpen_xor'_mem hxy with ⟨U, hUo, hU⟩ wlog h : x ∈ U ∧ y ∉ U @@ -361,7 +360,6 @@ theorem IsClosed.exists_closed_singleton [T0Space X] [CompactSpace X] {S : Set X theorem minimal_nonempty_open_subsingleton [T0Space X] {s : Set X} (hs : IsOpen s) (hmin : ∀ t, t ⊆ s → t.Nonempty → IsOpen t → t = s) : s.Subsingleton := by - clear Y -- Porting note: added refine fun x hx y hy => of_not_not fun hxy => ?_ rcases exists_isOpen_xor'_mem hxy with ⟨U, hUo, hU⟩ wlog h : x ∈ U ∧ y ∉ U @@ -922,16 +920,16 @@ theorem Set.Subsingleton.isGδ_compl {s : Set X} [T1Space X] (hs : s.Subsingleto theorem Finset.isGδ_compl [T1Space X] (s : Finset X) : IsGδ (sᶜ : Set X) := s.finite_toSet.isGδ_compl -variable [FirstCountableTopology X] - -protected theorem IsGδ.singleton [T1Space X] (x : X) : IsGδ ({x} : Set X) := by +protected theorem IsGδ.singleton [FirstCountableTopology X] [T1Space X] (x : X) : + IsGδ ({x} : Set X) := by rcases (nhds_basis_opens x).exists_antitone_subbasis with ⟨U, hU, h_basis⟩ rw [← biInter_basis_nhds h_basis.toHasBasis] exact .biInter (to_countable _) fun n _ => (hU n).2.isGδ @[deprecated (since := "2024-02-15")] alias isGδ_singleton := IsGδ.singleton -theorem Set.Finite.isGδ {s : Set X} [T1Space X] (hs : s.Finite) : IsGδ s := +theorem Set.Finite.isGδ [FirstCountableTopology X] {s : Set X} [T1Space X] (hs : s.Finite) : + IsGδ s := Finite.induction_on hs .empty fun _ _ ↦ .union (.singleton _) theorem SeparationQuotient.t1Space_iff : T1Space (SeparationQuotient X) ↔ R0Space X := by @@ -2026,7 +2024,7 @@ alias separatedNhds_of_isCompact_isClosed := SeparatedNhds.of_isCompact_isClosed /-- This technique to witness `HasSeparatingCover` in regular Lindelöf topological spaces will be used to prove regular Lindelöf spaces are normal. -/ -lemma IsClosed.HasSeparatingCover {s t : Set X} [r : RegularSpace X] [LindelofSpace X] +lemma IsClosed.HasSeparatingCover {s t : Set X} [LindelofSpace X] (s_cl : IsClosed s) (t_cl : IsClosed t) (st_dis : Disjoint s t) : HasSeparatingCover s t := by -- `IsLindelof.indexed_countable_subcover` requires the space be Nonempty rcases isEmpty_or_nonempty X with empty_X | nonempty_X @@ -2037,7 +2035,7 @@ lemma IsClosed.HasSeparatingCover {s t : Set X} [r : RegularSpace X] [LindelofSp have (a : X) : ∃ n : Set X, IsOpen n ∧ Disjoint (closure n) t ∧ (a ∈ s → a ∈ n) := by wlog ains : a ∈ s · exact ⟨∅, isOpen_empty, SeparatedNhds.empty_left t |>.disjoint_closure_left, fun a ↦ ains a⟩ - obtain ⟨n, nna, ncl, nsubkc⟩ := ((regularSpace_TFAE X).out 0 3 :).mp r a tᶜ <| + obtain ⟨n, nna, ncl, nsubkc⟩ := ((regularSpace_TFAE X).out 0 3 :).mp ‹RegularSpace X› a tᶜ <| t_cl.compl_mem_nhds (disjoint_left.mp st_dis ains) exact ⟨interior n, diff --git a/Mathlib/Topology/StoneCech.lean b/Mathlib/Topology/StoneCech.lean index 5d91050675249..629e0e1ffa49c 100644 --- a/Mathlib/Topology/StoneCech.lean +++ b/Mathlib/Topology/StoneCech.lean @@ -266,6 +266,7 @@ theorem preStoneCech_hom_ext {g₁ g₂ : PreStoneCech α → β} (h₁ : Contin variable [CompactSpace β] variable {g : α → β} (hg : Continuous g) +include hg lemma preStoneCechCompat {F G : Ultrafilter α} {x : α} (hF : ↑F ≤ 𝓝 x) (hG : ↑G ≤ 𝓝 x) : Ultrafilter.extend g F = Ultrafilter.extend g G := by diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index 83b0c3960e575..fbccca79fa36f 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -230,7 +230,11 @@ lemma isCompact_preimage [TopologicalSpace β] apply subset_mulTSupport aesop -variable [T2Space α'] (hf : HasCompactMulSupport f) {g : α → α'} (cont : Continuous g) +variable [T2Space α'] + +section +variable (hf : HasCompactMulSupport f) {g : α → α'} (cont : Continuous g) +include hf cont @[to_additive] theorem mulTSupport_extend_one_subset : @@ -250,6 +254,8 @@ theorem mulTSupport_extend_one (inj : g.Injective) : (image_closure_subset_closure_image cont).trans (closure_mono (mulSupport_extend_one inj).superset) +end + @[to_additive] theorem continuous_extend_one [TopologicalSpace β] {U : Set α'} (hU : IsOpen U) {f : U → β} (cont : Continuous f) (supp : HasCompactMulSupport f) : @@ -262,7 +268,7 @@ theorem continuous_extend_one [TopologicalSpace β] {U : Set α'} (hU : IsOpen U /-- If `f` has compact multiplicative support, then `f` tends to 1 at infinity. -/ @[to_additive "If `f` has compact support, then `f` tends to zero at infinity."] -theorem is_one_at_infty {f : α → γ} [TopologicalSpace γ] [One γ] +theorem is_one_at_infty {f : α → γ} [TopologicalSpace γ] (h : HasCompactMulSupport f) : Tendsto f (cocompact α) (𝓝 1) := by intro N hN rw [mem_map, mem_cocompact'] @@ -276,7 +282,7 @@ end HasCompactMulSupport section Compact -variable [CompactSpace α] [One γ] [TopologicalSpace γ] +variable [CompactSpace α] /-- In a compact space `α`, any function has compact support. -/ @[to_additive] @@ -344,7 +350,7 @@ end MulZeroClass section OrderedAddGroup -variable {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] [AddGroup β] [Lattice β] +variable {α β : Type*} [TopologicalSpace α] [AddGroup β] [Lattice β] [CovariantClass β β (· + ·) (· ≤ ·)] protected theorem HasCompactSupport.abs {f : α → β} (hf : HasCompactSupport f) : diff --git a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean index 1e7afc1c45c52..b674ec1daf22e 100644 --- a/Mathlib/Topology/UniformSpace/UniformEmbedding.lean +++ b/Mathlib/Topology/UniformSpace/UniformEmbedding.lean @@ -395,6 +395,7 @@ variable {α : Type*} {β : Type*} {γ : Type*} [UniformSpace α] [UniformSpace local notation "ψ" => DenseInducing.extend (UniformInducing.denseInducing h_e h_dense) f +include h_e h_dense h_f in theorem uniformly_extend_exists [CompleteSpace γ] (a : α) : ∃ c, Tendsto f (comap e (𝓝 a)) (𝓝 c) := let de := h_e.denseInducing h_dense have : Cauchy (𝓝 a) := cauchy_nhds @@ -419,10 +420,12 @@ theorem uniform_extend_subtype [CompleteSpace γ] {p : α → Prop} {e : α → rw [Subtype.range_coe_subtype] exact ⟨_, hb, by rwa [← de.toInducing.closure_eq_preimage_closure_image, hs.closure_eq]⟩ +include h_e h_f in theorem uniformly_extend_spec [CompleteSpace γ] (a : α) : Tendsto f (comap e (𝓝 a)) (𝓝 (ψ a)) := by simpa only [DenseInducing.extend] using tendsto_nhds_limUnder (uniformly_extend_exists h_e ‹_› h_f _) +include h_f in theorem uniformContinuous_uniformly_extend [CompleteSpace γ] : UniformContinuous ψ := fun d hd => let ⟨s, hs, hs_comp⟩ := comp3_mem_uniformity hd have h_pnt : ∀ {a m}, m ∈ 𝓝 a → ∃ c ∈ f '' (e ⁻¹' m), (c, ψ a) ∈ s ∧ (ψ a, c) ∈ s := @@ -451,6 +454,7 @@ theorem uniformContinuous_uniformly_extend [CompleteSpace γ] : UniformContinuou variable [T0Space γ] +include h_f in theorem uniformly_extend_of_ind (b : β) : ψ (e b) = f b := DenseInducing.extend_eq_at _ h_f.continuous.continuousAt diff --git a/lean-toolchain b/lean-toolchain index 64981ae5f58db..e7a4f40b892b4 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.11.0-rc1 +leanprover/lean4:v4.11.0-rc2 diff --git a/test/casesm.lean b/test/casesm.lean index b32734cae49da..5af4572b66196 100644 --- a/test/casesm.lean +++ b/test/casesm.lean @@ -82,6 +82,7 @@ example : True ∧ True ∧ True := by section AuxDecl variable {p q r : Prop} variable (h : p ∧ q ∨ p ∧ r) +include h -- Make sure that we don't try to work on auxiliary declarations. -- In this case, there will be an auxiliary recursive declaration for diff --git a/test/fun_prop_dev.lean b/test/fun_prop_dev.lean index fdcf619066278..c9443462421a9 100644 --- a/test/fun_prop_dev.lean +++ b/test/fun_prop_dev.lean @@ -31,13 +31,12 @@ instance [Obj α] [Obj β] : Obj (α × β) := ⟨⟩ instance [∀ x, Obj (E x)] : Obj ((x' : α) → E x') := ⟨⟩ instance : Obj Nat := ⟨⟩ -@[fun_prop] opaque Con {α β} [Obj α] [Obj β] (f : α → β) : Prop -@[fun_prop] opaque Lin {α β} [Obj α] [Obj β] (f : α → β) : Prop +@[fun_prop] opaque Con {α β} (f : α → β) : Prop +@[fun_prop] opaque Lin {α β} (f : α → β) : Prop -- state basic lambda calculus rules -- --------------------------------------- -variable [Obj α] [Obj β] [Obj γ] [Obj δ] [∀ x, Obj (E x)] @[fun_prop] theorem Con_id : Con (id : α → α) := silentSorry @[fun_prop] theorem Con_const (y : β) : Con (fun x : α => y) := silentSorry @@ -78,33 +77,32 @@ theorem prod_mk_Lin (fst : α → β) (snd : α → γ) (hfst : Lin fst) (hsnd : : Lin fun x => (fst x, snd x) := silentSorry -variable [Add α] [Add β] -- "simple form" of theorems @[fun_prop] theorem fst_Con : Con fun x : α×β => x.1 := silentSorry @[fun_prop] theorem snd_Con : Con fun x : α×β => x.2 := silentSorry -@[fun_prop] theorem add_Con : Con (Function.uncurry (fun x y : α => x + y)) := silentSorry -@[fun_prop] theorem add_Lin : Lin ↿(fun x y : α => x + y) := silentSorry +@[fun_prop] theorem add_Con [Add α] : Con (Function.uncurry (fun x y : α => x + y)) := silentSorry +@[fun_prop] theorem add_Lin [Add α] : Lin ↿(fun x y : α => x + y) := silentSorry -- "compositional form" of theorems @[fun_prop] theorem fst_Con' (self : α → β×γ) (hself : Con self) : Con fun x => (self x).1 := by fun_prop @[fun_prop] theorem snd_Con' (self : α → β×γ) (hself : Con self) : Con fun x => (self x).2 := by fun_prop -@[fun_prop] theorem add_Con' (x y : α → β) (hx : Con x) (hy : Con y) : Con (fun w => x w + y w) := by fun_prop -@[fun_prop] theorem add_Lin' (x y : α → β) (hx : Lin x) (hy : Lin y) : Lin (fun w => x w + y w) := by fun_prop +@[fun_prop] theorem add_Con' [Add β] (x y : α → β) (hx : Con x) (hy : Con y) : Con (fun w => x w + y w) := by fun_prop +@[fun_prop] theorem add_Lin' [Add β] (x y : α → β) (hx : Lin x) (hy : Lin y) : Lin (fun w => x w + y w) := by fun_prop -- set up hom objects/bundled morphisms -- ------------------------------------------ -structure ConHom (α β) [Obj α] [Obj β] where +structure ConHom (α β) where toFun : α → β con : Con toFun infixr:25 " ->> " => ConHom -structure LinHom (α β) [Obj α] [Obj β] where +structure LinHom (α β) where toFun : α → β lin : Lin toFun @@ -119,7 +117,7 @@ instance : FunLike (α -o β) α β where #eval Lean.Elab.Command.liftTermElabM do Lean.Meta.registerCoercion ``ConHom.toFun - (some { numArgs := 5, coercee := 4, type := .coeFun }) + (some { numArgs := 3, coercee := 2, type := .coeFun }) instance : HasUncurry (α ->> β) α β := ⟨fun f x => f x⟩ @@ -161,10 +159,10 @@ macro "fun" xs:explicitBinders " -o " b:term : term => expandExplicitBinders ``L end Notation -example (f : α → β → γ) (hx : ∀ y, Lin (f · y)) (hy : ∀ x, Lin (f x ·)) : +example [Add β] (f : α → β → γ) (hx : ∀ y, Lin (f · y)) (hy : ∀ x, Lin (f x ·)) : Lin (fun x => fun y ⊸ f y (x+x)) := by fun_prop -example (f : α → α → α → α) (hx : ∀ x y, Lin (f x y ·)) (hy : ∀ x z, Lin (f x · z)) (hz : ∀ y z, Lin (f · y z)) : +example [Add α] (f : α → α → α → α) (hx : ∀ x y, Lin (f x y ·)) (hy : ∀ x z, Lin (f x · z)) (hz : ∀ y z, Lin (f · y z)) : Lin (fun x => fun y z ⊸ f z (x+x) y) := by fun_prop -- the only analoge is this theorem but that is alredy provable @@ -190,15 +188,12 @@ example (x : α) [Add α] : Con (let y := x + x; fun x' : α => x' + y) := by fu example (f : β → γ) (g : α → β) (hf : Con f) (hg : Con g) : Con (fun x => f (g x)) := by fun_prop example (f : α → β → γ) (g : α → β) (hf : Con (fun (x,y) => f x y)) (hg : Con g) : Con (fun x => f x (g x)) := by fun_prop example (f : α → β → γ) (g : α → β) (hf : Con (fun (x,y) => f x y)) (hg : Con g) : Con (fun x => let y := g x; f x y) := by fun_prop -example {ι} (f : α → ι → γ) (hf : ∀ i, Con (fun x => f x i)) : Con (fun x i => f x i) := by fun_prop +example {ι : Type _} (f : α → ι → γ) (hf : ∀ i, Con (fun x => f x i)) : Con (fun x i => f x i) := by fun_prop example : Con (fun (f : α → β → γ) x y => f x y) := by fun_prop example : Con (fun (f : α → β → γ) y x => f x y) := by fun_prop example : Con (fun (f : α → α → α → α → α) y x => f x y x y) := by fun_prop --- set_option pp.notation false - - -- local hypothesis are assumed to be always in fully applied form -- so `(hf : Con f)` is not considered valid -- is this valid assumption? @@ -256,6 +251,9 @@ example (f : α → α ->> (α → α)) (hf : Con ↿f) : Con fun x y => f y x x example (f : α → β ->> γ) (hf : Con f) : Con ↿f := by fun_prop +section WithAdd +variable [Add α] + example : Con (HAdd.hAdd : α → α → α) := by fun_prop -- under applied constant example : Con (fun x => (HAdd.hAdd : α → α → α) x) := by fun_prop -- under applied constant @@ -265,8 +263,9 @@ example : Con (fun x y i => (HAdd.hAdd : ((ι→α) → (ι→α) → (ι→α)) example (y) : Con (fun x i => (HAdd.hAdd : ((ι→α) → (ι→α) → (ι→α))) x y i) := by fun_prop example (y i) : Con (fun x => (HAdd.hAdd : ((ι→α) → (ι→α) → (ι→α))) x y i) := by fun_prop --- example (f : β → γ) (x) (hf : Lin f) : Lin (fun (g : α → β) => f (g x)) := by fun_prop +end WithAdd +example (f : β → γ) (x) (hf : Lin f) : Lin (fun (g : α → β) => f (g x)) := by fun_prop -- apply theorems about FunLike.coe example (f : α ->> β) : Con f := by fun_prop @@ -276,7 +275,6 @@ example (f : β → γ) (g : α ->> β) (hf: Con f) : Con (fun x => f (g x)) := example (f : β ->> γ) (g : α → β) (hg: Con g) : Con (fun x => f (g x)) := by fun_prop example (f : β -o γ) (g : α → β) (hg : Con g) : Con fun x => f (g x) := by fun_prop --- set_option trace.Meta.Tactic.fun_prop true in example (f : α → β ->> γ) (hf : Con f) (g : α → β) (hg : Lin g) : Con (fun x => f x (g x)) := by fun_prop example (f : α → β ->> γ) (hf : Lin (fun (x,y) => f x y)) (g : α → β) (hg : Lin g) : Con (fun x => f x (g x)) := by fun_prop example (f : α → β ->> γ) (hf : Lin (fun (x,y) => f x y)) (g : α → β) (hg : Lin g) : Lin (fun x => f x (g x)) := by fun_prop @@ -319,7 +317,7 @@ example (x) : Con fun (f : α ->> α) => f (f x) := by fun_prop example (x) : Con fun (f : α ->> α) => f (f (f x)) := by fun_prop -example [Zero α] : Lin (fun x : α => (0 : α) + x + (0 : α) + (0 : α) + x) := by fun_prop +example [Zero α] [Obj α] [Add α] : Lin (fun x : α => (0 : α) + x + (0 : α) + (0 : α) + x) := by fun_prop noncomputable def foo : α ->> α ->> α := silentSorry @@ -358,12 +356,12 @@ theorem iterate_con (n : Nat) (f : α → α) (hf : Con f) : Con (iterate n f) : example : let f := fun x : α => x; Con f := by fun_prop -example : let f := fun x => x + y; ∀ y : α, ∀ z : α, Con fun x => x + f x + z := by fun_prop -example : ∀ y : α, let f := fun x => x + y; ∀ z : α, Con fun x => x + f x + z := by fun_prop +example [Add α] : let f := fun x => x + y; ∀ y : α, ∀ z : α, Con fun x => x + f x + z := by fun_prop +example [Add α] : ∀ y : α, let f := fun x => x + y; ∀ z : α, Con fun x => x + f x + z := by fun_prop -- this is still broken -- example : ∀ y : α, ∀ z : α, let f := fun x => x + y; Con fun x => x + f x + z := by fun_prop -example (f g : α → β) (hf : Con f := by fun_prop) (hg : outParam (Con g)) : +example [Add β] (f g : α → β) (hf : Con f := by fun_prop) (hg : outParam (Con g)) : Con (fun x => f x + g x) := by fun_prop opaque foo1 : α → α := id @@ -377,8 +375,8 @@ theorem foo2_lin : Lin (foo2 : α → α) := silentSorry example : Con (fun x : α => foo1 (foo2 x)) := by fun_prop -def foo3 (x : α) := x + x -example : Con (fun x : α => foo3 x) := by fun_prop [foo3] +def foo3 [Add α] (x : α) := x + x +example [Add α] : Con (fun x : α => foo3 x) := by fun_prop [foo3] def myUncurry (f : α → β → γ) : α×β → γ := fun (x,y) => f x y def diag (f : α → α → α) (x : α) := f x x @@ -411,7 +409,7 @@ example : Con fun ((x, _, _) : α × α × α) => x := by fun_prop example : Con fun ((_, x, _) : α × α × α) => x := by fun_prop example : Con fun ((_, _, x) : α × α × α) => x := by fun_prop -example : let f := (by exact (fun x : α => x+x)); Con f := by +example [Add α] : let f := (by exact (fun x : α => x+x)); Con f := by intro f; let F := fun x : α => x+x have : Con F := by fun_prop -- this used to be problematic @@ -428,7 +426,7 @@ Issues: No theorems found for `f1` in order to prove `Con fun a => f1 a` -/ #guard_msgs in -example : Con (fun x : α => x + f1 x) := by fun_prop +example [Add α] : Con (fun x : α => x + f1 x) := by fun_prop /-- error: `fun_prop` was unable to prove `Con fun x => f1 x + f1 x` @@ -437,7 +435,7 @@ Issues: No theorems found for `f1` in order to prove `Con fun a => f1 a` -/ #guard_msgs in -example : Con (fun x : α => f1 x + f1 x) := by fun_prop +example [Add α] : Con (fun x : α => f1 x + f1 x) := by fun_prop /-- error: `fun_prop` was unable to prove `Con fun x => f2 x + f1 x` @@ -446,7 +444,7 @@ Issues: No theorems found for `f2` in order to prove `Con fun a => f2 a` -/ #guard_msgs in -example : Con (fun x : α => f2 x + f1 x) := by fun_prop +example [Add α] : Con (fun x : α => f2 x + f1 x) := by fun_prop def f3 (a : α) := a From 9e8eaf98be987b1ee7a826a4efa1de7ceb692883 Mon Sep 17 00:00:00 2001 From: D-Thomine <100795491+D-Thomine@users.noreply.github.com> Date: Mon, 12 Aug 2024 19:34:20 +0000 Subject: [PATCH 222/359] feat(Dynamics): add TopologicalEntropy.DynamicalUniformity (#14938) First PR in a sequence to implement the notion of topological entropy for maps using Bowen-Dinaburg's formalism. This file describes a notion of _dynamical uniformities_, which generalizes dynamical balls in metric spaces. - [x] Dynamical uniformities - [ ] Topological entropy via covers - [ ] Topological entropy via nets - [ ] Behaviour for subsets - [ ] Behaviour under morphisms - [ ] Behaviour under union - [ ] Behaviour under products - [ ] Full shift This file is pretty short; the next ones are more involved. Paging @pitmonticone @sgouezel Co-authored-by: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> --- Mathlib.lean | 1 + .../DynamicalEntourage.lean | 128 ++++++++++++++++++ 2 files changed, 129 insertions(+) create mode 100644 Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean diff --git a/Mathlib.lean b/Mathlib.lean index 54c7940c8bc63..05b79411aac95 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2559,6 +2559,7 @@ import Mathlib.Dynamics.Minimal import Mathlib.Dynamics.Newton import Mathlib.Dynamics.OmegaLimit import Mathlib.Dynamics.PeriodicPts +import Mathlib.Dynamics.TopologicalEntropy.DynamicalEntourage import Mathlib.FieldTheory.AbelRuffini import Mathlib.FieldTheory.AbsoluteGaloisGroup import Mathlib.FieldTheory.Adjoin diff --git a/Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean b/Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean new file mode 100644 index 0000000000000..ad6c52b6be339 --- /dev/null +++ b/Mathlib/Dynamics/TopologicalEntropy/DynamicalEntourage.lean @@ -0,0 +1,128 @@ +/- +Copyright (c) 2024 Damien Thomine. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damien Thomine, Pietro Monticone +-/ +import Mathlib.Order.Interval.Finset.Nat +import Mathlib.Order.OmegaCompletePartialOrder +import Mathlib.Topology.UniformSpace.Basic + +/-! +# Dynamical entourages +Bowen-Dinaburg's definition of topological entropy of a transformation `T` in a metric space +`(X, d)` relies on the so-called dynamical balls. These balls are sets +`B (x, ε, n) = { y | ∀ k < n, d(T^[k] x, T^[k] y) < ε }`. + +We implement Bowen-Dinaburg's definitions in the more general context of uniform spaces. Dynamical +balls are replaced by what we call dynamical entourages. This file collects all general lemmas +about these objects. + +## Main definitions +- `dynEntourage`: dynamical entourage associated with a given transformation `T`, entourage `U` +and time `n`. + +## Tags +entropy + +## TODO +Once #PR14718 has passed, add product of entourages. + +In the context of (pseudo-e)metric spaces, relate the usual definition of dynamical balls with +these dynamical entourages. +-/ + +namespace Dynamics + +open Prod Set Uniformity UniformSpace + +variable {X : Type*} + +/-- The dynamical entourage associated to a transformation `T`, entourage `U` and time `n` +is the set of points `(x, y)` such that `(T^[k] x, T^[k] y) ∈ U` for all `k < n`, i.e. +which are `U`-close up to time `n`.-/ +def dynEntourage (T : X → X) (U : Set (X × X)) (n : ℕ) : Set (X × X) := + ⋂ k < n, (map T T)^[k] ⁻¹' U + +lemma dynEntourage_eq_inter_Ico (T : X → X) (U : Set (X × X)) (n : ℕ) : + dynEntourage T U n = ⋂ k : Ico 0 n, (map T T)^[k] ⁻¹' U := by + simp [dynEntourage] + +lemma mem_dynEntourage {T : X → X} {U : Set (X × X)} {n : ℕ} {x y : X} : + (x, y) ∈ dynEntourage T U n ↔ ∀ k < n, (T^[k] x, T^[k] y) ∈ U := by + simp [dynEntourage] + +lemma mem_ball_dynEntourage {T : X → X} {U : Set (X × X)} {n : ℕ} {x y : X} : + y ∈ ball x (dynEntourage T U n) ↔ ∀ k < n, T^[k] y ∈ ball (T^[k] x) U := by + simp only [ball, mem_preimage]; exact mem_dynEntourage + +lemma dynEntourage_mem_uniformity [UniformSpace X] {T : X → X} (h : UniformContinuous T) + {U : Set (X × X)} (U_uni : U ∈ 𝓤 X) (n : ℕ) : + dynEntourage T U n ∈ 𝓤 X := by + rw [dynEntourage_eq_inter_Ico T U n] + refine Filter.iInter_mem.2 fun k ↦ ?_ + rw [map_iterate T T k] + exact uniformContinuous_def.1 (UniformContinuous.iterate T k h) U U_uni + +lemma idRel_subset_dynEntourage (T : X → X) {U : Set (X × X)} (h : idRel ⊆ U) (n : ℕ) : + idRel ⊆ (dynEntourage T U n) := by + simp only [dynEntourage, map_iterate, subset_iInter_iff, idRel_subset, mem_preimage, map_apply] + exact fun _ _ _ ↦ h rfl + +lemma _root_.SymmetricRel.dynEntourage (T : X → X) {U : Set (X × X)} (h : SymmetricRel U) (n : ℕ) : + SymmetricRel (dynEntourage T U n) := by + ext xy + simp only [Dynamics.dynEntourage, map_iterate, mem_preimage, mem_iInter] + refine forall₂_congr fun k _ ↦ ?_ + exact map_apply' _ _ _ ▸ SymmetricRel.mk_mem_comm h + +lemma dynEntourage_comp_subset (T : X → X) (U V : Set (X × X)) (n : ℕ) : + (dynEntourage T U n) ○ (dynEntourage T V n) ⊆ dynEntourage T (U ○ V) n := by + simp only [dynEntourage, map_iterate, subset_iInter_iff] + intro k k_n xy xy_comp + simp only [compRel, mem_iInter, mem_preimage, map_apply, mem_setOf_eq] at xy_comp ⊢ + rcases xy_comp with ⟨z, hz1, hz2⟩ + exact mem_ball_comp (hz1 k k_n) (hz2 k k_n) + +lemma _root_.isOpen.dynEntourage [TopologicalSpace X] {T : X → X} (T_cont : Continuous T) + {U : Set (X × X)} (U_open : IsOpen U) (n : ℕ) : + IsOpen (dynEntourage T U n) := by + rw [dynEntourage_eq_inter_Ico T U n] + refine isOpen_iInter_of_finite fun k ↦ ?_ + exact continuous_def.1 ((T_cont.prod_map T_cont).iterate k) U U_open + +lemma dynEntourage_monotone (T : X → X) (n : ℕ) : + Monotone (fun U : Set (X × X) ↦ dynEntourage T U n) := + fun _ _ h ↦ iInter₂_mono fun _ _ ↦ preimage_mono h + +lemma dynEntourage_antitone (T : X → X) (U : Set (X × X)) : + Antitone (fun n : ℕ ↦ dynEntourage T U n) := + fun m n m_n ↦ iInter₂_mono' fun k k_m ↦ by use k, lt_of_lt_of_le k_m m_n + +@[simp] +lemma dynEntourage_zero {T : X → X} {U : Set (X × X)} : + dynEntourage T U 0 = univ := by simp [dynEntourage] + +@[simp] +lemma dynEntourage_one {T : X → X} {U : Set (X × X)} : + dynEntourage T U 1 = U := by simp [dynEntourage] + +@[simp] +lemma dynEntourage_univ {T : X → X} {n : ℕ} : + dynEntourage T univ n = univ := by simp [dynEntourage] + +lemma mem_ball_dynEntourage_comp (T : X → X) (n : ℕ) {U : Set (X × X)} (U_symm : SymmetricRel U) + (x y : X) (h : (ball x (dynEntourage T U n) ∩ ball y (dynEntourage T U n)).Nonempty) : + x ∈ ball y (dynEntourage T (U ○ U) n) := by + rcases h with ⟨z, z_Bx, z_By⟩ + rw [mem_ball_symmetry (SymmetricRel.dynEntourage T U_symm n)] at z_Bx + exact dynEntourage_comp_subset T U U n (mem_ball_comp z_By z_Bx) + +lemma _root_.Function.Semiconj.preimage_dynEntourage {Y : Type*} {S : X → X} {T : Y → Y} {φ : X → Y} + (h : Function.Semiconj φ S T) (U : Set (Y × Y)) (n : ℕ) : + (map φ φ)⁻¹' (dynEntourage T U n) = dynEntourage S ((map φ φ)⁻¹' U) n := by + rw [dynEntourage, preimage_iInter₂] + refine iInter₂_congr fun k _ ↦ ?_ + rw [← preimage_comp, ← preimage_comp, map_iterate S S k, map_iterate T T k, map_comp_map, + map_comp_map, (Function.Semiconj.iterate_right h k).comp_eq] + +end Dynamics From 1c910395b91975683fb7d115c3d243537d3fe451 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Mon, 12 Aug 2024 19:34:21 +0000 Subject: [PATCH 223/359] perf(CharP/CharZero): scope simp theorems with weak keys (#15631) These have keys like `HAdd.hAdd _ _ _ _ _ _` but are specific to characteristic two and zero. --- Mathlib/Algebra/CharP/Two.lean | 6 +++--- Mathlib/Algebra/CharZero/Lemmas.lean | 4 ++-- Mathlib/Analysis/Complex/Isometry.lean | 2 ++ Mathlib/Geometry/Euclidean/Triangle.lean | 2 ++ Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean | 2 +- Mathlib/NumberTheory/LSeries/RiemannZeta.lean | 2 +- Mathlib/NumberTheory/LegendreSymbol/Basic.lean | 4 ++-- Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean | 2 +- .../NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean | 2 +- Mathlib/NumberTheory/Pell.lean | 2 +- 10 files changed, 16 insertions(+), 12 deletions(-) diff --git a/Mathlib/Algebra/CharP/Two.lean b/Mathlib/Algebra/CharP/Two.lean index 7984a9aad9a6d..0f91ff1fa1e5b 100644 --- a/Mathlib/Algebra/CharP/Two.lean +++ b/Mathlib/Algebra/CharP/Two.lean @@ -26,7 +26,7 @@ variable [Semiring R] [CharP R 2] theorem two_eq_zero : (2 : R) = 0 := by rw [← Nat.cast_two, CharP.cast_eq_zero] -@[simp] +@[scoped simp] theorem add_self_eq_zero (x : R) : x + x = 0 := by rw [← two_smul R x, two_eq_zero, zero_smul] end Semiring @@ -35,14 +35,14 @@ section Ring variable [Ring R] [CharP R 2] -@[simp] +@[scoped simp] theorem neg_eq (x : R) : -x = x := by rw [neg_eq_iff_add_eq_zero, ← two_smul R x, two_eq_zero, zero_smul] theorem neg_eq' : Neg.neg = (id : R → R) := funext neg_eq -@[simp] +@[scoped simp] theorem sub_eq_add (x y : R) : x - y = x + y := by rw [sub_eq_add_neg, neg_eq] theorem sub_eq_add' : HSub.hSub = ((· + ·) : R → R → R) := diff --git a/Mathlib/Algebra/CharZero/Lemmas.lean b/Mathlib/Algebra/CharZero/Lemmas.lean index 0c637d4a69ff7..4a3176488e297 100644 --- a/Mathlib/Algebra/CharZero/Lemmas.lean +++ b/Mathlib/Algebra/CharZero/Lemmas.lean @@ -87,10 +87,10 @@ section variable {R : Type*} [NonAssocRing R] [NoZeroDivisors R] [CharZero R] -@[simp] theorem neg_eq_self_iff {a : R} : -a = a ↔ a = 0 := +@[scoped simp] theorem CharZero.neg_eq_self_iff {a : R} : -a = a ↔ a = 0 := neg_eq_iff_add_eq_zero.trans add_self_eq_zero -@[simp] theorem eq_neg_self_iff {a : R} : a = -a ↔ a = 0 := +@[scoped simp] theorem CharZero.eq_neg_self_iff {a : R} : a = -a ↔ a = 0 := eq_neg_iff_add_eq_zero.trans add_self_eq_zero theorem nat_mul_inj {n : ℕ} {a b : R} (h : (n : R) * a = (n : R) * b) : n = 0 ∨ a = b := by diff --git a/Mathlib/Analysis/Complex/Isometry.lean b/Mathlib/Analysis/Complex/Isometry.lean index 0224bd11c7fa4..379ef2487f8c3 100644 --- a/Mathlib/Analysis/Complex/Isometry.lean +++ b/Mathlib/Analysis/Complex/Isometry.lean @@ -30,6 +30,8 @@ noncomputable section open Complex +open CharZero + open ComplexConjugate local notation "|" x "|" => Complex.abs x diff --git a/Mathlib/Geometry/Euclidean/Triangle.lean b/Mathlib/Geometry/Euclidean/Triangle.lean index b91f8c483a927..c76fa197f9315 100644 --- a/Mathlib/Geometry/Euclidean/Triangle.lean +++ b/Mathlib/Geometry/Euclidean/Triangle.lean @@ -36,6 +36,8 @@ unnecessarily. noncomputable section +open scoped CharZero + open scoped Classical open scoped Real diff --git a/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean b/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean index a77aff55c0fd1..e14ca16ba1108 100644 --- a/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean +++ b/Mathlib/NumberTheory/LSeries/HurwitzZetaOdd.lean @@ -43,7 +43,7 @@ various versions of the Jacobi theta function. noncomputable section open Complex hiding abs_of_nonneg -open Filter Topology Asymptotics Real Set MeasureTheory +open CharZero Filter Topology Asymptotics Real Set MeasureTheory open scoped ComplexConjugate namespace HurwitzZeta diff --git a/Mathlib/NumberTheory/LSeries/RiemannZeta.lean b/Mathlib/NumberTheory/LSeries/RiemannZeta.lean index b63a86133b4f5..0afac073a7aa4 100644 --- a/Mathlib/NumberTheory/LSeries/RiemannZeta.lean +++ b/Mathlib/NumberTheory/LSeries/RiemannZeta.lean @@ -43,7 +43,7 @@ proved in `Mathlib.NumberTheory.LSeries.HurwitzZetaEven`. -/ -open MeasureTheory Set Filter Asymptotics TopologicalSpace Real Asymptotics +open CharZero MeasureTheory Set Filter Asymptotics TopologicalSpace Real Asymptotics Classical HurwitzZeta open Complex hiding exp norm_eq_abs abs_of_nonneg abs_two continuous_exp diff --git a/Mathlib/NumberTheory/LegendreSymbol/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/Basic.lean index f0a2686941e1c..dce141a92e9e5 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/Basic.lean @@ -235,9 +235,9 @@ theorem eq_zero_mod_of_eq_neg_one {p : ℕ} [Fact p.Prime] {a : ℤ} (h : legend simp at h by_contra hf cases' imp_iff_or_not.mp (not_and'.mp hf) with hx hy - · rw [eq_one_of_sq_sub_mul_sq_eq_zero' ha hx hxy, eq_neg_self_iff] at h + · rw [eq_one_of_sq_sub_mul_sq_eq_zero' ha hx hxy, CharZero.eq_neg_self_iff] at h exact one_ne_zero h - · rw [eq_one_of_sq_sub_mul_sq_eq_zero ha hy hxy, eq_neg_self_iff] at h + · rw [eq_one_of_sq_sub_mul_sq_eq_zero ha hy hxy, CharZero.eq_neg_self_iff] at h exact one_ne_zero h /-- If `legendreSym p a = -1` and `p` divides `x^2 - a*y^2`, then `p` must divide `x` and `y`. -/ diff --git a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean index ab60e7cabe93d..7002957721377 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean @@ -260,7 +260,7 @@ theorem eq_neg_one_at_prime_divisor_of_eq_neg_one {a : ℤ} {n : ℕ} (h : J(a | ∃ p : ℕ, p.Prime ∧ p ∣ n ∧ J(a | p) = -1 := by have hn₀ : n ≠ 0 := by rintro rfl - rw [zero_right, eq_neg_self_iff] at h + rw [zero_right, CharZero.eq_neg_self_iff] at h exact one_ne_zero h have hf₀ (p) (hp : p ∈ n.primeFactorsList) : p ≠ 0 := (Nat.pos_of_mem_primeFactorsList hp).ne.symm rw [← Nat.prod_primeFactorsList hn₀, list_prod_right hf₀] at h diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index b73957e3efc28..b9a2c36a497c9 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -213,7 +213,7 @@ when the domain has odd characteristic. -/ theorem quadraticChar_ne_one (hF : ringChar F ≠ 2) : quadraticChar F ≠ 1 := by rcases quadraticChar_exists_neg_one' hF with ⟨a, ha⟩ intro hχ - simp only [hχ, one_apply a.isUnit, eq_neg_self_iff, one_ne_zero] at ha + simp only [hχ, one_apply a.isUnit, one_ne_zero] at ha set_option linter.deprecated false in @[deprecated quadraticChar_ne_one (since := "2024-06-16")] diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index 3012b77db01b5..d6a8d14bb8b9e 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -72,7 +72,7 @@ We then set up an API for `Pell.Solution₁ d`. -/ -open Zsqrtd +open CharZero Zsqrtd /-- An element of `ℤ√d` has norm one (i.e., `a.re^2 - d*a.im^2 = 1`) if and only if it is contained in the submonoid of unitary elements. From 3697f5137c91b23ca54f02c5c4d03d643ea9485a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 12 Aug 2024 19:34:22 +0000 Subject: [PATCH 224/359] chore(MeasurePreserving): rename a lemma (#15742) There is no `volume` in the statement. --- Mathlib/Dynamics/Ergodic/MeasurePreserving.lean | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) diff --git a/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean b/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean index 39c16a8d90fa4..3a4d7509a3c58 100644 --- a/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean +++ b/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean @@ -146,7 +146,7 @@ lemma measure_symmDiff_preimage_iterate_le /-- If `μ univ < n * μ s` and `f` is a map preserving measure `μ`, then for some `x ∈ s` and `0 < m < n`, `f^[m] x ∈ s`. -/ -theorem exists_mem_iterate_mem_of_volume_lt_mul_volume (hf : MeasurePreserving f μ μ) +theorem exists_mem_iterate_mem_of_measure_univ_lt_mul_measure (hf : MeasurePreserving f μ μ) (hs : NullMeasurableSet s μ) {n : ℕ} (hvol : μ (Set.univ : Set α) < n * μ s) : ∃ x ∈ s, ∃ m ∈ Set.Ioo 0 n, f^[m] x ∈ s := by have A : ∀ m, NullMeasurableSet (f^[m] ⁻¹' s) μ := fun m ↦ @@ -161,6 +161,10 @@ theorem exists_mem_iterate_mem_of_volume_lt_mul_volume (hf : MeasurePreserving f refine ⟨f^[i] x, hxi, j - i, ⟨tsub_pos_of_lt hlt, lt_of_le_of_lt (j.sub_le i) hj⟩, ?_⟩ rwa [← iterate_add_apply, tsub_add_cancel_of_le hlt.le] +@[deprecated (since := "2024-08-12")] +alias exists_mem_iterate_mem_of_volume_lt_mul_volume := + exists_mem_iterate_mem_of_measure_univ_lt_mul_measure + /-- A self-map preserving a finite measure is conservative: if `μ s ≠ 0`, then at least one point `x ∈ s` comes back to `s` under iterations of `f`. Actually, a.e. point of `s` comes back to `s` infinitely many times, see `MeasureTheory.MeasurePreserving.conservative` and theorems about @@ -168,7 +172,7 @@ infinitely many times, see `MeasureTheory.MeasurePreserving.conservative` and th theorem exists_mem_iterate_mem [IsFiniteMeasure μ] (hf : MeasurePreserving f μ μ) (hs : NullMeasurableSet s μ) (hs' : μ s ≠ 0) : ∃ x ∈ s, ∃ m ≠ 0, f^[m] x ∈ s := by rcases ENNReal.exists_nat_mul_gt hs' (measure_ne_top μ (Set.univ : Set α)) with ⟨N, hN⟩ - rcases hf.exists_mem_iterate_mem_of_volume_lt_mul_volume hs hN with ⟨x, hx, m, hm, hmx⟩ + rcases hf.exists_mem_iterate_mem_of_measure_univ_lt_mul_measure hs hN with ⟨x, hx, m, hm, hmx⟩ exact ⟨x, hx, m, hm.1.ne', hmx⟩ end MeasurePreserving From 38dc174b2547f019db53a09849da9dc22b257003 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 12 Aug 2024 20:23:08 +0000 Subject: [PATCH 225/359] feat(StrongTopology): generalize `ContinuousLinearMap.restrictScalarsL` (#15285) --- Mathlib/Algebra/Group/Action/Pi.lean | 4 ++ Mathlib/Analysis/LocallyConvex/Bounded.lean | 46 ++++++++++++ .../NormedSpace/OperatorNorm/Basic.lean | 18 ----- .../Algebra/Module/StrongTopology.lean | 70 ++++++++++++++++++- Mathlib/Topology/Bornology/Absorbs.lean | 7 ++ 5 files changed, 124 insertions(+), 21 deletions(-) diff --git a/Mathlib/Algebra/Group/Action/Pi.lean b/Mathlib/Algebra/Group/Action/Pi.lean index 25c28ef657a00..d8c0c7b05deb3 100644 --- a/Mathlib/Algebra/Group/Action/Pi.lean +++ b/Mathlib/Algebra/Group/Action/Pi.lean @@ -28,6 +28,10 @@ namespace Pi @[to_additive] instance smul' [∀ i, SMul (α i) (β i)] : SMul (∀ i, α i) (∀ i, β i) where smul s x i := s i • x i +@[to_additive] +lemma smul_def' [∀ i, SMul (α i) (β i)] (s : ∀ i, α i) (x : ∀ i, β i) : s • x = fun i ↦ s i • x i := + rfl + @[to_additive (attr := simp)] lemma smul_apply' [∀ i, SMul (α i) (β i)] (s : ∀ i, α i) (x : ∀ i, β i) : (s • x) i = s i • x i := rfl diff --git a/Mathlib/Analysis/LocallyConvex/Bounded.lean b/Mathlib/Analysis/LocallyConvex/Bounded.lean index 1651870d65d89..680f38ea9ba4f 100644 --- a/Mathlib/Analysis/LocallyConvex/Bounded.lean +++ b/Mathlib/Analysis/LocallyConvex/Bounded.lean @@ -93,9 +93,16 @@ theorem isVonNBounded_union {s t : Set E} : theorem IsVonNBounded.union {s₁ s₂ : Set E} (hs₁ : IsVonNBounded 𝕜 s₁) (hs₂ : IsVonNBounded 𝕜 s₂) : IsVonNBounded 𝕜 (s₁ ∪ s₂) := isVonNBounded_union.2 ⟨hs₁, hs₂⟩ +@[nontriviality] theorem IsVonNBounded.of_boundedSpace [BoundedSpace 𝕜] {s : Set E} : IsVonNBounded 𝕜 s := fun _ _ ↦ .of_boundedSpace +@[nontriviality] +theorem IsVonNBounded.of_subsingleton [Subsingleton E] {s : Set E} : IsVonNBounded 𝕜 s := + fun U hU ↦ eventually_of_forall fun c ↦ calc + s ⊆ univ := subset_univ s + _ = c • U := .symm <| Subsingleton.eq_univ_of_nonempty <| (Filter.nonempty_of_mem hU).image _ + @[simp] theorem isVonNBounded_iUnion {ι : Sort*} [Finite ι] {s : ι → Set E} : IsVonNBounded 𝕜 (⋃ i, s i) ↔ ∀ i, IsVonNBounded 𝕜 (s i) := by @@ -237,6 +244,20 @@ theorem isVonNBounded_iff_smul_tendsto_zero {ε : ι → 𝕜} {l : Filter ι} [ end sequence +/-- If a set is von Neumann bounded with respect to a smaller field, +then it is also von Neumann bounded with respect to a larger field. +See also `Bornology.IsVonNBounded.restrict_scalars` below. -/ +theorem IsVonNBounded.extend_scalars [NontriviallyNormedField 𝕜] + {E : Type*} [AddCommGroup E] [Module 𝕜 E] + (𝕝 : Type*) [NontriviallyNormedField 𝕝] [NormedAlgebra 𝕜 𝕝] + [Module 𝕝 E] [TopologicalSpace E] [ContinuousSMul 𝕝 E] [IsScalarTower 𝕜 𝕝 E] + {s : Set E} (h : IsVonNBounded 𝕜 s) : IsVonNBounded 𝕝 s := by + obtain ⟨ε, hε, hε₀⟩ : ∃ ε : ℕ → 𝕜, Tendsto ε atTop (𝓝 0) ∧ ∀ᶠ n in atTop, ε n ≠ 0 := by + simpa only [tendsto_nhdsWithin_iff] using exists_seq_tendsto (𝓝[≠] (0 : 𝕜)) + refine isVonNBounded_of_smul_tendsto_zero (ε := (ε · • 1)) (by simpa) fun x hx ↦ ?_ + have := h.smul_tendsto_zero (eventually_of_forall hx) hε + simpa only [Pi.smul_def', smul_one_smul] + section NormedField variable [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] @@ -379,6 +400,31 @@ theorem Filter.Tendsto.isVonNBounded_range [NormedField 𝕜] [AddCommGroup E] [ haveI := comm_topologicalAddGroup_is_uniform (G := E) hf.cauchySeq.totallyBounded_range.isVonNBounded 𝕜 +variable (𝕜) in +protected theorem Bornology.IsVonNBounded.restrict_scalars_of_nontrivial + [NormedField 𝕜] [NormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] [Nontrivial 𝕜'] + [Zero E] [TopologicalSpace E] + [SMul 𝕜 E] [MulAction 𝕜' E] [IsScalarTower 𝕜 𝕜' E] {s : Set E} + (h : IsVonNBounded 𝕜' s) : IsVonNBounded 𝕜 s := by + intro V hV + refine (h hV).restrict_scalars <| AntilipschitzWith.tendsto_cobounded (K := ‖(1 : 𝕜')‖₊⁻¹) ?_ + refine AntilipschitzWith.of_le_mul_nndist fun x y ↦ ?_ + rw [nndist_eq_nnnorm, nndist_eq_nnnorm, ← sub_smul, nnnorm_smul, ← div_eq_inv_mul, + mul_div_cancel_right₀ _ (nnnorm_ne_zero_iff.2 one_ne_zero)] + +variable (𝕜) in +protected theorem Bornology.IsVonNBounded.restrict_scalars + [NormedField 𝕜] [NormedRing 𝕜'] [NormedAlgebra 𝕜 𝕜'] + [Zero E] [TopologicalSpace E] + [SMul 𝕜 E] [MulActionWithZero 𝕜' E] [IsScalarTower 𝕜 𝕜' E] {s : Set E} + (h : IsVonNBounded 𝕜' s) : IsVonNBounded 𝕜 s := + match subsingleton_or_nontrivial 𝕜' with + | .inl _ => + have : Subsingleton E := MulActionWithZero.subsingleton 𝕜' E + IsVonNBounded.of_subsingleton + | .inr _ => + h.restrict_scalars_of_nontrivial _ + section VonNBornologyEqMetric namespace NormedSpace diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean index 639f15c04f7ce..cb8d792f631d6 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Basic.lean @@ -413,24 +413,6 @@ theorem restrictScalarsIsometry_toLinearMap : (restrictScalarsIsometry 𝕜 E Fₗ 𝕜' 𝕜'').toLinearMap = restrictScalarsₗ 𝕜 E Fₗ 𝕜' 𝕜'' := rfl -variable (𝕜'') - - -/-- `ContinuousLinearMap.restrictScalars` as a `ContinuousLinearMap`. -/ -def restrictScalarsL : (E →L[𝕜] Fₗ) →L[𝕜''] E →L[𝕜'] Fₗ := - (restrictScalarsIsometry 𝕜 E Fₗ 𝕜' 𝕜'').toContinuousLinearMap - -variable {𝕜 E Fₗ 𝕜' 𝕜''} - -@[simp] -theorem coe_restrictScalarsL : (restrictScalarsL 𝕜 E Fₗ 𝕜' 𝕜'' : (E →L[𝕜] Fₗ) →ₗ[𝕜''] E →L[𝕜'] Fₗ) = - restrictScalarsₗ 𝕜 E Fₗ 𝕜' 𝕜'' := - rfl - -@[simp] -theorem coe_restrict_scalarsL' : ⇑(restrictScalarsL 𝕜 E Fₗ 𝕜' 𝕜'') = restrictScalars 𝕜' := - rfl - end RestrictScalars lemma norm_pi_le_of_le {ι : Type*} [Fintype ι] diff --git a/Mathlib/Topology/Algebra/Module/StrongTopology.lean b/Mathlib/Topology/Algebra/Module/StrongTopology.lean index 8dd679cb9f9d7..a66e7b2cf135a 100644 --- a/Mathlib/Topology/Algebra/Module/StrongTopology.lean +++ b/Mathlib/Topology/Algebra/Module/StrongTopology.lean @@ -275,6 +275,10 @@ protected theorem hasBasis_nhds_zero [TopologicalSpace F] [TopologicalAddGroup F fun SV => { f : E →SL[σ] F | ∀ x ∈ SV.1, f x ∈ SV.2 } := ContinuousLinearMap.hasBasis_nhds_zero_of_basis (𝓝 0).basis_sets +theorem uniformEmbedding_toUniformOnFun [UniformSpace F] [UniformAddGroup F] : + UniformEmbedding fun f : E →SL[σ] F ↦ UniformOnFun.ofFun {s | Bornology.IsVonNBounded 𝕜₁ s} f := + UniformConvergenceCLM.uniformEmbedding_coeFn .. + instance uniformContinuousConstSMul {M : Type*} [Monoid M] [DistribMulAction M F] [SMulCommClass 𝕜₂ M F] [UniformSpace F] [UniformAddGroup F] [UniformContinuousConstSMul M F] : @@ -345,6 +349,66 @@ def toLinearMap₂ (L : E →L[𝕜] F →L[𝕜] G) : E →ₗ[𝕜] F →ₗ[ end BilinearMaps +section RestrictScalars + +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {E : Type*} [AddCommGroup E] [TopologicalSpace E] [Module 𝕜 E] [ContinuousSMul 𝕜 E] + {F : Type*} [AddCommGroup F] + +section UniformSpace + +variable [UniformSpace F] [UniformAddGroup F] [Module 𝕜 F] + (𝕜' : Type*) [NontriviallyNormedField 𝕜'] [NormedAlgebra 𝕜' 𝕜] + [Module 𝕜' E] [IsScalarTower 𝕜' 𝕜 E] [Module 𝕜' F] [IsScalarTower 𝕜' 𝕜 F] + +theorem uniformEmbedding_restrictScalars : + UniformEmbedding (restrictScalars 𝕜' : (E →L[𝕜] F) → (E →L[𝕜'] F)) := by + rw [← uniformEmbedding_toUniformOnFun.of_comp_iff] + convert uniformEmbedding_toUniformOnFun using 4 with s + exact ⟨fun h ↦ h.extend_scalars _, fun h ↦ h.restrict_scalars _⟩ + +theorem uniformContinuous_restrictScalars : + UniformContinuous (restrictScalars 𝕜' : (E →L[𝕜] F) → (E →L[𝕜'] F)) := + (uniformEmbedding_restrictScalars 𝕜').uniformContinuous + +end UniformSpace + +variable [TopologicalSpace F] [TopologicalAddGroup F] [Module 𝕜 F] + (𝕜' : Type*) [NontriviallyNormedField 𝕜'] [NormedAlgebra 𝕜' 𝕜] + [Module 𝕜' E] [IsScalarTower 𝕜' 𝕜 E] [Module 𝕜' F] [IsScalarTower 𝕜' 𝕜 F] + +theorem embedding_restrictScalars : + Embedding (restrictScalars 𝕜' : (E →L[𝕜] F) → (E →L[𝕜'] F)) := + letI : UniformSpace F := TopologicalAddGroup.toUniformSpace F + haveI : UniformAddGroup F := comm_topologicalAddGroup_is_uniform + (uniformEmbedding_restrictScalars _).embedding + +@[continuity, fun_prop] +theorem continuous_restrictScalars : + Continuous (restrictScalars 𝕜' : (E →L[𝕜] F) → (E →L[𝕜'] F)) := + (embedding_restrictScalars _).continuous + +variable (𝕜 E F) +variable (𝕜'' : Type*) [Ring 𝕜''] + [Module 𝕜'' F] [ContinuousConstSMul 𝕜'' F] [SMulCommClass 𝕜 𝕜'' F] [SMulCommClass 𝕜' 𝕜'' F] + +/-- `ContinuousLinearMap.restrictScalars` as a `ContinuousLinearMap`. -/ +def restrictScalarsL : (E →L[𝕜] F) →L[𝕜''] E →L[𝕜'] F := + .mk <| restrictScalarsₗ 𝕜 E F 𝕜' 𝕜'' + +variable {𝕜 E F 𝕜' 𝕜''} + +@[simp] +theorem coe_restrictScalarsL : (restrictScalarsL 𝕜 E F 𝕜' 𝕜'' : (E →L[𝕜] F) →ₗ[𝕜''] E →L[𝕜'] F) = + restrictScalarsₗ 𝕜 E F 𝕜' 𝕜'' := + rfl + +@[simp] +theorem coe_restrict_scalarsL' : ⇑(restrictScalarsL 𝕜 E F 𝕜' 𝕜'') = restrictScalars 𝕜' := + rfl + +end RestrictScalars + end ContinuousLinearMap open ContinuousLinearMap @@ -357,8 +421,8 @@ section Semilinear variable {𝕜 : Type*} {𝕜₂ : Type*} {𝕜₃ : Type*} {𝕜₄ : Type*} {E : Type*} {F : Type*} {G : Type*} {H : Type*} [AddCommGroup E] [AddCommGroup F] [AddCommGroup G] [AddCommGroup H] - [NontriviallyNormedField 𝕜] [NontriviallyNormedField 𝕜₂] [NontriviallyNormedField 𝕜₃] - [NontriviallyNormedField 𝕜₄] [Module 𝕜 E] [Module 𝕜₂ F] [Module 𝕜₃ G] [Module 𝕜₄ H] + [NormedField 𝕜] [NormedField 𝕜₂] [NormedField 𝕜₃] [NormedField 𝕜₄] + [Module 𝕜 E] [Module 𝕜₂ F] [Module 𝕜₃ G] [Module 𝕜₄ H] [TopologicalSpace E] [TopologicalSpace F] [TopologicalSpace G] [TopologicalSpace H] [TopologicalAddGroup G] [TopologicalAddGroup H] [ContinuousConstSMul 𝕜₃ G] [ContinuousConstSMul 𝕜₄ H] {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₁ : 𝕜₂ →+* 𝕜} {σ₂₃ : 𝕜₂ →+* 𝕜₃} {σ₁₃ : 𝕜 →+* 𝕜₃} @@ -404,7 +468,7 @@ end Semilinear section Linear variable {𝕜 : Type*} {E : Type*} {F : Type*} {G : Type*} {H : Type*} [AddCommGroup E] - [AddCommGroup F] [AddCommGroup G] [AddCommGroup H] [NontriviallyNormedField 𝕜] [Module 𝕜 E] + [AddCommGroup F] [AddCommGroup G] [AddCommGroup H] [NormedField 𝕜] [Module 𝕜 E] [Module 𝕜 F] [Module 𝕜 G] [Module 𝕜 H] [TopologicalSpace E] [TopologicalSpace F] [TopologicalSpace G] [TopologicalSpace H] [TopologicalAddGroup G] [TopologicalAddGroup H] [ContinuousConstSMul 𝕜 G] [ContinuousConstSMul 𝕜 H] diff --git a/Mathlib/Topology/Bornology/Absorbs.lean b/Mathlib/Topology/Bornology/Absorbs.lean index 105cdec612a49..b4753b1d5ec86 100644 --- a/Mathlib/Topology/Bornology/Absorbs.lean +++ b/Mathlib/Topology/Bornology/Absorbs.lean @@ -244,3 +244,10 @@ lemma Absorbent.zero_mem [NeBot (cobounded G₀)] [AddMonoid E] [DistribMulActio absorbs_zero_iff.1 (hs 0) end GroupWithZero + +protected theorem Absorbs.restrict_scalars + {M N α : Type*} [Monoid N] [SMul M N] [SMul M α] [MulAction N α] + [IsScalarTower M N α] [Bornology M] [Bornology N] {s t : Set α} (h : Absorbs N s t) + (hbdd : Tendsto (· • 1 : M → N) (cobounded M) (cobounded N)) : + Absorbs M s t := + (hbdd.eventually h).mono <| fun x hx ↦ by rwa [smul_one_smul N x s] at hx From 3654a23e3f7e943473833547d4e77c1b73f1bffb Mon Sep 17 00:00:00 2001 From: grunweg Date: Mon, 12 Aug 2024 20:23:09 +0000 Subject: [PATCH 226/359] chore: more fixes for the flexible linter (#15318) --- Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean | 5 +++-- Mathlib/Algebra/Group/Subgroup/Finite.lean | 4 +--- Mathlib/Algebra/Module/Submodule/Map.lean | 2 +- Mathlib/Algebra/Order/Antidiag/Pi.lean | 4 +--- Mathlib/Algebra/Polynomial/Degree/Definitions.lean | 2 +- Mathlib/Algebra/Polynomial/Div.lean | 2 +- Mathlib/Analysis/Analytic/Inverse.lean | 3 ++- Mathlib/CategoryTheory/Adjunction/Mates.lean | 3 +-- Mathlib/CategoryTheory/Limits/Bicones.lean | 5 ++--- Mathlib/CategoryTheory/Limits/Final.lean | 4 ++-- Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean | 2 +- Mathlib/CategoryTheory/Localization/Bousfield.lean | 4 +--- Mathlib/CategoryTheory/Subobject/Lattice.lean | 2 +- Mathlib/Combinatorics/Enumerative/Composition.lean | 3 +-- Mathlib/Data/DFinsupp/Basic.lean | 2 +- Mathlib/Data/List/Cycle.lean | 2 +- Mathlib/Data/List/Sym.lean | 2 +- Mathlib/Data/PFunctor/Multivariate/M.lean | 3 ++- Mathlib/Data/PNat/Xgcd.lean | 2 +- Mathlib/Data/Set/Card.lean | 2 +- Mathlib/Data/ZMod/Basic.lean | 4 ++-- Mathlib/GroupTheory/CoprodI.lean | 2 +- Mathlib/GroupTheory/Coxeter/Inversion.lean | 2 +- Mathlib/GroupTheory/Perm/Sign.lean | 2 +- Mathlib/GroupTheory/Perm/Support.lean | 6 +++--- Mathlib/LinearAlgebra/Finsupp.lean | 2 +- Mathlib/LinearAlgebra/LinearPMap.lean | 2 +- Mathlib/LinearAlgebra/SesquilinearForm.lean | 2 +- Mathlib/LinearAlgebra/Span.lean | 2 +- Mathlib/LinearAlgebra/StdBasis.lean | 4 ++-- Mathlib/MeasureTheory/Function/SimpleFunc.lean | 3 ++- Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean | 5 +++-- Mathlib/MeasureTheory/Measure/AEDisjoint.lean | 2 +- Mathlib/MeasureTheory/Measure/MeasureSpace.lean | 2 +- Mathlib/NumberTheory/Dioph.lean | 2 +- Mathlib/NumberTheory/Divisors.lean | 3 +-- Mathlib/NumberTheory/PellMatiyasevic.lean | 9 ++++++--- Mathlib/Order/LiminfLimsup.lean | 4 ++-- Mathlib/RingTheory/Multiplicity.lean | 3 +-- Mathlib/RingTheory/MvPowerSeries/Trunc.lean | 2 +- Mathlib/RingTheory/PowerSeries/Order.lean | 2 +- Mathlib/SetTheory/Ordinal/Notation.lean | 6 +++--- Mathlib/Topology/Category/Profinite/Nobeling.lean | 2 +- Mathlib/Topology/Connected/Clopen.lean | 4 ++-- Mathlib/Topology/Connected/PathConnected.lean | 4 ++-- Mathlib/Topology/ContinuousOn.lean | 2 +- Mathlib/Topology/Order/LeftRightLim.lean | 3 +-- Mathlib/Topology/UniformSpace/Cauchy.lean | 4 ++-- Mathlib/Topology/UniformSpace/Equicontinuity.lean | 2 +- 49 files changed, 73 insertions(+), 77 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean b/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean index 8b608aa47fd69..f8a6bb3e6e61e 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Unitization.lean @@ -152,8 +152,9 @@ theorem _root_.AlgHomClass.unitization_injective' {F R S A : Type*} [CommRing R] simp_rw [map_add, hf, ← Unitization.algebraMap_eq_inl, AlgHomClass.commutes] at hx rw [add_eq_zero_iff_eq_neg] at hx ⊢ by_cases hr : r = 0 - · ext <;> simp [hr] at hx ⊢ - exact hx + · ext + · simp [hr] + · simpa [hr] using hx · exact (h r hr <| hx ▸ (neg_mem a.property)).elim /-- This is a generic version which allows us to prove both diff --git a/Mathlib/Algebra/Group/Subgroup/Finite.lean b/Mathlib/Algebra/Group/Subgroup/Finite.lean index 289bb71cc3b31..14f53f72374fd 100644 --- a/Mathlib/Algebra/Group/Subgroup/Finite.lean +++ b/Mathlib/Algebra/Group/Subgroup/Finite.lean @@ -179,9 +179,7 @@ theorem pi_mem_of_mulSingle_mem_aux [DecidableEq η] (I : Finset η) {H : Subgro by_cases heq : j = i · subst heq simp - · simp [heq] - apply h1 j - simpa [heq] using hj + · simpa [heq] using h1 j (by simpa [heq] using hj) · intro j hj have : j ≠ i := by rintro rfl diff --git a/Mathlib/Algebra/Module/Submodule/Map.lean b/Mathlib/Algebra/Module/Submodule/Map.lean index 7e0f78656d3e7..e6d75f1363cf8 100644 --- a/Mathlib/Algebra/Module/Submodule/Map.lean +++ b/Mathlib/Algebra/Module/Submodule/Map.lean @@ -299,7 +299,7 @@ variable [RingHomSurjective σ₁₂] {f : F} /-- `map f` and `comap f` form a `GaloisCoinsertion` when `f` is injective. -/ def gciMapComap (hf : Injective f) : GaloisCoinsertion (map f) (comap f) := (gc_map_comap f).toGaloisCoinsertion fun S x => by - simp [mem_comap, mem_map, forall_exists_index, and_imp] + simp only [mem_comap, mem_map, forall_exists_index, and_imp] intro y hy hxy rw [hf.eq_iff] at hxy rwa [← hxy] diff --git a/Mathlib/Algebra/Order/Antidiag/Pi.lean b/Mathlib/Algebra/Order/Antidiag/Pi.lean index 7af083a0c2b08..3bbf985a82d41 100644 --- a/Mathlib/Algebra/Order/Antidiag/Pi.lean +++ b/Mathlib/Algebra/Order/Antidiag/Pi.lean @@ -233,9 +233,7 @@ lemma map_nsmul_piAntidiag (s : Finset ι) (m : ℕ) {n : ℕ} (hn : n ≠ 0) : lemma nsmul_piAntidiag_univ [Fintype ι] (m : ℕ) {n : ℕ} (hn : n ≠ 0) : @SMul.smul _ _ Finset.smulFinset n (piAntidiag univ m) = (piAntidiag univ (n * m)).filter fun f : ι → ℕ ↦ ∀ i, n ∣ f i := by - have := nsmul_piAntidiag (univ : Finset ι) m hn - simp at this - convert this + simpa using nsmul_piAntidiag (univ : Finset ι) m hn lemma map_nsmul_piAntidiag_univ [Fintype ι] (m : ℕ) {n : ℕ} (hn : n ≠ 0) : (piAntidiag (univ : Finset ι) m).map diff --git a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean index 70d9b2bb4743e..3a4e2fc353d1c 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean @@ -524,7 +524,7 @@ theorem ne_zero_of_natDegree_gt {n : ℕ} (h : n < natDegree p) : p ≠ 0 := fun theorem degree_lt_degree (h : natDegree p < natDegree q) : degree p < degree q := by by_cases hp : p = 0 - · simp [hp] + · simp only [hp, degree_zero] rw [bot_lt_iff_ne_bot] intro hq simp [hp, degree_eq_bot.mp hq, lt_irrefl] at h diff --git a/Mathlib/Algebra/Polynomial/Div.lean b/Mathlib/Algebra/Polynomial/Div.lean index f50623b79b73a..61c5825557e4d 100644 --- a/Mathlib/Algebra/Polynomial/Div.lean +++ b/Mathlib/Algebra/Polynomial/Div.lean @@ -504,7 +504,7 @@ theorem rootMultiplicity_eq_multiplicity [DecidableEq R] [@DecidableRel R[X] (· (p : R[X]) (a : R) : rootMultiplicity a p = if h0 : p = 0 then 0 else (multiplicity (X - C a) p).get (multiplicity_X_sub_C_finite a h0) := - by simp [multiplicity, rootMultiplicity, Part.Dom]; congr; funext; congr + by simp only [rootMultiplicity, multiplicity, PartENat.find_get]; congr; funext; congr @[simp] theorem rootMultiplicity_zero {x : R} : rootMultiplicity x 0 = 0 := diff --git a/Mathlib/Analysis/Analytic/Inverse.lean b/Mathlib/Analysis/Analytic/Inverse.lean index dc8f10b6918b0..704910afccb1c 100644 --- a/Mathlib/Analysis/Analytic/Inverse.lean +++ b/Mathlib/Analysis/Analytic/Inverse.lean @@ -284,7 +284,8 @@ theorem leftInv_eq_rightInv (p : FormalMultilinearSeries 𝕜 E F) (i : E ≃L[ (h : p 1 = (continuousMultilinearCurryFin1 𝕜 E F).symm i) : leftInv p i = rightInv p i := calc leftInv p i = leftInv p.removeZero i := by rw [leftInv_removeZero] - _ = rightInv p.removeZero i := by apply leftInv_eq_rightInv_aux <;> simp; exact h + _ = rightInv p.removeZero i := by + apply leftInv_eq_rightInv_aux _ _ (by simpa using h) (by simp) _ = rightInv p i := by rw [rightInv_removeZero] /-! diff --git a/Mathlib/CategoryTheory/Adjunction/Mates.lean b/Mathlib/CategoryTheory/Adjunction/Mates.lean index bedcd72458b01..cacbc727a8ab8 100644 --- a/Mathlib/CategoryTheory/Adjunction/Mates.lean +++ b/Mathlib/CategoryTheory/Adjunction/Mates.lean @@ -432,8 +432,7 @@ theorem conjugateEquiv_comp (α : L₂ ⟶ L₁) (β : L₃ ⟶ L₂) : (L₃.leftUnitor.hom ≫ β ≫ L₂.rightUnitor.inv) have vcompd := congr_app vcomp d dsimp [mateEquiv, leftAdjointSquare.vcomp, rightAdjointSquare.vcomp] at vcompd - simp at vcompd - simp only [comp_id, id_comp, assoc, map_comp] + simp only [comp_id, id_comp, assoc, map_comp] at vcompd ⊢ rw [vcompd] theorem conjugateEquiv_symm_comp (α : R₁ ⟶ R₂) (β : R₂ ⟶ R₃) : diff --git a/Mathlib/CategoryTheory/Limits/Bicones.lean b/Mathlib/CategoryTheory/Limits/Bicones.lean index a928f42f0ad2c..a2be62de66858 100644 --- a/Mathlib/CategoryTheory/Limits/Bicones.lean +++ b/Mathlib/CategoryTheory/Limits/Bicones.lean @@ -48,8 +48,7 @@ instance : Inhabited (Bicone J) := instance finBicone [Fintype J] : Fintype (Bicone J) where elems := [Bicone.left, Bicone.right].toFinset ∪ Finset.image Bicone.diagram Fintype.elems complete j := by - cases j <;> simp - apply Fintype.complete + cases j <;> simp [Fintype.complete] variable [Category.{v₁} J] @@ -65,7 +64,7 @@ instance : Inhabited (BiconeHom J Bicone.left Bicone.left) := ⟨BiconeHom.left_id⟩ instance BiconeHom.decidableEq {j k : Bicone J} : DecidableEq (BiconeHom J j k) := fun f g => by - cases f <;> cases g <;> simp <;> infer_instance + cases f <;> cases g <;> simp only [diagram.injEq] <;> infer_instance @[simps] instance biconeCategoryStruct : CategoryStruct (Bicone J) where diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 88f2a7239bcc4..70bc2fb14a714 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -224,7 +224,7 @@ def extendCocone : Cocone (F ⋙ G) ⥤ Cocone G where ι := { app := fun X => G.map (homToLift F X) ≫ c.ι.app (lift F X) naturality := fun X Y f => by - dsimp; simp + dsimp; simp only [Category.comp_id] -- This would be true if we'd chosen `lift F X` to be `lift F Y` -- and `homToLift F X` to be `f ≫ homToLift F Y`. apply @@ -494,7 +494,7 @@ def extendCone : Cone (F ⋙ G) ⥤ Cone G where π := { app := fun d => c.π.app (lift F d) ≫ G.map (homToLift F d) naturality := fun X Y f => by - dsimp; simp + dsimp; simp only [Category.id_comp, Category.assoc] -- This would be true if we'd chosen `lift F Y` to be `lift F X` -- and `homToLift F Y` to be `homToLift F X ≫ f`. apply diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean index 6072892b6b6f8..3b1dc95db77ed 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Multiequalizer.lean @@ -325,7 +325,7 @@ def ofι (I : MulticospanIndex.{w} C) (P : C) (ι : ∀ a, P ⟶ I.left a) | WalkingMulticospan.right b => ι (I.fstTo b) ≫ I.fst b naturality := by rintro (_ | _) (_ | _) (_ | _ | _) <;> - dsimp <;> simp + dsimp <;> simp only [Category.id_comp, Category.comp_id] apply w } @[reassoc (attr := simp)] diff --git a/Mathlib/CategoryTheory/Localization/Bousfield.lean b/Mathlib/CategoryTheory/Localization/Bousfield.lean index 35fde70a59774..1dd1de3c1799e 100644 --- a/Mathlib/CategoryTheory/Localization/Bousfield.lean +++ b/Mathlib/CategoryTheory/Localization/Bousfield.lean @@ -74,9 +74,7 @@ lemma W_isoClosure : W (isoClosure P) = W P := by exact ⟨a ≫ e.inv, by simp only [reassoc_of% h, e.hom_inv_id, comp_id]⟩ instance : (W P).IsMultiplicative where - id_mem X Z _ := by - simp [id_comp] - exact Function.bijective_id + id_mem X Z _ := by simpa [id_comp] using Function.bijective_id comp_mem f g hf hg Z hZ := by simpa using Function.Bijective.comp (hf Z hZ) (hg Z hZ) diff --git a/Mathlib/CategoryTheory/Subobject/Lattice.lean b/Mathlib/CategoryTheory/Subobject/Lattice.lean index a70aade5ffa1b..ac8612b0e93bd 100644 --- a/Mathlib/CategoryTheory/Subobject/Lattice.lean +++ b/Mathlib/CategoryTheory/Subobject/Lattice.lean @@ -476,7 +476,7 @@ theorem finset_sup_factors {I : Type*} {A B : C} {s : Finset I} {P : I → Subob · rintro ⟨_, ⟨⟨⟩, _⟩⟩ · rintro ⟨j, ⟨m, h⟩⟩ simp only [Finset.sup_insert] - simp at m + simp only [Finset.mem_insert] at m rcases m with (rfl | m) · exact sup_factors_of_factors_left h · exact sup_factors_of_factors_right (ih ⟨j, ⟨m, h⟩⟩) diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index 87ab6914334a6..05f69e27409d1 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -761,8 +761,7 @@ def compositionAsSetEquiv (n : ℕ) : CompositionAsSet n ≃ Finset (Fin (n - 1) Finset.mem_univ, forall_true_left, Finset.mem_filter, add_eq_zero, and_false, add_left_inj, false_or, true_and] erw [Set.mem_setOf_eq] - simp [this, false_or_iff, add_right_inj, add_eq_zero, one_ne_zero, false_and_iff, - Fin.val_mk] + simp only [Finset.mem_val] constructor · intro h cases' h with n h diff --git a/Mathlib/Data/DFinsupp/Basic.lean b/Mathlib/Data/DFinsupp/Basic.lean index 4b0b99b3d2e74..8b5dbe5f93246 100644 --- a/Mathlib/Data/DFinsupp/Basic.lean +++ b/Mathlib/Data/DFinsupp/Basic.lean @@ -1039,7 +1039,7 @@ instance decidableZero [∀ (i) (x : β i), Decidable (x = 0)] (f : Π₀ i, β case neg => exact (s.prop i).resolve_left hs₂ theorem support_subset_iff {s : Set ι} {f : Π₀ i, β i} : ↑f.support ⊆ s ↔ ∀ i ∉ s, f i = 0 := by - simp [Set.subset_def]; exact forall_congr' fun i => not_imp_comm + simpa [Set.subset_def] using forall_congr' fun i => not_imp_comm theorem support_single_ne_zero {i : ι} {b : β i} (hb : b ≠ 0) : (single i b).support = {i} := by ext j; by_cases h : i = j diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index dcf48a84137c4..d16894eab2b25 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -462,7 +462,7 @@ instance : Inhabited (Cycle α) := theorem induction_on {C : Cycle α → Prop} (s : Cycle α) (H0 : C nil) (HI : ∀ (a) (l : List α), C ↑l → C ↑(a :: l)) : C s := Quotient.inductionOn' s fun l => by - refine List.recOn l ?_ ?_ <;> simp + refine List.recOn l ?_ ?_ <;> simp only [mk''_eq_coe, coe_nil] assumption' /-- For `x : α`, `s : Cycle α`, `x ∈ s` indicates that `x` occurs at least once in `s`. -/ diff --git a/Mathlib/Data/List/Sym.lean b/Mathlib/Data/List/Sym.lean index eaf32bb1f3904..80954c1ad9dc7 100644 --- a/Mathlib/Data/List/Sym.lean +++ b/Mathlib/Data/List/Sym.lean @@ -123,7 +123,7 @@ theorem map_mk_sublist_sym2 (x : α) (xs : List α) (h : x ∈ xs) : induction xs with | nil => simp | cons x' xs ih => - simp [List.sym2] + simp only [map_cons, List.sym2, cons_append] cases h with | head => exact (sublist_append_left _ _).cons₂ _ diff --git a/Mathlib/Data/PFunctor/Multivariate/M.lean b/Mathlib/Data/PFunctor/Multivariate/M.lean index 8ed6fc37adad7..8a40846b7c6fc 100644 --- a/Mathlib/Data/PFunctor/Multivariate/M.lean +++ b/Mathlib/Data/PFunctor/Multivariate/M.lean @@ -299,7 +299,8 @@ theorem M.map_dest {α β : TypeVec n} (g : (α ::: P.M α) ⟹ (β ::: P.M β)) (h : ∀ x : P.M α, lastFun g x = (dropFun g <$$> x : P.M β)) : g <$$> M.dest P x = M.dest P (dropFun g <$$> x) := by rw [M.dest_map]; congr - apply eq_of_drop_last_eq <;> simp + apply eq_of_drop_last_eq (by simp) + simp only [lastFun_appendFun] ext1; apply h end MvPFunctor diff --git a/Mathlib/Data/PNat/Xgcd.lean b/Mathlib/Data/PNat/Xgcd.lean index b7727098b0b6a..03c2d2cb0227d 100644 --- a/Mathlib/Data/PNat/Xgcd.lean +++ b/Mathlib/Data/PNat/Xgcd.lean @@ -133,7 +133,7 @@ def IsSpecial' : Prop := theorem isSpecial_iff : u.IsSpecial ↔ u.IsSpecial' := by dsimp [IsSpecial, IsSpecial'] let ⟨wp, x, y, zp, ap, bp⟩ := u - constructor <;> intro h <;> simp [w, z, succPNat] at * <;> + constructor <;> intro h <;> simp only [w, succPNat, succ_eq_add_one, z] at * <;> simp only [← coe_inj, mul_coe, mk_coe] at * · simp_all [← h]; ring · simp [Nat.mul_add, Nat.add_mul, ← Nat.add_assoc] at h; rw [← h]; ring diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index d8d7d4405bf96..115d6555eaac1 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -944,7 +944,7 @@ theorem exists_eq_insert_iff_ncard (hs : s.Finite := by toFinite_tac) : cases' t.finite_or_infinite with ht ht · rw [ncard_eq_toFinset_card _ hs, ncard_eq_toFinset_card _ ht, ← @Finite.toFinset_subset_toFinset _ _ _ hs ht, ← Finset.exists_eq_insert_iff] - convert Iff.rfl using 2; simp + convert Iff.rfl using 2; simp only [Finite.mem_toFinset] ext x simp [Finset.ext_iff, Set.ext_iff] simp only [ht.ncard, exists_prop, add_eq_zero, and_false, iff_false, not_exists, not_and] diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 6c696dc0794a1..0a65bb9d6fac8 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -729,8 +729,8 @@ theorem val_add_of_le {n : ℕ} [NeZero n] {a b : ZMod n} (h : n ≤ a.val + b.v theorem val_add_le {n : ℕ} (a b : ZMod n) : (a + b).val ≤ a.val + b.val := by cases n - · simp [ZMod.val]; apply Int.natAbs_add_le - · simp [ZMod.val_add]; apply Nat.mod_le + · simpa [ZMod.val] using Int.natAbs_add_le _ _ + · simpa [ZMod.val_add] using Nat.mod_le _ _ theorem val_mul {n : ℕ} (a b : ZMod n) : (a * b).val = a.val * b.val % n := by cases n diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index 547d6a74acec3..71732e3b60cb4 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -535,7 +535,7 @@ theorem mem_smul_iff {i j : ι} {m₁ : M i} {m₂ : M j} {w : Word M} : · rintro rfl exact Or.inl ⟨_, rfl, rfl⟩ · rintro (⟨_, h, rfl⟩ | hm') - · simp [Sigma.ext_iff] at h + · simp only [Sigma.ext_iff, heq_eq_eq, true_and] at h subst h rfl · simp only [fstIdx, Option.map_eq_some', Sigma.exists, diff --git a/Mathlib/GroupTheory/Coxeter/Inversion.lean b/Mathlib/GroupTheory/Coxeter/Inversion.lean index ff587ef4dcc2e..4901cbdf6ff8b 100644 --- a/Mathlib/GroupTheory/Coxeter/Inversion.lean +++ b/Mathlib/GroupTheory/Coxeter/Inversion.lean @@ -377,7 +377,7 @@ theorem prod_rightInvSeq (ω : List B) : prod (ris ω) = (π ω)⁻¹ := by · simp [rightInvSeq, ih, wordProd_cons] theorem prod_leftInvSeq (ω : List B) : prod (lis ω) = (π ω)⁻¹ := by - simp [leftInvSeq_eq_reverse_rightInvSeq_reverse, prod_reverse_noncomm] + simp only [leftInvSeq_eq_reverse_rightInvSeq_reverse, prod_reverse_noncomm, inv_inj] have : List.map (fun x ↦ x⁻¹) (ris ω.reverse) = ris ω.reverse := calc List.map (fun x ↦ x⁻¹) (ris ω.reverse) _ = List.map id (ris ω.reverse) := by diff --git a/Mathlib/GroupTheory/Perm/Sign.lean b/Mathlib/GroupTheory/Perm/Sign.lean index eecab704aa297..1f4c096e2337d 100644 --- a/Mathlib/GroupTheory/Perm/Sign.lean +++ b/Mathlib/GroupTheory/Perm/Sign.lean @@ -184,7 +184,7 @@ theorem signBijAux_injOn {n : ℕ} {f : Perm (Fin n)} : rw [Finset.mem_coe, mem_finPairsLT] at * have : ¬b₁ < b₂ := hb.le.not_lt split_ifs at h <;> - simp_all [(Equiv.injective f).eq_iff, eq_self_iff_true, and_self_iff, heq_iff_eq] + simp_all only [not_lt, Sigma.mk.inj_iff, (Equiv.injective f).eq_iff, heq_eq_eq] · exact absurd this (not_le.mpr ha) · exact absurd this (not_le.mpr ha) diff --git a/Mathlib/GroupTheory/Perm/Support.lean b/Mathlib/GroupTheory/Perm/Support.lean index e967d7aef1e73..63593ca5a6715 100644 --- a/Mathlib/GroupTheory/Perm/Support.lean +++ b/Mathlib/GroupTheory/Perm/Support.lean @@ -379,9 +379,9 @@ theorem support_swap {x y : α} (h : x ≠ y) : support (swap x y) = {x, y} := b ext z by_cases hx : z = x any_goals simpa [hx] using h.symm - by_cases hy : z = y <;> - · simp [swap_apply_of_ne_of_ne, hx, hy] <;> - exact h + by_cases hy : z = y + · simpa [swap_apply_of_ne_of_ne, hx, hy] using h + · simp [swap_apply_of_ne_of_ne, hx, hy] theorem support_swap_iff (x y : α) : support (swap x y) = {x, y} ↔ x ≠ y := by refine ⟨fun h => ?_, fun h => support_swap h⟩ diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index e03004b27ef2f..209908215216b 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -554,7 +554,7 @@ theorem lmapDomain_disjoint_ker (f : α → α') {s : Set α} rw [disjoint_iff_inf_le] rintro l ⟨h₁, h₂⟩ rw [SetLike.mem_coe, mem_ker, lmapDomain_apply, mapDomain] at h₂ - simp; ext x + simp only [mem_bot]; ext x haveI := Classical.decPred fun x => x ∈ s by_cases xs : x ∈ s · have : Finsupp.sum l (fun a => Finsupp.single (f a)) (f x) = 0 := by diff --git a/Mathlib/LinearAlgebra/LinearPMap.lean b/Mathlib/LinearAlgebra/LinearPMap.lean index a0c9f4fde5fa5..a22821e0e63c5 100644 --- a/Mathlib/LinearAlgebra/LinearPMap.lean +++ b/Mathlib/LinearAlgebra/LinearPMap.lean @@ -523,7 +523,7 @@ noncomputable def supSpanSingleton (f : E →ₗ.[K] F) (x : E) (y : F) (hx : x E →ₗ.[K] F := -- Porting note: `simpa [..]` → `simp [..]; exact ..` f.sup (mkSpanSingleton x y fun h₀ => hx <| h₀.symm ▸ f.domain.zero_mem) <| - sup_h_of_disjoint _ _ <| by simp [disjoint_span_singleton]; exact fun h => False.elim <| hx h + sup_h_of_disjoint _ _ <| by simpa [disjoint_span_singleton] using fun h ↦ False.elim <| hx h @[simp] theorem domain_supSpanSingleton (f : E →ₗ.[K] F) (x : E) (y : F) (hx : x ∉ f.domain) : diff --git a/Mathlib/LinearAlgebra/SesquilinearForm.lean b/Mathlib/LinearAlgebra/SesquilinearForm.lean index 8a2b3862b074a..f97b86492d2e9 100644 --- a/Mathlib/LinearAlgebra/SesquilinearForm.lean +++ b/Mathlib/LinearAlgebra/SesquilinearForm.lean @@ -112,7 +112,7 @@ theorem ortho_smul_right {B : V₁ →ₛₗ[I₁] V₂ →ₛₗ[I₂] V} {x y} · rw [map_smulₛₗ, H, smul_zero] · rw [map_smulₛₗ, smul_eq_zero] at H cases' H with H H - · simp at H + · simp only [map_eq_zero] at H exfalso exact ha H · exact H diff --git a/Mathlib/LinearAlgebra/Span.lean b/Mathlib/LinearAlgebra/Span.lean index c560119afc158..9e71038dfec4b 100644 --- a/Mathlib/LinearAlgebra/Span.lean +++ b/Mathlib/LinearAlgebra/Span.lean @@ -758,7 +758,7 @@ theorem prod_inf_prod : prod p q₁ ⊓ prod p' q₁' = prod (p ⊓ p') (q₁ theorem prod_sup_prod : prod p q₁ ⊔ prod p' q₁' = prod (p ⊔ p') (q₁ ⊔ q₁') := by refine le_antisymm (sup_le (prod_mono le_sup_left le_sup_left) (prod_mono le_sup_right le_sup_right)) ?_ - simp [SetLike.le_def]; intro xx yy hxx hyy + simp only [SetLike.le_def, mem_prod, and_imp, Prod.forall]; intro xx yy hxx hyy rcases mem_sup.1 hxx with ⟨x, hx, x', hx', rfl⟩ rcases mem_sup.1 hyy with ⟨y, hy, y', hy', rfl⟩ exact mem_sup.2 ⟨(x, y), ⟨hx, hy⟩, (x', y'), ⟨hx', hy'⟩, rfl⟩ diff --git a/Mathlib/LinearAlgebra/StdBasis.lean b/Mathlib/LinearAlgebra/StdBasis.lean index d058dc9b63955..bd602506eb518 100644 --- a/Mathlib/LinearAlgebra/StdBasis.lean +++ b/Mathlib/LinearAlgebra/StdBasis.lean @@ -302,8 +302,8 @@ theorem stdBasis_eq_stdBasisMatrix (i : m) (j : n) [DecidableEq m] [DecidableEq -- Porting note: `simp` fails to apply `Pi.basis_apply` ext a b by_cases hi : i = a <;> by_cases hj : j = b - · simp [stdBasis, hi, hj, Basis.coe_reindex, comp_apply, Equiv.sigmaEquivProd_symm_apply, - StdBasisMatrix.apply_same] + · simp only [stdBasis, hi, hj, Basis.coe_reindex, comp_apply, Equiv.sigmaEquivProd_symm_apply, + StdBasisMatrix.apply_same] erw [Pi.basis_apply] simp · simp only [stdBasis, hi, Basis.coe_reindex, comp_apply, Equiv.sigmaEquivProd_symm_apply, diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index ab03e2640207b..76cad3cf36a87 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -913,7 +913,8 @@ theorem const_lintegral (c : ℝ≥0∞) : (const α c).lintegral μ = c * μ un rw [lintegral] cases isEmpty_or_nonempty α · simp [μ.eq_zero_of_isEmpty] - · simp; unfold Function.const; rw [preimage_const_of_mem (mem_singleton c)] + · simp only [range_const, coe_const, Finset.sum_singleton] + unfold Function.const; rw [preimage_const_of_mem (mem_singleton c)] theorem const_lintegral_restrict (c : ℝ≥0∞) (s : Set α) : (const α c).lintegral (μ.restrict s) = c * μ s := by diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean index 0f0f93a47b486..ec89ef7faa51b 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean @@ -602,8 +602,9 @@ noncomputable def equivImage (s : Set α) (hf : MeasurableEmbedding f) : s ≃ toEquiv := Equiv.Set.image f s hf.injective measurable_toFun := (hf.measurable.comp measurable_id.subtype_val).subtype_mk measurable_invFun := by - rintro t ⟨u, hu, rfl⟩; simp [preimage_preimage, Set.image_symm_preimage hf.injective] - exact measurable_subtype_coe (hf.measurableSet_image' hu) + rintro t ⟨u, hu, rfl⟩ + simpa [preimage_preimage, Set.image_symm_preimage hf.injective] + using measurable_subtype_coe (hf.measurableSet_image' hu) /-- The domain of `f` is equivalent to its range as measurable spaces, if `f` is a measurable embedding -/ diff --git a/Mathlib/MeasureTheory/Measure/AEDisjoint.lean b/Mathlib/MeasureTheory/Measure/AEDisjoint.lean index bbf8c6bbd9839..4e7005ee1da8e 100644 --- a/Mathlib/MeasureTheory/Measure/AEDisjoint.lean +++ b/Mathlib/MeasureTheory/Measure/AEDisjoint.lean @@ -118,7 +118,7 @@ theorem exists_disjoint_diff (h : AEDisjoint μ s t) : ∃ u, MeasurableSet u ∧ μ u = 0 ∧ Disjoint (s \ u) t := ⟨toMeasurable μ (s ∩ t), measurableSet_toMeasurable _ _, (measure_toMeasurable _).trans h, disjoint_sdiff_self_left.mono_left fun x hx => by - simp; exact ⟨hx.1, fun hxt => hx.2 <| subset_toMeasurable _ _ ⟨hx.1, hxt⟩⟩⟩ + simpa using ⟨hx.1, fun hxt => hx.2 <| subset_toMeasurable _ _ ⟨hx.1, hxt⟩⟩⟩ theorem of_null_right (h : μ t = 0) : AEDisjoint μ s t := measure_mono_null inter_subset_right h diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 1e59d29ef22d1..9ada0295ca07b 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -500,7 +500,7 @@ theorem measure_iInter_eq_iInf' {α ι : Type*} [MeasurableSpace α] {μ : Measu μ (⋂ i, f i) = ⨅ i, μ (⋂ j ≤ i, f j) := by let s := fun i ↦ ⋂ j ≤ i, f j have iInter_eq : ⋂ i, f i = ⋂ i, s i := by - ext x; simp [s]; constructor + ext x; simp only [mem_iInter, s]; constructor · exact fun h _ j _ ↦ h j · intro h i rcases directed_of (· ≤ ·) i i with ⟨j, rij, -⟩ diff --git a/Mathlib/NumberTheory/Dioph.lean b/Mathlib/NumberTheory/Dioph.lean index b6602777b1e87..674ecffb62b68 100644 --- a/Mathlib/NumberTheory/Dioph.lean +++ b/Mathlib/NumberTheory/Dioph.lean @@ -594,7 +594,7 @@ theorem mod_dioph : DiophFn fun v => f v % g v := (vectorAll_iff_forall _).1 fun z x y => show ((y = 0 ∨ z < y) ∧ ∃ c, z + y * c = x) ↔ x % y = z from ⟨fun ⟨h, c, hc⟩ => by - rw [← hc]; simp; cases' h with x0 hl + rw [← hc]; simp only [add_mul_mod_self_left]; cases' h with x0 hl · rw [x0, mod_zero] exact mod_eq_of_lt hl, fun e => by rw [← e] diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 3c64e6b088196..78221e801b0c2 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -401,8 +401,7 @@ theorem properDivisors_eq_singleton_one_iff_prime : n.properDivisors = {1} ↔ n | Nat.succ (Nat.succ n) => simp [succ_le_succ] · rw [← mem_singleton, ← h, mem_properDivisors] have := Nat.le_of_dvd ?_ hdvd - · simp [hdvd, this] - exact (le_iff_eq_or_lt.mp this).symm + · simpa [hdvd, this] using (le_iff_eq_or_lt.mp this).symm · by_contra! simp only [nonpos_iff_eq_zero.mp this, this] at h contradiction diff --git a/Mathlib/NumberTheory/PellMatiyasevic.lean b/Mathlib/NumberTheory/PellMatiyasevic.lean index bd603798cbd3d..66f902be634a1 100644 --- a/Mathlib/NumberTheory/PellMatiyasevic.lean +++ b/Mathlib/NumberTheory/PellMatiyasevic.lean @@ -358,7 +358,7 @@ theorem strictMono_y : StrictMono (yn a1) have : yn a1 m ≤ yn a1 n := Or.elim (lt_or_eq_of_le <| Nat.le_of_succ_le_succ h) (fun hl => le_of_lt <| strictMono_y hl) fun e => by rw [e] - simp; refine lt_of_le_of_lt ?_ (Nat.lt_add_of_pos_left <| x_pos a1 n) + simp only [yn_succ, gt_iff_lt]; refine lt_of_le_of_lt ?_ (Nat.lt_add_of_pos_left <| x_pos a1 n) rw [← mul_one (yn a1 m)] exact mul_le_mul this (le_of_lt a1) (Nat.zero_le _) (Nat.zero_le _) @@ -368,7 +368,8 @@ theorem strictMono_x : StrictMono (xn a1) have : xn a1 m ≤ xn a1 n := Or.elim (lt_or_eq_of_le <| Nat.le_of_succ_le_succ h) (fun hl => le_of_lt <| strictMono_x hl) fun e => by rw [e] - simp; refine lt_of_lt_of_le (lt_of_le_of_lt this ?_) (Nat.le_add_right _ _) + simp only [xn_succ, gt_iff_lt] + refine lt_of_lt_of_le (lt_of_le_of_lt this ?_) (Nat.le_add_right _ _) have t := Nat.mul_lt_mul_of_pos_left a1 (x_pos a1 n) rwa [mul_one] at t @@ -816,7 +817,9 @@ theorem matiyasevic {a k x y} : (tk : yn b1 j ≡ k [MOD 4 * yn a1 i])⟩, (ky : k ≤ yn a1 i) => (Nat.eq_zero_or_pos i).elim - (fun i0 => by simp [i0] at ky; rw [i0, ky]; exact ⟨rfl, rfl⟩) fun ipos => by + (fun i0 => by + simp only [i0, yn_zero, nonpos_iff_eq_zero] at ky; rw [i0, ky]; exact ⟨rfl, rfl⟩) + fun ipos => by suffices i = k by rw [this]; exact ⟨rfl, rfl⟩ clear o rem xy uv st have iln : i ≤ n := diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 826d4e82c0cf3..daf5f974aacdb 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -710,11 +710,11 @@ theorem limsInf_bot : limsInf (⊥ : Filter α) = ⊤ := @[simp] theorem limsSup_top : limsSup (⊤ : Filter α) = ⊤ := - top_unique <| le_sInf <| by simp [eq_univ_iff_forall]; exact fun b hb => top_unique <| hb _ + top_unique <| le_sInf <| by simpa [eq_univ_iff_forall] using fun b hb => top_unique <| hb _ @[simp] theorem limsInf_top : limsInf (⊤ : Filter α) = ⊥ := - bot_unique <| sSup_le <| by simp [eq_univ_iff_forall]; exact fun b hb => bot_unique <| hb _ + bot_unique <| sSup_le <| by simpa [eq_univ_iff_forall] using fun b hb => bot_unique <| hb _ @[simp] theorem blimsup_false {f : Filter β} {u : β → α} : (blimsup u f fun _ => False) = ⊥ := by diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index 0ea5d8494b28e..effafa7779249 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -529,8 +529,7 @@ theorem Finset.prod {β : Type*} {p : α} (hp : Prime p) (s : Finset β) (f : β induction' s using Finset.induction with a s has ih h · simp only [Finset.sum_empty, Finset.prod_empty] convert one_right hp.not_unit - · simp [has, ← ih] - convert multiplicity.mul hp + · simpa [has, ← ih] using multiplicity.mul hp -- Porting note: with protected could not use pow' k in the succ branch protected theorem pow' {p a : α} (hp : Prime p) (ha : Finite p a) : diff --git a/Mathlib/RingTheory/MvPowerSeries/Trunc.lean b/Mathlib/RingTheory/MvPowerSeries/Trunc.lean index 66b722579e2c5..4c1ecd9d3176f 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Trunc.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Trunc.lean @@ -91,7 +91,7 @@ theorem trunc_c (n : σ →₀ ℕ) (hnn : n ≠ 0) (a : R) : trunc R n (C σ R MvPolynomial.ext _ _ fun m => by classical rw [coeff_trunc, coeff_C, MvPolynomial.coeff_C] - split_ifs with H <;> first |rfl|try simp_all + split_ifs with H <;> first |rfl|try simp_all only [ne_eq, not_true_eq_false] exfalso; apply H; subst m; exact Ne.bot_lt hnn end Trunc diff --git a/Mathlib/RingTheory/PowerSeries/Order.lean b/Mathlib/RingTheory/PowerSeries/Order.lean index b87f41f0ee1ca..f73b2bb21d977 100644 --- a/Mathlib/RingTheory/PowerSeries/Order.lean +++ b/Mathlib/RingTheory/PowerSeries/Order.lean @@ -349,8 +349,8 @@ theorem divided_by_X_pow_orderMul {f g : R⟦X⟧} (hf : f ≠ 0) (hg : g ≠ 0) _ = X ^ (df + dg) * divided_by_X_pow_order hf * divided_by_X_pow_order hg := by rw [pow_add] _ = X ^ dfg * divided_by_X_pow_order hf * divided_by_X_pow_order hg := by rw [H_add_d] _ = X ^ dfg * (divided_by_X_pow_order hf * divided_by_X_pow_order hg) := by rw [mul_assoc] - simp [← hdfg, this] at H refine (IsLeftCancelMulZero.mul_left_cancel_of_ne_zero (pow_ne_zero dfg X_ne_zero) ?_).symm + simp only [this] at H convert H end OrderIsDomain diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 9806ceba908c9..71c72e2dc3afd 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -639,7 +639,7 @@ theorem opow_def (o₁ o₂ : ONote) : o₁ ^ o₂ = opowAux2 o₂ (split o₁) theorem split_eq_scale_split' : ∀ {o o' m} [NF o], split' o = (o', m) → split o = (scale 1 o', m) | 0, o', m, _, p => by injection p; substs o' m; rfl | oadd e n a, o', m, h, p => by - by_cases e0 : e = 0 <;> simp [e0, split, split'] at p ⊢ + by_cases e0 : e = 0 <;> simp only [split', e0, ↓reduceIte, Prod.mk.injEq, split] at p ⊢ · rcases p with ⟨rfl, rfl⟩ exact ⟨rfl, rfl⟩ · revert p @@ -755,9 +755,9 @@ instance nf_opow (o₁ o₂) [NF o₁] [NF o₂] : NF (o₁ ^ o₂) := by decide · simp only [(· ^ ·), Pow.pow, pow, opow, opowAux2, mulNat_eq_mul, ofNat, *] infer_instance - · simp [(· ^ ·),Pow.pow,pow, opow, opowAux2, e₁, e₂, split_eq_scale_split' e₂] + · simp only [(· ^ ·), Pow.pow, opow, opowAux2, e₁, split_eq_scale_split' e₂, mulNat_eq_mul] have := na.fst - cases' k with k <;> simp + cases' k with k · infer_instance · cases k <;> cases m <;> infer_instance diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index 94b7c9055f65a..03bbf9e71818c 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -860,7 +860,7 @@ precomposition with the projections defined in the section `Projections`. theorem contained_eq_proj (o : Ordinal) (h : contained C o) : C = π C (ord I · < o) := by have := proj_prop_eq_self C (ord I · < o) - simp [π, Bool.not_eq_false] at this + simp only [ne_eq, Bool.not_eq_false, π] at this exact (this (fun i x hx ↦ h x hx i)).symm theorem isClosed_proj (o : Ordinal) (hC : IsClosed C) : IsClosed (π C (ord I · < o)) := diff --git a/Mathlib/Topology/Connected/Clopen.lean b/Mathlib/Topology/Connected/Clopen.lean index e99f549d58af3..cfba16bc1231f 100644 --- a/Mathlib/Topology/Connected/Clopen.lean +++ b/Mathlib/Topology/Connected/Clopen.lean @@ -255,7 +255,7 @@ theorem isPreconnected_iff_subset_of_disjoint {s : Set α} : · intro u v hu hv hs huv specialize h u v hu hv hs contrapose! huv - simp [not_subset] at huv + simp only [not_subset] at huv rcases huv with ⟨⟨x, hxs, hxu⟩, ⟨y, hys, hyv⟩⟩ have hxv : x ∈ v := or_iff_not_imp_left.mp (hs hxs) hxu have hyu : y ∈ u := or_iff_not_imp_right.mp (hs hys) hyv @@ -312,7 +312,7 @@ theorem isPreconnected_iff_subset_of_disjoint_closed : rw [isPreconnected_closed_iff] at h specialize h u v hu hv hs contrapose! huv - simp [not_subset] at huv + simp only [not_subset] at huv rcases huv with ⟨⟨x, hxs, hxu⟩, ⟨y, hys, hyv⟩⟩ have hxv : x ∈ v := or_iff_not_imp_left.mp (hs hxs) hxu have hyu : y ∈ u := or_iff_not_imp_right.mp (hs hys) hyv diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index dba770f9f562c..81f6d23e10a49 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -769,13 +769,13 @@ theorem JoinedIn.refl (h : x ∈ F) : JoinedIn F x x := @[symm] theorem JoinedIn.symm (h : JoinedIn F x y) : JoinedIn F y x := by cases' h.mem with hx hy - simp_all [joinedIn_iff_joined] + simp_all only [joinedIn_iff_joined] exact h.symm theorem JoinedIn.trans (hxy : JoinedIn F x y) (hyz : JoinedIn F y z) : JoinedIn F x z := by cases' hxy.mem with hx hy cases' hyz.mem with hx hy - simp_all [joinedIn_iff_joined] + simp_all only [joinedIn_iff_joined] exact hxy.trans hyz theorem Specializes.joinedIn (h : x ⤳ y) (hx : x ∈ F) (hy : y ∈ F) : JoinedIn F x y := by diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index f31cbf8ecdb0a..44d171d882c3f 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -1093,7 +1093,7 @@ theorem continuous_if {p : α → Prop} {f g : α → β} [∀ a, Decidable (p a (hg : ContinuousOn g (closure { x | ¬p x })) : Continuous fun a => if p a then f a else g a := by rw [continuous_iff_continuousOn_univ] - apply ContinuousOn.if <;> simp <;> assumption + apply ContinuousOn.if <;> simpa theorem Continuous.if {p : α → Prop} {f g : α → β} [∀ a, Decidable (p a)] (hp : ∀ a ∈ frontier { x | p x }, f a = g a) (hf : Continuous f) (hg : Continuous g) : diff --git a/Mathlib/Topology/Order/LeftRightLim.lean b/Mathlib/Topology/Order/LeftRightLim.lean index 09a573dd0a4b5..d6fc5b3889c57 100644 --- a/Mathlib/Topology/Order/LeftRightLim.lean +++ b/Mathlib/Topology/Order/LeftRightLim.lean @@ -149,8 +149,7 @@ theorem rightLim_le_leftLim (h : x < y) : rightLim f x ≤ leftLim f y := by letI : TopologicalSpace α := Preorder.topology α haveI : OrderTopology α := ⟨rfl⟩ rcases eq_or_ne (𝓝[<] y) ⊥ with (h' | h') - · simp [leftLim, h'] - exact rightLim_le hf h + · simpa [leftLim, h'] using rightLim_le hf h obtain ⟨a, ⟨xa, ay⟩⟩ : (Ioo x y).Nonempty := forall_mem_nonempty_iff_neBot.2 (neBot_iff.2 h') (Ioo x y) (Ioo_mem_nhdsWithin_Iio ⟨h, le_refl _⟩) diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index 36a5bea483aed..6ade1b68d70e8 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -563,8 +563,8 @@ theorem TotallyBounded.image [UniformSpace β] {f : α → β} {s : Set α} (hs preimage_iUnion, preimage_setOf_eq] simp? [subset_def] at hct says simp only [mem_setOf_eq, subset_def, mem_iUnion, exists_prop] at hct - intro x hx; simp - exact hct x hx⟩ + intro x hx + simpa using hct x hx⟩ theorem Ultrafilter.cauchy_of_totallyBounded {s : Set α} (f : Ultrafilter α) (hs : TotallyBounded s) (h : ↑f ≤ 𝓟 s) : Cauchy (f : Filter α) := diff --git a/Mathlib/Topology/UniformSpace/Equicontinuity.lean b/Mathlib/Topology/UniformSpace/Equicontinuity.lean index 69321becf013e..d047acb2e03e0 100644 --- a/Mathlib/Topology/UniformSpace/Equicontinuity.lean +++ b/Mathlib/Topology/UniformSpace/Equicontinuity.lean @@ -567,7 +567,7 @@ theorem uniformEquicontinuousOn_iInf_rng {u : κ → UniformSpace α'} {F : ι theorem equicontinuousWithinAt_iInf_dom {t : κ → TopologicalSpace X'} {F : ι → X' → α} {S : Set X'} {x₀ : X'} {k : κ} (hk : EquicontinuousWithinAt (tX := t k) F S x₀) : EquicontinuousWithinAt (tX := ⨅ k, t k) F S x₀ := by - simp [equicontinuousWithinAt_iff_continuousWithinAt (tX := _)] at hk ⊢ + simp only [equicontinuousWithinAt_iff_continuousWithinAt (tX := _)] at hk ⊢ unfold ContinuousWithinAt nhdsWithin at hk ⊢ rw [nhds_iInf] exact hk.mono_left <| inf_le_inf_right _ <| iInf_le _ k From 68d96e18b0bbca45b0f3bcfb4555f5766eff6a32 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 12 Aug 2024 21:17:59 +0000 Subject: [PATCH 227/359] chore: add comment to a borderline `set_option maxHeartbeats` (#15655) count_heartbeats reports `192105`, which is very close to the limit. Found by the linter #13653. --- Mathlib/CategoryTheory/Monoidal/Bimon_.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean index 2a5a87a7c6305..9d5818ef51531 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean @@ -68,6 +68,9 @@ def toComon_ : Bimon_ C ⥤ Comon_ C := (Mon_.forgetMonoidal C).toOplaxMonoidalF @[simp] theorem toComon_forget : toComon_ C ⋙ Comon_.forget C = forget C := rfl +-- TODO: the `set_option` is not strictly necessary, but the declaration is just a heartbeat +-- away from using too many heartbeats. Squeezing `(d)simp` improves the situation, but pulls +-- out too many lemmas set_option maxHeartbeats 400000 in /-- The object level part of the forward direction of `Comon_ (Mon_ C) ≌ Mon_ (Comon_ C)` -/ def toMon_Comon_obj (M : Bimon_ C) : Mon_ (Comon_ C) where From 7a462dc8edd53b2bd59710a4183d11a0fa10c5fb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 12 Aug 2024 22:22:14 +0000 Subject: [PATCH 228/359] feat(SetTheory/Game/Ordinal): game product of ordinals is natural product (#15690) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I also flipped and renamed `toPGame_nadd` at the request of Yaël. --- Mathlib/SetTheory/Game/Basic.lean | 24 ++++----- Mathlib/SetTheory/Game/Ordinal.lean | 77 +++++++++++++++++++---------- 2 files changed, 64 insertions(+), 37 deletions(-) diff --git a/Mathlib/SetTheory/Game/Basic.lean b/Mathlib/SetTheory/Game/Basic.lean index 13796f6637a38..2752abc297073 100644 --- a/Mathlib/SetTheory/Game/Basic.lean +++ b/Mathlib/SetTheory/Game/Basic.lean @@ -149,6 +149,8 @@ theorem lt_iff_game_lt {x y : PGame} : x < y ↔ (⟦x⟧ : Game) < ⟦y⟧ := theorem equiv_iff_game_eq {x y : PGame} : x ≈ y ↔ (⟦x⟧ : Game) = ⟦y⟧ := (@Quotient.eq' _ _ x y).symm +alias ⟨game_eq, _⟩ := equiv_iff_game_eq + theorem fuzzy_iff_game_fuzzy {x y : PGame} : x ‖ y ↔ Game.Fuzzy ⟦x⟧ ⟦y⟧ := Iff.rfl @@ -396,21 +398,19 @@ theorem quot_mul_comm (x y : PGame.{u}) : (⟦x * y⟧ : Game) = ⟦y * x⟧ := theorem mul_comm_equiv (x y : PGame) : x * y ≈ y * x := Quotient.exact <| quot_mul_comm _ _ -instance isEmpty_mul_zero_leftMoves (x : PGame.{u}) : IsEmpty (x * 0).LeftMoves := by - cases x - exact instIsEmptySum - -instance isEmpty_mul_zero_rightMoves (x : PGame.{u}) : IsEmpty (x * 0).RightMoves := by +instance isEmpty_leftMoves_mul (x y : PGame.{u}) + [IsEmpty (x.LeftMoves × y.LeftMoves ⊕ x.RightMoves × y.RightMoves)] : + IsEmpty (x * y).LeftMoves := by cases x - apply instIsEmptySum - -instance isEmpty_zero_mul_leftMoves (x : PGame.{u}) : IsEmpty (0 * x).LeftMoves := by - cases x - apply instIsEmptySum + cases y + assumption -instance isEmpty_zero_mul_rightMoves (x : PGame.{u}) : IsEmpty (0 * x).RightMoves := by +instance isEmpty_rightMoves_mul (x y : PGame.{u}) + [IsEmpty (x.LeftMoves × y.RightMoves ⊕ x.RightMoves × y.LeftMoves)] : + IsEmpty (x * y).RightMoves := by cases x - apply instIsEmptySum + cases y + assumption /-- `x * 0` has exactly the same moves as `0`. -/ def mulZeroRelabelling (x : PGame) : x * 0 ≡r 0 := diff --git a/Mathlib/SetTheory/Game/Ordinal.lean b/Mathlib/SetTheory/Game/Ordinal.lean index 48fc98aeacaca..f5a9f99f65e82 100644 --- a/Mathlib/SetTheory/Game/Ordinal.lean +++ b/Mathlib/SetTheory/Game/Ordinal.lean @@ -154,30 +154,57 @@ noncomputable def toPGameEmbedding : Ordinal.{u} ↪o PGame.{u} where inj' := toPGame_injective map_rel_iff' := @toPGame_le_iff -/-- The sum of ordinals as games corresponds to natural addition of ordinals. -/ -theorem toPGame_add : ∀ a b : Ordinal.{u}, a.toPGame + b.toPGame ≈ (a ♯ b).toPGame - | a, b => by - refine ⟨le_of_forall_lf (fun i => ?_) isEmptyElim, le_of_forall_lf (fun i => ?_) isEmptyElim⟩ - · apply leftMoves_add_cases i <;> - intro i <;> - let wf := toLeftMovesToPGame_symm_lt i <;> - (try rw [add_moveLeft_inl]) <;> - (try rw [add_moveLeft_inr]) <;> - rw [toPGame_moveLeft', lf_congr_left (toPGame_add _ _), toPGame_lf_iff] - · exact nadd_lt_nadd_right wf _ - · exact nadd_lt_nadd_left wf _ - · rw [toPGame_moveLeft'] - rcases lt_nadd_iff.1 (toLeftMovesToPGame_symm_lt i) with (⟨c, hc, hc'⟩ | ⟨c, hc, hc'⟩) <;> - rw [← toPGame_le_iff, ← le_congr_right (toPGame_add _ _)] at hc' <;> - apply lf_of_le_of_lf hc' - · apply add_lf_add_right - rwa [toPGame_lf_iff] - · apply add_lf_add_left - rwa [toPGame_lf_iff] -termination_by a b => (a, b) - -@[simp] -theorem toPGame_add_mk' (a b : Ordinal) : (⟦a.toPGame⟧ + ⟦b.toPGame⟧ : Game) = ⟦(a ♯ b).toPGame⟧ := - Quot.sound (toPGame_add a b) +/-- Converts an ordinal into the corresponding game. -/ +noncomputable abbrev toGame (o : Ordinal) : Game := ⟦o.toPGame⟧ + +/-- The natural addition of ordinals corresponds to their sum as games. -/ +theorem toPGame_nadd (a b : Ordinal.{u}) : (a ♯ b).toPGame ≈ a.toPGame + b.toPGame := by + refine ⟨le_of_forall_lf (fun i => ?_) isEmptyElim, le_of_forall_lf (fun i => ?_) isEmptyElim⟩ + · rw [toPGame_moveLeft'] + rcases lt_nadd_iff.1 (toLeftMovesToPGame_symm_lt i) with (⟨c, hc, hc'⟩ | ⟨c, hc, hc'⟩) <;> + rw [← toPGame_le_iff, le_congr_right (toPGame_nadd _ _)] at hc' <;> + apply lf_of_le_of_lf hc' + · apply add_lf_add_right + rwa [toPGame_lf_iff] + · apply add_lf_add_left + rwa [toPGame_lf_iff] + · apply leftMoves_add_cases i <;> + intro i <;> + let wf := toLeftMovesToPGame_symm_lt i <;> + (try rw [add_moveLeft_inl]) <;> + (try rw [add_moveLeft_inr]) <;> + rw [toPGame_moveLeft', ← lf_congr_left (toPGame_nadd _ _), toPGame_lf_iff] + · exact nadd_lt_nadd_right wf _ + · exact nadd_lt_nadd_left wf _ +termination_by (a, b) + +theorem toGame_nadd (a b : Ordinal) : (a ♯ b).toGame = a.toGame + b.toGame := + Quot.sound (toPGame_nadd a b) + +/-- The natural multiplication of ordinals corresponds to their product as pre-games. -/ +theorem toPGame_nmul (a b : Ordinal.{u}) : (a ⨳ b).toPGame ≈ a.toPGame * b.toPGame := by + refine ⟨le_of_forall_lf (fun i => ?_) isEmptyElim, le_of_forall_lf (fun i => ?_) isEmptyElim⟩ + · rw [toPGame_moveLeft'] + rcases lt_nmul_iff.1 (toLeftMovesToPGame_symm_lt i) with ⟨c, hc, d, hd, h⟩ + rw [← toPGame_le_iff, le_iff_game_le, ← toGame, ← toGame, toGame_nadd _ _, toGame_nadd _ _, + ← le_sub_iff_add_le] at h + refine lf_of_le_of_lf h <| (lf_congr_left ?_).1 <| moveLeft_lf <| toLeftMovesMul <| Sum.inl + ⟨toLeftMovesToPGame ⟨c, hc⟩, toLeftMovesToPGame ⟨d, hd⟩⟩ + simp only [mul_moveLeft_inl, toPGame_moveLeft', Equiv.symm_apply_apply, equiv_iff_game_eq, + quot_sub, quot_add] + repeat rw [← game_eq (toPGame_nmul _ _)] + rfl + · apply leftMoves_mul_cases i ?_ isEmptyElim + intro i j + rw [mul_moveLeft_inl, toPGame_moveLeft', toPGame_moveLeft', lf_iff_game_lf, + quot_sub, quot_add, ← Game.not_le, le_sub_iff_add_le] + repeat rw [← game_eq (toPGame_nmul _ _)] + repeat rw [← toGame_nadd] + apply toPGame_lf (nmul_nadd_lt _ _) <;> + exact toLeftMovesToPGame_symm_lt _ +termination_by (a, b) + +theorem toGame_nmul (a b : Ordinal) : (a ⨳ b).toGame = ⟦a.toPGame * b.toPGame⟧ := + Quot.sound (toPGame_nmul a b) end Ordinal From 577bed1b4cd26b7a3138f67bfdef4ab08cd11f56 Mon Sep 17 00:00:00 2001 From: Nicolas Rolland Date: Mon, 12 Aug 2024 22:49:32 +0000 Subject: [PATCH 229/359] feat(CategoryTheory/Adjunction): `typeToCat` is right adjoint to `connectedComponents` functor (#15614) The embedding of `Type` into `Cat`, which views a set as a discrete category, is right adjoint to the functor `connectedComponents : Cat -> Set` which maps a category to its set of connected components. --- .../Category/Cat/Adjunction.lean | 35 ++++++++++++++----- .../CategoryTheory/ConnectedComponents.lean | 27 +++++++++++++- Mathlib/CategoryTheory/IsConnected.lean | 10 ++++++ 3 files changed, 63 insertions(+), 9 deletions(-) diff --git a/Mathlib/CategoryTheory/Category/Cat/Adjunction.lean b/Mathlib/CategoryTheory/Category/Cat/Adjunction.lean index b4d89ce7a89ce..8d930249e4120 100644 --- a/Mathlib/CategoryTheory/Category/Cat/Adjunction.lean +++ b/Mathlib/CategoryTheory/Category/Cat/Adjunction.lean @@ -5,25 +5,23 @@ Authors: Nicolas Rolland -/ import Mathlib.CategoryTheory.Category.Cat import Mathlib.CategoryTheory.Adjunction.Basic +import Mathlib.CategoryTheory.ConnectedComponents + /-! # Adjunctions related to Cat, the category of categories The embedding `typeToCat: Type ⥤ Cat`, mapping a type to the corresponding discrete category, is left adjoint to the functor `Cat.objects`, which maps a category to its set of objects. - +Another functor `connectedComponents : Cat ⥤ Type` maps a category to the set of its connected +components and functors to functions between those sets. ## Notes All this could be made with 2-functors -## TODO - -Define the left adjoint `Cat.connectedComponents`, which map -a category to its set of connected components. - -/ -universe u +universe v u namespace CategoryTheory.Cat variable (X : Type u) (C : Cat) @@ -42,7 +40,7 @@ private def typeToCatObjectsAdjCounitApp : (Cat.objects ⋙ typeToCat).obj C ⥤ /-- `typeToCat : Type ⥤ Cat` is left adjoint to `Cat.objects : Cat ⥤ Type` -/ def typeToCatObjectsAdj : typeToCat ⊣ Cat.objects where - homEquiv := typeToCatObjectsAdjHomEquiv + homEquiv := typeToCatObjectsAdjHomEquiv unit := { app:= fun _ ↦ Discrete.mk } counit := { app := typeToCatObjectsAdjCounitApp @@ -51,4 +49,25 @@ def typeToCatObjectsAdj : typeToCat ⊣ Cat.objects where obtain rfl := Discrete.eq_of_hom f aesop_cat ) } +/-- The connected components functor -/ +def connectedComponents : Cat.{v, u} ⥤ Type u where + obj C := ConnectedComponents C + map F := + Quotient.lift (Quotient.mk (Zigzag.setoid _) ∘ F.obj) + (fun _ _ ↦ Quot.sound ∘ zigzag_obj_of_zigzag F) + map_id _ := funext fun x ↦ (Quotient.exists_rep x).elim (fun _ h ↦ by simp [<- h]; rfl) + map_comp _ _ := funext fun x ↦ (Quotient.exists_rep x).elim (fun _ h => by simp [<- h]) + +/-- `typeToCat : Type ⥤ Cat` is right adjoint to `connectedComponents : Cat ⥤ Type` -/ +def connectedComponentsTypeToCatAdj : connectedComponents ⊣ typeToCat where + homEquiv C X := ConnectedComponents.typeToCatHomEquiv C X + unit := { app:= fun C ↦ ConnectedComponents.functorToDiscrete _ (𝟙 (connectedComponents.obj C)) } + counit := { + app := fun X => ConnectedComponents.liftFunctor _ (𝟙 typeToCat.obj X) + naturality := fun _ _ _ => + funext (fun xcc => by + obtain ⟨x,h⟩ := Quotient.exists_rep xcc + aesop_cat) } + homEquiv_counit := fun {C X G} => by funext cc;obtain ⟨_,_⟩ := Quotient.exists_rep cc; aesop_cat + end CategoryTheory.Cat diff --git a/Mathlib/CategoryTheory/ConnectedComponents.lean b/Mathlib/CategoryTheory/ConnectedComponents.lean index 4f198a18289bb..6f018d34805a0 100644 --- a/Mathlib/CategoryTheory/ConnectedComponents.lean +++ b/Mathlib/CategoryTheory/ConnectedComponents.lean @@ -31,7 +31,6 @@ namespace CategoryTheory attribute [instance 100] IsConnected.is_nonempty variable {J : Type u₁} [Category.{v₁} J] -variable {C : Type u₂} [Category.{u₁} C] /-- This type indexes the connected components of the category `J`. -/ def ConnectedComponents (J : Type u₁) [Category.{v₁} J] : Type u₁ := @@ -40,6 +39,32 @@ def ConnectedComponents (J : Type u₁) [Category.{v₁} J] : Type u₁ := instance [Inhabited J] : Inhabited (ConnectedComponents J) := ⟨Quotient.mk'' default⟩ +/-- Every function from connected components of a category gives a functor to discrete category -/ +def ConnectedComponents.functorToDiscrete (X : Type*) + (f : ConnectedComponents J → X) : J ⥤ Discrete X where + obj Y := Discrete.mk (f (Quotient.mk (Zigzag.setoid _) Y)) + map g := Discrete.eqToHom (congrArg f (Quotient.sound (Zigzag.of_hom g))) + +/-- Every functor to a discrete category gives a function from connected components -/ +def ConnectedComponents.liftFunctor (J) [Category J] {X : Type*} (F :J ⥤ Discrete X) : + (ConnectedComponents J → X) := + Quotient.lift (fun c => (F.obj c).as) + (fun _ _ h => eq_of_zigzag X (zigzag_obj_of_zigzag F h)) + +/-- Functions from connected components and functors to discrete category are in bijection -/ +def ConnectedComponents.typeToCatHomEquiv (J) [Category J] (X : Type*) : + (ConnectedComponents J → X) ≃ (J ⥤ Discrete X) where + toFun := ConnectedComponents.functorToDiscrete _ + invFun := ConnectedComponents.liftFunctor _ + left_inv := fun f ↦ funext fun x ↦ by + obtain ⟨x, h⟩ := Quotient.exists_rep x + rw [← h] + rfl + right_inv := fun fctr ↦ + Functor.hext (fun _ ↦ rfl) (fun c d f ↦ + have : Subsingleton (fctr.obj c ⟶ fctr.obj d) := Discrete.instSubsingletonDiscreteHom _ _ + (Subsingleton.elim (fctr.map f) _).symm.heq) + /-- Given an index for a connected component, produce the actual component as a full subcategory. -/ def Component (j : ConnectedComponents J) : Type u₁ := FullSubcategory fun k => Quotient.mk'' k = j diff --git a/Mathlib/CategoryTheory/IsConnected.lean b/Mathlib/CategoryTheory/IsConnected.lean index 88545d2a24ec0..4f4fd20d39dbe 100644 --- a/Mathlib/CategoryTheory/IsConnected.lean +++ b/Mathlib/CategoryTheory/IsConnected.lean @@ -331,6 +331,16 @@ theorem zigzag_obj_of_zigzag (F : J ⥤ K) {j₁ j₂ : J} (h : Zigzag j₁ j₂ Zigzag (F.obj j₁) (F.obj j₂) := h.lift _ fun _ _ => Or.imp (Nonempty.map fun f => F.map f) (Nonempty.map fun f => F.map f) +/-- A Zag in a discrete category entails an equality of its extremities -/ +lemma eq_of_zag (X) {a b : Discrete X} (h : Zag a b) : a.as = b.as := + h.elim (fun ⟨f⟩ ↦ Discrete.eq_of_hom f) (fun ⟨f⟩ ↦ (Discrete.eq_of_hom f).symm) + +/-- A zigzag in a discrete category entails an equality of its extremities -/ +lemma eq_of_zigzag (X) {a b : Discrete X} (h : Zigzag a b) : a.as = b.as := by + induction h with + | refl => rfl + | tail _ h eq => exact eq.trans (eq_of_zag _ h) + -- TODO: figure out the right way to generalise this to `Zigzag`. theorem zag_of_zag_obj (F : J ⥤ K) [F.Full] {j₁ j₂ : J} (h : Zag (F.obj j₁) (F.obj j₂)) : Zag j₁ j₂ := From 08570c45d1aa03a1fbd90726896570534d03275f Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 12 Aug 2024 23:05:08 +0000 Subject: [PATCH 230/359] feat: `Matrix.PosDef` on diagonal and numeric matrices (#13974) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main result here is `PosDef (diagonal d) ↔ ∀ i, 0 < d i`, everything else is just special or less interesting cases. --- Mathlib/LinearAlgebra/Matrix/DotProduct.lean | 43 ++++++++- Mathlib/LinearAlgebra/Matrix/Hermitian.lean | 13 +++ Mathlib/LinearAlgebra/Matrix/PosDef.lean | 98 +++++++++++++++++++- 3 files changed, 144 insertions(+), 10 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean index 4b648de325167..a549bac8773ce 100644 --- a/Mathlib/LinearAlgebra/Matrix/DotProduct.lean +++ b/Mathlib/LinearAlgebra/Matrix/DotProduct.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ +import Mathlib.Algebra.Ring.Regular import Mathlib.Algebra.Star.Order import Mathlib.Data.Matrix.Basic import Mathlib.LinearAlgebra.StdBasis @@ -91,18 +92,30 @@ theorem dotProduct_self_eq_zero [LinearOrderedRing R] {v : n → R} : dotProduct section StarOrderedRing -variable [PartialOrder R] [NonUnitalRing R] [StarRing R] [StarOrderedRing R] [NoZeroDivisors R] +variable [PartialOrder R] [NonUnitalRing R] [StarRing R] [StarOrderedRing R] -/-- Note that this applies to `ℂ` via `Complex.strictOrderedCommRing`. -/ +/-- Note that this applies to `ℂ` via `RCLike.toStarOrderedRing`. -/ +@[simp] +theorem dotProduct_star_self_nonneg (v : n → R) : 0 ≤ dotProduct (star v) v := + Fintype.sum_nonneg fun _ => star_mul_self_nonneg _ + +/-- Note that this applies to `ℂ` via `RCLike.toStarOrderedRing`. -/ +@[simp] +theorem dotProduct_self_star_nonneg (v : n → R) : 0 ≤ dotProduct v (star v) := + Fintype.sum_nonneg fun _ => mul_star_self_nonneg _ + +variable [NoZeroDivisors R] + +/-- Note that this applies to `ℂ` via `RCLike.toStarOrderedRing`. -/ @[simp] theorem dotProduct_star_self_eq_zero {v : n → R} : dotProduct (star v) v = 0 ↔ v = 0 := - (Finset.sum_eq_zero_iff_of_nonneg fun i _ => (star_mul_self_nonneg (r := v i) : _)).trans <| + (Fintype.sum_eq_zero_iff_of_nonneg fun i => star_mul_self_nonneg _).trans <| by simp [Function.funext_iff, mul_eq_zero] -/-- Note that this applies to `ℂ` via `Complex.strictOrderedCommRing`. -/ +/-- Note that this applies to `ℂ` via `RCLike.toStarOrderedRing`. -/ @[simp] theorem dotProduct_self_star_eq_zero {v : n → R} : dotProduct v (star v) = 0 ↔ v = 0 := - (Finset.sum_eq_zero_iff_of_nonneg fun i _ => (mul_star_self_nonneg (r := v i) : _)).trans <| + (Fintype.sum_eq_zero_iff_of_nonneg fun i => mul_star_self_nonneg _).trans <| by simp [Function.funext_iff, mul_eq_zero] @[simp] @@ -155,6 +168,26 @@ lemma vecMul_self_mul_conjTranspose_eq_zero (A : Matrix m n R) (v : m → R) : v ᵥ* (A * Aᴴ) = 0 ↔ v ᵥ* A = 0 := by simpa only [conjTranspose_conjTranspose] using vecMul_conjTranspose_mul_self_eq_zero Aᴴ _ +/-- Note that this applies to `ℂ` via `RCLike.toStarOrderedRing`. -/ +@[simp] +theorem dotProduct_star_self_pos_iff {v : n → R} : + 0 < dotProduct (star v) v ↔ v ≠ 0 := by + cases subsingleton_or_nontrivial R + · obtain rfl : v = 0 := Subsingleton.elim _ _ + simp + refine (Fintype.sum_pos_iff_of_nonneg fun i => star_mul_self_nonneg _).trans ?_ + simp_rw [Pi.lt_def, Function.ne_iff, Pi.zero_apply] + refine (and_iff_right fun i => star_mul_self_nonneg (v i)).trans <| exists_congr fun i => ?_ + constructor + · rintro h hv + simp [hv] at h + · exact (star_mul_self_pos <| isRegular_of_ne_zero ·) + +/-- Note that this applies to `ℂ` via `RCLike.toStarOrderedRing`. -/ +@[simp] +theorem dotProduct_self_star_pos_iff {v : n → R} : 0 < dotProduct v (star v) ↔ v ≠ 0 := by + simpa using dotProduct_star_self_pos_iff (v := star v) + end StarOrderedRing end Self diff --git a/Mathlib/LinearAlgebra/Matrix/Hermitian.lean b/Mathlib/LinearAlgebra/Matrix/Hermitian.lean index 0cbcc086ea586..0201300fe7f22 100644 --- a/Mathlib/LinearAlgebra/Matrix/Hermitian.lean +++ b/Mathlib/LinearAlgebra/Matrix/Hermitian.lean @@ -210,11 +210,24 @@ require `Fintype n`, which is necessary for `Monoid (Matrix n n R)`. -/ theorem isHermitian_one [DecidableEq n] : (1 : Matrix n n α).IsHermitian := conjTranspose_one +@[simp] +theorem isHermitian_natCast [DecidableEq n] (d : ℕ) : (d : Matrix n n α).IsHermitian := + conjTranspose_natCast _ + theorem IsHermitian.pow [Fintype n] [DecidableEq n] {A : Matrix n n α} (h : A.IsHermitian) (k : ℕ) : (A ^ k).IsHermitian := IsSelfAdjoint.pow h _ end Semiring +section Ring +variable [Ring α] [StarRing α] + +@[simp] +theorem isHermitian_intCast [DecidableEq n] (d : ℤ) : (d : Matrix n n α).IsHermitian := + conjTranspose_intCast _ + +end Ring + section CommRing variable [CommRing α] [StarRing α] diff --git a/Mathlib/LinearAlgebra/Matrix/PosDef.lean b/Mathlib/LinearAlgebra/Matrix/PosDef.lean index fe70520844931..5aee7b5572d59 100644 --- a/Mathlib/LinearAlgebra/Matrix/PosDef.lean +++ b/Mathlib/LinearAlgebra/Matrix/PosDef.lean @@ -45,13 +45,17 @@ nonnegative for all `x`. -/ def PosSemidef (M : Matrix n n R) := M.IsHermitian ∧ ∀ x : n → R, 0 ≤ dotProduct (star x) (M *ᵥ x) +protected theorem PosSemidef.diagonal [StarOrderedRing R] [DecidableEq n] {d : n → R} (h : 0 ≤ d) : + PosSemidef (diagonal d) := + ⟨isHermitian_diagonal_of_self_adjoint _ <| funext fun i => IsSelfAdjoint.of_nonneg (h i), + fun x => by + refine Fintype.sum_nonneg fun i => ?_ + simpa only [mulVec_diagonal, ← mul_assoc] using conjugate_nonneg (h i) _⟩ + /-- A diagonal matrix is positive semidefinite iff its diagonal entries are nonnegative. -/ lemma posSemidef_diagonal_iff [StarOrderedRing R] [DecidableEq n] {d : n → R} : - PosSemidef (diagonal d) ↔ (∀ i : n, 0 ≤ d i) := by - refine ⟨fun ⟨_, hP⟩ i ↦ by simpa using hP (Pi.single i 1), ?_⟩ - refine fun hd ↦ ⟨isHermitian_diagonal_iff.2 fun i ↦ IsSelfAdjoint.of_nonneg (hd i), ?_⟩ - refine fun x ↦ Finset.sum_nonneg fun i _ ↦ ?_ - simpa only [mulVec_diagonal, mul_assoc] using conjugate_nonneg (hd i) _ + PosSemidef (diagonal d) ↔ (∀ i : n, 0 ≤ d i) := + ⟨fun ⟨_, hP⟩ i ↦ by simpa using hP (Pi.single i 1), .diagonal⟩ namespace PosSemidef @@ -97,6 +101,31 @@ protected lemma one [StarOrderedRing R] [DecidableEq n] : PosSemidef (1 : Matrix ⟨isHermitian_one, fun x => by rw [one_mulVec]; exact Fintype.sum_nonneg fun i => star_mul_self_nonneg _⟩ +protected theorem natCast [StarOrderedRing R] [DecidableEq n] (d : ℕ) : + PosSemidef (d : Matrix n n R) := + ⟨isHermitian_natCast _, fun x => by + simp only [natCast_mulVec, dotProduct_smul] + rw [Nat.cast_smul_eq_nsmul] + refine nsmul_nonneg (dotProduct_star_self_nonneg _) _⟩ + +-- See note [no_index around OfNat.ofNat] +protected theorem ofNat [StarOrderedRing R] [DecidableEq n] (d : ℕ) [d.AtLeastTwo] : + PosSemidef (no_index (OfNat.ofNat d) : Matrix n n R) := + .natCast d + +protected theorem intCast [StarOrderedRing R] [DecidableEq n] (d : ℤ) (hd : 0 ≤ d) : + PosSemidef (d : Matrix n n R) := + ⟨isHermitian_intCast _, fun x => by + simp only [intCast_mulVec, dotProduct_smul] + rw [Int.cast_smul_eq_zsmul] + refine zsmul_nonneg (dotProduct_star_self_nonneg _) hd⟩ + +@[simp] +protected theorem _root_.Matrix.posSemidef_intCast_iff + [StarOrderedRing R] [DecidableEq n] [Nonempty n] [Nontrivial R] (d : ℤ) : + PosSemidef (d : Matrix n n R) ↔ 0 ≤ d := + posSemidef_diagonal_iff.trans <| by simp [Pi.le_def] + protected lemma pow [StarOrderedRing R] [DecidableEq n] {M : Matrix n n R} (hM : M.PosSemidef) (k : ℕ) : PosSemidef (M ^ k) := @@ -318,6 +347,65 @@ theorem transpose {M : Matrix n n R} (hM : M.PosDef) : Mᵀ.PosDef := by convert hM.2 (star x) (star_ne_zero.2 hx) using 1 rw [mulVec_transpose, Matrix.dotProduct_mulVec, star_star, dotProduct_comm] +protected theorem diagonal [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] + {d : n → R} (h : ∀ i, 0 < d i) : + PosDef (diagonal d) := + ⟨isHermitian_diagonal_of_self_adjoint _ <| funext fun i => IsSelfAdjoint.of_nonneg (h i).le, + fun x hx => by + refine Fintype.sum_pos ?_ + simp_rw [mulVec_diagonal, ← mul_assoc, Pi.lt_def] + obtain ⟨i, hi⟩ := Function.ne_iff.mp hx + exact ⟨fun i => conjugate_nonneg (h i).le _, + i, conjugate_pos (h _) (isRegular_of_ne_zero hi)⟩⟩ + +@[simp] +theorem _root_.Matrix.posDef_diagonal_iff + [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] [Nontrivial R] {d : n → R} : + PosDef (diagonal d) ↔ ∀ i, 0 < d i := by + refine ⟨fun h i => ?_, .diagonal⟩ + have := h.2 (Pi.single i 1) + simp only [mulVec_single, mul_one, dotProduct_diagonal', Pi.star_apply, Pi.single_eq_same, + star_one, one_mul, Function.ne_iff, Pi.zero_apply] at this + exact this ⟨i, by simp⟩ + +protected theorem one [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] : + PosDef (1 : Matrix n n R) := + ⟨isHermitian_one, fun x hx => by simpa only [one_mulVec, dotProduct_star_self_pos_iff]⟩ + +protected theorem natCast [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] + (d : ℕ) (hd : d ≠ 0) : + PosDef (d : Matrix n n R) := + ⟨isHermitian_natCast _, fun x hx => by + simp only [natCast_mulVec, dotProduct_smul] + rw [Nat.cast_smul_eq_nsmul] + refine nsmul_pos (dotProduct_star_self_pos_iff.mpr hx) hd⟩ + +@[simp] +theorem _root_.Matrix.posDef_natCast_iff [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] + [Nonempty n] [Nontrivial R] {d : ℕ} : + PosDef (d : Matrix n n R) ↔ 0 < d := + posDef_diagonal_iff.trans <| by simp + +-- See note [no_index around OfNat.ofNat] +protected theorem ofNat [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] + (d : ℕ) [d.AtLeastTwo] : + PosDef (no_index (OfNat.ofNat d) : Matrix n n R) := + .natCast d (NeZero.ne _) + +protected theorem intCast [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] + (d : ℤ) (hd : 0 < d) : + PosDef (d : Matrix n n R) := + ⟨isHermitian_intCast _, fun x hx => by + simp only [intCast_mulVec, dotProduct_smul] + rw [Int.cast_smul_eq_zsmul] + refine zsmul_pos (dotProduct_star_self_pos_iff.mpr hx) hd⟩ + +@[simp] +theorem _root_.Matrix.posDef_intCast_iff [StarOrderedRing R] [DecidableEq n] [NoZeroDivisors R] + [Nonempty n] [Nontrivial R] {d : ℤ} : + PosDef (d : Matrix n n R) ↔ 0 < d := + posDef_diagonal_iff.trans <| by simp + protected lemma add_posSemidef [CovariantClass R R (· + ·) (· ≤ · )] {A : Matrix m m R} {B : Matrix m m R} (hA : A.PosDef) (hB : B.PosSemidef) : (A + B).PosDef := From 08d60f60deac394734356b6ac9aa73ac701d5877 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 13 Aug 2024 06:06:56 +0000 Subject: [PATCH 231/359] chore: move more `assert_not_exists` up (#15747) Co-authored-by: Kim Morrison Co-authored-by: Parcly Taxel --- Mathlib/Algebra/Ring/Center.lean | 1 - Mathlib/CategoryTheory/ConcreteCategory/Basic.lean | 4 ++-- Mathlib/CategoryTheory/Monoidal/Tor.lean | 4 ++-- Mathlib/Computability/TuringMachine.lean | 4 ++-- Mathlib/Data/Fin/Basic.lean | 1 - Mathlib/Data/Fintype/BigOperators.lean | 4 ++-- Mathlib/Data/List/InsertNth.lean | 4 ++-- Mathlib/Data/Multiset/FinsetOps.lean | 8 ++++---- Mathlib/Data/Set/Pointwise/Basic.lean | 4 ++-- Mathlib/Data/Vector/Defs.lean | 2 +- Mathlib/NumberTheory/Wilson.lean | 2 +- Mathlib/Topology/Bornology/Absorbs.lean | 4 ++-- 12 files changed, 20 insertions(+), 22 deletions(-) diff --git a/Mathlib/Algebra/Ring/Center.lean b/Mathlib/Algebra/Ring/Center.lean index d28091e849958..73f4fbcbad95f 100644 --- a/Mathlib/Algebra/Ring/Center.lean +++ b/Mathlib/Algebra/Ring/Center.lean @@ -15,7 +15,6 @@ import Mathlib.Data.Int.Cast.Lemmas assert_not_exists Finset assert_not_exists Subsemigroup - variable {M : Type*} namespace Set diff --git a/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean b/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean index ccef7b95c0f52..6482fc883f061 100644 --- a/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean +++ b/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean @@ -34,11 +34,11 @@ related work. -/ -universe w w' v v' v'' u u' u'' - assert_not_exists CategoryTheory.CommSq assert_not_exists CategoryTheory.Adjunction +universe w w' v v' v'' u u' u'' + namespace CategoryTheory /-- A concrete category is a category `C` with a fixed faithful functor `Forget : C ⥤ Type`. diff --git a/Mathlib/CategoryTheory/Monoidal/Tor.lean b/Mathlib/CategoryTheory/Monoidal/Tor.lean index f7b6014ae640b..4f9a2de2e1361 100644 --- a/Mathlib/CategoryTheory/Monoidal/Tor.lean +++ b/Mathlib/CategoryTheory/Monoidal/Tor.lean @@ -22,6 +22,8 @@ Possibly it's best to axiomatize delta functors, and obtain a unique characteris -/ +assert_not_exists ModuleCat.abelian + noncomputable section open CategoryTheory.Limits @@ -70,5 +72,3 @@ lemma isZero_Tor'_succ_of_projective (X Y : C) [Projective X] (n : ℕ) : apply Functor.isZero_leftDerived_obj_projective_succ end CategoryTheory - -assert_not_exists Module.abelian diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 020bd9ea089ba..70ae4513ecba5 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -58,14 +58,14 @@ Given these parameters, there are a few common structures for the model that ari formalizes "essentially finite" mentioned above. -/ +assert_not_exists MonoidWithZero + -- After https://github.com/leanprover/lean4/pull/4400 -- the simp normal forms for `List` lookup use the `GetElem` typeclass, rather than `List.get?`. -- This file has not been updated to reflect that change, so uses a number of deprecated lemmas. -- Updating this file to allow restoring the deprecation linter would be much appreciated. set_option linter.deprecated false -assert_not_exists MonoidWithZero - open Mathlib (Vector) open Relation diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 0e9ba529e710b..937ec7c6da20d 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -64,7 +64,6 @@ This file expands on the development in the core library. * `Fin.revPerm : Equiv.Perm (Fin n)` : `Fin.rev` as an `Equiv.Perm`, the antitone involution given by `i ↦ n-(i+1)` - -/ assert_not_exists Monoid diff --git a/Mathlib/Data/Fintype/BigOperators.lean b/Mathlib/Data/Fintype/BigOperators.lean index a999dd19f13ea..02838b7dad046 100644 --- a/Mathlib/Data/Fintype/BigOperators.lean +++ b/Mathlib/Data/Fintype/BigOperators.lean @@ -23,10 +23,10 @@ However many of the results here really belong in `Algebra.BigOperators.Group.Fi and should be moved at some point. -/ -open Mathlib - assert_not_exists MulAction +open Mathlib + universe u v variable {α : Type*} {β : Type*} {γ : Type*} diff --git a/Mathlib/Data/List/InsertNth.lean b/Mathlib/Data/List/InsertNth.lean index 98977de0577e2..33044a0a6908c 100644 --- a/Mathlib/Data/List/InsertNth.lean +++ b/Mathlib/Data/List/InsertNth.lean @@ -11,12 +11,12 @@ import Mathlib.Data.List.Basic Proves various lemmas about `List.insertNth`. -/ +assert_not_exists Set.range + open Function open Nat hiding one_pos -assert_not_exists Set.range - namespace List universe u v w diff --git a/Mathlib/Data/Multiset/FinsetOps.lean b/Mathlib/Data/Multiset/FinsetOps.lean index 0e2b4e57ed01b..8cec332c637a5 100644 --- a/Mathlib/Data/Multiset/FinsetOps.lean +++ b/Mathlib/Data/Multiset/FinsetOps.lean @@ -14,6 +14,10 @@ and preparatory for defining the corresponding operations on `Finset`. -/ +-- Assert that we define `Finset` without the material on the set lattice. +-- Note that we cannot put this in `Data.Finset.Basic` because we proved relevant lemmas there. +assert_not_exists Set.sInter + namespace Multiset open List @@ -248,7 +252,3 @@ theorem Subset.ndinter_eq_left {s t : Multiset α} (h : s ⊆ t) : s.ndinter t = rw [quot_mk_to_coe'', quot_mk_to_coe'', coe_ndinter, List.Subset.inter_eq_left h] end Multiset - --- Assert that we define `Finset` without the material on the set lattice. --- Note that we cannot put this in `Data.Finset.Basic` because we proved relevant lemmas there. -assert_not_exists Set.sInter diff --git a/Mathlib/Data/Set/Pointwise/Basic.lean b/Mathlib/Data/Set/Pointwise/Basic.lean index 369ed71edf29c..bc6cb8dcea2ee 100644 --- a/Mathlib/Data/Set/Pointwise/Basic.lean +++ b/Mathlib/Data/Set/Pointwise/Basic.lean @@ -51,6 +51,8 @@ pointwise subtraction -/ +assert_not_exists OrderedAddCommMonoid + library_note "pointwise nat action"/-- Pointwise monoids (`Set`, `Finset`, `Filter`) have derived pointwise actions of the form `SMul α β → SMul α (Set β)`. When `α` is `ℕ` or `ℤ`, this action conflicts with the @@ -72,8 +74,6 @@ namespace Set /-! ### `0`/`1` as sets -/ -assert_not_exists OrderedAddCommMonoid - section One variable [One α] {s : Set α} {a : α} diff --git a/Mathlib/Data/Vector/Defs.lean b/Mathlib/Data/Vector/Defs.lean index 3dbc67d6b93b1..c25f510f74bd1 100644 --- a/Mathlib/Data/Vector/Defs.lean +++ b/Mathlib/Data/Vector/Defs.lean @@ -10,8 +10,8 @@ import Mathlib.Tactic.Common The type `Vector` represents lists with fixed length. -/ -namespace Mathlib assert_not_exists Monoid +namespace Mathlib universe u v w /-- `Vector α n` is the type of lists of length `n` with elements of type `α`. -/ diff --git a/Mathlib/NumberTheory/Wilson.lean b/Mathlib/NumberTheory/Wilson.lean index a5d360e59a890..5832256155ede 100644 --- a/Mathlib/NumberTheory/Wilson.lean +++ b/Mathlib/NumberTheory/Wilson.lean @@ -24,7 +24,7 @@ This could be generalized to similar results about finite abelian groups. * Give `wilsons_lemma` a descriptive name. -/ -assert_not_exists legendre_sym.quadratic_reciprocity +assert_not_exists legendreSym.quadratic_reciprocity open Finset Nat FiniteField ZMod diff --git a/Mathlib/Topology/Bornology/Absorbs.lean b/Mathlib/Topology/Bornology/Absorbs.lean index b4753b1d5ec86..ea21c3717d826 100644 --- a/Mathlib/Topology/Bornology/Absorbs.lean +++ b/Mathlib/Topology/Bornology/Absorbs.lean @@ -33,11 +33,11 @@ They can be added later when someone needs them. absorbs, absorbent -/ +assert_not_exists Real + open Set Bornology Filter open scoped Pointwise -assert_not_exists Real - section Defs variable (M : Type*) {α : Type*} [Bornology M] [SMul M α] From 34f13c7489a482e92c4b4d7bf0e91df69aed4b73 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 13 Aug 2024 07:46:35 +0000 Subject: [PATCH 232/359] chore: use down-stream existing declarations when `assert_not_exists`ing them (#15732) Another PR stemming from the future `assertNotExists` linter. --- Mathlib/Algebra/Field/Defs.lean | 2 +- Mathlib/Data/Finset/Basic.lean | 3 ++- Mathlib/Data/Nat/Cast/Basic.lean | 5 +++-- 3 files changed, 6 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Field/Defs.lean b/Mathlib/Algebra/Field/Defs.lean index f7a7742c83c9d..c29f6e33d0c67 100644 --- a/Mathlib/Algebra/Field/Defs.lean +++ b/Mathlib/Algebra/Field/Defs.lean @@ -47,7 +47,7 @@ field, division ring, skew field, skew-field, skewfield assert_not_exists NeZero -- Check that we have not imported `Mathlib.Tactic.Common` yet. -assert_not_exists Mathlib.Tactic.LibrarySearch.librarySearch +assert_not_exists Mathlib.Tactic.scopedNS assert_not_exists MonoidHom diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 2eff3772f98da..c010b93b78035 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -118,7 +118,8 @@ finite sets, finset -- Assert that we define `Finset` without the material on `List.sublists`. -- Note that we cannot use `List.sublists` itself as that is defined very early. assert_not_exists List.sublistsLen -assert_not_exists Multiset.Powerset +assert_not_exists Multiset.powerset + assert_not_exists CompleteLattice open Multiset Subtype Nat Function diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index 57acfff19ec9f..8e33f0b37f992 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -23,8 +23,9 @@ assert_not_exists OrderedCommGroup assert_not_exists Commute.zero_right assert_not_exists Commute.add_right assert_not_exists abs_eq_max_neg -assert_not_exists natCast_ne -assert_not_exists MulOpposite.natCast +assert_not_exists NeZero.natCast_ne +-- TODO: `MulOpposite.op_natCast` was not intended to be imported +-- assert_not_exists MulOpposite.op_natCast -- Porting note: There are many occasions below where we need `simp [map_zero f]` -- where `simp [map_zero]` should suffice. (Similarly for `map_one`.) From ee3fc238d3735b9285cfe9dd1071444fd0ce9a7a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 13 Aug 2024 08:09:53 +0000 Subject: [PATCH 233/359] chore(Algebra/Group/Aut): Do not import `MonoidWithZero` (#15430) Move what needs it to a new file `Algebra.GroupWithZero.Action.Basic` which will eventually contain more content coming from `Algebra.SMulWithZero`. --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Aut.lean | 10 +- .../Algebra/GroupWithZero/Action/Basic.lean | 175 ++++++++++++++++++ Mathlib/Algebra/Module/Equiv/Basic.lean | 4 +- Mathlib/Algebra/Ring/Action/Group.lean | 2 +- Mathlib/Algebra/Ring/AddAut.lean | 2 +- Mathlib/Algebra/SMulWithZero.lean | 9 +- Mathlib/Data/DFinsupp/Basic.lean | 1 + Mathlib/Data/Set/Pointwise/SMul.lean | 1 + Mathlib/GroupTheory/GroupAction/Group.lean | 84 +-------- Mathlib/RingTheory/Nilpotent/Basic.lean | 1 + 11 files changed, 187 insertions(+), 103 deletions(-) create mode 100644 Mathlib/Algebra/GroupWithZero/Action/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index 05b79411aac95..bd642926af4a3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -290,6 +290,7 @@ import Mathlib.Algebra.Group.WithOne.Basic import Mathlib.Algebra.Group.WithOne.Defs import Mathlib.Algebra.Group.ZeroOne import Mathlib.Algebra.GroupPower.IterateHom +import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.Algebra.GroupWithZero.Action.Defs import Mathlib.Algebra.GroupWithZero.Action.Opposite import Mathlib.Algebra.GroupWithZero.Action.Pi diff --git a/Mathlib/Algebra/Group/Aut.lean b/Mathlib/Algebra/Group/Aut.lean index 9a89372872dd6..cb169b83f1a64 100644 --- a/Mathlib/Algebra/Group/Aut.lean +++ b/Mathlib/Algebra/Group/Aut.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ -import Mathlib.Algebra.GroupWithZero.Action.Defs import Mathlib.GroupTheory.Perm.Basic /-! @@ -27,7 +26,6 @@ MulAut, AddAut -/ assert_not_exists MonoidWithZero -assert_not_exists Ring variable {A : Type*} {M : Type*} {G : Type*} @@ -101,12 +99,10 @@ def toPerm : MulAut M →* Equiv.Perm M where /-- The tautological action by `MulAut M` on `M`. This generalizes `Function.End.applyMulAction`. -/ -instance applyMulDistribMulAction {M} [Monoid M] : MulDistribMulAction (MulAut M) M where +instance applyMulAction {M} [Monoid M] : MulAction (MulAut M) M where smul := (· <| ·) one_smul _ := rfl mul_smul _ _ _ := rfl - smul_one := map_one - smul_mul := map_mul @[simp] protected theorem smul_def {M} [Monoid M] (f : MulAut M) (a : M) : f • a = f a := @@ -208,10 +204,8 @@ def toPerm : AddAut A →* Equiv.Perm A where /-- The tautological action by `AddAut A` on `A`. This generalizes `Function.End.applyMulAction`. -/ -instance applyDistribMulAction {A} [AddMonoid A] : DistribMulAction (AddAut A) A where +instance applyMulAction {A} [AddMonoid A] : MulAction (AddAut A) A where smul := (· <| ·) - smul_zero := map_zero - smul_add := map_add one_smul _ := rfl mul_smul _ _ _ := rfl diff --git a/Mathlib/Algebra/GroupWithZero/Action/Basic.lean b/Mathlib/Algebra/GroupWithZero/Action/Basic.lean new file mode 100644 index 0000000000000..616b74daf1af9 --- /dev/null +++ b/Mathlib/Algebra/GroupWithZero/Action/Basic.lean @@ -0,0 +1,175 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Yury Kudryashov +-/ +import Mathlib.Algebra.Group.Action.Basic +import Mathlib.Algebra.Group.Action.Prod +import Mathlib.Algebra.Group.Aut +import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Prod +import Mathlib.Algebra.SMulWithZero + +/-! +# Definitions of group actions + +This file defines a hierarchy of group action type-classes on top of the previously defined +notation classes `SMul` and its additive version `VAdd`: + +* `MulAction M α` and its additive version `AddAction G P` are typeclasses used for + actions of multiplicative and additive monoids and groups; they extend notation classes + `SMul` and `VAdd` that are defined in `Algebra.Group.Defs`; +* `DistribMulAction M A` is a typeclass for an action of a multiplicative monoid on + an additive monoid such that `a • (b + c) = a • b + a • c` and `a • 0 = 0`. + +The hierarchy is extended further by `Module`, defined elsewhere. + +Also provided are typeclasses for faithful and transitive actions, and typeclasses regarding the +interaction of different group actions, + +* `SMulCommClass M N α` and its additive version `VAddCommClass M N α`; +* `IsScalarTower M N α` and its additive version `VAddAssocClass M N α`; +* `IsCentralScalar M α` and its additive version `IsCentralVAdd M N α`. + +## Notation + +- `a • b` is used as notation for `SMul.smul a b`. +- `a +ᵥ b` is used as notation for `VAdd.vadd a b`. + +## Implementation details + +This file should avoid depending on other parts of `GroupTheory`, to avoid import cycles. +More sophisticated lemmas belong in `GroupTheory.GroupAction`. + +## Tags + +group action +-/ + +-- TODO: +-- assert_not_exists Ring + +open Function + +variable {G G₀ A M N M₀ N₀ R α : Type*} + +section GroupWithZero +variable [GroupWithZero G₀] [MulAction G₀ α] {a : G₀} + +protected lemma MulAction.bijective₀ (ha : a ≠ 0) : Bijective (a • · : α → α) := + MulAction.bijective <| Units.mk0 a ha + +protected lemma MulAction.injective₀ (ha : a ≠ 0) : Injective (a • · : α → α) := + (MulAction.bijective₀ ha).injective + +protected lemma MulAction.surjective₀ (ha : a ≠ 0) : Surjective (a • · : α → α) := + (MulAction.bijective₀ ha).surjective + +end GroupWithZero + +section DistribMulAction +variable [Group G] [Monoid M] [AddMonoid A] [DistribMulAction M A] +variable (A) + +/-- Each element of the group defines an additive monoid isomorphism. + +This is a stronger version of `MulAction.toPerm`. -/ +@[simps (config := { simpRhs := true })] +def DistribMulAction.toAddEquiv [DistribMulAction G A] (x : G) : A ≃+ A where + __ := toAddMonoidHom A x + __ := MulAction.toPermHom G A x + +variable (G) + +/-- Each element of the group defines an additive monoid isomorphism. + +This is a stronger version of `MulAction.toPermHom`. -/ +@[simps] +def DistribMulAction.toAddAut [DistribMulAction G A] : G →* AddAut A where + toFun := toAddEquiv _ + map_one' := AddEquiv.ext (one_smul _) + map_mul' _ _ := AddEquiv.ext (mul_smul _ _) + +end DistribMulAction + +/-- Scalar multiplication as a monoid homomorphism with zero. -/ +@[simps] +def smulMonoidWithZeroHom [MonoidWithZero M₀] [MulZeroOneClass N₀] [MulActionWithZero M₀ N₀] + [IsScalarTower M₀ N₀ N₀] [SMulCommClass M₀ N₀ N₀] : M₀ × N₀ →*₀ N₀ := + { smulMonoidHom with map_zero' := smul_zero _ } + +section MulDistribMulAction +variable [Group G] [Monoid M] [MulDistribMulAction G M] +variable (M) + +/-- Each element of the group defines a multiplicative monoid isomorphism. + +This is a stronger version of `MulAction.toPerm`. -/ +@[simps (config := { simpRhs := true })] +def MulDistribMulAction.toMulEquiv (x : G) : M ≃* M := + { MulDistribMulAction.toMonoidHom M x, MulAction.toPermHom G M x with } + +variable (G) + +/-- Each element of the group defines a multiplicative monoid isomorphism. + +This is a stronger version of `MulAction.toPermHom`. -/ +@[simps] +def MulDistribMulAction.toMulAut : G →* MulAut M where + toFun := MulDistribMulAction.toMulEquiv M + map_one' := MulEquiv.ext (one_smul _) + map_mul' _ _ := MulEquiv.ext (mul_smul _ _) + +end MulDistribMulAction + + +namespace MulAut + +/-- The tautological action by `MulAut M` on `M`. + +This generalizes `Function.End.applyMulAction`. -/ +instance applyMulDistribMulAction [Monoid M] : MulDistribMulAction (MulAut M) M where + smul := (· <| ·) + one_smul _ := rfl + mul_smul _ _ _ := rfl + smul_one := map_one + smul_mul := map_mul + +end MulAut + +namespace AddAut + +/-- The tautological action by `AddAut A` on `A`. + +This generalizes `Function.End.applyMulAction`. -/ +instance applyDistribMulAction [AddMonoid A] : DistribMulAction (AddAut A) A where + smul := (· <| ·) + smul_zero := map_zero + smul_add := map_add + one_smul _ := rfl + mul_smul _ _ _ := rfl + +end AddAut + +section Arrow +variable [Group G] [MulAction G A] [Monoid M] + +attribute [local instance] arrowAction + +/-- When `M` is a monoid, `ArrowAction` is additionally a `MulDistribMulAction`. -/ +def arrowMulDistribMulAction : MulDistribMulAction G (A → M) where + smul_one _ := rfl + smul_mul _ _ _ := rfl + +attribute [local instance] arrowMulDistribMulAction + +/-- Given groups `G H` with `G` acting on `A`, `G` acts by +multiplicative automorphisms on `A → H`. -/ +@[simps!] def mulAutArrow : G →* MulAut (A → M) := MulDistribMulAction.toMulAut _ _ + +end Arrow + +lemma IsUnit.smul_sub_iff_sub_inv_smul [Group G] [Monoid R] [AddGroup R] [DistribMulAction G R] + [IsScalarTower G R R] [SMulCommClass G R R] (r : G) (a : R) : + IsUnit (r • (1 : R) - a) ↔ IsUnit (1 - r⁻¹ • a) := by + rw [← isUnit_smul_iff r (1 - r⁻¹ • a), smul_sub, smul_inv_smul] diff --git a/Mathlib/Algebra/Module/Equiv/Basic.lean b/Mathlib/Algebra/Module/Equiv/Basic.lean index fcb550321d2d3..1197686e29602 100644 --- a/Mathlib/Algebra/Module/Equiv/Basic.lean +++ b/Mathlib/Algebra/Module/Equiv/Basic.lean @@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Baanen, Frédéric Dupuis, Heather Macbeth -/ -import Mathlib.Algebra.Module.Equiv.Defs import Mathlib.Algebra.Field.Defs +import Mathlib.Algebra.GroupWithZero.Action.Basic +import Mathlib.Algebra.Module.Equiv.Defs import Mathlib.Algebra.Module.Hom import Mathlib.Algebra.Module.LinearMap.End import Mathlib.Algebra.Module.Pi -import Mathlib.GroupTheory.GroupAction.Group /-! # Further results on (semi)linear equivalences. diff --git a/Mathlib/Algebra/Ring/Action/Group.lean b/Mathlib/Algebra/Ring/Action/Group.lean index 1e1876407e31c..4317fba7f1f4a 100644 --- a/Mathlib/Algebra/Ring/Action/Group.lean +++ b/Mathlib/Algebra/Ring/Action/Group.lean @@ -3,8 +3,8 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ +import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.Algebra.Ring.Action.Basic -import Mathlib.GroupTheory.GroupAction.Group import Mathlib.Algebra.Ring.Equiv /-! diff --git a/Mathlib/Algebra/Ring/AddAut.lean b/Mathlib/Algebra/Ring/AddAut.lean index dfddf66ce51dd..f593f12d75bda 100644 --- a/Mathlib/Algebra/Ring/AddAut.lean +++ b/Mathlib/Algebra/Ring/AddAut.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.GroupTheory.GroupAction.Group +import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.Algebra.Module.Defs /-! diff --git a/Mathlib/Algebra/SMulWithZero.lean b/Mathlib/Algebra/SMulWithZero.lean index bba8bf4abb805..3ef4bb96b5b9f 100644 --- a/Mathlib/Algebra/SMulWithZero.lean +++ b/Mathlib/Algebra/SMulWithZero.lean @@ -4,10 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ import Mathlib.Algebra.Group.Action.Opposite -import Mathlib.Algebra.Group.Action.Prod import Mathlib.Algebra.GroupWithZero.Action.Defs +import Mathlib.Algebra.GroupWithZero.Hom import Mathlib.Algebra.GroupWithZero.Opposite -import Mathlib.Algebra.GroupWithZero.Prod import Mathlib.Algebra.Ring.Defs /-! @@ -201,12 +200,6 @@ theorem smul_inv₀ [SMulCommClass α β β] [IsScalarTower α β β] (c : α) ( end GroupWithZero -/-- Scalar multiplication as a monoid homomorphism with zero. -/ -@[simps] -def smulMonoidWithZeroHom {α β : Type*} [MonoidWithZero α] [MulZeroOneClass β] - [MulActionWithZero α β] [IsScalarTower α β β] [SMulCommClass α β β] : α × β →*₀ β := - { smulMonoidHom with map_zero' := smul_zero _ } - -- This instance seems a bit incongruous in this file, but `#find_home!` told me to put it here. instance NonUnitalNonAssocSemiring.toDistribSMul [NonUnitalNonAssocSemiring R] : DistribSMul R R where diff --git a/Mathlib/Data/DFinsupp/Basic.lean b/Mathlib/Data/DFinsupp/Basic.lean index 8b5dbe5f93246..2a240cc29359d 100644 --- a/Mathlib/Data/DFinsupp/Basic.lean +++ b/Mathlib/Data/DFinsupp/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Kenny Lau -/ import Mathlib.Algebra.BigOperators.GroupWithZero.Finset +import Mathlib.Algebra.Group.Action.Prod import Mathlib.Algebra.Group.Submonoid.Membership import Mathlib.Algebra.GroupWithZero.Action.Pi import Mathlib.Algebra.Module.LinearMap.Defs diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 37f8bfda969d9..fd2bf0cf217e8 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Floris van Doorn -/ import Mathlib.Algebra.Group.Pi.Basic +import Mathlib.Algebra.GroupWithZero.Action.Basic import Mathlib.Algebra.Module.Defs import Mathlib.Data.Set.Pairwise.Basic import Mathlib.Data.Set.Pointwise.Basic diff --git a/Mathlib/GroupTheory/GroupAction/Group.lean b/Mathlib/GroupTheory/GroupAction/Group.lean index 600fba33cc68c..ae6ba4516d404 100644 --- a/Mathlib/GroupTheory/GroupAction/Group.lean +++ b/Mathlib/GroupTheory/GroupAction/Group.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import Mathlib.Algebra.Group.Action.Basic -import Mathlib.Algebra.Group.Aut import Mathlib.Algebra.Group.Invertible.Basic import Mathlib.Algebra.GroupWithZero.Action.Units import Mathlib.Algebra.GroupWithZero.Units.Basic +import Mathlib.Algebra.SMulWithZero /-! # Group actions applied to various types of group @@ -62,15 +62,6 @@ theorem Commute.smul_left_iff₀ [Mul β] [SMulCommClass α β β] [IsScalarTowe left_inv := inv_smul_smul₀ ha right_inv := smul_inv_smul₀ ha -protected theorem MulAction.bijective₀ (ha : a ≠ 0) : Function.Bijective (a • · : β → β) := - MulAction.bijective <| Units.mk0 a ha - -protected theorem MulAction.injective₀ (ha : a ≠ 0) : Function.Injective (a • · : β → β) := - (MulAction.bijective₀ ha).injective - -protected theorem MulAction.surjective₀ (ha : a ≠ 0) : Function.Surjective (a • · : β → β) := - (MulAction.bijective₀ ha).surjective - end Gwz end MulAction @@ -82,24 +73,8 @@ section Group variable [Group α] [AddMonoid β] [DistribMulAction α β] variable (β) -/-- Each element of the group defines an additive monoid isomorphism. - -This is a stronger version of `MulAction.toPerm`. -/ -@[simps (config := { simpRhs := true })] -def DistribMulAction.toAddEquiv (x : α) : β ≃+ β := - { DistribMulAction.toAddMonoidHom β x, MulAction.toPermHom α β x with } - variable (α) -/-- Each element of the group defines an additive monoid isomorphism. - -This is a stronger version of `MulAction.toPermHom`. -/ -@[simps] -def DistribMulAction.toAddAut : α →* AddAut β where - toFun := DistribMulAction.toAddEquiv β - map_one' := AddEquiv.ext (one_smul _) - map_mul' _ _ := AddEquiv.ext (mul_smul _ _) - /-- Each non-zero element of a `GroupWithZero` defines an additive monoid isomorphism of an `AddMonoid` on which it acts distributively. This is a stronger version of `DistribMulAction.toAddMonoidHom`. -/ @@ -122,53 +97,6 @@ end Group end DistribMulAction -section MulDistribMulAction - -variable [Group α] [Monoid β] [MulDistribMulAction α β] -variable (β) - -/-- Each element of the group defines a multiplicative monoid isomorphism. - -This is a stronger version of `MulAction.toPerm`. -/ -@[simps (config := { simpRhs := true })] -def MulDistribMulAction.toMulEquiv (x : α) : β ≃* β := - { MulDistribMulAction.toMonoidHom β x, MulAction.toPermHom α β x with } - -variable (α) - -/-- Each element of the group defines a multiplicative monoid isomorphism. - -This is a stronger version of `MulAction.toPermHom`. -/ -@[simps] -def MulDistribMulAction.toMulAut : α →* MulAut β where - toFun := MulDistribMulAction.toMulEquiv β - map_one' := MulEquiv.ext (one_smul _) - map_mul' _ _ := MulEquiv.ext (mul_smul _ _) - -variable {α β} - -end MulDistribMulAction - -section Arrow - -attribute [local instance] arrowAction - -/-- When `B` is a monoid, `ArrowAction` is additionally a `MulDistribMulAction`. -/ -def arrowMulDistribMulAction {G A B : Type*} [Group G] [MulAction G A] [Monoid B] : - MulDistribMulAction G (A → B) where - smul_one _ := rfl - smul_mul _ _ _ := rfl - -attribute [local instance] arrowMulDistribMulAction - -/-- Given groups `G H` with `G` acting on `A`, `G` acts by - multiplicative automorphisms on `A → H`. -/ -@[simps!] -def mulAutArrow {G A H} [Group G] [MulAction G A] [Monoid H] : G →* MulAut (A → H) := - MulDistribMulAction.toMulAut _ _ - -end Arrow - namespace IsUnit section DistribMulAction @@ -182,13 +110,3 @@ theorem smul_eq_zero {u : α} (hu : IsUnit u) {x : β} : u • x = 0 ↔ x = 0 : end DistribMulAction end IsUnit - -section SMul - -variable [Group α] [Monoid β] - -theorem IsUnit.smul_sub_iff_sub_inv_smul [AddGroup β] [DistribMulAction α β] [IsScalarTower α β β] - [SMulCommClass α β β] (r : α) (a : β) : IsUnit (r • (1 : β) - a) ↔ IsUnit (1 - r⁻¹ • a) := by - rw [← isUnit_smul_iff r (1 - r⁻¹ • a), smul_sub, smul_inv_smul] - -end SMul diff --git a/Mathlib/RingTheory/Nilpotent/Basic.lean b/Mathlib/RingTheory/Nilpotent/Basic.lean index 3480196f77b5d..0080e768a4300 100644 --- a/Mathlib/RingTheory/Nilpotent/Basic.lean +++ b/Mathlib/RingTheory/Nilpotent/Basic.lean @@ -5,6 +5,7 @@ Authors: Oliver Nash -/ import Mathlib.Algebra.Associated.Basic import Mathlib.Algebra.GeomSum +import Mathlib.Algebra.Group.Action.Prod import Mathlib.Algebra.GroupWithZero.NonZeroDivisors import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.SMulWithZero From b11309201674465698a9716724a2a99a80f5197e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 13 Aug 2024 09:00:40 +0000 Subject: [PATCH 234/359] feat: Scaling of affine bases (#14863) Given a group `G` acting on the ring `k`, provide the `G`-action over the `k`-affine bases of a `k`-module `V`. Co-authored-by: Eric Wieser --- Mathlib/LinearAlgebra/AffineSpace/Basis.lean | 54 ++++++++++- .../AffineSpace/Combination.lean | 11 +++ .../AffineSpace/Independent.lean | 8 ++ .../LinearAlgebra/AffineSpace/Pointwise.lean | 97 ++++++++++++++++--- 4 files changed, 150 insertions(+), 20 deletions(-) diff --git a/Mathlib/LinearAlgebra/AffineSpace/Basis.lean b/Mathlib/LinearAlgebra/AffineSpace/Basis.lean index 7d6f380b7c9dc..878ac8fe1f82c 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Basis.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Basis.lean @@ -38,10 +38,8 @@ barycentric coordinate of `q : P` is `1 - fᵢ (q -ᵥ p i)`. -/ - -open Affine - -open Set +open Affine Set +open scoped Pointwise universe u₁ u₂ u₃ u₄ @@ -52,7 +50,7 @@ structure AffineBasis (ι : Type u₁) (k : Type u₂) {V : Type u₃} (P : Type protected ind' : AffineIndependent k toFun protected tot' : affineSpan k (range toFun) = ⊤ -variable {ι ι' k V P : Type*} [AddCommGroup V] [AffineSpace V P] +variable {ι ι' G G' k V P : Type*} [AddCommGroup V] [AffineSpace V P] namespace AffineBasis @@ -275,6 +273,52 @@ instance instAddAction : AddAction V (AffineBasis ι k P) := congr! 1 rw [vadd_vsub_assoc, neg_add_eq_sub, vsub_vadd_eq_vsub_sub] +section SMul +variable [Group G] [Group G'] +variable [DistribMulAction G V] [DistribMulAction G' V] +variable [SMulCommClass G k V] [SMulCommClass G' k V] + +/-- In an affine space that is also a vector space, an `AffineBasis` can be scaled. + +TODO: generalize to include `SMul (P ≃ᵃ[k] P) (AffineBasis ι k P)`, which acts on `P` with a `VAdd` +version of a `DistribMulAction`. -/ +instance instSMul : SMul G (AffineBasis ι k V) where + smul a b := + { toFun := a • ⇑b, + ind' := b.ind'.smul, + tot' := by + rw [Pi.smul_def, ← smul_set_range, ← AffineSubspace.smul_span, b.tot, + AffineSubspace.smul_top (Group.isUnit a)] } + +@[simp, norm_cast] lemma coe_smul (a : G) (b : AffineBasis ι k V) : ⇑(a • b) = a • ⇑b := rfl + +/-- TODO: generalize to include `SMul (P ≃ᵃ[k] P) (AffineBasis ι k P)`, which acts on `P` with a +`VAdd` version of a `DistribMulAction`. -/ +instance [SMulCommClass G G' V] : SMulCommClass G G' (AffineBasis ι k V) where + smul_comm _g _g' _b := DFunLike.ext _ _ fun _ => smul_comm _ _ _ + +/-- TODO: generalize to include `SMul (P ≃ᵃ[k] P) (AffineBasis ι k P)`, which acts on `P` with a +`VAdd` version of a `DistribMulAction`. -/ +instance [SMul G G'] [IsScalarTower G G' V] : IsScalarTower G G' (AffineBasis ι k V) where + smul_assoc _g _g' _b := DFunLike.ext _ _ fun _ => smul_assoc _ _ _ + +@[simp] lemma basisOf_smul (a : G) (b : AffineBasis ι k V) (i : ι) : + (a • b).basisOf i = a • b.basisOf i := by ext j; simp [smul_sub] + +@[simp] lemma reindex_smul (a : G) (b : AffineBasis ι k V) (e : ι ≃ ι') : + (a • b).reindex e = a • b.reindex e := + rfl + +@[simp] lemma coord_smul (a : G) (b : AffineBasis ι k V) (i : ι) : + (a • b).coord i = (b.coord i).comp (DistribMulAction.toLinearEquiv _ _ a).symm.toAffineMap := by + ext v; simp [coord] + +/-- TODO: generalize to include `SMul (P ≃ᵃ[k] P) (AffineBasis ι k P)`, which acts on `P` with a +`VAdd` version of a `DistribMulAction`. -/ +instance instMulAction : MulAction G (AffineBasis ι k V) := + DFunLike.coe_injective.mulAction _ coe_smul + +end SMul end Ring section DivisionRing diff --git a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean index 621360ad2235f..e1d31428561b4 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean @@ -79,6 +79,11 @@ lemma weightedVSubOfPoint_vadd (s : Finset ι) (w : ι → k) (p : ι → P) (b s.weightedVSubOfPoint (v +ᵥ p) b w = s.weightedVSubOfPoint p (-v +ᵥ b) w := by simp [vadd_vsub_assoc, vsub_vadd_eq_vsub_sub, add_comm] +lemma weightedVSubOfPoint_smul {G : Type*} [Group G] [DistribMulAction G V] [SMulCommClass G k V] + (s : Finset ι) (w : ι → k) (p : ι → V) (b : V) (a : G) : + s.weightedVSubOfPoint (a • p) b w = a • s.weightedVSubOfPoint p (a⁻¹ • b) w := by + simp [smul_sum, smul_sub, smul_comm a (w _)] + /-- `weightedVSubOfPoint` gives equal results for two families of weights and two families of points that are equal on `s`. -/ theorem weightedVSubOfPoint_congr {w₁ w₂ : ι → k} (hw : ∀ i ∈ s, w₁ i = w₂ i) {p₁ p₂ : ι → P} @@ -261,6 +266,12 @@ lemma weightedVSub_vadd {s : Finset ι} {w : ι → k} (h : ∑ i ∈ s, w i = 0 rw [weightedVSub, weightedVSubOfPoint_vadd, weightedVSub_eq_weightedVSubOfPoint_of_sum_eq_zero _ _ _ h] +lemma weightedVSub_smul {G : Type*} [Group G] [DistribMulAction G V] [SMulCommClass G k V] + {s : Finset ι} {w : ι → k} (h : ∑ i ∈ s, w i = 0) (p : ι → V) (a : G) : + s.weightedVSub (a • p) w = a • s.weightedVSub p w := by + rw [weightedVSub, weightedVSubOfPoint_smul, + weightedVSub_eq_weightedVSubOfPoint_of_sum_eq_zero _ _ _ h] + /-- `weightedVSub` gives equal results for two families of weights and two families of points that are equal on `s`. -/ theorem weightedVSub_congr {w₁ w₂ : ι → k} (hw : ∀ i ∈ s, w₁ i = w₂ i) {p₁ p₂ : ι → P} diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index 826df0fc1aed6..0adb6c8ef435b 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -81,6 +81,14 @@ theorem affineIndependent_iff_of_fintype [Fintype ι] (p : ι → P) : protected alias ⟨AffineIndependent.of_vadd, AffineIndependent.vadd⟩ := affineIndependent_vadd +@[simp] lemma affineIndependent_smul {G : Type*} [Group G] [DistribMulAction G V] + [SMulCommClass G k V] {p : ι → V} {a : G} : + AffineIndependent k (a • p) ↔ AffineIndependent k p := by + simp (config := { contextual := true }) [AffineIndependent, weightedVSub_smul, + ← smul_comm (α := V) a, ← smul_sum, smul_eq_zero_iff_eq] + +protected alias ⟨AffineIndependent.of_smul, AffineIndependent.smul⟩ := affineIndependent_smul + /-- A family is affinely independent if and only if the differences from a base point in that family are linearly independent. -/ theorem affineIndependent_iff_linearIndependent_vsub (p : ι → P) (i1 : ι) : diff --git a/Mathlib/LinearAlgebra/AffineSpace/Pointwise.lean b/Mathlib/LinearAlgebra/AffineSpace/Pointwise.lean index 239e40097e013..d7a79d2cfd7b2 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Pointwise.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Pointwise.lean @@ -17,36 +17,38 @@ open Affine Pointwise open Set -variable {k : Type*} [Ring k] -variable {V P V₁ P₁ V₂ P₂ : Type*} +variable {M k V P V₁ P₁ V₂ P₂ : Type*} + +namespace AffineSubspace +section Ring +variable [Ring k] variable [AddCommGroup V] [Module k V] [AffineSpace V P] variable [AddCommGroup V₁] [Module k V₁] [AddTorsor V₁ P₁] variable [AddCommGroup V₂] [Module k V₂] [AddTorsor V₂ P₂] -namespace AffineSubspace +/-- The additive action on an affine subspace corresponding to applying the action to every element. + +This is available as an instance in the `Pointwise` locale. -/ +protected def pointwiseVAdd : VAdd V (AffineSubspace k P) where + vadd x s := s.map (AffineEquiv.constVAdd k P x) + +scoped[Pointwise] attribute [instance] AffineSubspace.pointwiseVAdd + +@[simp, norm_cast] lemma coe_pointwise_vadd (v : V) (s : AffineSubspace k P) : + ((v +ᵥ s : AffineSubspace k P) : Set P) = v +ᵥ (s : Set P) := rfl /-- The additive action on an affine subspace corresponding to applying the action to every element. This is available as an instance in the `Pointwise` locale. -/ -protected def pointwiseAddAction : AddAction V (AffineSubspace k P) where - vadd x S := S.map (AffineEquiv.constVAdd k P x) - zero_vadd p := ((congr_arg fun f => p.map f) <| AffineMap.ext <| zero_vadd _).trans p.map_id - add_vadd _ _ p := - ((congr_arg fun f => p.map f) <| AffineMap.ext <| add_vadd _ _).trans (p.map_map _ _).symm +protected def pointwiseAddAction : AddAction V (AffineSubspace k P) := + SetLike.coe_injective.addAction _ coe_pointwise_vadd scoped[Pointwise] attribute [instance] AffineSubspace.pointwiseAddAction -open Pointwise - theorem pointwise_vadd_eq_map (v : V) (s : AffineSubspace k P) : v +ᵥ s = s.map (AffineEquiv.constVAdd k P v) := rfl -@[simp] -theorem coe_pointwise_vadd (v : V) (s : AffineSubspace k P) : - ((v +ᵥ s : AffineSubspace k P) : Set P) = v +ᵥ (s : Set P) := - rfl - theorem vadd_mem_pointwise_vadd_iff {v : V} {s : AffineSubspace k P} {p : P} : v +ᵥ p ∈ v +ᵥ s ↔ p ∈ s := vadd_mem_vadd_set_iff @@ -72,4 +74,69 @@ theorem map_pointwise_vadd (f : P₁ →ᵃ[k] P₂) (v : V₁) (s : AffineSubsp ext exact f.map_vadd _ _ +section SMul +variable [Monoid M] [DistribMulAction M V] [SMulCommClass M k V] {a : M} {s : AffineSubspace k V} + {p : V} + +/-- The multiplicative action on an affine subspace corresponding to applying the action to every +element. + +This is available as an instance in the `Pointwise` locale. + +TODO: generalize to include `SMul (P ≃ᵃ[k] P) (AffineSubspace k P)`, which acts on `P` with a +`VAdd` version of a `DistribMulAction`. -/ +protected def pointwiseSMul : SMul M (AffineSubspace k V) where + smul a s := s.map (DistribMulAction.toLinearMap k _ a).toAffineMap + +scoped[Pointwise] attribute [instance] AffineSubspace.pointwiseSMul + +@[simp, norm_cast] +lemma coe_smul (a : M) (s : AffineSubspace k V) : ↑(a • s) = a • (s : Set V) := rfl + +/-- The multiplicative action on an affine subspace corresponding to applying the action to every +element. + +This is available as an instance in the `Pointwise` locale. + +TODO: generalize to include `SMul (P ≃ᵃ[k] P) (AffineSubspace k P)`, which acts on `P` with a +`VAdd` version of a `DistribMulAction`. -/ +protected def mulAction : MulAction M (AffineSubspace k V) := + SetLike.coe_injective.mulAction _ coe_smul + +scoped[Pointwise] attribute [instance] AffineSubspace.mulAction + +lemma smul_eq_map (a : M) (s : AffineSubspace k V) : + a • s = s.map (DistribMulAction.toLinearMap k _ a).toAffineMap := rfl + +lemma smul_mem_smul_iff {G : Type*} [Group G] [DistribMulAction G V] [SMulCommClass G k V] {a : G} : + a • p ∈ a • s ↔ p ∈ s := smul_mem_smul_set_iff + +lemma smul_mem_smul_iff_of_isUnit (ha : IsUnit a) : a • p ∈ a • s ↔ p ∈ s := + smul_mem_smul_iff (a := ha.unit) + +lemma smul_mem_smul_iff₀ {G₀ : Type*} [GroupWithZero G₀] [DistribMulAction G₀ V] + [SMulCommClass G₀ k V] {a : G₀} (ha : a ≠ 0) : a • p ∈ a • s ↔ p ∈ s := + smul_mem_smul_iff_of_isUnit ha.isUnit + +@[simp] lemma smul_bot (a : M) : a • (⊥ : AffineSubspace k V) = ⊥ := by + ext; simp [smul_eq_map, map_bot] + +@[simp] lemma smul_top (ha : IsUnit a) : a • (⊤ : AffineSubspace k V) = ⊤ := by + ext x; simpa [smul_eq_map, map_top] using ⟨ha.unit⁻¹ • x, smul_inv_smul ha.unit _⟩ + +lemma smul_span (a : M) (s : Set V) : a • affineSpan k s = affineSpan k (a • s) := map_span _ s + +end SMul +end Ring + +section Field +variable [Field k] [AddCommGroup V] [Module k V] {a : k} + +@[simp] +lemma direction_smul (ha : a ≠ 0) (s : AffineSubspace k V) : (a • s).direction = s.direction := by + rw [smul_eq_map, map_direction] + change Submodule.map (a • LinearMap.id) _ = _ + simp [Submodule.map_smul, ha] + +end Field end AffineSubspace From 4c9b452c851ab595abb572428fa540c565dc7584 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 13 Aug 2024 09:26:51 +0000 Subject: [PATCH 235/359] refactor: Make `circle` a type `Circle` (#15116) See https://github.com/leanprover-community/mathlib4/wiki/Tradeoffs-of-concrete-types-defined-as-subobjects for context --- Mathlib/Analysis/Complex/Circle.lean | 161 ++++++++++------- Mathlib/Analysis/Complex/Isometry.lean | 30 ++-- Mathlib/Analysis/Complex/UnitDisc/Basic.lean | 16 +- Mathlib/Analysis/Fourier/AddCircle.lean | 10 +- .../Analysis/Fourier/FourierTransform.lean | 33 ++-- .../Fourier/FourierTransformDeriv.lean | 2 +- .../Analysis/Fourier/PoissonSummation.lean | 2 +- .../Fourier/RiemannLebesgueLemma.lean | 4 +- Mathlib/Analysis/Normed/Field/UnitBall.lean | 1 + .../SpecialFunctions/Complex/Circle.lean | 169 +++++++++--------- .../Complex/CircleAddChar.lean | 20 ++- .../Euclidean/Angle/Oriented/Rotation.lean | 28 +-- .../Geometry/Manifold/Instances/Sphere.lean | 24 +-- Mathlib/Topology/Algebra/PontryaginDual.lean | 40 ++--- test/TCSynth.lean | 2 +- 15 files changed, 293 insertions(+), 249 deletions(-) diff --git a/Mathlib/Analysis/Complex/Circle.lean b/Mathlib/Analysis/Complex/Circle.lean index 0a4bff315a88a..aab56f8eec106 100644 --- a/Mathlib/Analysis/Complex/Circle.lean +++ b/Mathlib/Analysis/Complex/Circle.lean @@ -38,105 +38,142 @@ open Complex Metric open ComplexConjugate -/-- The unit circle in `ℂ`, here given the structure of a submonoid of `ℂ`. -/ +/-- 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")] def circle : Submonoid ℂ := Submonoid.unitSphere ℂ -@[simp] +set_option linter.deprecated false in +@[deprecated (since := "2024-07-24")] theorem mem_circle_iff_abs {z : ℂ} : z ∈ circle ↔ abs z = 1 := mem_sphere_zero_iff_norm -theorem circle_def : ↑circle = { z : ℂ | abs z = 1 } := - Set.ext fun _ => mem_circle_iff_abs - -@[simp] -theorem abs_coe_circle (z : circle) : abs z = 1 := - mem_circle_iff_abs.mp z.2 +set_option linter.deprecated false in +@[deprecated (since := "2024-07-24")] +theorem mem_circle_iff_normSq {z : ℂ} : z ∈ circle ↔ normSq z = 1 := by + simp [Complex.abs, mem_circle_iff_abs] -theorem mem_circle_iff_normSq {z : ℂ} : z ∈ circle ↔ normSq z = 1 := by simp [Complex.abs] +/-- The unit circle in `ℂ`. -/ +def Circle : Type := Submonoid.unitSphere ℂ +deriving TopologicalSpace -@[simp] -theorem normSq_eq_of_mem_circle (z : circle) : normSq z = 1 := by simp [normSq_eq_abs] +namespace Circle -theorem ne_zero_of_mem_circle (z : circle) : (z : ℂ) ≠ 0 := - ne_zero_of_mem_unit_sphere z +instance instCoeOut : CoeOut Circle ℂ := subtypeCoe -instance commGroup : CommGroup circle := - Metric.sphere.commGroup +instance instCommGroup : CommGroup Circle := Metric.sphere.commGroup +instance instMetricSpace : MetricSpace Circle := Subtype.metricSpace -@[simp] -theorem coe_inv_circle (z : circle) : ↑z⁻¹ = (z : ℂ)⁻¹ := - rfl +@[simp] lemma abs_coe (z : Circle) : abs z = 1 := mem_sphere_zero_iff_norm.1 z.2 +@[simp] lemma normSq_coe (z : Circle) : normSq z = 1 := by simp [normSq_eq_abs] +@[simp] lemma coe_ne_zero (z : Circle) : (z : ℂ) ≠ 0 := ne_zero_of_mem_unit_sphere z +@[simp, norm_cast] lemma coe_one : ↑(1 : Circle) = (1 : ℂ) := rfl +@[simp, norm_cast] lemma coe_mul (z w : Circle) : ↑(z * w) = (z : ℂ) * w := rfl +@[simp, norm_cast] lemma coe_inv (z : Circle) : ↑z⁻¹ = (z : ℂ)⁻¹ := rfl +lemma coe_inv_eq_conj (z : Circle) : ↑z⁻¹ = conj (z : ℂ) := by + rw [coe_inv, inv_def, normSq_coe, inv_one, ofReal_one, mul_one] -theorem coe_inv_circle_eq_conj (z : circle) : ↑z⁻¹ = conj (z : ℂ) := by - rw [coe_inv_circle, inv_def, normSq_eq_of_mem_circle, inv_one, ofReal_one, mul_one] +@[simp, norm_cast] lemma coe_div (z w : Circle) : ↑(z / w) = (z : ℂ) / w := rfl -@[simp] -theorem coe_div_circle (z w : circle) : ↑(z / w) = (z : ℂ) / w := - circle.subtype.map_div z w +/-- The coercion `Circle → ℂ` as a monoid homomorphism. -/ +@[simps] +def coeHom : Circle →* ℂ where + toFun := (↑) + map_one' := coe_one + map_mul' := coe_mul + +@[deprecated (since := "2024-07-24")] alias _root_.abs_coe_circle := abs_coe +@[deprecated (since := "2024-07-24")] alias _root_.normSq_eq_of_mem_circle := normSq_coe +@[deprecated (since := "2024-07-24")] alias _root_.ne_zero_of_mem_circle := coe_ne_zero +@[deprecated (since := "2024-07-24")] alias _root_.coe_inv_circle := coe_inv +@[deprecated (since := "2024-07-24")] alias _root_.coe_inv_circle_eq_conj := coe_inv_eq_conj +@[deprecated (since := "2024-07-24")] alias _root_.coe_div_circle := coe_div /-- The elements of the circle embed into the units. -/ -def circle.toUnits : circle →* Units ℂ := - unitSphereToUnits ℂ - --- written manually because `@[simps]` was slow and generated the wrong lemma -@[simp] -theorem circle.toUnits_apply (z : circle) : - circle.toUnits z = Units.mk0 ↑z (ne_zero_of_mem_circle z) := - rfl - -instance : CompactSpace circle := - Metric.sphere.compactSpace _ _ +def toUnits : Circle →* Units ℂ := unitSphereToUnits ℂ -instance : TopologicalGroup circle := - Metric.sphere.topologicalGroup +-- written manually because `@[simps]` generated the wrong lemma +@[simp] lemma toUnits_apply (z : Circle) : toUnits z = Units.mk0 ↑z z.coe_ne_zero := rfl -instance : UniformGroup circle := by - convert topologicalGroup_is_uniform_of_compactSpace circle +instance : CompactSpace Circle := Metric.sphere.compactSpace _ _ +instance : TopologicalGroup Circle := Metric.sphere.topologicalGroup +instance instUniformSpace : UniformSpace Circle := instUniformSpaceSubtype +instance : UniformGroup Circle := by + convert topologicalGroup_is_uniform_of_compactSpace Circle exact unique_uniformity_of_compact rfl rfl /-- If `z` is a nonzero complex number, then `conj z / z` belongs to the unit circle. -/ @[simps] -def circle.ofConjDivSelf (z : ℂ) (hz : z ≠ 0) : circle := - ⟨conj z / z, - mem_circle_iff_abs.2 <| by rw [map_div₀, abs_conj, div_self]; exact Complex.abs.ne_zero hz⟩ +def ofConjDivSelf (z : ℂ) (hz : z ≠ 0) : Circle where + val := conj z / z + property := mem_sphere_zero_iff_norm.2 <| by + rw [norm_div, RCLike.norm_conj, div_self]; exact Complex.abs.ne_zero hz /-- The map `fun t => exp (t * I)` from `ℝ` to the unit circle in `ℂ`. -/ -def expMapCircle : C(ℝ, circle) where - toFun t := ⟨exp (t * I), by simp [exp_mul_I, abs_cos_add_sin_mul_I]⟩ +def exp : C(ℝ, Circle) where + toFun t := ⟨(t * I).exp, by simp [Submonoid.unitSphere, exp_mul_I, abs_cos_add_sin_mul_I]⟩ + continuous_toFun := Continuous.subtype_mk (by fun_prop) + (by simp [Submonoid.unitSphere, exp_mul_I, abs_cos_add_sin_mul_I]) @[simp] -theorem expMapCircle_apply (t : ℝ) : ↑(expMapCircle t) = Complex.exp (t * Complex.I) := - rfl +theorem exp_apply (t : ℝ) : ↑(exp t) = Complex.exp (t * Complex.I) := rfl @[simp] -theorem expMapCircle_zero : expMapCircle 0 = 1 := - Subtype.ext <| by - rw [expMapCircle_apply, ofReal_zero, zero_mul, exp_zero, Submonoid.coe_one] +theorem exp_zero : exp 0 = 1 := + Subtype.ext <| by rw [exp_apply, ofReal_zero, zero_mul, Complex.exp_zero, coe_one] @[simp] -theorem expMapCircle_add (x y : ℝ) : expMapCircle (x + y) = expMapCircle x * expMapCircle y := +theorem exp_add (x y : ℝ) : exp (x + y) = exp x * exp y := Subtype.ext <| by - simp only [expMapCircle_apply, Submonoid.coe_mul, ofReal_add, add_mul, Complex.exp_add] + simp only [exp_apply, Submonoid.coe_mul, ofReal_add, add_mul, Complex.exp_add, coe_mul] /-- The map `fun t => exp (t * I)` from `ℝ` to the unit circle in `ℂ`, considered as a homomorphism of groups. -/ @[simps] -def expMapCircleHom : ℝ →+ Additive circle where - toFun := Additive.ofMul ∘ expMapCircle - map_zero' := expMapCircle_zero - map_add' := expMapCircle_add +def expHom : ℝ →+ Additive Circle where + toFun := Additive.ofMul ∘ exp + map_zero' := exp_zero + map_add' := exp_add -@[simp] -theorem expMapCircle_sub (x y : ℝ) : expMapCircle (x - y) = expMapCircle x / expMapCircle y := - expMapCircleHom.map_sub x y +@[simp] lemma exp_sub (x y : ℝ) : exp (x - y) = exp x / exp y := expHom.map_sub x y +@[simp] lemma exp_neg (x : ℝ) : exp (-x) = (exp x)⁻¹ := expHom.map_neg x -@[simp] -theorem expMapCircle_neg (x : ℝ) : expMapCircle (-x) = (expMapCircle x)⁻¹ := - expMapCircleHom.map_neg x +variable {α β M : Type*} + +instance instSMul [SMul ℂ α] : SMul Circle α := Submonoid.smul _ + +instance instSMulCommClass_left [SMul ℂ β] [SMul α β] [SMulCommClass ℂ α β] : + SMulCommClass Circle α β := Submonoid.smulCommClass_left _ + +instance instSMulCommClass_right [SMul ℂ β] [SMul α β] [SMulCommClass α ℂ β] : + SMulCommClass α Circle β := Submonoid.smulCommClass_right _ + +instance instIsScalarTower [SMul ℂ α] [SMul ℂ β] [SMul α β] [IsScalarTower ℂ α β] : + IsScalarTower Circle α β := Submonoid.isScalarTower _ + +instance instMulAction [MulAction ℂ α] : MulAction Circle α := Submonoid.mulAction _ + +instance instDistribMulAction [AddMonoid M] [DistribMulAction ℂ M] : + DistribMulAction Circle M := Submonoid.distribMulAction _ + +lemma smul_def [SMul ℂ α] (z : Circle) (a : α) : z • a = (z : ℂ) • a := rfl + +instance instContinuousSMul [TopologicalSpace α] [MulAction ℂ α] [ContinuousSMul ℂ α] : + ContinuousSMul Circle α := Submonoid.continuousSMul @[simp] -lemma norm_circle_smul {E : Type*} [SeminormedAddCommGroup E] [NormedSpace ℂ E] - (u : circle) (v : E) : +protected lemma norm_smul {E : Type*} [SeminormedAddCommGroup E] [NormedSpace ℂ E] + (u : Circle) (v : E) : ‖u • v‖ = ‖v‖ := by rw [Submonoid.smul_def, norm_smul, norm_eq_of_mem_sphere, one_mul] + +@[deprecated (since := "2024-07-24")] noncomputable alias _root_.expMapCircle := exp +@[deprecated (since := "2024-07-24")] noncomputable alias _root_.expMapCircle_apply := exp_apply +@[deprecated (since := "2024-07-24")] noncomputable alias _root_.expMapCircle_zero := exp_zero +@[deprecated (since := "2024-07-24")] noncomputable alias _root_.expMapCircle_sub := exp_sub +@[deprecated (since := "2024-07-24")] noncomputable alias _root_.norm_circle_smul := + Circle.norm_smul + +end Circle diff --git a/Mathlib/Analysis/Complex/Isometry.lean b/Mathlib/Analysis/Complex/Isometry.lean index 379ef2487f8c3..57ac151acc4e2 100644 --- a/Mathlib/Analysis/Complex/Isometry.lean +++ b/Mathlib/Analysis/Complex/Isometry.lean @@ -38,27 +38,27 @@ local notation "|" x "|" => Complex.abs x /-- An element of the unit circle defines a `LinearIsometryEquiv` from `ℂ` to itself, by rotation. -/ -def rotation : circle →* ℂ ≃ₗᵢ[ℝ] ℂ where +def rotation : Circle →* ℂ ≃ₗᵢ[ℝ] ℂ where toFun a := { DistribMulAction.toLinearEquiv ℝ ℂ a with - norm_map' := fun x => show |a * x| = |x| by rw [map_mul, abs_coe_circle, one_mul] } - map_one' := LinearIsometryEquiv.ext <| one_smul circle + norm_map' := fun x => show |a * x| = |x| by rw [map_mul, Circle.abs_coe, one_mul] } + map_one' := LinearIsometryEquiv.ext <| by simp map_mul' a b := LinearIsometryEquiv.ext <| mul_smul a b @[simp] -theorem rotation_apply (a : circle) (z : ℂ) : rotation a z = a * z := +theorem rotation_apply (a : Circle) (z : ℂ) : rotation a z = a * z := rfl @[simp] -theorem rotation_symm (a : circle) : (rotation a).symm = rotation a⁻¹ := +theorem rotation_symm (a : Circle) : (rotation a).symm = rotation a⁻¹ := LinearIsometryEquiv.ext fun _ => rfl @[simp] -theorem rotation_trans (a b : circle) : (rotation a).trans (rotation b) = rotation (b * a) := by +theorem rotation_trans (a b : Circle) : (rotation a).trans (rotation b) = rotation (b * a) := by ext1 simp -theorem rotation_ne_conjLIE (a : circle) : rotation a ≠ conjLIE := by +theorem rotation_ne_conjLIE (a : Circle) : rotation a ≠ conjLIE := by intro h have h1 : rotation a 1 = conj 1 := LinearIsometryEquiv.congr_fun h 1 have hI : rotation a I = conj I := LinearIsometryEquiv.congr_fun h I @@ -69,11 +69,11 @@ theorem rotation_ne_conjLIE (a : circle) : rotation a ≠ conjLIE := by /-- Takes an element of `ℂ ≃ₗᵢ[ℝ] ℂ` and checks if it is a rotation, returns an element of the unit circle. -/ @[simps] -def rotationOf (e : ℂ ≃ₗᵢ[ℝ] ℂ) : circle := - ⟨e 1 / Complex.abs (e 1), by simp⟩ +def rotationOf (e : ℂ ≃ₗᵢ[ℝ] ℂ) : Circle := + ⟨e 1 / Complex.abs (e 1), by simp [Submonoid.unitSphere, ← Complex.norm_eq_abs]⟩ @[simp] -theorem rotationOf_rotation (a : circle) : rotationOf (rotation a) = a := +theorem rotationOf_rotation (a : Circle) : rotationOf (rotation a) = a := Subtype.ext <| by simp theorem rotation_injective : Function.Injective rotation := @@ -127,8 +127,8 @@ theorem linear_isometry_complex_aux {f : ℂ ≃ₗᵢ[ℝ] ℂ} (h : f 1 = 1) : fin_cases i <;> simp [h, h'] theorem linear_isometry_complex (f : ℂ ≃ₗᵢ[ℝ] ℂ) : - ∃ a : circle, f = rotation a ∨ f = conjLIE.trans (rotation a) := by - let a : circle := ⟨f 1, by rw [mem_circle_iff_abs, ← Complex.norm_eq_abs, f.norm_map, norm_one]⟩ + ∃ a : Circle, f = rotation a ∨ f = conjLIE.trans (rotation a) := by + let a : Circle := ⟨f 1, by simp [Submonoid.unitSphere, ← Complex.norm_eq_abs, f.norm_map]⟩ use a have : (f.trans (rotation a).symm) 1 = 1 := by simpa using rotation_apply a⁻¹ (f 1) refine (linear_isometry_complex_aux this).imp (fun h₁ => ?_) fun h₂ => ?_ @@ -137,7 +137,7 @@ theorem linear_isometry_complex (f : ℂ ≃ₗᵢ[ℝ] ℂ) : /-- The matrix representation of `rotation a` is equal to the conformal matrix `!![re a, -im a; im a, re a]`. -/ -theorem toMatrix_rotation (a : circle) : +theorem toMatrix_rotation (a : Circle) : LinearMap.toMatrix basisOneI basisOneI (rotation a).toLinearEquiv = Matrix.planeConformalMatrix (re a) (im a) (by simp [pow_two, ← normSq_apply]) := by ext i j @@ -149,11 +149,11 @@ theorem toMatrix_rotation (a : circle) : /-- The determinant of `rotation` (as a linear map) is equal to `1`. -/ @[simp] -theorem det_rotation (a : circle) : LinearMap.det ((rotation a).toLinearEquiv : ℂ →ₗ[ℝ] ℂ) = 1 := by +theorem det_rotation (a : Circle) : LinearMap.det ((rotation a).toLinearEquiv : ℂ →ₗ[ℝ] ℂ) = 1 := by rw [← LinearMap.det_toMatrix basisOneI, toMatrix_rotation, Matrix.det_fin_two] simp [← normSq_apply] /-- The determinant of `rotation` (as a linear equiv) is equal to `1`. -/ @[simp] -theorem linearEquiv_det_rotation (a : circle) : LinearEquiv.det (rotation a).toLinearEquiv = 1 := by +theorem linearEquiv_det_rotation (a : Circle) : LinearEquiv.det (rotation a).toLinearEquiv = 1 := by rw [← Units.eq_iff, LinearEquiv.coe_det, det_rotation, Units.val_one] diff --git a/Mathlib/Analysis/Complex/UnitDisc/Basic.lean b/Mathlib/Analysis/Complex/UnitDisc/Basic.lean index cee2d76eede82..adaa04bfb5b12 100644 --- a/Mathlib/Analysis/Complex/UnitDisc/Basic.lean +++ b/Mathlib/Analysis/Complex/UnitDisc/Basic.lean @@ -97,23 +97,23 @@ theorem coe_eq_zero {z : 𝔻} : (z : ℂ) = 0 ↔ z = 0 := instance : Inhabited 𝔻 := ⟨0⟩ -instance circleAction : MulAction circle 𝔻 := +instance circleAction : MulAction Circle 𝔻 := mulActionSphereBall -instance isScalarTower_circle_circle : IsScalarTower circle circle 𝔻 := +instance isScalarTower_circle_circle : IsScalarTower Circle Circle 𝔻 := isScalarTower_sphere_sphere_ball -instance isScalarTower_circle : IsScalarTower circle 𝔻 𝔻 := +instance isScalarTower_circle : IsScalarTower Circle 𝔻 𝔻 := isScalarTower_sphere_ball_ball -instance instSMulCommClass_circle : SMulCommClass circle 𝔻 𝔻 := +instance instSMulCommClass_circle : SMulCommClass Circle 𝔻 𝔻 := instSMulCommClass_sphere_ball_ball -instance instSMulCommClass_circle' : SMulCommClass 𝔻 circle 𝔻 := +instance instSMulCommClass_circle' : SMulCommClass 𝔻 Circle 𝔻 := SMulCommClass.symm _ _ _ @[simp, norm_cast] -theorem coe_smul_circle (z : circle) (w : 𝔻) : ↑(z • w) = (z * w : ℂ) := +theorem coe_smul_circle (z : Circle) (w : 𝔻) : ↑(z • w) = (z * w : ℂ) := rfl instance closedBallAction : MulAction (closedBall (0 : ℂ) 1) 𝔻 := @@ -132,10 +132,10 @@ instance instSMulCommClass_closedBall : SMulCommClass (closedBall (0 : ℂ) 1) instance instSMulCommClass_closedBall' : SMulCommClass 𝔻 (closedBall (0 : ℂ) 1) 𝔻 := SMulCommClass.symm _ _ _ -instance instSMulCommClass_circle_closedBall : SMulCommClass circle (closedBall (0 : ℂ) 1) 𝔻 := +instance instSMulCommClass_circle_closedBall : SMulCommClass Circle (closedBall (0 : ℂ) 1) 𝔻 := instSMulCommClass_sphere_closedBall_ball -instance instSMulCommClass_closedBall_circle : SMulCommClass (closedBall (0 : ℂ) 1) circle 𝔻 := +instance instSMulCommClass_closedBall_circle : SMulCommClass (closedBall (0 : ℂ) 1) Circle 𝔻 := SMulCommClass.symm _ _ _ @[simp, norm_cast] diff --git a/Mathlib/Analysis/Fourier/AddCircle.lean b/Mathlib/Analysis/Fourier/AddCircle.lean index 3de5271afa8a5..fe1f6f77fa8b1 100644 --- a/Mathlib/Analysis/Fourier/AddCircle.lean +++ b/Mathlib/Analysis/Fourier/AddCircle.lean @@ -111,7 +111,7 @@ theorem fourier_apply {n : ℤ} {x : AddCircle T} : fourier n x = toCircle (n theorem fourier_coe_apply {n : ℤ} {x : ℝ} : fourier n (x : AddCircle T) = Complex.exp (2 * π * Complex.I * n * x / T) := by rw [fourier_apply, ← QuotientAddGroup.mk_zsmul, toCircle, Function.Periodic.lift_coe, - expMapCircle_apply, Complex.ofReal_mul, Complex.ofReal_div, Complex.ofReal_mul, zsmul_eq_mul, + Circle.exp_apply, Complex.ofReal_mul, Complex.ofReal_div, Complex.ofReal_mul, zsmul_eq_mul, Complex.ofReal_mul, Complex.ofReal_intCast] norm_num congr 1; ring @@ -145,7 +145,7 @@ theorem fourier_neg {n : ℤ} {x : AddCircle T} : fourier (-n) x = conj (fourier induction x using QuotientAddGroup.induction_on simp_rw [fourier_apply, toCircle] rw [← QuotientAddGroup.mk_zsmul, ← QuotientAddGroup.mk_zsmul] - simp_rw [Function.Periodic.lift_coe, ← coe_inv_circle_eq_conj, ← expMapCircle_neg, + simp_rw [Function.Periodic.lift_coe, ← Circle.coe_inv_eq_conj, ← Circle.exp_neg, neg_smul, mul_neg] @[simp] @@ -154,7 +154,7 @@ theorem fourier_neg' {n : ℤ} {x : AddCircle T} : @toCircle T (-(n • x)) = co -- @[simp] -- Porting note: simp normal form is `fourier_add'` theorem fourier_add {m n : ℤ} {x : AddCircle T} : fourier (m+n) x = fourier m x * fourier n x := by - simp_rw [fourier_apply, add_zsmul, toCircle_add, coe_mul_unitSphere] + simp_rw [fourier_apply, add_zsmul, toCircle_add, Circle.coe_mul] @[simp] theorem fourier_add' {m n : ℤ} {x : AddCircle T} : @@ -163,7 +163,7 @@ theorem fourier_add' {m n : ℤ} {x : AddCircle T} : theorem fourier_norm [Fact (0 < T)] (n : ℤ) : ‖@fourier T n‖ = 1 := by rw [ContinuousMap.norm_eq_iSup_norm] - have : ∀ x : AddCircle T, ‖fourier n x‖ = 1 := fun x => abs_coe_circle _ + have : ∀ x : AddCircle T, ‖fourier n x‖ = 1 := fun x => Circle.abs_coe _ simp_rw [this] exact @ciSup_const _ _ _ Zero.instNonempty _ @@ -173,7 +173,7 @@ theorem fourier_add_half_inv_index {n : ℤ} (hn : n ≠ 0) (hT : 0 < T) (x : Ad rw [fourier_apply, zsmul_add, ← QuotientAddGroup.mk_zsmul, toCircle_add, coe_mul_unitSphere] have : (n : ℂ) ≠ 0 := by simpa using hn have : (@toCircle T (n • (T / 2 / n) : ℝ) : ℂ) = -1 := by - rw [zsmul_eq_mul, toCircle, Function.Periodic.lift_coe, expMapCircle_apply] + rw [zsmul_eq_mul, toCircle, Function.Periodic.lift_coe, Circle.exp_apply] replace hT := Complex.ofReal_ne_zero.mpr hT.ne' convert Complex.exp_pi_mul_I using 3 field_simp; ring diff --git a/Mathlib/Analysis/Fourier/FourierTransform.lean b/Mathlib/Analysis/Fourier/FourierTransform.lean index 57b7a7debd124..562e533a3b0f4 100644 --- a/Mathlib/Analysis/Fourier/FourierTransform.lean +++ b/Mathlib/Analysis/Fourier/FourierTransform.lean @@ -21,7 +21,7 @@ We set up the Fourier transform for complex-valued functions on finite-dimension In namespace `VectorFourier`, we define the Fourier integral in the following context: * `𝕜` is a commutative ring. * `V` and `W` are `𝕜`-modules. -* `e` is a unitary additive character of `𝕜`, i.e. an `AddChar 𝕜 circle`. +* `e` is a unitary additive character of `𝕜`, i.e. an `AddChar 𝕜 Circle`. * `μ` is a measure on `V`. * `L` is a `𝕜`-bilinear form `V × W → 𝕜`. * `E` is a complete normed `ℂ`-vector space. @@ -55,7 +55,7 @@ Fourier transform of an integrable function is continuous (under mild assumption noncomputable section -local notation "𝕊" => circle +local notation "𝕊" => Circle open MeasureTheory Filter @@ -93,7 +93,7 @@ theorem norm_fourierIntegral_le_integral_norm (e : AddChar 𝕜 𝕊) (μ : Meas (L : V →ₗ[𝕜] W →ₗ[𝕜] 𝕜) (f : V → E) (w : W) : ‖fourierIntegral e μ L f w‖ ≤ ∫ v : V, ‖f v‖ ∂μ := by refine (norm_integral_le_integral_norm _).trans (le_of_eq ?_) - simp_rw [norm_circle_smul] + simp_rw [Circle.norm_smul] /-- The Fourier integral converts right-translation into scalar multiplication by a phase factor. -/ theorem fourierIntegral_comp_add_right [MeasurableAdd V] (e : AddChar 𝕜 𝕊) (μ : Measure V) @@ -101,11 +101,11 @@ theorem fourierIntegral_comp_add_right [MeasurableAdd V] (e : AddChar 𝕜 𝕊) fourierIntegral e μ L (f ∘ fun v ↦ v + v₀) = fun w ↦ e (L v₀ w) • fourierIntegral e μ L f w := by ext1 w - dsimp only [fourierIntegral, Function.comp_apply, Submonoid.smul_def] + dsimp only [fourierIntegral, Function.comp_apply, Circle.smul_def] conv in L _ => rw [← add_sub_cancel_right v v₀] rw [integral_add_right_eq_self fun v : V ↦ (e (-L (v - v₀) w) : ℂ) • f v, ← integral_smul] congr 1 with v - rw [← smul_assoc, smul_eq_mul, ← Submonoid.coe_mul, ← e.map_add_eq_mul, ← LinearMap.neg_apply, + rw [← smul_assoc, smul_eq_mul, ← Circle.coe_mul, ← e.map_add_eq_mul, ← LinearMap.neg_apply, ← sub_eq_add_neg, ← LinearMap.sub_apply, LinearMap.map_sub, neg_sub] end Defs @@ -131,7 +131,7 @@ theorem fourierIntegral_convergent_iff (he : Continuous e) Integrable (fun v : V ↦ e (-L v x) • g v) μ := by have c : Continuous fun v ↦ e (-L v x) := he.comp (hL.comp (continuous_prod_mk.mpr ⟨continuous_id, continuous_const⟩)).neg - simp_rw [← integrable_norm_iff (c.aestronglyMeasurable.smul hg.1), norm_circle_smul] + simp_rw [← integrable_norm_iff (c.aestronglyMeasurable.smul hg.1), Circle.norm_smul] exact hg.norm -- then use it for both directions refine ⟨fun hf ↦ ?_, fun hf ↦ aux hf w⟩ @@ -159,7 +159,7 @@ theorem fourierIntegral_continuous [FirstCountableTopology W] (he : Continuous e Continuous (fourierIntegral e μ L f) := by apply continuous_of_dominated · exact fun w ↦ ((fourierIntegral_convergent_iff he hL w).2 hf).1 - · exact fun w ↦ ae_of_all _ fun v ↦ le_of_eq (norm_circle_smul _ _) + · exact fun w ↦ ae_of_all _ fun v ↦ le_of_eq (Circle.norm_smul _ _) · exact hf.norm · refine ae_of_all _ fun v ↦ (he.comp ?_).smul continuous_const exact (hL.comp (continuous_prod_mk.mpr ⟨continuous_const, continuous_id⟩)).neg @@ -206,7 +206,7 @@ theorem integral_bilin_fourierIntegral_eq_flip apply B.comp_aestronglyMeasurable A' -- `exact` works, but `apply` is 10x faster! · filter_upwards with ⟨ξ, x⟩ rw [Function.uncurry_apply_pair, Submonoid.smul_def, (M.flip (g ξ)).map_smul, - ← Submonoid.smul_def, norm_circle_smul, ContinuousLinearMap.flip_apply, + ← Submonoid.smul_def, Circle.norm_smul, ContinuousLinearMap.flip_apply, norm_mul, norm_norm M, norm_mul, norm_norm, norm_norm, mul_comm (‖g _‖), ← mul_assoc] exact M.le_opNorm₂ (f x) (g ξ) _ = ∫ x, (∫ ξ, M (f x) (e (-L.flip ξ x) • g ξ) ∂ν) ∂μ := by @@ -307,9 +307,9 @@ namespace Real /-- The standard additive character of `ℝ`, given by `fun x ↦ exp (2 * π * x * I)`. -/ def fourierChar : AddChar ℝ 𝕊 where - toFun z := expMapCircle (2 * π * z) - map_zero_eq_one' := by simp only; rw [mul_zero, expMapCircle_zero] - map_add_eq_mul' x y := by simp only; rw [mul_add, expMapCircle_add] + toFun z := .exp (2 * π * z) + map_zero_eq_one' := by simp only; rw [mul_zero, Circle.exp_zero] + map_add_eq_mul' x y := by simp only; rw [mul_add, Circle.exp_add] @[inherit_doc] scoped[FourierTransform] notation "𝐞" => Real.fourierChar @@ -319,8 +319,7 @@ theorem fourierChar_apply (x : ℝ) : 𝐞 x = Complex.exp (↑(2 * π * x) * Co rfl @[continuity] -theorem continuous_fourierChar : Continuous 𝐞 := - (map_continuous expMapCircle).comp (continuous_mul_left _) +theorem continuous_fourierChar : Continuous 𝐞 := Circle.exp.continuous.comp (continuous_mul_left _) variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] @@ -329,7 +328,7 @@ theorem vector_fourierIntegral_eq_integral_exp_smul {V : Type*} [AddCommGroup V] (μ : Measure V) (f : V → E) (w : W) : VectorFourier.fourierIntegral fourierChar μ L f w = ∫ v : V, Complex.exp (↑(-2 * π * L v w) * Complex.I) • f v ∂μ := by - simp_rw [VectorFourier.fourierIntegral, Submonoid.smul_def, Real.fourierChar_apply, mul_neg, + simp_rw [VectorFourier.fourierIntegral, Circle.smul_def, Real.fourierChar_apply, mul_neg, neg_mul] /-- The Fourier integral is well defined iff the function is integrable. Version with a general @@ -397,7 +396,7 @@ lemma fourierIntegral_eq (f : V → E) (w : V) : lemma fourierIntegral_eq' (f : V → E) (w : V) : 𝓕 f w = ∫ v, Complex.exp ((↑(-2 * π * ⟪v, w⟫) * Complex.I)) • f v := by - simp_rw [fourierIntegral_eq, Submonoid.smul_def, Real.fourierChar_apply, mul_neg, neg_mul] + simp_rw [fourierIntegral_eq, Circle.smul_def, Real.fourierChar_apply, mul_neg, neg_mul] lemma fourierIntegralInv_eq (f : V → E) (w : V) : 𝓕⁻ f w = ∫ v, 𝐞 ⟪v, w⟫ • f v := by @@ -405,7 +404,7 @@ lemma fourierIntegralInv_eq (f : V → E) (w : V) : lemma fourierIntegralInv_eq' (f : V → E) (w : V) : 𝓕⁻ f w = ∫ v, Complex.exp ((↑(2 * π * ⟪v, w⟫) * Complex.I)) • f v := by - simp_rw [fourierIntegralInv_eq, Submonoid.smul_def, Real.fourierChar_apply] + simp_rw [fourierIntegralInv_eq, Circle.smul_def, Real.fourierChar_apply] lemma fourierIntegral_comp_linearIsometry (A : W ≃ₗᵢ[ℝ] V) (f : V → E) (w : W) : 𝓕 (f ∘ A) w = (𝓕 f) (A w) := by @@ -440,7 +439,7 @@ theorem fourierIntegral_real_eq (f : ℝ → E) (w : ℝ) : theorem fourierIntegral_real_eq_integral_exp_smul (f : ℝ → E) (w : ℝ) : 𝓕 f w = ∫ v : ℝ, Complex.exp (↑(-2 * π * v * w) * Complex.I) • f v := by - simp_rw [fourierIntegral_real_eq, Submonoid.smul_def, Real.fourierChar_apply, mul_neg, neg_mul, + simp_rw [fourierIntegral_real_eq, Circle.smul_def, Real.fourierChar_apply, mul_neg, neg_mul, mul_assoc] @[deprecated (since := "2024-02-21")] diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 226ce61296678..a4616a71f06cb 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -225,7 +225,7 @@ theorem hasFDerivAt_fourierIntegral exact (L.continuous₂.comp (Continuous.Prod.mk_left w)).neg have h4 : (∀ᵐ v ∂μ, ∀ (w' : W), w' ∈ Metric.ball w 1 → ‖F' w' v‖ ≤ B v) := by filter_upwards with v w' _ - rw [norm_circle_smul _ (fourierSMulRight L f v)] + rw [Circle.norm_smul _ (fourierSMulRight L f v)] exact norm_fourierSMulRight_le L f v have h5 : Integrable B μ := by simpa only [← mul_assoc] using hf'.const_mul (2 * π * ‖L‖) have h6 : ∀ᵐ v ∂μ, ∀ w', w' ∈ Metric.ball w 1 → HasFDerivAt (fun x ↦ F x v) (F' w' v) w' := diff --git a/Mathlib/Analysis/Fourier/PoissonSummation.lean b/Mathlib/Analysis/Fourier/PoissonSummation.lean index 848611bbd9fac..f88b25200c60c 100644 --- a/Mathlib/Analysis/Fourier/PoissonSummation.lean +++ b/Mathlib/Analysis/Fourier/PoissonSummation.lean @@ -58,7 +58,7 @@ theorem Real.fourierCoeff_tsum_comp_add {f : C(ℝ, ℂ)} -- block, but I think it's more legible this way. We start with preliminaries about the integrand. let e : C(ℝ, ℂ) := (fourier (-m)).comp ⟨((↑) : ℝ → UnitAddCircle), continuous_quotient_mk'⟩ have neK : ∀ (K : Compacts ℝ) (g : C(ℝ, ℂ)), ‖(e * g).restrict K‖ = ‖g.restrict K‖ := by - have (x : ℝ) : ‖e x‖ = 1 := abs_coe_circle (AddCircle.toCircle (-m • x)) + have (x : ℝ) : ‖e x‖ = 1 := (AddCircle.toCircle (-m • x)).abs_coe intro K g simp_rw [norm_eq_iSup_norm, restrict_apply, mul_apply, norm_mul, this, one_mul] have eadd : ∀ (n : ℤ), e.comp (ContinuousMap.addRight n) = e := by diff --git a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean index c8a6fff1cb363..59c58e215be50 100644 --- a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean +++ b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean @@ -69,7 +69,7 @@ theorem fourierIntegral_half_period_translate {w : V} (hw : w ≠ 0) : (fun v : V => 𝐞 (-⟪v, w⟫) • f (v + i w)) = fun v : V => (fun x : V => -(𝐞 (-⟪x, w⟫) • f x)) (v + i w) := by ext1 v - simp_rw [inner_add_left, hiw, Submonoid.smul_def, Real.fourierChar_apply, neg_add, mul_add, + simp_rw [inner_add_left, hiw, Circle.smul_def, Real.fourierChar_apply, neg_add, mul_add, ofReal_add, add_mul, exp_add] have : 2 * π * -(1 / 2) = -π := by field_simp; ring rw [this, ofReal_neg, neg_mul, exp_neg, exp_pi_mul_I, inv_neg, inv_one, mul_neg_one, neg_smul, @@ -142,7 +142,7 @@ theorem tendsto_integral_exp_inner_smul_cocompact_of_continuous_compact_support (hf1.integrable_of_hasCompactSupport hf2), norm_smul, this, inv_mul_eq_div, div_lt_iff' two_pos] refine lt_of_le_of_lt (norm_integral_le_integral_norm _) ?_ - simp_rw [norm_circle_smul] + simp_rw [Circle.norm_smul] --* Show integral can be taken over A only. have int_A : ∫ v : V, ‖f v - f (v + i w)‖ = ∫ v in A, ‖f v - f (v + i w)‖ := by refine (setIntegral_eq_integral_of_forall_compl_eq_zero fun v hv => ?_).symm diff --git a/Mathlib/Analysis/Normed/Field/UnitBall.lean b/Mathlib/Analysis/Normed/Field/UnitBall.lean index 68d1cb5b9eda5..8652d79164427 100644 --- a/Mathlib/Analysis/Normed/Field/UnitBall.lean +++ b/Mathlib/Analysis/Normed/Field/UnitBall.lean @@ -95,6 +95,7 @@ theorem coe_pow_unitClosedBall [SeminormedRing 𝕜] [NormOneClass 𝕜] (x : cl rfl /-- Unit sphere in a normed division ring as a bundled `Submonoid`. -/ +@[simps] def Submonoid.unitSphere (𝕜 : Type*) [NormedDivisionRing 𝕜] : Submonoid 𝕜 where carrier := sphere (0 : 𝕜) 1 mul_mem' hx hy := by diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean index 1dde0cd951592..0b7ee63908505 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Circle.lean @@ -19,177 +19,175 @@ open Complex Function Set open Real -namespace circle +namespace Circle -theorem injective_arg : Injective fun z : circle => arg z := fun z w h => - Subtype.ext <| ext_abs_arg ((abs_coe_circle z).trans (abs_coe_circle w).symm) h +theorem injective_arg : Injective fun z : Circle => arg z := fun z w h => + Subtype.ext <| ext_abs_arg (z.abs_coe.trans w.abs_coe.symm) h @[simp] -theorem arg_eq_arg {z w : circle} : arg z = arg w ↔ z = w := +theorem arg_eq_arg {z w : Circle} : arg z = arg w ↔ z = w := injective_arg.eq_iff -end circle - -theorem arg_expMapCircle {x : ℝ} (h₁ : -π < x) (h₂ : x ≤ π) : arg (expMapCircle x) = x := by - rw [expMapCircle_apply, exp_mul_I, arg_cos_add_sin_mul_I ⟨h₁, h₂⟩] +theorem arg_exp {x : ℝ} (h₁ : -π < x) (h₂ : x ≤ π) : arg (exp x) = x := by + rw [exp_apply, exp_mul_I, arg_cos_add_sin_mul_I ⟨h₁, h₂⟩] @[simp] -theorem expMapCircle_arg (z : circle) : expMapCircle (arg z) = z := - circle.injective_arg <| arg_expMapCircle (neg_pi_lt_arg _) (arg_le_pi _) +theorem exp_arg (z : Circle) : exp (arg z) = z := + injective_arg <| arg_exp (neg_pi_lt_arg _) (arg_le_pi _) -namespace circle +@[deprecated (since := "2024-07-25")] alias _root_.arg_expMapCircle := arg_exp +@[deprecated (since := "2024-07-25")] alias _root_.expMapCircle_arg := exp_arg /-- `Complex.arg ∘ (↑)` and `expMapCircle` define a partial equivalence between `circle` and `ℝ` with `source = Set.univ` and `target = Set.Ioc (-π) π`. -/ @[simps (config := .asFn)] -noncomputable def argPartialEquiv : PartialEquiv circle ℝ where +noncomputable def argPartialEquiv : PartialEquiv Circle ℝ where toFun := arg ∘ (↑) - invFun := expMapCircle + invFun := exp source := univ target := Ioc (-π) π map_source' _ _ := ⟨neg_pi_lt_arg _, arg_le_pi _⟩ map_target' := mapsTo_univ _ _ - left_inv' z _ := expMapCircle_arg z - right_inv' _ hx := arg_expMapCircle hx.1 hx.2 + left_inv' z _ := exp_arg z + right_inv' _ hx := arg_exp hx.1 hx.2 /-- `Complex.arg` and `expMapCircle` define an equivalence between `circle` and `(-π, π]`. -/ @[simps (config := .asFn)] -noncomputable def argEquiv : circle ≃ Ioc (-π) π where +noncomputable def argEquiv : Circle ≃ Ioc (-π) π where toFun z := ⟨arg z, neg_pi_lt_arg _, arg_le_pi _⟩ - invFun := expMapCircle ∘ (↑) + invFun := exp ∘ (↑) left_inv _ := argPartialEquiv.left_inv trivial right_inv x := Subtype.ext <| argPartialEquiv.right_inv x.2 -end circle - -theorem leftInverse_expMapCircle_arg : LeftInverse expMapCircle (arg ∘ (↑)) := - expMapCircle_arg - -theorem invOn_arg_expMapCircle : InvOn (arg ∘ (↑)) expMapCircle (Ioc (-π) π) univ := - circle.argPartialEquiv.symm.invOn +lemma leftInverse_exp_arg : LeftInverse exp (arg ∘ (↑)) := exp_arg +lemma invOn_arg_exp : InvOn (arg ∘ (↑)) exp (Ioc (-π) π) univ := argPartialEquiv.symm.invOn +lemma surjOn_exp_neg_pi_pi : SurjOn exp (Ioc (-π) π) univ := argPartialEquiv.symm.surjOn -theorem surjOn_expMapCircle_neg_pi_pi : SurjOn expMapCircle (Ioc (-π) π) univ := - circle.argPartialEquiv.symm.surjOn - -theorem expMapCircle_eq_expMapCircle {x y : ℝ} : - expMapCircle x = expMapCircle y ↔ ∃ m : ℤ, x = y + m * (2 * π) := by - rw [Subtype.ext_iff, expMapCircle_apply, expMapCircle_apply, exp_eq_exp_iff_exists_int] +lemma exp_eq_exp {x y : ℝ} : exp x = exp y ↔ ∃ m : ℤ, x = y + m * (2 * π) := by + rw [Subtype.ext_iff, exp_apply, exp_apply, exp_eq_exp_iff_exists_int] refine exists_congr fun n => ?_ rw [← mul_assoc, ← add_mul, mul_left_inj' I_ne_zero] norm_cast -theorem periodic_expMapCircle : Periodic expMapCircle (2 * π) := fun z => - expMapCircle_eq_expMapCircle.2 ⟨1, by rw [Int.cast_one, one_mul]⟩ +lemma periodic_exp : Periodic exp (2 * π) := fun z ↦ exp_eq_exp.2 ⟨1, by rw [Int.cast_one, one_mul]⟩ -@[simp] -theorem expMapCircle_two_pi : expMapCircle (2 * π) = 1 := - periodic_expMapCircle.eq.trans expMapCircle_zero +@[simp] lemma exp_two_pi : exp (2 * π) = 1 := periodic_exp.eq.trans exp_zero -theorem expMapCircle_sub_two_pi (x : ℝ) : expMapCircle (x - 2 * π) = expMapCircle x := - periodic_expMapCircle.sub_eq x +lemma exp_sub_two_pi (x : ℝ) : exp (x - 2 * π) = exp x := periodic_exp.sub_eq x +lemma exp_add_two_pi (x : ℝ) : exp (x + 2 * π) = exp x := periodic_exp x -theorem expMapCircle_add_two_pi (x : ℝ) : expMapCircle (x + 2 * π) = expMapCircle x := - periodic_expMapCircle x +@[deprecated (since := "2024-07-25")] +alias _root_.leftInverse_expMapCircle_arg := leftInverse_exp_arg -/-- `expMapCircle`, applied to a `Real.Angle`. -/ -noncomputable def Real.Angle.expMapCircle (θ : Real.Angle) : circle := - periodic_expMapCircle.lift θ +@[deprecated (since := "2024-07-25")] +alias _root_.surjOn_expMapCircle_neg_pi_pi := surjOn_exp_neg_pi_pi -@[simp] -theorem Real.Angle.expMapCircle_coe (x : ℝ) : Real.Angle.expMapCircle x = _root_.expMapCircle x := - rfl +@[deprecated (since := "2024-07-25")] alias _root_.invOn_arg_expMapCircle := invOn_arg_exp +@[deprecated (since := "2024-07-25")] alias _root_.expMapCircle_eq_expMapCircle := exp_eq_exp +@[deprecated (since := "2024-07-25")] alias _root_.periodic_expMapCircle := periodic_exp +@[deprecated (since := "2024-07-25")] alias _root_.expMapCircle_two_pi := exp_two_pi +@[deprecated (since := "2024-07-25")] alias _root_.expMapCircle_sub_two_pi := exp_sub_two_pi +@[deprecated (since := "2024-07-25")] alias _root_.expMapCircle_add_two_pi := exp_add_two_pi + +end Circle + +namespace Real.Angle + +/-- `Circle.exp`, applied to a `Real.Angle`. -/ +noncomputable def toCircle (θ : Angle) : Circle := Circle.periodic_exp.lift θ -theorem Real.Angle.coe_expMapCircle (θ : Real.Angle) : - (θ.expMapCircle : ℂ) = θ.cos + θ.sin * I := by +@[simp] lemma toCircle_coe (x : ℝ) : toCircle x = .exp x := rfl + +lemma coe_toCircle (θ : Angle) : (θ.toCircle : ℂ) = θ.cos + θ.sin * I := by induction θ using Real.Angle.induction_on simp [exp_mul_I] -@[simp] -theorem Real.Angle.expMapCircle_zero : Real.Angle.expMapCircle 0 = 1 := by - rw [← Real.Angle.coe_zero, Real.Angle.expMapCircle_coe, _root_.expMapCircle_zero] +@[simp] lemma toCircle_zero : toCircle 0 = 1 := by rw [← coe_zero, toCircle_coe, Circle.exp_zero] -@[simp] -theorem Real.Angle.expMapCircle_neg (θ : Real.Angle) : - Real.Angle.expMapCircle (-θ) = (Real.Angle.expMapCircle θ)⁻¹ := by +@[simp] lemma toCircle_neg (θ : Angle) : toCircle (-θ) = (toCircle θ)⁻¹ := by induction θ using Real.Angle.induction_on - simp_rw [← Real.Angle.coe_neg, Real.Angle.expMapCircle_coe, _root_.expMapCircle_neg] + simp_rw [← coe_neg, toCircle_coe, Circle.exp_neg] -@[simp] -theorem Real.Angle.expMapCircle_add (θ₁ θ₂ : Real.Angle) : Real.Angle.expMapCircle (θ₁ + θ₂) = - Real.Angle.expMapCircle θ₁ * Real.Angle.expMapCircle θ₂ := by +@[simp] lemma toCircle_add (θ₁ θ₂ : Angle) : toCircle (θ₁ + θ₂) = toCircle θ₁ * toCircle θ₂ := by induction θ₁ using Real.Angle.induction_on induction θ₂ using Real.Angle.induction_on - exact _root_.expMapCircle_add _ _ + exact Circle.exp_add _ _ -@[simp] -theorem Real.Angle.arg_expMapCircle (θ : Real.Angle) : - (arg (Real.Angle.expMapCircle θ) : Real.Angle) = θ := by +@[simp] lemma arg_toCircle (θ : Real.Angle) : (arg θ.toCircle : Angle) = θ := by induction θ using Real.Angle.induction_on - rw [Real.Angle.expMapCircle_coe, expMapCircle_apply, exp_mul_I, ← ofReal_cos, ← ofReal_sin, ← + rw [toCircle_coe, Circle.exp_apply, exp_mul_I, ← ofReal_cos, ← ofReal_sin, ← Real.Angle.cos_coe, ← Real.Angle.sin_coe, arg_cos_add_sin_mul_I_coe_angle] +@[deprecated (since := "2024-07-25")] alias expMapCircle := toCircle +@[deprecated (since := "2024-07-25")] alias expMapCircle_coe := toCircle_coe +@[deprecated (since := "2024-07-25")] alias coe_expMapCircle := coe_toCircle +@[deprecated (since := "2024-07-25")] alias expMapCircle_zero := toCircle_zero +@[deprecated (since := "2024-07-25")] alias expMapCircle_neg := toCircle_neg +@[deprecated (since := "2024-07-25")] alias expMapCircle_add := toCircle_add +@[deprecated (since := "2024-07-25")] alias arg_expMapCircle := arg_toCircle + +end Real.Angle + namespace AddCircle variable {T : ℝ} /-! ### Map from `AddCircle` to `Circle` -/ -theorem scaled_exp_map_periodic : Function.Periodic (fun x => expMapCircle (2 * π / T * x)) T := by +theorem scaled_exp_map_periodic : Function.Periodic (fun x => Circle.exp (2 * π / T * x)) T := by -- The case T = 0 is not interesting, but it is true, so we prove it to save hypotheses rcases eq_or_ne T 0 with (rfl | hT) · intro x; simp - · intro x; simp_rw [mul_add]; rw [div_mul_cancel₀ _ hT, periodic_expMapCircle] + · intro x; simp_rw [mul_add]; rw [div_mul_cancel₀ _ hT, Circle.periodic_exp] /-- The canonical map `fun x => exp (2 π i x / T)` from `ℝ / ℤ • T` to the unit circle in `ℂ`. If `T = 0` we understand this as the constant function 1. -/ -noncomputable def toCircle : AddCircle T → circle := +noncomputable def toCircle : AddCircle T → Circle := (@scaled_exp_map_periodic T).lift -theorem toCircle_apply_mk (x : ℝ) : @toCircle T x = expMapCircle (2 * π / T * x) := +theorem toCircle_apply_mk (x : ℝ) : @toCircle T x = Circle.exp (2 * π / T * x) := rfl theorem toCircle_add (x : AddCircle T) (y : AddCircle T) : @toCircle T (x + y) = toCircle x * toCircle y := by induction x using QuotientAddGroup.induction_on induction y using QuotientAddGroup.induction_on - simp_rw [← coe_add, toCircle_apply_mk, mul_add, expMapCircle_add] + simp_rw [← coe_add, toCircle_apply_mk, mul_add, Circle.exp_add] lemma toCircle_zero : toCircle (0 : AddCircle T) = 1 := by - rw [← QuotientAddGroup.mk_zero, toCircle_apply_mk, mul_zero, expMapCircle_zero] + rw [← QuotientAddGroup.mk_zero, toCircle_apply_mk, mul_zero, Circle.exp_zero] theorem continuous_toCircle : Continuous (@toCircle T) := - continuous_coinduced_dom.mpr (expMapCircle.continuous.comp <| continuous_const.mul continuous_id') + continuous_coinduced_dom.mpr (Circle.exp.continuous.comp <| continuous_const.mul continuous_id') theorem injective_toCircle (hT : T ≠ 0) : Function.Injective (@toCircle T) := by intro a b h induction a using QuotientAddGroup.induction_on induction b using QuotientAddGroup.induction_on simp_rw [toCircle_apply_mk] at h - obtain ⟨m, hm⟩ := expMapCircle_eq_expMapCircle.mp h.symm + obtain ⟨m, hm⟩ := Circle.exp_eq_exp.mp h.symm rw [QuotientAddGroup.eq]; simp_rw [AddSubgroup.mem_zmultiples_iff, zsmul_eq_mul] use m field_simp at hm rw [← mul_right_inj' Real.two_pi_pos.ne'] linarith -/-- The homeomorphism between `AddCircle (2 * π)` and `circle`. -/ -@[simps] noncomputable def homeomorphCircle' : AddCircle (2 * π) ≃ₜ circle where - toFun := Angle.expMapCircle +/-- The homeomorphism between `AddCircle (2 * π)` and `Circle`. -/ +@[simps] noncomputable def homeomorphCircle' : AddCircle (2 * π) ≃ₜ Circle where + toFun := Angle.toCircle invFun := fun x ↦ arg x - left_inv := Angle.arg_expMapCircle - right_inv := expMapCircle_arg - continuous_toFun := continuous_coinduced_dom.mpr expMapCircle.continuous + left_inv := Angle.arg_toCircle + right_inv := Circle.exp_arg + continuous_toFun := continuous_coinduced_dom.mpr Circle.exp.continuous continuous_invFun := by rw [continuous_iff_continuousAt] intro x - apply (continuousAt_arg_coe_angle (ne_zero_of_mem_circle x)).comp - continuousAt_subtype_val + exact (continuousAt_arg_coe_angle x.coe_ne_zero).comp continuousAt_subtype_val -theorem homeomorphCircle'_apply_mk (x : ℝ) : homeomorphCircle' x = expMapCircle x := - rfl +theorem homeomorphCircle'_apply_mk (x : ℝ) : homeomorphCircle' x = Circle.exp x := rfl -/-- The homeomorphism between `AddCircle` and `circle`. -/ -noncomputable def homeomorphCircle (hT : T ≠ 0) : AddCircle T ≃ₜ circle := +/-- The homeomorphism between `AddCircle` and `Circle`. -/ +noncomputable def homeomorphCircle (hT : T ≠ 0) : AddCircle T ≃ₜ Circle := (homeomorphAddCircle T (2 * π) hT (by positivity)).trans homeomorphCircle' theorem homeomorphCircle_apply (hT : T ≠ 0) (x : AddCircle T) : @@ -203,7 +201,10 @@ end AddCircle open AddCircle --- todo: upgrade this to `IsCoveringMap expMapCircle`. -lemma isLocalHomeomorph_expMapCircle : IsLocalHomeomorph expMapCircle := by +-- todo: upgrade this to `IsCoveringMap Circle.exp`. +lemma isLocalHomeomorph_circleExp : IsLocalHomeomorph Circle.exp := by have : Fact (0 < 2 * π) := ⟨by positivity⟩ exact homeomorphCircle'.isLocalHomeomorph.comp (isLocalHomeomorph_coe (2 * π)) + +@[deprecated (since := "2024-07-25")] +alias isLocalHomeomorph_expMapCircle := isLocalHomeomorph_circleExp diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/CircleAddChar.lean b/Mathlib/Analysis/SpecialFunctions/Complex/CircleAddChar.lean index a688be59e21a2..02804c44f2602 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/CircleAddChar.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/CircleAddChar.lean @@ -22,7 +22,7 @@ open Complex Function open scoped Real /-- The canonical map from the additive to the multiplicative circle, as an `AddChar`. -/ -noncomputable def AddCircle.toCircle_addChar {T : ℝ} : AddChar (AddCircle T) circle where +noncomputable def AddCircle.toCircle_addChar {T : ℝ} : AddChar (AddCircle T) Circle where toFun := toCircle map_zero_eq_one' := toCircle_zero map_add_eq_mul' := toCircle_add @@ -31,17 +31,23 @@ open AddCircle namespace ZMod +/-! +### Additive characters valued in the complex circle +-/ + +open scoped Real + variable {N : ℕ} [NeZero N] /-- The additive character from `ZMod N` to the unit circle in `ℂ`, sending `j mod N` to `exp (2 * π * I * j / N)`. -/ -noncomputable def toCircle : AddChar (ZMod N) circle := +noncomputable def toCircle : AddChar (ZMod N) Circle := toCircle_addChar.compAddMonoidHom toAddCircle lemma toCircle_intCast (j : ℤ) : toCircle (j : ZMod N) = exp (2 * π * I * j / N) := by rw [toCircle, AddChar.compAddMonoidHom_apply, toCircle_addChar, AddChar.coe_mk, - AddCircle.toCircle, toAddCircle_intCast, Function.Periodic.lift_coe, expMapCircle_apply] + AddCircle.toCircle, toAddCircle_intCast, Function.Periodic.lift_coe, Circle.exp_apply] push_cast ring_nf @@ -57,16 +63,14 @@ lemma toCircle_apply (j : ZMod N) : toCircle j = exp (2 * π * I * j.val / N) := by rw [← toCircle_natCast, natCast_zmod_val] -lemma injective_toCircle : Injective (toCircle : ZMod N → circle) := +lemma injective_toCircle : Injective (toCircle : ZMod N → Circle) := (AddCircle.injective_toCircle one_ne_zero).comp (toAddCircle_injective N) /-- The additive character from `ZMod N` to `ℂ`, sending `j mod N` to `exp (2 * π * I * j / N)`. -/ -noncomputable def stdAddChar : AddChar (ZMod N) ℂ := circle.subtype.compAddChar toCircle +noncomputable def stdAddChar : AddChar (ZMod N) ℂ := Circle.coeHom.compAddChar toCircle lemma stdAddChar_coe (j : ℤ) : - stdAddChar (j : ZMod N) = exp (2 * π * I * j / N) := by - simp only [stdAddChar, MonoidHom.coe_compAddChar, Function.comp_apply, - Submonoid.coe_subtype, toCircle_intCast] + stdAddChar (j : ZMod N) = exp (2 * π * I * j / N) := by simp [stdAddChar, toCircle_intCast] lemma stdAddChar_apply (j : ZMod N) : stdAddChar j = ↑(toCircle j) := rfl diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean index f98682ac579f3..b933762e6b4a6 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Rotation.lean @@ -161,13 +161,13 @@ theorem rotation_trans (θ₁ θ₂ : Real.Angle) : /-- Rotating the first of two vectors by `θ` scales their Kahler form by `cos θ - sin θ * I`. -/ @[simp] theorem kahler_rotation_left (x y : V) (θ : Real.Angle) : - o.kahler (o.rotation θ x) y = conj (θ.expMapCircle : ℂ) * o.kahler x y := by + o.kahler (o.rotation θ x) y = conj (θ.toCircle : ℂ) * o.kahler x y := by -- Porting note: this needed the `Complex.conj_ofReal` instead of `RCLike.conj_ofReal`; -- I believe this is because the respective coercions are no longer defeq, and - -- `Real.Angle.coe_expMapCircle` uses the `Complex` version. + -- `Real.Angle.coe_toCircle` uses the `Complex` version. simp only [o.rotation_apply, map_add, map_mul, LinearMap.map_smulₛₗ, RingHom.id_apply, LinearMap.add_apply, LinearMap.smul_apply, real_smul, kahler_rightAngleRotation_left, - Real.Angle.coe_expMapCircle, Complex.conj_ofReal, conj_I] + Real.Angle.coe_toCircle, Complex.conj_ofReal, conj_I] ring /-- Negating a rotation is equivalent to rotation by π plus the angle. -/ @@ -187,15 +187,15 @@ theorem neg_rotation_pi_div_two (x : V) : -o.rotation (π / 2 : ℝ) x = o.rotat /-- Rotating the first of two vectors by `θ` scales their Kahler form by `cos (-θ) + sin (-θ) * I`. -/ theorem kahler_rotation_left' (x y : V) (θ : Real.Angle) : - o.kahler (o.rotation θ x) y = (-θ).expMapCircle * o.kahler x y := by - simp only [Real.Angle.expMapCircle_neg, coe_inv_circle_eq_conj, kahler_rotation_left] + o.kahler (o.rotation θ x) y = (-θ).toCircle * o.kahler x y := by + simp only [Real.Angle.toCircle_neg, Circle.coe_inv_eq_conj, kahler_rotation_left] /-- Rotating the second of two vectors by `θ` scales their Kahler form by `cos θ + sin θ * I`. -/ @[simp] theorem kahler_rotation_right (x y : V) (θ : Real.Angle) : - o.kahler x (o.rotation θ y) = θ.expMapCircle * o.kahler x y := by + o.kahler x (o.rotation θ y) = θ.toCircle * o.kahler x y := by simp only [o.rotation_apply, map_add, LinearMap.map_smulₛₗ, RingHom.id_apply, real_smul, - kahler_rightAngleRotation_right, Real.Angle.coe_expMapCircle] + kahler_rightAngleRotation_right, Real.Angle.coe_toCircle] ring /-- Rotating the first vector by `θ` subtracts `θ` from the angle between two vectors. -/ @@ -203,9 +203,9 @@ theorem kahler_rotation_right (x y : V) (θ : Real.Angle) : theorem oangle_rotation_left {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : Real.Angle) : o.oangle (o.rotation θ x) y = o.oangle x y - θ := by simp only [oangle, o.kahler_rotation_left'] - rw [Complex.arg_mul_coe_angle, Real.Angle.arg_expMapCircle] + rw [Complex.arg_mul_coe_angle, Real.Angle.arg_toCircle] · abel - · exact ne_zero_of_mem_circle _ + · exact Circle.coe_ne_zero _ · exact o.kahler_ne_zero hx hy /-- Rotating the second vector by `θ` adds `θ` to the angle between two vectors. -/ @@ -213,9 +213,9 @@ theorem oangle_rotation_left {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : Real. theorem oangle_rotation_right {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : Real.Angle) : o.oangle x (o.rotation θ y) = o.oangle x y + θ := by simp only [oangle, o.kahler_rotation_right] - rw [Complex.arg_mul_coe_angle, Real.Angle.arg_expMapCircle] + rw [Complex.arg_mul_coe_angle, Real.Angle.arg_toCircle] · abel - · exact ne_zero_of_mem_circle _ + · exact Circle.coe_ne_zero _ · exact o.kahler_ne_zero hx hy /-- The rotation of a vector by `θ` has an angle of `-θ` from that vector. -/ @@ -360,15 +360,15 @@ theorem rotation_map (θ : Real.Angle) (f : V ≃ₗᵢ[ℝ] V') (x : V') : @[simp] protected theorem _root_.Complex.rotation (θ : Real.Angle) (z : ℂ) : - Complex.orientation.rotation θ z = θ.expMapCircle * z := by - simp only [rotation_apply, Complex.rightAngleRotation, Real.Angle.coe_expMapCircle, real_smul] + Complex.orientation.rotation θ z = θ.toCircle * z := by + simp only [rotation_apply, Complex.rightAngleRotation, Real.Angle.coe_toCircle, real_smul] ring /-- Rotation in an oriented real inner product space of dimension 2 can be evaluated in terms of a complex-number representation of the space. -/ theorem rotation_map_complex (θ : Real.Angle) (f : V ≃ₗᵢ[ℝ] ℂ) (hf : Orientation.map (Fin 2) f.toLinearEquiv o = Complex.orientation) (x : V) : - f (o.rotation θ x) = θ.expMapCircle * f x := by + f (o.rotation θ x) = θ.toCircle * f x := by rw [← Complex.rotation, ← hf, o.rotation_map, LinearIsometryEquiv.symm_apply_apply] /-- Negating the orientation negates the angle in `rotation`. -/ diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index 83c790d52ad36..41a32a5f566cb 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -37,13 +37,13 @@ We prove two lemmas about smooth maps: As an application we prove `contMdiffNegSphere`, that the antipodal map is smooth. -Finally, we equip the `circle` (defined in `Analysis.Complex.Circle` to be the sphere in `ℂ` +Finally, we equip the `Circle` (defined in `Analysis.Complex.Circle` to be the sphere in `ℂ` centred at `0` of radius `1`) with the following structure: * a charted space with model space `EuclideanSpace ℝ (Fin 1)` (inherited from `Metric.Sphere`) * a Lie group with model with corners `𝓡 1` -We furthermore show that `expMapCircle` (defined in `Analysis.Complex.Circle` to be the natural -map `fun t ↦ exp (t * I)` from `ℝ` to `circle`) is smooth. +We furthermore show that `Circle.exp` (defined in `Analysis.Complex.Circle` to be the natural +map `fun t ↦ exp (t * I)` from `ℝ` to `Circle`) is smooth. ## Implementation notes @@ -550,7 +550,7 @@ theorem mfderiv_coe_sphere_injective {n : ℕ} [Fact (finrank ℝ E = n + 1)] (v end SmoothManifold -section circle +section Circle open Complex @@ -562,17 +562,17 @@ attribute [local instance] finrank_real_complex_fact' /-- The unit circle in `ℂ` is a charted space modelled on `EuclideanSpace ℝ (Fin 1)`. This follows by definition from the corresponding result for `Metric.Sphere`. -/ -instance : ChartedSpace (EuclideanSpace ℝ (Fin 1)) circle := +instance : ChartedSpace (EuclideanSpace ℝ (Fin 1)) Circle := EuclideanSpace.instChartedSpaceSphere -instance : SmoothManifoldWithCorners (𝓡 1) circle := +instance : SmoothManifoldWithCorners (𝓡 1) Circle := EuclideanSpace.instSmoothManifoldWithCornersSphere (E := ℂ) /-- The unit circle in `ℂ` is a Lie group. -/ -instance : LieGroup (𝓡 1) circle where +instance : LieGroup (𝓡 1) Circle where smooth_mul := by apply ContMDiff.codRestrict_sphere - let c : circle → ℂ := (↑) + let c : Circle → ℂ := (↑) have h₂ : ContMDiff (𝓘(ℝ, ℂ).prod 𝓘(ℝ, ℂ)) 𝓘(ℝ, ℂ) ∞ fun z : ℂ × ℂ => z.fst * z.snd := by rw [contMDiff_iff] exact ⟨continuous_mul, fun x y => contDiff_mul.contDiffOn⟩ @@ -583,11 +583,13 @@ instance : LieGroup (𝓡 1) circle where exact contMDiff_coe_sphere smooth_inv := by apply ContMDiff.codRestrict_sphere - simp only [← coe_inv_circle, coe_inv_circle_eq_conj] + simp only [← Circle.coe_inv, Circle.coe_inv_eq_conj] 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_expMapCircle : ContMDiff 𝓘(ℝ, ℝ) (𝓡 1) ∞ expMapCircle := +theorem contMDiff_circleExp : ContMDiff 𝓘(ℝ, ℝ) (𝓡 1) ∞ Circle.exp := (contDiff_exp.comp (contDiff_id.smul contDiff_const)).contMDiff.codRestrict_sphere _ -end circle +@[deprecated (since := "2024-07-25")] alias contMDiff_expMapCircle := contMDiff_circleExp + +end Circle diff --git a/Mathlib/Topology/Algebra/PontryaginDual.lean b/Mathlib/Topology/Algebra/PontryaginDual.lean index a07848b1564fa..afe0be4038552 100644 --- a/Mathlib/Topology/Algebra/PontryaginDual.lean +++ b/Mathlib/Topology/Algebra/PontryaginDual.lean @@ -10,14 +10,14 @@ import Mathlib.Topology.Algebra.ContinuousMonoidHom # Pontryagin dual This file defines the Pontryagin dual of a topological group. The Pontryagin dual of a topological -group `A` is the topological group of continuous homomorphisms `A →* circle` with the compact-open -topology. For example, `ℤ` and `circle` are Pontryagin duals of each other. This is an example of +group `A` is the topological group of continuous homomorphisms `A →* Circle` with the compact-open +topology. For example, `ℤ` and `Circle` are Pontryagin duals of each other. This is an example of Pontryagin duality, which states that a locally compact abelian topological group is canonically isomorphic to its double dual. ## Main definitions -* `PontryaginDual A`: The group of continuous homomorphisms `A →* circle`. +* `PontryaginDual A`: The group of continuous homomorphisms `A →* Circle`. -/ open Pointwise Function @@ -26,49 +26,49 @@ variable (A B C D E G : Type*) [Monoid A] [Monoid B] [Monoid C] [Monoid D] [Comm [TopologicalSpace A] [TopologicalSpace B] [TopologicalSpace C] [TopologicalSpace D] [TopologicalSpace E] [TopologicalSpace G] [TopologicalGroup E] [TopologicalGroup G] -/-- The Pontryagin dual of `A` is the group of continuous homomorphism `A → circle`. -/ +/-- The Pontryagin dual of `A` is the group of continuous homomorphism `A → Circle`. -/ def PontryaginDual := - ContinuousMonoidHom A circle + ContinuousMonoidHom A Circle -- Porting note: `deriving` doesn't derive these instances instance : TopologicalSpace (PontryaginDual A) := - (inferInstance : TopologicalSpace (ContinuousMonoidHom A circle)) + (inferInstance : TopologicalSpace (ContinuousMonoidHom A Circle)) instance : T2Space (PontryaginDual A) := - (inferInstance : T2Space (ContinuousMonoidHom A circle)) + (inferInstance : T2Space (ContinuousMonoidHom A Circle)) -- Porting note: instance is now noncomputable noncomputable instance : CommGroup (PontryaginDual A) := - (inferInstance : CommGroup (ContinuousMonoidHom A circle)) + (inferInstance : CommGroup (ContinuousMonoidHom A Circle)) instance : TopologicalGroup (PontryaginDual A) := - (inferInstance : TopologicalGroup (ContinuousMonoidHom A circle)) + (inferInstance : TopologicalGroup (ContinuousMonoidHom A Circle)) -- Porting note: instance is now noncomputable noncomputable instance : Inhabited (PontryaginDual A) := - (inferInstance : Inhabited (ContinuousMonoidHom A circle)) + (inferInstance : Inhabited (ContinuousMonoidHom A Circle)) instance [LocallyCompactSpace G] : LocallyCompactSpace (PontryaginDual G) := by - let Vn : ℕ → Set circle := - fun n ↦ expMapCircle '' { x | |x| < Real.pi / 2 ^ (n + 1)} + let Vn : ℕ → Set Circle := + fun n ↦ Circle.exp '' { x | |x| < Real.pi / 2 ^ (n + 1)} have hVn : ∀ n x, x ∈ Vn n ↔ |Complex.arg x| < Real.pi / 2 ^ (n + 1) := by - refine fun n x ↦ ⟨?_, fun hx ↦ ⟨Complex.arg x, hx, expMapCircle_arg x⟩⟩ + refine fun n x ↦ ⟨?_, fun hx ↦ ⟨Complex.arg x, hx, Circle.exp_arg x⟩⟩ rintro ⟨t, ht : |t| < _, rfl⟩ have ht' := ht.trans_le (div_le_self Real.pi_nonneg (one_le_pow_of_one_le one_le_two (n + 1))) - rwa [arg_expMapCircle (neg_lt_of_abs_lt ht') (lt_of_abs_lt ht').le] + rwa [Circle.arg_exp (neg_lt_of_abs_lt ht') (lt_of_abs_lt ht').le] refine ContinuousMonoidHom.locallyCompactSpace_of_hasBasis Vn ?_ ?_ · intro n x h1 h2 rw [hVn] at h1 h2 ⊢ - rwa [Submonoid.coe_mul, Complex.arg_mul (ne_zero_of_mem_circle x) (ne_zero_of_mem_circle x), + rwa [Circle.coe_mul, Complex.arg_mul x.coe_ne_zero x.coe_ne_zero, ← two_mul, abs_mul, abs_two, ← lt_div_iff' two_pos, div_div, ← pow_succ] at h2 apply Set.Ioo_subset_Ioc_self rw [← two_mul, Set.mem_Ioo, ← abs_lt, abs_mul, abs_two, ← lt_div_iff' two_pos] exact h1.trans_le (div_le_div_of_nonneg_left Real.pi_nonneg two_pos (le_self_pow one_le_two n.succ_ne_zero)) - · rw [← expMapCircle_zero, ← isLocalHomeomorph_expMapCircle.map_nhds_eq 0] + · rw [← Circle.exp_zero, ← isLocalHomeomorph_circleExp.map_nhds_eq 0] refine ((nhds_basis_zero_abs_sub_lt ℝ).to_hasBasis (fun x hx ↦ ⟨Nat.ceil (Real.pi / x), trivial, fun t ht ↦ ?_⟩) - fun k _ ↦ ⟨Real.pi / 2 ^ (k + 1), by positivity, le_rfl⟩).map expMapCircle + fun k _ ↦ ⟨Real.pi / 2 ^ (k + 1), by positivity, le_rfl⟩).map Circle.exp rw [Set.mem_setOf_eq] at ht ⊢ refine lt_of_lt_of_le ht ?_ rw [div_le_iff' (pow_pos two_pos _), ← div_le_iff hx] @@ -81,16 +81,16 @@ namespace PontryaginDual open ContinuousMonoidHom -instance : FunLike (PontryaginDual A) A circle := +instance : FunLike (PontryaginDual A) A Circle := ContinuousMonoidHom.funLike -noncomputable instance : ContinuousMonoidHomClass (PontryaginDual A) A circle := +noncomputable instance : ContinuousMonoidHomClass (PontryaginDual A) A Circle := ContinuousMonoidHom.ContinuousMonoidHomClass /-- `PontryaginDual` is a contravariant functor. -/ noncomputable def map (f : ContinuousMonoidHom A B) : ContinuousMonoidHom (PontryaginDual B) (PontryaginDual A) := - f.compLeft circle + f.compLeft Circle @[simp] theorem map_apply (f : ContinuousMonoidHom A B) (x : PontryaginDual B) (y : A) : diff --git a/test/TCSynth.lean b/test/TCSynth.lean index b29a76f58971b..391a3a3bd3101 100644 --- a/test/TCSynth.lean +++ b/test/TCSynth.lean @@ -54,7 +54,7 @@ section open Real in set_option synthInstance.maxHeartbeats 10000 in -example : expMapCircle (2 * π) = 1 := by simp +example : Circle.exp (2 * π) = 1 := by simp end From c074e33e291a55f9eaa48b6ceb2a04a0d728e8b1 Mon Sep 17 00:00:00 2001 From: grunweg Date: Tue, 13 Aug 2024 09:26:52 +0000 Subject: [PATCH 236/359] chore(Topology): remove more bare open Classical (#15671) The only remaining occurrences are in Topology/Algebra. --- Mathlib/Topology/Category/Compactum.lean | 4 +--- .../Category/Profinite/CofilteredLimit.lean | 10 +++++----- Mathlib/Topology/Connected/PathConnected.lean | 2 -- Mathlib/Topology/ContinuousFunction/Bounded.lean | 3 +-- Mathlib/Topology/ContinuousFunction/Compact.lean | 9 ++------- Mathlib/Topology/EMetricSpace/Basic.lean | 6 ++---- Mathlib/Topology/FiberBundle/Constructions.lean | 5 ----- Mathlib/Topology/Instances/EReal.lean | 3 +-- Mathlib/Topology/Instances/NNReal.lean | 9 +++------ Mathlib/Topology/Instances/Real.lean | 2 -- Mathlib/Topology/LocallyConstant/Basic.lean | 8 ++------ Mathlib/Topology/MetricSpace/CantorScheme.lean | 6 +----- Mathlib/Topology/MetricSpace/CauSeqFilter.lean | 6 +----- Mathlib/Topology/MetricSpace/Closeds.lean | 7 +------ Mathlib/Topology/MetricSpace/Contracting.lean | 2 -- Mathlib/Topology/MetricSpace/Dilation.lean | 5 ++--- Mathlib/Topology/MetricSpace/Gluing.lean | 3 +-- Mathlib/Topology/MetricSpace/PiNat.lean | 16 +++++++++++----- .../Topology/MetricSpace/ThickenedIndicator.lean | 6 +----- Mathlib/Topology/OmegaCompletePartialOrder.lean | 4 ---- Mathlib/Topology/Order/MonotoneConvergence.lean | 10 ++++------ Mathlib/Topology/UniformSpace/Compact.lean | 3 --- Mathlib/Topology/UniformSpace/Completion.lean | 3 +-- Mathlib/Topology/UniformSpace/Separation.lean | 2 +- Mathlib/Topology/VectorBundle/Basic.lean | 9 +++++++-- Mathlib/Topology/VectorBundle/Constructions.lean | 2 -- 26 files changed, 48 insertions(+), 97 deletions(-) diff --git a/Mathlib/Topology/Category/Compactum.lean b/Mathlib/Topology/Category/Compactum.lean index 0dab814bb83e6..7cd821aa6d9d6 100644 --- a/Mathlib/Topology/Category/Compactum.lean +++ b/Mathlib/Topology/Category/Compactum.lean @@ -72,9 +72,7 @@ We also add wrappers around structures which already exist. Here are the main on universe u open CategoryTheory Filter Ultrafilter TopologicalSpace CategoryTheory.Limits FiniteInter - -open scoped Classical -open Topology +open scoped Topology local notation "β" => ofTypeMonad Ultrafilter diff --git a/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean b/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean index 7e55e93416f41..acb77ad51c59d 100644 --- a/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean +++ b/Mathlib/Topology/Category/Profinite/CofilteredLimit.lean @@ -24,11 +24,7 @@ This file contains some theorems about cofiltered limits of profinite sets. namespace Profinite -open scoped Classical - -open CategoryTheory - -open CategoryTheory.Limits +open CategoryTheory Limits -- This was a global instance prior to #13170. We may experiment with removing it. attribute [local instance] ConcreteCategory.instFunLike @@ -84,6 +80,7 @@ theorem exists_isClopen_of_cofiltered {U : Set C.pt} (hC : IsLimit C) (hU : IsCl -- Since `J` is cofiltered, we can find a single `j0` dominating all the `j ∈ G`. -- Pulling back all of the sets from the previous step to `F.obj j0` and taking a union, -- we obtain a clopen set in `F.obj j0` which works. + classical obtain ⟨j0, hj0⟩ := IsCofiltered.inf_objs_exists (G.image j) let f : ∀ s ∈ G, j0 ⟶ j s := fun s hs => (hj0 (Finset.mem_image.mpr ⟨s, hs, rfl⟩)).some let W : S → Set (F.obj j0) := fun s => if hs : s ∈ G then F.map (f s hs) ⁻¹' V s else Set.univ @@ -114,6 +111,7 @@ theorem exists_locallyConstant_fin_two (hC : IsLimit C) (f : LocallyConstant C.p let U := f ⁻¹' {0} have hU : IsClopen U := f.isLocallyConstant.isClopen_fiber _ obtain ⟨j, V, hV, h⟩ := exists_isClopen_of_cofiltered C hC hU + classical use j, LocallyConstant.ofIsClopen hV apply LocallyConstant.locallyConstant_eq_of_fiber_zero_eq simp only [Fin.isValue, Functor.const_obj_obj, LocallyConstant.coe_comap, Set.preimage_comp, @@ -121,6 +119,7 @@ theorem exists_locallyConstant_fin_two (hC : IsLimit C) (f : LocallyConstant C.p -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [← h] +open Classical in theorem exists_locallyConstant_finite_aux {α : Type*} [Finite α] (hC : IsLimit C) (f : LocallyConstant C.pt α) : ∃ (j : J) (g : LocallyConstant (F.obj j) (α → Fin 2)), (f.map fun a b => if a = b then (0 : Fin 2) else 1) = g.comap (C.π.app _) := by @@ -159,6 +158,7 @@ theorem exists_locallyConstant_finite_nonempty {α : Type*} [Finite α] [Nonempt ∃ (j : J) (g : LocallyConstant (F.obj j) α), f = g.comap (C.π.app _) := by inhabit α obtain ⟨j, gg, h⟩ := exists_locallyConstant_finite_aux _ hC f + classical let ι : α → α → Fin 2 := fun a b => if a = b then 0 else 1 let σ : (α → Fin 2) → α := fun f => if h : ∃ a : α, ι a = f then h.choose else default refine ⟨j, gg.map σ, ?_⟩ diff --git a/Mathlib/Topology/Connected/PathConnected.lean b/Mathlib/Topology/Connected/PathConnected.lean index 81f6d23e10a49..07fea44bddd1e 100644 --- a/Mathlib/Topology/Connected/PathConnected.lean +++ b/Mathlib/Topology/Connected/PathConnected.lean @@ -58,10 +58,8 @@ This is used to define `Path.extend` that turns `γ : Path x y` into a continuou on `(-∞, 0]` and to `y` on `[1, +∞)`. -/ - noncomputable section -open scoped Classical open Topology Filter unitInterval Set Function variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] {x y z : X} {ι : Type*} diff --git a/Mathlib/Topology/ContinuousFunction/Bounded.lean b/Mathlib/Topology/ContinuousFunction/Bounded.lean index a76e6a27a8e44..fe002c06a38ed 100644 --- a/Mathlib/Topology/ContinuousFunction/Bounded.lean +++ b/Mathlib/Topology/ContinuousFunction/Bounded.lean @@ -19,10 +19,8 @@ the uniform distance. -/ - noncomputable section -open scoped Classical open Topology Bornology NNReal uniformity UniformConvergence open Set Filter Metric Function @@ -491,6 +489,7 @@ theorem arzela_ascoli₁ [CompactSpace β] (A : Set (α →ᵇ β)) (closed : Is -- `F : β → β`, `hF : ∀ (y : β), F y ∈ tβ ∧ dist y (F y) < ε₂` /- Associate to every function a discrete approximation, mapping each point in `tα` to a point in `tβ` close to its true image by the function. -/ + classical refine ⟨tα → tβ, by infer_instance, fun f a => ⟨F (f.1 a), (hF (f.1 a)).1⟩, ?_⟩ rintro ⟨f, hf⟩ ⟨g, hg⟩ f_eq_g -- If two functions have the same approximation, then they are within distance `ε` diff --git a/Mathlib/Topology/ContinuousFunction/Compact.lean b/Mathlib/Topology/ContinuousFunction/Compact.lean index 5257cc87a452a..77c4cda47593d 100644 --- a/Mathlib/Topology/ContinuousFunction/Compact.lean +++ b/Mathlib/Topology/ContinuousFunction/Compact.lean @@ -21,17 +21,11 @@ characterising these structures. If you need a lemma which is proved about `α →ᵇ β` but not for `C(α, β)` when `α` is compact, you should restate it here. You can also use `ContinuousMap.equivBoundedOfCompact` to move functions back and forth. - -/ noncomputable section -open scoped Classical -open Topology NNReal BoundedContinuousFunction - -open Set Filter Metric - -open BoundedContinuousFunction +open NNReal BoundedContinuousFunction Set Metric namespace ContinuousMap @@ -434,6 +428,7 @@ variable {E : Type*} [NormedAddCommGroup E] [CompleteSpace E] theorem summable_of_locally_summable_norm {ι : Type*} {F : ι → C(X, E)} (hF : ∀ K : Compacts X, Summable fun i => ‖(F i).restrict K‖) : Summable F := by + classical refine (ContinuousMap.exists_tendsto_compactOpen_iff_forall _).2 fun K hK => ?_ lift K to Compacts X using hK have A : ∀ s : Finset ι, restrict (↑K) (∑ i ∈ s, F i) = ∑ i ∈ s, restrict K (F i) := by diff --git a/Mathlib/Topology/EMetricSpace/Basic.lean b/Mathlib/Topology/EMetricSpace/Basic.lean index 93c6d661659c0..7845b2fb5b6dd 100644 --- a/Mathlib/Topology/EMetricSpace/Basic.lean +++ b/Mathlib/Topology/EMetricSpace/Basic.lean @@ -27,10 +27,8 @@ theory of `PseudoEMetricSpace`, where we don't require `edist x y = 0 → x = y` to `EMetricSpace` at the end. -/ - -open Set Filter Classical - -open scoped Uniformity Topology Filter NNReal ENNReal Pointwise +open Set Filter +open scoped Uniformity Topology NNReal ENNReal Pointwise universe u v w diff --git a/Mathlib/Topology/FiberBundle/Constructions.lean b/Mathlib/Topology/FiberBundle/Constructions.lean index 437065450fbd1..01fe742ebee19 100644 --- a/Mathlib/Topology/FiberBundle/Constructions.lean +++ b/Mathlib/Topology/FiberBundle/Constructions.lean @@ -25,15 +25,10 @@ fiber bundle, fibre bundle, fiberwise product, pullback -/ - open TopologicalSpace Filter Set Bundle -open scoped Classical -open Topology Bundle - /-! ### The trivial bundle -/ - namespace Bundle namespace Trivial diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index d699574ac88c4..9da4bbc6a3b9b 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -27,9 +27,8 @@ Most proofs are adapted from the corresponding proofs on `ℝ≥0∞`. noncomputable section -open scoped Classical open Set Filter Metric TopologicalSpace Topology -open scoped ENNReal NNReal Filter +open scoped ENNReal variable {α : Type*} [TopologicalSpace α] diff --git a/Mathlib/Topology/Instances/NNReal.lean b/Mathlib/Topology/Instances/NNReal.lean index 22e81af331ff4..e12bda18c19cc 100644 --- a/Mathlib/Topology/Instances/NNReal.lean +++ b/Mathlib/Topology/Instances/NNReal.lean @@ -48,11 +48,9 @@ a few of which rely on the fact that subtraction is continuous. -/ - noncomputable section open Set TopologicalSpace Metric Filter - open scoped Topology namespace NNReal @@ -185,11 +183,10 @@ theorem summable_mk {f : α → ℝ} (hf : ∀ n, 0 ≤ f n) : (@Summable ℝ≥0 _ _ _ fun n => ⟨f n, hf n⟩) ↔ Summable f := Iff.symm <| summable_coe (f := fun x => ⟨f x, hf x⟩) -open scoped Classical - @[norm_cast] -theorem coe_tsum {f : α → ℝ≥0} : ↑(∑' a, f a) = ∑' a, (f a : ℝ) := - if hf : Summable f then Eq.symm <| (hasSum_coe.2 <| hf.hasSum).tsum_eq +theorem coe_tsum {f : α → ℝ≥0} : ↑(∑' a, f a) = ∑' a, (f a : ℝ) := by + classical + exact if hf : Summable f then Eq.symm <| (hasSum_coe.2 <| hf.hasSum).tsum_eq else by simp [tsum_def, hf, mt summable_coe.1 hf] theorem coe_tsum_of_nonneg {f : α → ℝ} (hf₁ : ∀ n, 0 ≤ f n) : diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index 764564cb7bbfe..139a67e9753bb 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -17,10 +17,8 @@ import Mathlib.Topology.Order.Bornology # Topological properties of ℝ -/ - noncomputable section -open scoped Classical open Filter Int Metric Set TopologicalSpace Bornology open scoped Topology Uniformity Interval diff --git a/Mathlib/Topology/LocallyConstant/Basic.lean b/Mathlib/Topology/LocallyConstant/Basic.lean index cd4f5070a4e86..4e3ac23a0989d 100644 --- a/Mathlib/Topology/LocallyConstant/Basic.lean +++ b/Mathlib/Topology/LocallyConstant/Basic.lean @@ -19,15 +19,12 @@ This file sets up the theory of locally constant function from a topological spa * `LocallyConstant X Y` : the type of locally constant maps from `X` to `Y` * `LocallyConstant.map` : push-forward of locally constant maps * `LocallyConstant.comap` : pull-back of locally constant maps - -/ - variable {X Y Z α : Type*} [TopologicalSpace X] open Set Filter - -open Topology +open scoped Topology /-- A function between topological spaces is locally constant if the preimage of any set is open. -/ def IsLocallyConstant (f : X → Y) : Prop := @@ -434,8 +431,6 @@ section Indicator variable {R : Type*} [One R] {U : Set X} (f : LocallyConstant X R) -open scoped Classical - /-- Given a clopen set `U` and a locally constant function `f`, `LocallyConstant.mulIndicator` returns the locally constant function that is `f` on `U` and `1` otherwise. -/ @[to_additive (attr := simps) "Given a clopen set `U` and a locally constant function `f`, @@ -449,6 +444,7 @@ noncomputable def mulIndicator (hU : IsClopen U) : LocallyConstant X R where variable (a : X) +open Classical in @[to_additive] theorem mulIndicator_apply_eq_if (hU : IsClopen U) : mulIndicator f hU a = if a ∈ U then f a else 1 := diff --git a/Mathlib/Topology/MetricSpace/CantorScheme.lean b/Mathlib/Topology/MetricSpace/CantorScheme.lean index f576872f29ffe..da60c00142aad 100644 --- a/Mathlib/Topology/MetricSpace/CantorScheme.lean +++ b/Mathlib/Topology/MetricSpace/CantorScheme.lean @@ -38,13 +38,9 @@ scheme, cantor scheme, lusin scheme, approximation. -/ - namespace CantorScheme -open List Function Filter Set PiNat - -open scoped Classical -open Topology +open List Function Filter Set PiNat Topology variable {β α : Type*} (A : List β → Set α) diff --git a/Mathlib/Topology/MetricSpace/CauSeqFilter.lean b/Mathlib/Topology/MetricSpace/CauSeqFilter.lean index e7b3634c22f86..8bf22f3af496b 100644 --- a/Mathlib/Topology/MetricSpace/CauSeqFilter.lean +++ b/Mathlib/Topology/MetricSpace/CauSeqFilter.lean @@ -13,13 +13,9 @@ is complete in terms of `Cauchy` filter if and only if it is complete in terms of `CauSeq` Cauchy sequences. -/ - universe u v -open Set Filter - -open scoped Classical -open Topology +open Set Filter Topology variable {β : Type v} diff --git a/Mathlib/Topology/MetricSpace/Closeds.lean b/Mathlib/Topology/MetricSpace/Closeds.lean index 5f7dca32b13d8..0c09e5375464a 100644 --- a/Mathlib/Topology/MetricSpace/Closeds.lean +++ b/Mathlib/Topology/MetricSpace/Closeds.lean @@ -22,16 +22,11 @@ inherits a metric space structure from the Hausdorff distance, as the Hausdorff always finite in this context. -/ - noncomputable section -open scoped Classical -open Topology ENNReal - universe u -open scoped Classical -open Set Function TopologicalSpace Filter +open Set Function TopologicalSpace Filter Topology ENNReal namespace EMetric diff --git a/Mathlib/Topology/MetricSpace/Contracting.lean b/Mathlib/Topology/MetricSpace/Contracting.lean index 5831e19b36935..23dc6110e8e0c 100644 --- a/Mathlib/Topology/MetricSpace/Contracting.lean +++ b/Mathlib/Topology/MetricSpace/Contracting.lean @@ -28,8 +28,6 @@ of convergence, and some properties of the map sending a contracting map to its contracting map, fixed point, Banach fixed point theorem -/ - -open scoped Classical open NNReal Topology ENNReal Filter Function variable {α : Type*} diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index ffa01e46f437f..c68a245846375 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -48,12 +48,10 @@ needed. - [Marcel Berger, *Geometry*][berger1987] -/ - noncomputable section open Function Set Bornology - -open scoped Topology ENNReal NNReal Classical +open scoped Topology ENNReal NNReal section Defs @@ -126,6 +124,7 @@ theorem copy_eq_self (f : α →ᵈ β) {f' : α → β} (h : f' = f) : f.copy f variable [FunLike F α β] +open Classical in /-- The ratio of a dilation `f`. If the ratio is undefined (i.e., the distance between any two points in `α` is either zero or infinity), then we choose one as the ratio. -/ def ratio [DilationClass F α β] (f : F) : ℝ≥0 := diff --git a/Mathlib/Topology/MetricSpace/Gluing.lean b/Mathlib/Topology/MetricSpace/Gluing.lean index 41444fef866b6..9091ef227ff5c 100644 --- a/Mathlib/Topology/MetricSpace/Gluing.lean +++ b/Mathlib/Topology/MetricSpace/Gluing.lean @@ -294,8 +294,7 @@ namespace Sigma of two spaces. I.e., work with sigma types instead of sum types. -/ variable {ι : Type*} {E : ι → Type*} [∀ i, MetricSpace (E i)] -open scoped Classical - +open Classical in /-- Distance on a disjoint union. There are many (noncanonical) ways to put a distance compatible with each factor. We choose a construction that works for unbounded spaces, but requires basepoints, diff --git a/Mathlib/Topology/MetricSpace/PiNat.lean b/Mathlib/Topology/MetricSpace/PiNat.lean index 1c0e6daf3ca63..b001d3c706ab2 100644 --- a/Mathlib/Topology/MetricSpace/PiNat.lean +++ b/Mathlib/Topology/MetricSpace/PiNat.lean @@ -47,13 +47,9 @@ in general), and `ι` is countable. the uniformity is definitionally the product uniformity. Not registered as an instance. -/ - noncomputable section -open scoped Classical -open Topology Filter - -open TopologicalSpace Set Metric Filter Function +open Topology TopologicalSpace Set Metric Filter Function attribute [local simp] pow_le_pow_iff_right one_lt_two inv_le_inv zero_le_two zero_lt_two @@ -63,6 +59,7 @@ namespace PiNat /-! ### The firstDiff function -/ +open Classical in /-- In a product space `Π n, E n`, then `firstDiff x y` is the first index at which `x` and `y` differ. If `x = y`, then by convention we set `firstDiff x x = 0`. -/ irreducible_def firstDiff (x y : ∀ n, E n) : ℕ := @@ -71,6 +68,7 @@ irreducible_def firstDiff (x y : ∀ n, E n) : ℕ := theorem apply_firstDiff_ne {x y : ∀ n, E n} (h : x ≠ y) : x (firstDiff x y) ≠ y (firstDiff x y) := by rw [firstDiff_def, dif_pos h] + classical exact Nat.find_spec (ne_iff.1 h) theorem apply_eq_of_lt_firstDiff {x y : ∀ n, E n} {n : ℕ} (hn : n < firstDiff x y) : x n = y n := by @@ -81,6 +79,7 @@ theorem apply_eq_of_lt_firstDiff {x y : ∀ n, E n} {n : ℕ} (hn : n < firstDif · exact (not_lt_zero' hn).elim theorem firstDiff_comm (x y : ∀ n, E n) : firstDiff x y = firstDiff y x := by + classical simp only [firstDiff_def, ne_comm] theorem min_firstDiff_le (x y z : ∀ n, E n) (h : x ≠ z) : @@ -236,6 +235,7 @@ a `MetricSpace` instance, as other distances may be used on these spaces, but we local instances in this section. -/ +open Classical in /-- The distance function on a product space `Π n, E n`, given by `dist x y = (1/2)^n` where `n` is the first index at which `x` and `y` differ. -/ protected def dist : Dist (∀ n, E n) := @@ -249,6 +249,7 @@ theorem dist_eq_of_ne {x y : ∀ n, E n} (h : x ≠ y) : dist x y = (1 / 2 : ℝ protected theorem dist_self (x : ∀ n, E n) : dist x x = 0 := by simp [dist] protected theorem dist_comm (x y : ∀ n, E n) : dist x y = dist y x := by + classical simp [dist, @eq_comm _ x y, firstDiff_comm] protected theorem dist_nonneg (x y : ∀ n, E n) : 0 ≤ dist x y := by @@ -452,6 +453,7 @@ theorem exists_disjoint_cylinder {s : Set (∀ n, E n)} (hs : IsClosed s) {x : exact mem_cylinder_iff_dist_le.1 hy _ < infDist x s := hn +open Classical in /-- Given a point `x` in a product space `Π (n : ℕ), E n`, and `s` a subset of this space, then `shortestPrefixDiff x s` if the smallest `n` for which there is no element of `s` having the same prefix of length `n` as `x`. If there is no such `n`, then use `0` by convention. -/ @@ -462,6 +464,7 @@ theorem firstDiff_lt_shortestPrefixDiff {s : Set (∀ n, E n)} (hs : IsClosed s) (hx : x ∉ s) (hy : y ∈ s) : firstDiff x y < shortestPrefixDiff x s := by have A := exists_disjoint_cylinder hs hx rw [shortestPrefixDiff, dif_pos A] + classical have B := Nat.find_spec A contrapose! B rw [not_disjoint_iff_nonempty_inter] @@ -494,6 +497,7 @@ theorem inter_cylinder_longestPrefix_nonempty {s : Set (∀ n, E n)} (hs : IsClo have B : longestPrefix x s < shortestPrefixDiff x s := Nat.pred_lt (shortestPrefixDiff_pos hs hne hx).ne' rw [longestPrefix, shortestPrefixDiff, dif_pos A] at B ⊢ + classical obtain ⟨y, ys, hy⟩ : ∃ y : ∀ n : ℕ, E n, y ∈ s ∧ x ∈ cylinder y (Nat.find A - 1) := by simpa only [not_disjoint_iff, mem_cylinder_comm] using Nat.find_min A B refine ⟨y, ys, ?_⟩ @@ -552,6 +556,7 @@ theorem exists_lipschitz_retraction_of_isClosed {s : Set (∀ n, E n)} (hs : IsC length is `< n`, then it is also the longest prefix of `y`, and we get `f x = f y = z_w`. Otherwise, `f x` remains in the same `n`-cylinder as `x`. Similarly for `y`. Finally, `f x` and `f y` are again in the same `n`-cylinder, as desired. -/ + classical set f := fun x => if x ∈ s then x else (inter_cylinder_longestPrefix_nonempty hs hne x).some have fs : ∀ x ∈ s, f x = x := fun x xs => by simp [f, xs] refine ⟨f, fs, ?_, ?_⟩ @@ -827,6 +832,7 @@ protected def metricSpace : MetricSpace (∀ i, F i) where apply le_antisymm · simp only [le_iInf_iff, le_principal_iff] intro ε εpos + classical obtain ⟨K, hK⟩ : ∃ K : Finset ι, (∑' i : { j // j ∉ K }, (1 / 2 : ℝ) ^ encode (i : ι)) < ε / 2 := ((tendsto_order.1 (tendsto_tsum_compl_atTop_zero fun i : ι => (1 / 2 : ℝ) ^ encode i)).2 _ diff --git a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean index 715f0ae201e9a..c274fbbc8e376 100644 --- a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean +++ b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean @@ -33,11 +33,7 @@ members of the approximating sequence are nonnegative bounded continuous functio -/ - -open scoped Classical -open NNReal ENNReal Topology BoundedContinuousFunction - -open NNReal ENNReal Set Metric EMetric Filter +open NNReal ENNReal Topology BoundedContinuousFunction Set Metric EMetric Filter noncomputable section thickenedIndicator diff --git a/Mathlib/Topology/OmegaCompletePartialOrder.lean b/Mathlib/Topology/OmegaCompletePartialOrder.lean index 95303763f6bda..7f45468c25312 100644 --- a/Mathlib/Topology/OmegaCompletePartialOrder.lean +++ b/Mathlib/Topology/OmegaCompletePartialOrder.lean @@ -19,15 +19,11 @@ of continuity is equivalent to continuity in ωCPOs. -/ - open Set OmegaCompletePartialOrder -open scoped Classical - universe u -- "Scott", "ωSup" - namespace Scott /-- `x` is an `ω`-Sup of a chain `c` if it is the least upper bound of the range of `c`. -/ diff --git a/Mathlib/Topology/Order/MonotoneConvergence.lean b/Mathlib/Topology/Order/MonotoneConvergence.lean index 9a5a74631dced..320b7751381fa 100644 --- a/Mathlib/Topology/Order/MonotoneConvergence.lean +++ b/Mathlib/Topology/Order/MonotoneConvergence.lean @@ -25,11 +25,8 @@ then `f n ≤ a` for all `n`. monotone convergence -/ - open Filter Set Function - -open scoped Classical -open Filter Topology +open scoped Topology variable {α β : Type*} @@ -201,8 +198,9 @@ instance Pi.infConvergenceClass' {ι : Type*} [Preorder α] [TopologicalSpace α theorem tendsto_of_monotone {ι α : Type*} [Preorder ι] [TopologicalSpace α] [ConditionallyCompleteLinearOrder α] [OrderTopology α] {f : ι → α} (h_mono : Monotone f) : - Tendsto f atTop atTop ∨ ∃ l, Tendsto f atTop (𝓝 l) := - if H : BddAbove (range f) then Or.inr ⟨_, tendsto_atTop_ciSup h_mono H⟩ + Tendsto f atTop atTop ∨ ∃ l, Tendsto f atTop (𝓝 l) := by + classical + exact if H : BddAbove (range f) then Or.inr ⟨_, tendsto_atTop_ciSup h_mono H⟩ else Or.inl <| tendsto_atTop_atTop_of_monotone' h_mono H theorem tendsto_of_antitone {ι α : Type*} [Preorder ι] [TopologicalSpace α] diff --git a/Mathlib/Topology/UniformSpace/Compact.lean b/Mathlib/Topology/UniformSpace/Compact.lean index 0e371638327a8..6c6f0ab1644d7 100644 --- a/Mathlib/Topology/UniformSpace/Compact.lean +++ b/Mathlib/Topology/UniformSpace/Compact.lean @@ -33,8 +33,6 @@ loop. uniform space, uniform continuity, compact space -/ - -open scoped Classical open Uniformity Topology Filter UniformSpace Set variable {α β γ : Type*} [UniformSpace α] [UniformSpace β] @@ -43,7 +41,6 @@ variable {α β γ : Type*} [UniformSpace α] [UniformSpace β] ### Uniformity on compact spaces -/ - /-- On a compact uniform space, the topology determines the uniform structure, entourages are exactly the neighborhoods of the diagonal. -/ theorem nhdsSet_diagonal_eq_uniformity [CompactSpace α] : 𝓝ˢ (diagonal α) = 𝓤 α := by diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index c85211f14a68a..6e6085fcbee0d 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -40,14 +40,12 @@ This formalization is mostly based on From a slightly different perspective in order to reuse material in `Topology.UniformSpace.Basic`. -/ - noncomputable section open Filter Set universe u v w x -open scoped Classical open Uniformity Topology Filter /-- Space of Cauchy filters @@ -222,6 +220,7 @@ instance [h : Nonempty α] : Nonempty (CauchyFilter α) := section Extend +open Classical in /-- Extend a uniformly continuous function `α → β` to a function `CauchyFilter α → β`. Outputs junk when `f` is not uniformly continuous. -/ def extend (f : α → β) : CauchyFilter α → β := diff --git a/Mathlib/Topology/UniformSpace/Separation.lean b/Mathlib/Topology/UniformSpace/Separation.lean index 4513d502f08bc..9d11da63e1ac6 100644 --- a/Mathlib/Topology/UniformSpace/Separation.lean +++ b/Mathlib/Topology/UniformSpace/Separation.lean @@ -96,7 +96,6 @@ uniform space, separated space, Hausdorff space, separation quotient -/ open Filter Set Function Topology Uniformity UniformSpace -open scoped Classical noncomputable section @@ -248,6 +247,7 @@ theorem uniformContinuous_uncurry_lift₂ {f : α → β → γ} theorem comap_mk_uniformity : (𝓤 (SeparationQuotient α)).comap (Prod.map mk mk) = 𝓤 α := comap_map_mk_uniformity +open Classical in /-- Factoring functions to a separated space through the separation quotient. TODO: unify with `SeparationQuotient.lift`. -/ diff --git a/Mathlib/Topology/VectorBundle/Basic.lean b/Mathlib/Topology/VectorBundle/Basic.lean index c6c5091999c82..f7f9be5e65667 100644 --- a/Mathlib/Topology/VectorBundle/Basic.lean +++ b/Mathlib/Topology/VectorBundle/Basic.lean @@ -54,7 +54,6 @@ Vector bundle noncomputable section -open scoped Classical open Bundle Set open scoped Topology @@ -104,12 +103,14 @@ def linearEquivAt (e : Pretrivialization F (π F E)) [e.IsLinear R] (b : B) (hb map_add' v w := (e.linear R hb).map_add v w map_smul' c v := (e.linear R hb).map_smul c v +open Classical in /-- A fiberwise linear map equal to `e` on `e.baseSet`. -/ protected def linearMapAt (e : Pretrivialization F (π F E)) [e.IsLinear R] (b : B) : E b →ₗ[R] F := if hb : b ∈ e.baseSet then e.linearEquivAt R b hb else 0 variable {R} +open Classical in theorem coe_linearMapAt (e : Pretrivialization F (π F E)) [e.IsLinear R] (b : B) : ⇑(e.linearMapAt R b) = fun y => if b ∈ e.baseSet then (e ⟨b, y⟩).2 else 0 := by rw [Pretrivialization.linearMapAt] @@ -119,6 +120,7 @@ theorem coe_linearMapAt_of_mem (e : Pretrivialization F (π F E)) [e.IsLinear R] (hb : b ∈ e.baseSet) : ⇑(e.linearMapAt R b) = fun y => (e ⟨b, y⟩).2 := by simp_rw [coe_linearMapAt, if_pos hb] +open Classical in theorem linearMapAt_apply (e : Pretrivialization F (π F E)) [e.IsLinear R] {b : B} (y : E b) : e.linearMapAt R b y = if b ∈ e.baseSet then (e ⟨b, y⟩).2 else 0 := by rw [coe_linearMapAt] @@ -208,6 +210,7 @@ protected def linearMapAt (e : Trivialization F (π F E)) [e.IsLinear R] (b : B) variable {R} +open Classical in theorem coe_linearMapAt (e : Trivialization F (π F E)) [e.IsLinear R] (b : B) : ⇑(e.linearMapAt R b) = fun y => if b ∈ e.baseSet then (e ⟨b, y⟩).2 else 0 := e.toPretrivialization.coe_linearMapAt b @@ -216,6 +219,7 @@ theorem coe_linearMapAt_of_mem (e : Trivialization F (π F E)) [e.IsLinear R] {b (hb : b ∈ e.baseSet) : ⇑(e.linearMapAt R b) = fun y => (e ⟨b, y⟩).2 := by simp_rw [coe_linearMapAt, if_pos hb] +open Classical in theorem linearMapAt_apply (e : Trivialization F (π F E)) [e.IsLinear R] {b : B} (y : E b) : e.linearMapAt R b y = if b ∈ e.baseSet then (e ⟨b, y⟩).2 else 0 := by rw [coe_linearMapAt] @@ -238,7 +242,7 @@ theorem linearMapAt_symmₗ (e : Trivialization F (π F E)) [e.IsLinear R] {b : variable (R) - +open Classical in /-- A coordinate change function between two trivializations, as a continuous linear equivalence. Defined to be the identity when `b` does not lie in the base set of both trivializations. -/ def coordChangeL (e e' : Trivialization F (π F E)) [e.IsLinear R] [e'.IsLinear R] (b : B) : @@ -378,6 +382,7 @@ def continuousLinearMapAt (e : Trivialization F (π F E)) [e.IsLinear R] (b : B) cont := by dsimp rw [e.coe_linearMapAt b] + classical refine continuous_if_const _ (fun hb => ?_) fun _ => continuous_zero exact (e.continuousOn.comp_continuous (FiberBundle.totalSpaceMk_inducing F E b).continuous fun x => e.mem_source.mpr hb).snd } diff --git a/Mathlib/Topology/VectorBundle/Constructions.lean b/Mathlib/Topology/VectorBundle/Constructions.lean index da81f821ecc48..93bb83d0965f4 100644 --- a/Mathlib/Topology/VectorBundle/Constructions.lean +++ b/Mathlib/Topology/VectorBundle/Constructions.lean @@ -26,10 +26,8 @@ This file contains several standard constructions on vector bundles: Vector bundle, direct sum, pullback -/ - noncomputable section -open scoped Classical open Bundle Set FiberBundle /-! ### The trivial vector bundle -/ From 421a092f37065ad242eeb3a5602b31eaad70d4bd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 13 Aug 2024 09:26:53 +0000 Subject: [PATCH 237/359] chore(SetTheory/Game/PGame): simplify ofLists API (#15766) The equation `(ofLists L R).LeftMoves = ULift (Fin L.length)` (ditto for `RightMoves`) is entirely def-eq in Lean 4. As such, `toOfListsLeftMoves` doesn't solve the same problem that other `toXLeftMoves` equivalences do, which is provide this correspondence without having to tediously figure out how to rewrite the equality. We instead turn this into an abbreviation for `ULift.up` and `ULift.down`, which has the nice side effect of making `simp` much more capable of simplifying equations involving `ofLists`. --- Mathlib/SetTheory/Game/PGame.lean | 30 +++++++++++++++++------------- 1 file changed, 17 insertions(+), 13 deletions(-) diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index b7f067d54e9eb..c17af8b22e725 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -160,30 +160,34 @@ theorem leftMoves_ofLists (L R : List PGame) : (ofLists L R).LeftMoves = ULift ( theorem rightMoves_ofLists (L R : List PGame) : (ofLists L R).RightMoves = ULift (Fin R.length) := rfl -/-- Converts a number into a left move for `ofLists`. -/ -def toOfListsLeftMoves {L R : List PGame} : Fin L.length ≃ (ofLists L R).LeftMoves := - ((Equiv.cast (leftMoves_ofLists L R).symm).trans Equiv.ulift).symm +/-- Converts a number into a left move for `ofLists`. -/-- Converts a number into a right move for `ofLists`. -/ -def toOfListsRightMoves {L R : List PGame} : Fin R.length ≃ (ofLists L R).RightMoves := - ((Equiv.cast (rightMoves_ofLists L R).symm).trans Equiv.ulift).symm +This is just an abbreviation for `Equiv.ulift.symm` -/ +abbrev toOfListsLeftMoves {L R : List PGame} : Fin L.length ≃ (ofLists L R).LeftMoves := + Equiv.ulift.symm -theorem ofLists_moveLeft {L R : List PGame} (i : Fin L.length) : - (ofLists L R).moveLeft (toOfListsLeftMoves i) = L[i] := - rfl +/-- Converts a number into a right move for `ofLists`. + +This is just an abbreviation for `Equiv.ulift.symm` -/ +abbrev toOfListsRightMoves {L R : List PGame} : Fin R.length ≃ (ofLists L R).RightMoves := + Equiv.ulift.symm @[simp] theorem ofLists_moveLeft' {L R : List PGame} (i : (ofLists L R).LeftMoves) : - (ofLists L R).moveLeft i = L[toOfListsLeftMoves.symm i] := + (ofLists L R).moveLeft i = L[i.down.val] := rfl -theorem ofLists_moveRight {L R : List PGame} (i : Fin R.length) : - (ofLists L R).moveRight (toOfListsRightMoves i) = R[i] := +theorem ofLists_moveLeft {L R : List PGame} (i : Fin L.length) : + (ofLists L R).moveLeft (ULift.up i) = L[i] := rfl @[simp] theorem ofLists_moveRight' {L R : List PGame} (i : (ofLists L R).RightMoves) : - (ofLists L R).moveRight i = R[toOfListsRightMoves.symm i] := + (ofLists L R).moveRight i = R[i.down.val] := + rfl + +theorem ofLists_moveRight {L R : List PGame} (i : Fin R.length) : + (ofLists L R).moveRight (ULift.up i) = R[i] := rfl /-- A variant of `PGame.recOn` expressed in terms of `PGame.moveLeft` and `PGame.moveRight`. From 42db034ebc27f0f3796efe4036dd513e93dd948f Mon Sep 17 00:00:00 2001 From: grunweg Date: Tue, 13 Aug 2024 11:31:29 +0000 Subject: [PATCH 238/359] chore: use double backticks for explicit names, if possible (#15173) This PR is very likely not exhaustive: the next step to finding these more exhaustively would be to write an intelligent linter. (I found these using a simple text-based linter, which has still too many false positive). --- Mathlib/Lean/Meta/DiscrTree.lean | 2 +- Mathlib/Tactic/Abel.lean | 2 +- Mathlib/Tactic/DeriveFintype.lean | 2 +- Mathlib/Tactic/DeriveToExpr.lean | 2 +- Mathlib/Tactic/Widget/CongrM.lean | 2 +- Mathlib/Tactic/Widget/GCongr.lean | 2 +- Mathlib/Tactic/Widget/SelectInsertParamsClass.lean | 2 +- 7 files changed, 7 insertions(+), 7 deletions(-) diff --git a/Mathlib/Lean/Meta/DiscrTree.lean b/Mathlib/Lean/Meta/DiscrTree.lean index ccd004b7f56ca..04468c7d9eba0 100644 --- a/Mathlib/Lean/Meta/DiscrTree.lean +++ b/Mathlib/Lean/Meta/DiscrTree.lean @@ -44,6 +44,6 @@ Check if a `keys : Array DiscTree.Key` is "specific", i.e. something other than `[*]` or `[=, *, *, *]`. -/ def keysSpecific (keys : Array DiscrTree.Key) : Bool := - keys != #[Key.star] && keys != #[Key.const `Eq 3, Key.star, Key.star, Key.star] + keys != #[Key.star] && keys != #[Key.const ``Eq 3, Key.star, Key.star, Key.star] end Lean.Meta.DiscrTree diff --git a/Mathlib/Tactic/Abel.lean b/Mathlib/Tactic/Abel.lean index decb801e22837..105a5e5f2c9bb 100644 --- a/Mathlib/Tactic/Abel.lean +++ b/Mathlib/Tactic/Abel.lean @@ -340,7 +340,7 @@ partial def eval (e : Expr) : M (NormalExpr × Expr) := do let (e₁, p₁) ← eval e let (e₂, p₂) ← evalNeg e₁ return (e₂, ← iapp `Mathlib.Tactic.Abel.subst_into_neg #[e, e₁, e₂, p₁, p₂]) - | (`AddMonoid.nsmul, #[_, _, e₁, e₂]) => do + | (``AddMonoid.nsmul, #[_, _, e₁, e₂]) => do let n ← if (← read).isGroup then mkAppM ``Int.ofNat #[e₁] else pure e₁ let (e', p) ← eval <| ← iapp ``smul #[n, e₂] return (e', ← iapp ``unfold_smul #[e₁, e₂, e', p]) diff --git a/Mathlib/Tactic/DeriveFintype.lean b/Mathlib/Tactic/DeriveFintype.lean index d475463b92670..694e0c7bd89fe 100644 --- a/Mathlib/Tactic/DeriveFintype.lean +++ b/Mathlib/Tactic/DeriveFintype.lean @@ -188,7 +188,7 @@ def mkFintypeInstanceHandler (declNames : Array Name) : CommandElabM Bool := do mkFintype declName initialize - registerDerivingHandler `Fintype mkFintypeInstanceHandler + registerDerivingHandler ``Fintype mkFintypeInstanceHandler registerTraceClass `Elab.Deriving.fintype end Mathlib.Deriving.Fintype diff --git a/Mathlib/Tactic/DeriveToExpr.lean b/Mathlib/Tactic/DeriveToExpr.lean index 7dccf4d691c16..a7b7b1a781888 100644 --- a/Mathlib/Tactic/DeriveToExpr.lean +++ b/Mathlib/Tactic/DeriveToExpr.lean @@ -223,7 +223,7 @@ def mkToExprInstanceHandler (declNames : Array Name) : CommandElabM Bool := do return false initialize - registerDerivingHandler `Lean.ToExpr mkToExprInstanceHandler + registerDerivingHandler ``Lean.ToExpr mkToExprInstanceHandler registerTraceClass `Elab.Deriving.toExpr end Mathlib.Deriving.ToExpr diff --git a/Mathlib/Tactic/Widget/CongrM.lean b/Mathlib/Tactic/Widget/CongrM.lean index c82f6139d99af..9c259c3daaf30 100644 --- a/Mathlib/Tactic/Widget/CongrM.lean +++ b/Mathlib/Tactic/Widget/CongrM.lean @@ -23,7 +23,7 @@ open Lean Meta Server ProofWidgets def makeCongrMString (pos : Array Lean.SubExpr.GoalsLocation) (goalType : Expr) (_ : SelectInsertParams) : MetaM (String × String × Option (String.Pos × String.Pos)) := do let subexprPos := getGoalLocations pos - unless goalType.isAppOf `Eq || goalType.isAppOf `Iff do + unless goalType.isAppOf ``Eq || goalType.isAppOf ``Iff do throwError "The goal must be an equality or iff." let mut goalTypeWithMetaVars := goalType for pos in subexprPos do diff --git a/Mathlib/Tactic/Widget/GCongr.lean b/Mathlib/Tactic/Widget/GCongr.lean index 20afd69a7cc8b..bd389f9f33d89 100644 --- a/Mathlib/Tactic/Widget/GCongr.lean +++ b/Mathlib/Tactic/Widget/GCongr.lean @@ -19,7 +19,7 @@ open Lean Meta Server ProofWidgets def makeGCongrString (pos : Array Lean.SubExpr.GoalsLocation) (goalType : Expr) (_ : SelectInsertParams) : MetaM (String × String × Option (String.Pos × String.Pos)) := do let subexprPos := getGoalLocations pos -unless goalType.isAppOf `LE.le || goalType.isAppOf `LT.lt || goalType.isAppOf `Int.ModEq do +unless goalType.isAppOf ``LE.le || goalType.isAppOf ``LT.lt || goalType.isAppOf `Int.ModEq do panic! "The goal must be a ≤ or < or ≡." let mut goalTypeWithMetaVars := goalType for pos in subexprPos do diff --git a/Mathlib/Tactic/Widget/SelectInsertParamsClass.lean b/Mathlib/Tactic/Widget/SelectInsertParamsClass.lean index 611b7ea94a05d..127473b69b56c 100644 --- a/Mathlib/Tactic/Widget/SelectInsertParamsClass.lean +++ b/Mathlib/Tactic/Widget/SelectInsertParamsClass.lean @@ -43,5 +43,5 @@ def mkSelectInsertParamsInstanceHandler (declNames : Array Name) : CommandElabM else return false -initialize registerDerivingHandler `SelectInsertParamsClass mkSelectInsertParamsInstanceHandler +initialize registerDerivingHandler ``SelectInsertParamsClass mkSelectInsertParamsInstanceHandler end Lean.Elab From 2ea5e248d072864d1fcde45c6b98a06a77a0dbbe Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 13 Aug 2024 13:38:58 +0000 Subject: [PATCH 239/359] chore: backports from nightly-testing-2024-08-12 (#15756) --- Cache/Hashing.lean | 8 ++++---- Cache/IO.lean | 6 +++--- Mathlib/Analysis/BoundedVariation.lean | 3 ++- Mathlib/Analysis/BoxIntegral/Partition/Filter.lean | 3 +-- .../CategoryTheory/Limits/Shapes/Countable.lean | 3 ++- Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean | 5 ++++- Mathlib/LinearAlgebra/Pi.lean | 2 +- Mathlib/RingTheory/Coprime/Lemmas.lean | 5 +++-- Mathlib/RingTheory/Localization/Basic.lean | 4 +--- Mathlib/Tactic/FunProp/RefinedDiscrTree.lean | 14 +++++++------- Mathlib/Tactic/FunProp/Theorems.lean | 6 +++--- Mathlib/Topology/Instances/RatLemmas.lean | 3 ++- 12 files changed, 33 insertions(+), 29 deletions(-) diff --git a/Cache/Hashing.lean b/Cache/Hashing.lean index e0e878935586a..f50466b1c1980 100644 --- a/Cache/Hashing.lean +++ b/Cache/Hashing.lean @@ -14,14 +14,14 @@ open System IO structure HashMemo where rootHash : UInt64 - depsMap : Lean.HashMap FilePath (Array FilePath) := {} - cache : Lean.HashMap FilePath (Option UInt64) := {} + depsMap : Std.HashMap FilePath (Array FilePath) := {} + cache : Std.HashMap FilePath (Option UInt64) := {} hashMap : HashMap := {} deriving Inhabited partial def insertDeps (hashMap : HashMap) (path : FilePath) (hashMemo : HashMemo) : HashMap := if hashMap.contains path then hashMap else - match (hashMemo.depsMap.find? path, hashMemo.hashMap.find? path) with + match (hashMemo.depsMap[path]?, hashMemo.hashMap[path]?) with | (some deps, some hash) => deps.foldl (insertDeps · · hashMemo) (hashMap.insert path hash) | _ => hashMap @@ -85,7 +85,7 @@ Computes the hash of a file, which mixes: * The hashes of the imported files that are part of `Mathlib` -/ partial def getFileHash (filePath : FilePath) : HashM <| Option UInt64 := do - match (← get).cache.find? filePath with + match (← get).cache[filePath]? with | some hash? => return hash? | none => let fixedPath := (← IO.getPackageDir filePath) / filePath diff --git a/Cache/IO.lean b/Cache/IO.lean index 7cb14dd575b26..850ce0a414f8c 100644 --- a/Cache/IO.lean +++ b/Cache/IO.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino -/ -import Lean.Data.HashMap +import Std.Data.HashMap import Lean.Data.RBMap import Lean.Data.RBTree import Lean.Data.Json.Printer @@ -256,7 +256,7 @@ partial def getFilesWithExtension (← fp.readDir).foldlM (fun acc dir => getFilesWithExtension dir.path extension acc) acc else return if fp.extension == some extension then acc.push fp else acc -abbrev HashMap := Lean.HashMap FilePath UInt64 +abbrev HashMap := Std.HashMap FilePath UInt64 namespace HashMap @@ -383,7 +383,7 @@ file) regarding the files with specified paths. -/ def lookup (hashMap : HashMap) (paths : List FilePath) : IO Unit := do let mut err := false for path in paths do - let some hash := hashMap.find? path | err := true + let some hash := hashMap[path]? | err := true let ltar := CACHEDIR / hash.asLTar IO.println s!"{path}: {ltar}" for line in (← runCmd (← getLeanTar) #["-k", ltar.toString]).splitOn "\n" |>.dropLast do diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index c9d1ff9e06ab5..fb2d9051dd7aa 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -640,7 +640,8 @@ protected theorem add {f : α → E} {s : Set α} (hf : LocallyBoundedVariationO variationOnFromTo.eq_of_le f s (xy.trans yz), ← ENNReal.toReal_add (hf x y xs ys) (hf y z ys zs), eVariationOn.Icc_add_Icc f xy yz ys] -protected theorem edist_zero_of_eq_zero {f : α → E} {s : Set α} (hf : LocallyBoundedVariationOn f s) +variable {f s} in +protected theorem edist_zero_of_eq_zero (hf : LocallyBoundedVariationOn f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) (h : variationOnFromTo f s a b = 0) : edist (f a) (f b) = 0 := by wlog h' : a ≤ b diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean index 40711ef598fae..caa6545c2bab8 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean @@ -343,14 +343,13 @@ theorem MemBaseSet.mono (h : l₁ ≤ l₂) (hc : c₁ ≤ c₂) hπ.mono' I h hc fun J _ => hr _ <| π.tag_mem_Icc J theorem MemBaseSet.exists_common_compl - {l : IntegrationParams} (h₁ : l.MemBaseSet I c₁ r₁ π₁) (h₂ : l.MemBaseSet I c₂ r₂ π₂) (hU : π₁.iUnion = π₂.iUnion) : ∃ π : Prepartition I, π.iUnion = ↑I \ π₁.iUnion ∧ (l.bDistortion → π.distortion ≤ c₁) ∧ (l.bDistortion → π.distortion ≤ c₂) := by wlog hc : c₁ ≤ c₂ with H · simpa [hU, _root_.and_comm] using - @H _ _ I c₂ c₁ r₂ r₁ π₂ π₁ _ h₂ h₁ hU.symm (le_of_not_le hc) + @H _ _ I c₂ c₁ l r₂ r₁ π₂ π₁ h₂ h₁ hU.symm (le_of_not_le hc) by_cases hD : (l.bDistortion : Prop) · rcases h₁.4 hD with ⟨π, hπU, hπc⟩ exact ⟨π, hπU, fun _ => hπc, fun _ => hπc.trans hc⟩ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean b/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean index d09a0b8fba24e..4893ab194160b 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean @@ -114,8 +114,9 @@ instance sequentialFunctor_initial : (sequentialFunctor J).Initial where 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'] + clear! C wlog h : (unop i.left) ≤ (unop j.left) - · exact or_comm.1 (this (C := C) J d n g inferInstance j i (le_of_lt (not_le.mp h))) + · exact or_comm.1 (this J d n g inferInstance j i (le_of_lt (not_le.mp h))) · right exact ⟨CostructuredArrow.homMk (homOfLE h).op rfl⟩ diff --git a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean index 25f5e0bd5dd28..74f30e36f4cda 100644 --- a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean @@ -86,6 +86,7 @@ variable {R} theorem ι_sq_zero (m : M) : ι R m * ι R m = 0 := (CliffordAlgebra.ι_sq_scalar _ m).trans <| map_zero _ +section variable {A : Type*} [Semiring A] [Algebra R A] -- @[simp] -- Porting note (#10618): simp can prove this @@ -255,6 +256,8 @@ theorem ι_mul_prod_list {n : ℕ} (f : Fin n → M) (i : Fin n) : refine (eq_zero_iff_eq_zero_of_add_eq_zero ?_).mp hn rw [← add_mul, ι_add_mul_swap, zero_mul] +end + variable (R) /-- The product of `n` terms of the form `ι R m` is an alternating map. @@ -268,7 +271,7 @@ def ιMulti (n : ℕ) : M [⋀^Fin n]→ₗ[R] ExteriorAlgebra R M := dsimp [F] clear F wlog h : x < y - · exact this R (A := A) n f y x hfxy.symm hxy.symm (hxy.lt_or_lt.resolve_left h) + · exact this R n f y x hfxy.symm hxy.symm (hxy.lt_or_lt.resolve_left h) clear hxy induction' n with n hn · exact x.elim0 diff --git a/Mathlib/LinearAlgebra/Pi.lean b/Mathlib/LinearAlgebra/Pi.lean index 002a80ea59395..d1342d034d19e 100644 --- a/Mathlib/LinearAlgebra/Pi.lean +++ b/Mathlib/LinearAlgebra/Pi.lean @@ -197,7 +197,7 @@ theorem disjoint_single_single (I J : Set ι) (h : Disjoint I J) : /-- The linear equivalence between linear functions on a finite product of modules and families of functions on these modules. See note [bundled maps over different rings]. -/ @[simps symm_apply] -def lsum (S) [AddCommMonoid M] [Module R M] [Fintype ι] [DecidableEq ι] [Semiring S] [Module S M] +def lsum (S) [AddCommMonoid M] [Module R M] [Fintype ι] [Semiring S] [Module S M] [SMulCommClass R S M] : ((i : ι) → φ i →ₗ[R] M) ≃ₗ[S] ((i : ι) → φ i) →ₗ[R] M where toFun f := ∑ i : ι, (f i).comp (proj i) invFun f i := f.comp (single R φ i) diff --git a/Mathlib/RingTheory/Coprime/Lemmas.lean b/Mathlib/RingTheory/Coprime/Lemmas.lean index 37284d7b39900..0034b62252647 100644 --- a/Mathlib/RingTheory/Coprime/Lemmas.lean +++ b/Mathlib/RingTheory/Coprime/Lemmas.lean @@ -151,8 +151,9 @@ theorem exists_sum_eq_one_iff_pairwise_coprime [DecidableEq I] (h : t.Nonempty) use fun i ↦ if i = a then u else v * μ i have hμ' : (∑ i ∈ t, v * ((μ i * ∏ j ∈ t \ {i}, s j) * s a)) = v * s a := by rw [← mul_sum, ← sum_mul, hμ, one_mul] - rw [sum_cons, cons_eq_insert, sdiff_singleton_eq_erase, erase_insert hat, if_pos rfl, - ← huv, ← hμ', sum_congr rfl] + rw [sum_cons, cons_eq_insert, sdiff_singleton_eq_erase, erase_insert hat] + simp only [↓reduceIte, ite_mul] + rw [← huv, ← hμ', sum_congr rfl] intro x hx rw [mul_assoc, if_neg fun ha : x = a ↦ hat (ha.casesOn hx)] rw [mul_assoc] diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index b9d1f5ee13512..ec5333ad7be07 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -1098,9 +1098,7 @@ section variable [Algebra Rₘ Sₘ] [Algebra R Sₘ] [IsScalarTower R Rₘ Sₘ] [IsScalarTower R S Sₘ] variable (S Rₘ Sₘ) -theorem IsLocalization.map_units_map_submonoid - (y : M) : - IsUnit (algebraMap R Sₘ y) := by +theorem IsLocalization.map_units_map_submonoid (y : M) : IsUnit (algebraMap R Sₘ y) := by rw [IsScalarTower.algebraMap_apply _ S] exact IsLocalization.map_units Sₘ ⟨algebraMap R S y, Algebra.mem_algebraMapSubmonoid_of_mem y⟩ diff --git a/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean b/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean index 0667d7a24122c..64e4a3bbc7639 100644 --- a/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean +++ b/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean @@ -333,7 +333,7 @@ partial def DTExpr.size : DTExpr → Nat private def DTExpr.eqv (a b : DTExpr) : Bool := (go a b).run' {} where - go (a b : DTExpr) : StateM (HashMap MVarId MVarId) Bool := + go (a b : DTExpr) : StateM (Std.HashMap MVarId MVarId) Bool := match a, b with | .opaque , .opaque => pure true | .const n₁ as₁ , .const n₂ as₂ => pure (n₁ == n₂) <&&> goArray as₁ as₂ @@ -346,12 +346,12 @@ where | .proj n₁ i₁ a₁ as₁, .proj n₂ i₂ a₂ as₂ => pure (n₁ == n₂ && i₁ == i₂) <&&> go a₁ a₂ <&&> goArray as₁ as₂ | .star none , .star none => pure true - | .star (some id₁) , .star (some id₂) => modifyGet fun map => match map.find? id₁ with + | .star (some id₁) , .star (some id₂) => modifyGet fun map => match map[id₁]? with | some id => (id == id₂, map) | none => (true, map.insert id₁ id₂) | _ , _ => return false - goArray (as bs : Array DTExpr) : StateM (HashMap MVarId MVarId) Bool := do + goArray (as bs : Array DTExpr) : StateM (Std.HashMap MVarId MVarId) Bool := do if h : as.size = bs.size then for g : i in [:as.size] do unless ← go as[i] (bs[i]'(h ▸ g.2)) do @@ -977,9 +977,9 @@ private structure State where /-- Score representing how good the match is. -/ score : Nat := 0 /-- Metavariable assignments for the `Key.star` patterns in the `RefinedDiscrTree`. -/ - starAssignments : HashMap Nat DTExpr := {} + starAssignments : Std.HashMap Nat DTExpr := {} /-- Metavariable assignments for the `Expr.mvar` in the expression. -/ - mvarAssignments : HashMap MVarId (Array Key) := {} + mvarAssignments : Std.HashMap MVarId (Array Key) := {} private abbrev M := ReaderT Context $ StateListM State @@ -1000,7 +1000,7 @@ private def insertStarAssignment (n : Nat) (e : DTExpr) : M Unit := /-- Log a metavariable assignment in the `State`. -/ private def assignMVar (mvarId : MVarId) (e : Array Key) : M Unit := do let { mvarAssignments, .. } ← get - match mvarAssignments.find? mvarId with + match mvarAssignments[mvarId]? with | some e' => guard (e == e') | none => modify fun s => { s with mvarAssignments := s.mvarAssignments.insert mvarId e } @@ -1033,7 +1033,7 @@ def matchTreeStars (e : DTExpr) (t : Trie α) : M (Trie α) := do so this loops through all of them. -/ for (k, c) in t.children! do let .star i := k | break - if let some assignment := starAssignments.find? i then + if let some assignment := starAssignments[i]? then if e == assignment then result := (incrementScore e.size *> pure c) <|> result else diff --git a/Mathlib/Tactic/FunProp/Theorems.lean b/Mathlib/Tactic/FunProp/Theorems.lean index 4faaad35f30b2..26571a80035b4 100644 --- a/Mathlib/Tactic/FunProp/Theorems.lean +++ b/Mathlib/Tactic/FunProp/Theorems.lean @@ -97,7 +97,7 @@ structure LambdaTheorem where /-- Collection of lambda theorems -/ structure LambdaTheorems where /-- map: function property name × theorem type → lambda theorem -/ - theorems : HashMap (Name × LambdaTheoremType) (Array LambdaTheorem) := {} + theorems : Std.HashMap (Name × LambdaTheoremType) (Array LambdaTheorem) := {} deriving Inhabited @@ -115,14 +115,14 @@ initialize lambdaTheoremsExt : LambdaTheoremsExt ← initial := {} addEntry := fun d e => {d with theorems := - let es := d.theorems.findD (e.funPropName, e.thmArgs.type) #[] + let es := d.theorems.getD (e.funPropName, e.thmArgs.type) #[] d.theorems.insert (e.funPropName, e.thmArgs.type) (es.push e)} } /-- Get lambda theorems for particular function property `funPropName`. -/ def getLambdaTheorems (funPropName : Name) (type : LambdaTheoremType) : CoreM (Array LambdaTheorem) := do - return (lambdaTheoremsExt.getState (← getEnv)).theorems.findD (funPropName,type) #[] + return (lambdaTheoremsExt.getState (← getEnv)).theorems.getD (funPropName,type) #[] -------------------------------------------------------------------------------- diff --git a/Mathlib/Topology/Instances/RatLemmas.lean b/Mathlib/Topology/Instances/RatLemmas.lean index 5841b1bc67e5b..b85ff2b228d17 100644 --- a/Mathlib/Topology/Instances/RatLemmas.lean +++ b/Mathlib/Topology/Instances/RatLemmas.lean @@ -71,10 +71,11 @@ theorem not_secondCountableTopology_opc : ¬SecondCountableTopology ℚ∞ := by exact not_firstCountableTopology_opc inferInstance instance : TotallyDisconnectedSpace ℚ := by + clear p q s t refine ⟨fun s hsu hs x hx y hy => ?_⟩; clear hsu by_contra! H : x ≠ y wlog hlt : x < y - · apply this s hs y hy x hx H.symm <| H.lt_or_lt.resolve_left hlt <;> assumption + · apply this s hs y hy x hx H.symm <| H.lt_or_lt.resolve_left hlt rcases exists_irrational_btwn (Rat.cast_lt.2 hlt) with ⟨z, hz, hxz, hzy⟩ have := hs.image _ continuous_coe_real.continuousOn rw [isPreconnected_iff_ordConnected] at this From ae4c62fa29b46dd14829a95cc52c5867cb304d3e Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 13 Aug 2024 14:20:40 +0000 Subject: [PATCH 240/359] =?UTF-8?q?chore:=20move=20`Analysis.CStarAlgebra.?= =?UTF-8?q?{Module=20=E2=87=92=20Module.Defs}`=20(#15761)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit we will shortly be adding more files to this folder. --- Mathlib.lean | 2 +- Mathlib/Analysis/CStarAlgebra/{Module.lean => Module/Defs.lean} | 0 2 files changed, 1 insertion(+), 1 deletion(-) rename Mathlib/Analysis/CStarAlgebra/{Module.lean => Module/Defs.lean} (100%) diff --git a/Mathlib.lean b/Mathlib.lean index bd642926af4a3..fc8bbfa934733 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -920,7 +920,7 @@ import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Unitary import Mathlib.Analysis.CStarAlgebra.Exponential import Mathlib.Analysis.CStarAlgebra.GelfandDuality import Mathlib.Analysis.CStarAlgebra.Matrix -import Mathlib.Analysis.CStarAlgebra.Module +import Mathlib.Analysis.CStarAlgebra.Module.Defs import Mathlib.Analysis.CStarAlgebra.Multiplier import Mathlib.Analysis.CStarAlgebra.Spectrum import Mathlib.Analysis.CStarAlgebra.Unitization diff --git a/Mathlib/Analysis/CStarAlgebra/Module.lean b/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean similarity index 100% rename from Mathlib/Analysis/CStarAlgebra/Module.lean rename to Mathlib/Analysis/CStarAlgebra/Module/Defs.lean From 7034a1cbaa30a3bc3c272157be910338a19f9d6b Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Tue, 13 Aug 2024 14:39:33 +0000 Subject: [PATCH 241/359] feat(CategoryTheory): Being representably flat is closed under isos (#15772) --- Mathlib/CategoryTheory/Functor/Flat.lean | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/Mathlib/CategoryTheory/Functor/Flat.lean b/Mathlib/CategoryTheory/Functor/Flat.lean index b1f602a3779b9..81deab9624ac3 100644 --- a/Mathlib/CategoryTheory/Functor/Flat.lean +++ b/Mathlib/CategoryTheory/Functor/Flat.lean @@ -64,13 +64,14 @@ class RepresentablyFlat (F : C ⥤ D) : Prop where attribute [instance] RepresentablyFlat.cofiltered -instance RepresentablyFlat.of_isRightAdjoint (F : C ⥤ D) [F.IsRightAdjoint] : - RepresentablyFlat F where +variable (F : C ⥤ D) + +instance RepresentablyFlat.of_isRightAdjoint [F.IsRightAdjoint] : RepresentablyFlat F where cofiltered _ := IsCofiltered.of_isInitial _ (mkInitialOfLeftAdjoint _ (.ofIsRightAdjoint F) _) theorem RepresentablyFlat.id : RepresentablyFlat (𝟭 C) := inferInstance -instance RepresentablyFlat.comp (F : C ⥤ D) (G : D ⥤ E) [RepresentablyFlat F] +instance RepresentablyFlat.comp (G : D ⥤ E) [RepresentablyFlat F] [RepresentablyFlat G] : RepresentablyFlat (F ⋙ G) := by refine ⟨fun X => IsCofiltered.of_cone_nonempty.{0} _ (fun {J} _ _ H => ?_)⟩ obtain ⟨c₁⟩ := IsCofiltered.cone_nonempty (H ⋙ StructuredArrow.pre X F G) @@ -83,6 +84,13 @@ instance RepresentablyFlat.comp (F : C ⥤ D) (G : D ⥤ E) [RepresentablyFlat F ⟨fun j => StructuredArrow.homMk (c₂.π.app j).right (by simp [← G.map_comp, (c₂.π.app j).w]), fun j j' f => by simpa using (c₂.w f).symm⟩⟩⟩ +variable {F} + +/-- Being a representably flat functor is closed under natural isomorphisms. -/ +theorem RepresentablyFlat.of_iso [RepresentablyFlat F] {G : C ⥤ D} (α : F ≅ G) : + RepresentablyFlat G where + cofiltered _ := IsCofiltered.of_equivalence (StructuredArrow.mapNatIso α) + end RepresentablyFlat section HasLimit From 6fdfc95972f2cef56de3675c04e01a42c5d157d7 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 13 Aug 2024 17:15:38 +0000 Subject: [PATCH 242/359] feat: results on the spectrum of elements in closed subalgebras of Banach algebras (#14925) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main result is well-known. Let `S` be a closed subalgebra of a Banach algebra `A`. If `x : S`, then the spectrum of `x` (in `S`) is the spectrum of `↑x : A` (in `A`) along with whichever, necessarily bounded, components of the complement (of the `A` spectrum) it happens to intersect. Along the way we prove a few other useful results. For example, the boundary of the spectrum of `x` (in `S`) is a subset of the boundary of the spectrum of `↑x : A` (in `A`). --- Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 109 ++++++++++++++++++ 1 file changed, 109 insertions(+) diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index b09a207deb0d5..560f636a6afa2 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -569,6 +569,115 @@ end CharacterSpace end WeakDual +section BoundarySpectrum + +local notation "σ" => spectrum + +variable {𝕜 A SA : Type*} [NormedField 𝕜] [NormedRing A] [NormedAlgebra 𝕜 A] [CompleteSpace A] +variable [SetLike SA A] [SubringClass SA A] [instSMulMem : SMulMemClass SA 𝕜 A] +variable (S : SA) [hS : IsClosed (S : Set A)] (x : S) + +open Topology Filter + +open SubalgebraClass in +include instSMulMem in +/-- Let `S` be a closed subalgebra of a Banach algebra `A`. If `a : S` is invertible in `A`, +and for all `x : S` sufficiently close to `a` within some filter `l`, `x` is invertible in `S`, +then `a` is invertible in `S` as well. -/ +lemma _root_.Subalgebra.isUnit_of_isUnit_val_of_eventually {l : Filter S} {a : S} + (ha : IsUnit (a : A)) (hla : l ≤ 𝓝 a) (hl : ∀ᶠ x in l, IsUnit x) (hl' : l.NeBot) : + IsUnit a := by + have hla₂ : Tendsto Ring.inverse (map (val S) l) (𝓝 (↑ha.unit⁻¹ : A)) := by + rw [← Ring.inverse_unit] + exact (NormedRing.inverse_continuousAt _).tendsto.comp <| + continuousAt_subtype_val.tendsto.comp <| map_mono hla + suffices mem : (↑ha.unit⁻¹ : A) ∈ S by + refine ⟨⟨a, ⟨(↑ha.unit⁻¹ : A), mem⟩, ?_, ?_⟩, rfl⟩ + all_goals ext; simp + apply hS.mem_of_tendsto hla₂ + rw [Filter.eventually_map] + apply hl.mp <| eventually_of_forall fun x hx ↦ ?_ + suffices Ring.inverse (val S x) = (val S ↑hx.unit⁻¹) from this ▸ Subtype.property _ + rw [← (hx.map (val S)).unit_spec, Ring.inverse_unit (hx.map (val S)).unit, val] + apply Units.mul_eq_one_iff_inv_eq.mp + simpa [-IsUnit.mul_val_inv] using congr(($hx.mul_val_inv : A)) + +open Set + +/-- If `S` is a closed subalgebra of a Banach algebra `A`, then for any +`x : S`, the boundary of the spectrum of `x` relative to `S` is a subset of the spectrum of +`↑x : A` relative to `A`. -/ +lemma _root_.Subalgebra.frontier_spectrum : frontier (σ 𝕜 x) ⊆ σ 𝕜 (x : A) := by + have : CompleteSpace S := hS.completeSpace_coe + intro μ hμ + by_contra h + rw [spectrum.not_mem_iff] at h + rw [← frontier_compl, (spectrum.isClosed _).isOpen_compl.frontier_eq, mem_diff] at hμ + obtain ⟨hμ₁, hμ₂⟩ := hμ + rw [mem_closure_iff_clusterPt] at hμ₁ + apply hμ₂ + rw [mem_compl_iff, spectrum.not_mem_iff] + refine Subalgebra.isUnit_of_isUnit_val_of_eventually S h ?_ ?_ <| .map hμ₁ (algebraMap 𝕜 S · - x) + · calc + _ ≤ map _ (𝓝 μ) := map_mono (by simp) + _ ≤ _ := by rw [← Filter.Tendsto, ← ContinuousAt]; fun_prop + · rw [eventually_map] + apply Eventually.filter_mono inf_le_right + simp [spectrum.not_mem_iff] + +/-- If `S` is a closed subalgebra of a Banach algebra `A`, then for any `x : S`, the boundary of +the spectrum of `x` relative to `S` is a subset of the boundary of the spectrum of `↑x : A` +relative to `A`. -/ +lemma Subalgebra.frontier_subset_frontier : + frontier (σ 𝕜 x) ⊆ frontier (σ 𝕜 (x : A)) := by + rw [frontier_eq_closure_inter_closure (s := σ 𝕜 (x : A)), + (spectrum.isClosed (x : A)).closure_eq] + apply subset_inter (frontier_spectrum S x) + rw [frontier_eq_closure_inter_closure] + exact inter_subset_right |>.trans <| + closure_mono <| compl_subset_compl.mpr <| spectrum.subset_subalgebra x + +open Set Notation + +/-- If `S` is a closed subalgebra of a Banach algebra `A`, then for any `x : S`, the spectrum of `x` +is the spectrum of `↑x : A` along with the connected components of the complement of the spectrum of +`↑x : A` which contain an element of the spectrum of `x : S`. -/ +lemma Subalgebra.spectrum_sUnion_connectedComponentIn : + σ 𝕜 x = σ 𝕜 (x : A) ∪ (⋃ z ∈ (σ 𝕜 x \ σ 𝕜 (x : A)), connectedComponentIn (σ 𝕜 (x : A))ᶜ z) := by + suffices IsClopen ((σ 𝕜 (x : A))ᶜ ↓∩ (σ 𝕜 x \ σ 𝕜 (x : A))) by + rw [← this.biUnion_connectedComponentIn (diff_subset_compl _ _), + union_diff_cancel (spectrum.subset_subalgebra x)] + have : CompleteSpace S := hS.completeSpace_coe + have h_open : IsOpen (σ 𝕜 x \ σ 𝕜 (x : A)) := by + rw [← (spectrum.isClosed (𝕜 := 𝕜) x).closure_eq, closure_eq_interior_union_frontier, + union_diff_distrib, diff_eq_empty.mpr (frontier_spectrum S x), + diff_eq_compl_inter, union_empty] + exact (spectrum.isClosed _).isOpen_compl.inter isOpen_interior + apply isClopen_preimage_val h_open + suffices h_frontier : frontier (σ 𝕜 x \ σ 𝕜 (x : A)) ⊆ frontier (σ 𝕜 (x : A)) from + disjoint_of_subset_left h_frontier <| disjoint_compl_right.frontier_left + (spectrum.isClosed _).isOpen_compl + rw [diff_eq_compl_inter] + apply (frontier_inter_subset _ _).trans + rw [frontier_compl] + apply union_subset <| inter_subset_left + refine inter_subset_inter_right _ ?_ |>.trans <| inter_subset_right + exact frontier_subset_frontier S x + +/-- Let `S` be a closed subalgebra of a Banach algebra `A`, and let `x : S`. If `z` is in the +spectrum of `x`, then the connected component of `z` in the complement of the spectrum of `↑x : A` +is bounded (or else `z` actually belongs to the spectrum of `↑x : A`). -/ +lemma Subalgebra.spectrum_isBounded_connectedComponentIn {z : 𝕜} (hz : z ∈ σ 𝕜 x) : + Bornology.IsBounded (connectedComponentIn (σ 𝕜 (x : A))ᶜ z) := by + by_cases hz' : z ∈ σ 𝕜 (x : A) + · simp [connectedComponentIn_eq_empty (show z ∉ (σ 𝕜 (x : A))ᶜ from not_not.mpr hz')] + · have : CompleteSpace S := hS.completeSpace_coe + suffices connectedComponentIn (σ 𝕜 (x : A))ᶜ z ⊆ σ 𝕜 x from spectrum.isBounded x |>.subset this + rw [spectrum_sUnion_connectedComponentIn S] + exact subset_biUnion_of_mem (mem_diff_of_mem hz hz') |>.trans subset_union_right + +end BoundarySpectrum + namespace SpectrumRestricts open NNReal ENNReal From 173d2e189fd79bd332fb19bffac0dce43ec8acb8 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 13 Aug 2024 17:25:05 +0000 Subject: [PATCH 243/359] feat: convenience lemmas for replacing the uniformity and bornology via bilipschitz equivalence (#15760) --- Mathlib.lean | 1 + Mathlib/Topology/MetricSpace/Bilipschitz.lean | 97 +++++++++++++++++++ 2 files changed, 98 insertions(+) create mode 100644 Mathlib/Topology/MetricSpace/Bilipschitz.lean diff --git a/Mathlib.lean b/Mathlib.lean index fc8bbfa934733..40bc9f91fafba 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4490,6 +4490,7 @@ import Mathlib.Topology.Maps.Proper.UniversallyClosed import Mathlib.Topology.MetricSpace.Algebra import Mathlib.Topology.MetricSpace.Antilipschitz import Mathlib.Topology.MetricSpace.Basic +import Mathlib.Topology.MetricSpace.Bilipschitz import Mathlib.Topology.MetricSpace.Bounded import Mathlib.Topology.MetricSpace.CantorScheme import Mathlib.Topology.MetricSpace.CauSeqFilter diff --git a/Mathlib/Topology/MetricSpace/Bilipschitz.lean b/Mathlib/Topology/MetricSpace/Bilipschitz.lean new file mode 100644 index 0000000000000..432fad174bbe3 --- /dev/null +++ b/Mathlib/Topology/MetricSpace/Bilipschitz.lean @@ -0,0 +1,97 @@ +/- +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.Topology.MetricSpace.Antilipschitz +import Mathlib.Topology.MetricSpace.Lipschitz + +/-! # Bilipschitz equivalence + +A common pattern in Mathlib is to replace the topology, uniformity and bornology on a type +synonym with those of the underlying type. + +The most common way to do this is to activate a local instance for something which puts a +`PseudoMetricSpace` structure on the type synonym, prove that this metric is bilipschitz equivalent +to the metric on the underlying type, and then use this to show that the uniformities and +bornologies agree, which can then be used with `PseudoMetricSpace.replaceUniformity` or +`PseudoMetricSpace.replaceBornology`. + +With the tooling outside this file, this can be a bit cumbersome, especially when it occurs +repeatedly, and moreover it can lend itself to abuse of the definitional equality inherent in the +type synonym. In this file, we make this pattern more convenient by providing lemmas which take +directly the conditions that the map is bilipschitz, and then prove the relevant equalities. +Moreover, because there are no type synonyms here, it is necessary to phrase these equalities in +terms of the induced uniformity and bornology, which means users will need to do the same if they +choose to use these convenience lemmas. This encourages good hygiene in the development of type +synonyms. +-/ + +open NNReal + +section Uniformity + +open Uniformity + +variable {α β : Type*} [PseudoEMetricSpace α] [PseudoEMetricSpace β] +variable {K₁ K₂ : ℝ≥0} {f : α → β} + +/-- If `f : α → β` is bilipschitz, then the pullback of the uniformity on `β` through `f` agrees +with the uniformity on `α`. + +This can be used to provide the replacement equality when applying +`PseudoMetricSpace.replaceUniformity`, which can be useful when following the forgetful inheritance +pattern when creating type synonyms. + +Important Note: if `α` is some synonym of a type `β` (at default transparency), and `f : α ≃ β` is +some bilipschitz equivalence, then instead of writing: +``` +instance : UniformSpace α := inferInstanceAs (UniformSpace β) +``` +Users should instead write something like: +``` +instance : UniformSpace α := (inferInstance : UniformSpace β).comap f +``` +in order to avoid abuse of the definitional equality `α := β`. -/ +lemma uniformity_eq_of_bilipschitz (hf₁ : AntilipschitzWith K₁ f) (hf₂ : LipschitzWith K₂ f) : + 𝓤[(inferInstance : UniformSpace β).comap f] = 𝓤 α := + hf₁.uniformInducing hf₂.uniformContinuous |>.comap_uniformity + +end Uniformity + +section Bornology + +open Bornology Filter + +variable {α β : Type*} [PseudoMetricSpace α] [PseudoMetricSpace β] +variable {K₁ K₂ : ℝ≥0} {f : α → β} + +/-- If `f : α → β` is bilipschitz, then the pullback of the bornology on `β` through `f` agrees +with the bornology on `α`. -/ +lemma bornology_eq_of_bilipschitz (hf₁ : AntilipschitzWith K₁ f) (hf₂ : LipschitzWith K₂ f) : + @cobounded _ (induced f) = cobounded α := + le_antisymm hf₂.comap_cobounded_le hf₁.tendsto_cobounded.le_comap + + +/-- If `f : α → β` is bilipschitz, then the pullback of the bornology on `β` through `f` agrees +with the bornology on `α`. + +This can be used to provide the replacement equality when applying +`PseudoMetricSpace.replaceBornology`, which can be useful when following the forgetful inheritance +pattern when creating type synonyms. + +Important Note: if `α` is some synonym of a type `β` (at default transparency), and `f : α ≃ β` is +some bilipschitz equivalence, then instead of writing: +``` +instance : Bornology α := inferInstanceAs (Bornology β) +``` +Users should instead write something like: +``` +instance : Bornology α := Bornology.induced (f : α → β) +``` +in order to avoid abuse of the definitional equality `α := β`. -/ +lemma isBounded_iff_of_bilipschitz (hf₁ : AntilipschitzWith K₁ f) (hf₂ : LipschitzWith K₂ f) + (s : Set α) : @IsBounded _ (induced f) s ↔ Bornology.IsBounded s := + Filter.ext_iff.1 (bornology_eq_of_bilipschitz hf₁ hf₂) (sᶜ) + +end Bornology From d7ada26614c81bc1ac3c8135fadb4494e1019bec Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 13 Aug 2024 19:42:50 +0000 Subject: [PATCH 244/359] feat: add `assert_not_imported` command (#15769) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR introduces the new command `assert_not_imported m₁ m₂ ... mₙ`. The command checks that each one of the modules `m₁ m₂ ... mₙ` are not among the transitive imports of the current file. It also checks that each one of `m₁ m₂ ... mₙ` is actually the name of an existing module, just one that is not currently imported! [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/assert_not_imported) --- Mathlib/Util/AssertExists.lean | 13 +++++++++++++ test/AssertExists.lean | 12 ++++++++++++ 2 files changed, 25 insertions(+) create mode 100644 test/AssertExists.lean diff --git a/Mathlib/Util/AssertExists.lean b/Mathlib/Util/AssertExists.lean index 747726c2b139b..88160a05dbf0a 100644 --- a/Mathlib/Util/AssertExists.lean +++ b/Mathlib/Util/AssertExists.lean @@ -68,3 +68,16 @@ elab "assert_not_exists " n:ident : command => do These invariants are maintained by `assert_not_exists` statements, \ and exist in order to ensure that \"complicated\" parts of the library \ are not accidentally introduced as dependencies of \"simple\" parts of the library." + +/-- `assert_not_imported m₁ m₂ ... mₙ` checks that each one of the modules `m₁ m₂ ... mₙ` is not +among the transitive imports of the current file. + +It also checks that each one of `m₁ m₂ ... mₙ` is actually the name of an existing module, just +one that is not currently imported! +-/ +elab "assert_not_imported " ids:ident* : command => do + let mods := (← getEnv).allImportedModuleNames + for id in ids do + if mods.contains id.getId then logWarningAt id m!"'{id}' is imported" + if let none ← (← searchPathRef.get).findModuleWithExt "olean" id.getId then + logWarningAt id m!"'{id}' does not exist" diff --git a/test/AssertExists.lean b/test/AssertExists.lean new file mode 100644 index 0000000000000..a15d0d8607539 --- /dev/null +++ b/test/AssertExists.lean @@ -0,0 +1,12 @@ +import Mathlib.Util.AssertExists + +/-- +warning: 'Lean.Elab.Command' is imported +--- +warning: 'I_do_not_exist' does not exist +-/ +#guard_msgs in +assert_not_imported + Mathlib.Tactic.Common + Lean.Elab.Command + I_do_not_exist From cf3f2a5bba7afe8d2ca9deafeee1e0ea831ce972 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 13 Aug 2024 21:38:08 +0000 Subject: [PATCH 245/359] chore: backports for leanprover/lean4#4814 (part 35) (#15606) This one is a bit awkward / done badly. I reordered theorems in order to avoid pulling in `SmoothManifoldWithCorners` unnecessarily, but I didn't preserve the section headings properly. :-( From d694ae6caaa2c0c7d98cd73a995c6a7d808c0c25 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 13 Aug 2024 22:03:43 +0000 Subject: [PATCH 246/359] chore: update Mathlib dependencies 2024-08-13 (#15778) 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 d771a7dee19b1..a136169bb5368 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "ad26fe1ebccc9d5b7ca9111d5daf9b4488374415", + "rev": "1d25ec7ec98d6d9fb526c997aa014bcabbad8b72", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 15f92d5ddaf35f96f002e041f5cf7def7d2c5c1c Mon Sep 17 00:00:00 2001 From: Rida Hamadani Date: Wed, 14 Aug 2024 00:10:36 +0000 Subject: [PATCH 247/359] chore(Combinatorics/SimpleGraph): rename and add `gcongr` tag to distance lemmas (#15780) --- Mathlib/Combinatorics/SimpleGraph/Metric.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Mathlib/Combinatorics/SimpleGraph/Metric.lean b/Mathlib/Combinatorics/SimpleGraph/Metric.lean index 0e09bff03a186..17d5d69884b4e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Metric.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Metric.lean @@ -142,7 +142,8 @@ lemma edist_top [DecidableEq V] : (⊤ : SimpleGraph V).edist u v = (if u = v th by_cases h : u = v <;> simp [h] /-- Supergraphs have smaller or equal extended distances to their subgraphs. -/ -theorem edist_le_subgraph_edist {G' : SimpleGraph V} (h : G ≤ G') : +@[gcongr] +theorem edist_anti {G' : SimpleGraph V} (h : G ≤ G') : G'.edist u v ≤ G.edist u v := by by_cases hr : G.Reachable u v · obtain ⟨_, hw⟩ := hr.exists_walk_length_eq_edist @@ -267,7 +268,8 @@ lemma dist_top [DecidableEq V] : (⊤ : SimpleGraph V).dist u v = (if u = v then by_cases h : u = v <;> simp [h] /-- Supergraphs have smaller or equal distances to their subgraphs. -/ -theorem dist_le_subgraph_dist {G' : SimpleGraph V} (h : G ≤ G') (hr : G.Reachable u v) : +@[gcongr] +protected theorem Reachable.dist_anti {G' : SimpleGraph V} (h : G ≤ G') (hr : G.Reachable u v) : G'.dist u v ≤ G.dist u v := by obtain ⟨_, hw⟩ := hr.exists_walk_length_eq_dist rw [← hw, ← Walk.length_map (Hom.mapSpanningSubgraphs h)] From 4525315b58685eca9ac0ae8d391093967b272331 Mon Sep 17 00:00:00 2001 From: Jeremy Tan Jie Rui Date: Wed, 14 Aug 2024 01:35:47 +0000 Subject: [PATCH 248/359] chore: delete `Init.Data.Fin.Basic` (#15762) --- Mathlib.lean | 1 - Mathlib/Data/Fin/Basic.lean | 5 +++++ Mathlib/Data/List/Nodup.lean | 3 +-- Mathlib/Data/List/Range.lean | 2 +- Mathlib/Init/Data/Fin/Basic.lean | 34 -------------------------------- 5 files changed, 7 insertions(+), 38 deletions(-) delete mode 100644 Mathlib/Init/Data/Fin/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index 40bc9f91fafba..7b8acb9dff65c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2791,7 +2791,6 @@ import Mathlib.InformationTheory.Hamming import Mathlib.Init.Algebra.Classes import Mathlib.Init.Classical import Mathlib.Init.Core -import Mathlib.Init.Data.Fin.Basic import Mathlib.Init.Data.Int.Order import Mathlib.Init.Data.List.Basic import Mathlib.Init.Data.List.Instances diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 937ec7c6da20d..bb29c0f7e1ad9 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -78,6 +78,11 @@ def finZeroElim {α : Fin 0 → Sort*} (x : Fin 0) : α x := namespace Fin +@[deprecated (since := "2024-02-15")] alias eq_of_veq := eq_of_val_eq +@[deprecated (since := "2024-02-15")] alias veq_of_eq := val_eq_of_eq +@[deprecated (since := "2024-08-13")] alias ne_of_vne := ne_of_val_ne +@[deprecated (since := "2024-08-13")] alias vne_of_ne := val_ne_of_ne + instance {n : ℕ} : CanLift ℕ (Fin n) Fin.val (· < n) where prf k hk := ⟨⟨k, hk⟩, rfl⟩ diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index 5bfb47206bb95..f5d371b85b7ad 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro, Kenny Lau -/ import Mathlib.Data.List.Forall2 import Mathlib.Data.Set.Pairwise.Basic -import Mathlib.Init.Data.Fin.Basic /-! # Lists with no duplicates @@ -84,7 +83,7 @@ theorem nodup_iff_nthLe_inj {l : List α} : Nodup l ↔ ∀ i j h₁ h₂, nthLe l i h₁ = nthLe l j h₂ → i = j := nodup_iff_injective_get.trans ⟨fun hinj _ _ _ _ h => congr_arg Fin.val (hinj h), - fun hinj i j h => Fin.eq_of_veq (hinj i j i.2 j.2 h)⟩ + fun hinj i j h => Fin.eq_of_val_eq (hinj i j i.2 j.2 h)⟩ theorem Nodup.get_inj_iff {l : List α} (h : Nodup l) {i j : Fin l.length} : l.get i = l.get j ↔ i = j := diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index 64aec7b5b00f9..4f1ef5b6229a7 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -67,7 +67,7 @@ theorem mem_finRange {n : ℕ} (a : Fin n) : a ∈ finRange n := rfl⟩ theorem nodup_finRange (n : ℕ) : (finRange n).Nodup := - (Pairwise.pmap (nodup_range n) _) fun _ _ _ _ => @Fin.ne_of_vne _ ⟨_, _⟩ ⟨_, _⟩ + (Pairwise.pmap (nodup_range n) _) fun _ _ _ _ => @Fin.ne_of_val_ne _ ⟨_, _⟩ ⟨_, _⟩ @[simp] theorem length_finRange (n : ℕ) : (finRange n).length = n := by diff --git a/Mathlib/Init/Data/Fin/Basic.lean b/Mathlib/Init/Data/Fin/Basic.lean deleted file mode 100644 index 65540086b8caa..0000000000000 --- a/Mathlib/Init/Data/Fin/Basic.lean +++ /dev/null @@ -1,34 +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 Mathlib.Data.Nat.Notation - -/-! -# Note about `Mathlib/Init/` -The files in `Mathlib/Init` are leftovers from the port from Mathlib3. -(They contain content moved from lean3 itself that Mathlib needed but was not moved to lean4.) - -We intend to move all the content of these files out into the main `Mathlib` directory structure. -Contributions assisting with this are appreciated. - -# Theorems about equality in `Fin`. --/ - -namespace Fin - -variable {n : ℕ} {i j : Fin n} - -@[deprecated eq_of_val_eq (since := "2024-02-15")] -theorem eq_of_veq : i.val = j.val → i = j := eq_of_val_eq - -@[deprecated val_eq_of_eq (since := "2024-02-15")] -theorem veq_of_eq : i = j → i.val = j.val := val_eq_of_eq - --- These two aren't deprecated because `ne_of_val_ne` and `val_ne_of_ne` --- use `¬a = b` instead of `a ≠ b`. TODO: fix or rename in Lean core. -theorem ne_of_vne (h : i.val ≠ j.val) : i ≠ j := ne_of_val_ne h -theorem vne_of_ne (h : i ≠ j) : i.val ≠ j.val := val_ne_of_ne h - -end Fin From f4aca625b698c49108b0e87b19c45fdf24338a30 Mon Sep 17 00:00:00 2001 From: grunweg Date: Wed, 14 Aug 2024 02:32:35 +0000 Subject: [PATCH 249/359] chore: remove now-obsolete nolint entries for mathport syntax stubs (#15752) With `Mathport/Syntax.lean` being removed, there is no reason to keep the corresponding nolint entries. --- scripts/nolints.json | 92 -------------------------------------------- 1 file changed, 92 deletions(-) diff --git a/scripts/nolints.json b/scripts/nolints.json index 9642b4626c92c..28ccef88aecff 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -531,115 +531,23 @@ ["docBlame", "Mathlib.Notation3.expandFoldl"], ["docBlame", "Mathlib.Notation3.expandFoldr"], ["docBlame", "Mathlib.Notation3.prettyPrintOpt"], - ["docBlame", "Mathlib.Tactic.abstract"], - ["docBlame", "Mathlib.Tactic.acMono"], - ["docBlame", "Mathlib.Tactic.addTacticDoc"], - ["docBlame", "Mathlib.Tactic.applyField"], - ["docBlame", "Mathlib.Tactic.applyNormed"], - ["docBlame", "Mathlib.Tactic.assertInstance"], - ["docBlame", "Mathlib.Tactic.assertNoInstance"], - ["docBlame", "Mathlib.Tactic.assocRw"], - ["docBlame", "Mathlib.Tactic.async"], ["docBlame", "Mathlib.Tactic.cases'"], - ["docBlame", "Mathlib.Tactic.cases''"], - ["docBlame", "Mathlib.Tactic.clarify"], - ["docBlame", "Mathlib.Tactic.compVal"], - ["docBlame", "Mathlib.Tactic.computeDegreeLE"], - ["docBlame", "Mathlib.Tactic.continue"], - ["docBlame", "Mathlib.Tactic.decide!"], - ["docBlame", "Mathlib.Tactic.defReplacer"], - ["docBlame", "Mathlib.Tactic.deltaInstance"], - ["docBlame", "Mathlib.Tactic.deriveElementwiseProof"], - ["docBlame", "Mathlib.Tactic.deriveReassocProof"], - ["docBlame", "Mathlib.Tactic.dsimpResult"], ["docBlame", "Mathlib.Tactic.elabConfig"], ["docBlame", "Mathlib.Tactic.elabVariables"], - ["docBlame", "Mathlib.Tactic.elide"], - ["docBlame", "Mathlib.Tactic.equivRw"], - ["docBlame", "Mathlib.Tactic.equivRwType"], ["docBlame", "Mathlib.Tactic.evalIntrov"], - ["docBlame", "Mathlib.Tactic.expandExists"], - ["docBlame", "Mathlib.Tactic.extractGoal!"], - ["docBlame", "Mathlib.Tactic.failIfSuccess?"], - ["docBlame", "Mathlib.Tactic.field"], - ["docBlame", "Mathlib.Tactic.finish"], - ["docBlame", "Mathlib.Tactic.fixingClause"], - ["docBlame", "Mathlib.Tactic.generalizes"], - ["docBlame", "Mathlib.Tactic.generalizesArg"], - ["docBlame", "Mathlib.Tactic.generalizingClause"], - ["docBlame", "Mathlib.Tactic.guardProofTerm"], - ["docBlame", "Mathlib.Tactic.guardTags"], - ["docBlame", "Mathlib.Tactic.hGeneralize"], - ["docBlame", "Mathlib.Tactic.hGeneralize!"], - ["docBlame", "Mathlib.Tactic.haveField"], - ["docBlame", "Mathlib.Tactic.include"], ["docBlame", "Mathlib.Tactic.induction'"], - ["docBlame", "Mathlib.Tactic.induction''"], - ["docBlame", "Mathlib.Tactic.injectionsAndClear"], - ["docBlame", "Mathlib.Tactic.interactive"], - ["docBlame", "Mathlib.Tactic.intro"], - ["docBlame", "Mathlib.Tactic.intro!"], - ["docBlame", "Mathlib.Tactic.isBounded_default"], - ["docBlame", "Mathlib.Tactic.listUnusedDecls"], - ["docBlame", "Mathlib.Tactic.mapply"], - ["docBlame", "Mathlib.Tactic.matchHyp"], - ["docBlame", "Mathlib.Tactic.mkDecorations"], - ["docBlame", "Mathlib.Tactic.mvBisim"], - ["docBlame", "Mathlib.Tactic.notationClass"], - ["docBlame", "Mathlib.Tactic.nthRwLHS"], - ["docBlame", "Mathlib.Tactic.nthRwRHS"], - ["docBlame", "Mathlib.Tactic.obviously"], - ["docBlame", "Mathlib.Tactic.omit"], - ["docBlame", "Mathlib.Tactic.padicIndexSimp"], - ["docBlame", "Mathlib.Tactic.parameter"], - ["docBlame", "Mathlib.Tactic.piInstance"], - ["docBlame", "Mathlib.Tactic.piInstanceDeriveField"], - ["docBlame", "Mathlib.Tactic.prettyCases"], - ["docBlame", "Mathlib.Tactic.printSorryIn"], - ["docBlame", "Mathlib.Tactic.propagateTags"], - ["docBlame", "Mathlib.Tactic.protectProj"], - ["docBlame", "Mathlib.Tactic.rcases?"], - ["docBlame", "Mathlib.Tactic.reassoc"], - ["docBlame", "Mathlib.Tactic.reassoc!"], - ["docBlame", "Mathlib.Tactic.reassocAxiom"], - ["docBlame", "Mathlib.Tactic.refineStruct"], ["docBlame", "Mathlib.Tactic.renameArg"], - ["docBlame", "Mathlib.Tactic.revertAfter"], - ["docBlame", "Mathlib.Tactic.revertDeps"], - ["docBlame", "Mathlib.Tactic.revertTargetDeps"], - ["docBlame", "Mathlib.Tactic.rintro?"], - ["docBlame", "Mathlib.Tactic.rsimp"], - ["docBlame", "Mathlib.Tactic.safe"], - ["docBlame", "Mathlib.Tactic.sample"], ["docBlame", "Mathlib.Tactic.setArgsRest"], ["docBlame", "Mathlib.Tactic.setTactic"], - ["docBlame", "Mathlib.Tactic.simpResult"], - ["docBlame", "Mathlib.Tactic.subtypeInstance"], - ["docBlame", "Mathlib.Tactic.suggest"], - ["docBlame", "Mathlib.Tactic.tacticDestruct_"], ["docBlame", "Mathlib.Tactic.tacticHave_"], ["docBlame", "Mathlib.Tactic.tacticLet_"], ["docBlame", "Mathlib.Tactic.tacticMatch_target_"], ["docBlame", "Mathlib.Tactic.tacticSet!_"], ["docBlame", "Mathlib.Tactic.tacticSuffices_"], ["docBlame", "Mathlib.Tactic.tacticTransitivity___"], - ["docBlame", "Mathlib.Tactic.tidy"], - ["docBlame", "Mathlib.Tactic.tidy?"], - ["docBlame", "Mathlib.Tactic.transport"], - ["docBlame", "Mathlib.Tactic.truncCases"], - ["docBlame", "Mathlib.Tactic.tryFor"], - ["docBlame", "Mathlib.Tactic.unelide"], - ["docBlame", "Mathlib.Tactic.unfoldAux"], - ["docBlame", "Mathlib.Tactic.unfoldCases"], - ["docBlame", "Mathlib.Tactic.unfoldCoes"], - ["docBlame", "Mathlib.Tactic.unfoldWf"], - ["docBlame", "Mathlib.Tactic.uniqueDiffWithinAt_Ici_Iic_univ"], - ["docBlame", "Mathlib.Tactic.unitInterval"], ["docBlame", "Mathlib.Tactic.usingArg"], ["docBlame", "Mathlib.Tactic.variables"], ["docBlame", "Mathlib.Tactic.withArgs"], - ["docBlame", "Mathlib.Tactic.withPattern"], - ["docBlame", "Mathlib.Tactic.wittTruncateFunTac"], ["docBlame", "Mathlib.WhatsNew.diffExtension"], ["docBlame", "Mathlib.WhatsNew.whatsNew"], ["docBlame", "Matrix.TransvectionStruct.c"], From 8c531e2aab7ddfdad569698c5521bc5dde8fa645 Mon Sep 17 00:00:00 2001 From: grunweg Date: Wed, 14 Aug 2024 02:32:36 +0000 Subject: [PATCH 250/359] chore(assert_not_exists): tweak message (#15783) Also, upgrade the warning about a non-existing module to an error. --- Mathlib/Util/AssertExists.lean | 4 ++-- test/AssertExists.lean | 4 ++-- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Util/AssertExists.lean b/Mathlib/Util/AssertExists.lean index 88160a05dbf0a..81b63120ee16e 100644 --- a/Mathlib/Util/AssertExists.lean +++ b/Mathlib/Util/AssertExists.lean @@ -78,6 +78,6 @@ one that is not currently imported! elab "assert_not_imported " ids:ident* : command => do let mods := (← getEnv).allImportedModuleNames for id in ids do - if mods.contains id.getId then logWarningAt id m!"'{id}' is imported" + if mods.contains id.getId then logWarningAt id m!"the module '{id}' is (transitively) imported" if let none ← (← searchPathRef.get).findModuleWithExt "olean" id.getId then - logWarningAt id m!"'{id}' does not exist" + logErrorAt id m!"the module '{id}' does not exist" diff --git a/test/AssertExists.lean b/test/AssertExists.lean index a15d0d8607539..8e40f4c5aaf7c 100644 --- a/test/AssertExists.lean +++ b/test/AssertExists.lean @@ -1,9 +1,9 @@ import Mathlib.Util.AssertExists /-- -warning: 'Lean.Elab.Command' is imported +warning: the module 'Lean.Elab.Command' is (transitively) imported --- -warning: 'I_do_not_exist' does not exist +error: the module 'I_do_not_exist' does not exist -/ #guard_msgs in assert_not_imported From 86a406ac1fb424e698654169fc87f60848a2a78f Mon Sep 17 00:00:00 2001 From: Lorenzo Luccioli Date: Wed, 14 Aug 2024 03:37:18 +0000 Subject: [PATCH 251/359] feat (MeasureTheory.Decomposition.RadonNikodym): `rnDeriv_le_one_iff_le` (#15473) Add `rnDeriv_le_one_iff_le` and `rnDeriv_eq_one_iff_eq`. --- .../Decomposition/RadonNikodym.lean | 25 +++++++++++++------ 1 file changed, 18 insertions(+), 7 deletions(-) diff --git a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean index da28661366bca..04bc751e124ce 100644 --- a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean +++ b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean @@ -78,7 +78,7 @@ lemma rnDeriv_pos [HaveLebesgueDecomposition μ ν] (hμν : μ ≪ ν) : ae_withDensity_iff (Measure.measurable_rnDeriv _ _), Measure.withDensity_rnDeriv_eq _ _ hμν] exact ae_of_all _ (fun x hx ↦ lt_of_le_of_ne (zero_le _) hx.symm) -lemma rnDeriv_pos' [SigmaFinite μ] [SFinite ν] (hμν : μ ≪ ν) : +lemma rnDeriv_pos' [HaveLebesgueDecomposition ν μ] [SigmaFinite μ] (hμν : μ ≪ ν) : ∀ᵐ x ∂μ, 0 < ν.rnDeriv μ x := by refine (absolutelyContinuous_withDensity_rnDeriv hμν).ae_le ?_ filter_upwards [Measure.rnDeriv_pos (withDensity_absolutelyContinuous μ (ν.rnDeriv μ)), @@ -143,8 +143,8 @@ lemma rnDeriv_withDensity_left {μ ν : Measure α} [SigmaFinite μ] [SigmaFinit rw [← hx2, hx, hx1] /-- Auxiliary lemma for `rnDeriv_withDensity_right`. -/ -lemma rnDeriv_withDensity_right_of_absolutelyContinuous {ν : Measure α} [SigmaFinite μ] - [SigmaFinite ν] (hμν : μ ≪ ν) (hf : AEMeasurable f ν) +lemma rnDeriv_withDensity_right_of_absolutelyContinuous {ν : Measure α} + [HaveLebesgueDecomposition μ ν] [SigmaFinite ν] (hμν : μ ≪ ν) (hf : AEMeasurable f ν) (hf_ne_zero : ∀ᵐ x ∂ν, f x ≠ 0) (hf_ne_top : ∀ᵐ x ∂ν, f x ≠ ∞) : μ.rnDeriv (ν.withDensity f) =ᵐ[ν] fun x ↦ (f x)⁻¹ * μ.rnDeriv ν x := by have : SigmaFinite (ν.withDensity f) := SigmaFinite.withDensity_of_ne_top hf_ne_top @@ -173,7 +173,6 @@ lemma rnDeriv_withDensity_right (μ ν : Measure α) [SigmaFinite μ] [SigmaFini rnDeriv_withDensity_withDensity_rnDeriv_right μ ν hf hf_ne_zero hf_ne_top have h₂ : μ.rnDeriv ν =ᵐ[ν] μ'.rnDeriv ν := (Measure.rnDeriv_withDensity _ (Measure.measurable_rnDeriv _ _)).symm - have : SigmaFinite μ' := SigmaFinite.withDensity_of_ne_top (Measure.rnDeriv_ne_top _ _) have hμ' := rnDeriv_withDensity_right_of_absolutelyContinuous (withDensity_absolutelyContinuous ν (μ.rnDeriv ν)) hf hf_ne_zero hf_ne_top filter_upwards [h₁, h₂, hμ'] with x hx₁ hx₂ hx_eq @@ -181,7 +180,7 @@ lemma rnDeriv_withDensity_right (μ ν : Measure α) [SigmaFinite μ] [SigmaFini end rnDeriv_withDensity_leftRight -lemma rnDeriv_eq_zero_of_mutuallySingular [SigmaFinite μ] {ν' : Measure α} +lemma rnDeriv_eq_zero_of_mutuallySingular {ν' : Measure α} [HaveLebesgueDecomposition μ ν'] [SigmaFinite ν'] (h : μ ⟂ₘ ν) (hνν' : ν ≪ ν') : μ.rnDeriv ν' =ᵐ[ν] 0 := by let t := h.nullSet @@ -203,7 +202,7 @@ lemma rnDeriv_eq_zero_of_mutuallySingular [SigmaFinite μ] {ν' : Measure α} /-- Auxiliary lemma for `rnDeriv_add_right_of_mutuallySingular`. -/ lemma rnDeriv_add_right_of_absolutelyContinuous_of_mutuallySingular {ν' : Measure α} - [SigmaFinite μ] [SigmaFinite ν] [SigmaFinite ν'] + [HaveLebesgueDecomposition μ ν] [HaveLebesgueDecomposition μ (ν + ν')] [SigmaFinite ν] (hμν : μ ≪ ν) (hνν' : ν ⟂ₘ ν') : μ.rnDeriv (ν + ν') =ᵐ[ν] μ.rnDeriv ν := by let t := hνν'.nullSet @@ -265,7 +264,8 @@ lemma rnDeriv_withDensity_rnDeriv [SigmaFinite μ] [SigmaFinite ν] (hμν : μ (Measure.mutuallySingular_singularPart ν μ).symm.withDensity).symm /-- Auxiliary lemma for `inv_rnDeriv`. -/ -lemma inv_rnDeriv_aux [SigmaFinite μ] [SigmaFinite ν] (hμν : μ ≪ ν) (hνμ : ν ≪ μ) : +lemma inv_rnDeriv_aux [HaveLebesgueDecomposition μ ν] [HaveLebesgueDecomposition ν μ] + [SigmaFinite μ] (hμν : μ ≪ ν) (hνμ : ν ≪ μ) : (μ.rnDeriv ν)⁻¹ =ᵐ[μ] ν.rnDeriv μ := by suffices μ.withDensity (μ.rnDeriv ν)⁻¹ = μ.withDensity (ν.rnDeriv μ) by calc (μ.rnDeriv ν)⁻¹ =ᵐ[μ] (μ.withDensity (μ.rnDeriv ν)⁻¹).rnDeriv μ := @@ -405,6 +405,17 @@ lemma rnDeriv_le_one_of_le (hμν : μ ≤ ν) [SigmaFinite ν] : μ.rnDeriv ν simp only [Pi.one_apply, MeasureTheory.setLIntegral_one] exact (Measure.setLIntegral_rnDeriv_le s).trans (hμν s) +lemma rnDeriv_le_one_iff_le [HaveLebesgueDecomposition μ ν] [SigmaFinite ν] (hμν : μ ≪ ν) : + μ.rnDeriv ν ≤ᵐ[ν] 1 ↔ μ ≤ ν := by + refine ⟨fun h s ↦ ?_, fun h ↦ rnDeriv_le_one_of_le h⟩ + rw [← withDensity_rnDeriv_eq _ _ hμν, withDensity_apply', ← setLIntegral_one] + exact setLIntegral_mono_ae aemeasurable_const (h.mono fun _ hh _ ↦ hh) + +lemma rnDeriv_eq_one_iff_eq [HaveLebesgueDecomposition μ ν] [SigmaFinite ν] (hμν : μ ≪ ν) : + μ.rnDeriv ν =ᵐ[ν] 1 ↔ μ = ν := by + refine ⟨fun h ↦ ?_, fun h ↦ h ▸ ν.rnDeriv_self⟩ + rw [← withDensity_rnDeriv_eq _ _ hμν, withDensity_congr_ae h, withDensity_one] + section MeasurableEmbedding variable {mβ : MeasurableSpace β} {f : α → β} From 81ab12fba29a10129e60076208e81cbeec69b8f9 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Wed, 14 Aug 2024 07:09:09 +0000 Subject: [PATCH 252/359] perf(CategoryTheory.ComposableArrows): use `suppress_compilation` on `opEquivalence` (#15593) This declaration now takes ~4s to compile on my machine after moving to `v4.11.0-rc1`. --- Mathlib/CategoryTheory/ComposableArrows.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/CategoryTheory/ComposableArrows.lean b/Mathlib/CategoryTheory/ComposableArrows.lean index d7b5a70f49f70..45883d4350c3a 100644 --- a/Mathlib/CategoryTheory/ComposableArrows.lean +++ b/Mathlib/CategoryTheory/ComposableArrows.lean @@ -8,6 +8,7 @@ import Mathlib.CategoryTheory.EqToHom import Mathlib.CategoryTheory.Functor.Const import Mathlib.Order.Fin.Basic import Mathlib.Tactic.FinCases +import Mathlib.Tactic.SuppressCompilation /-! # Composable arrows @@ -888,6 +889,7 @@ lemma mkOfObjOfMapSucc_arrow (i : ℕ) (hi : i < n := by valid) : end mkOfObjOfMapSucc +suppress_compilation in variable (C n) in /-- The equivalence `(ComposableArrows C n)ᵒᵖ ≌ ComposableArrows Cᵒᵖ n` obtained by reversing the arrows. -/ @@ -914,6 +916,7 @@ def Functor.mapComposableArrows : ComposableArrows C n ⥤ ComposableArrows D n := (whiskeringRight _ _ _).obj G +suppress_compilation in /-- The functor `ComposableArrows C n ⥤ ComposableArrows D n` induced by `G : C ⥤ D` commutes with `opEquivalence`. -/ def Functor.mapComposableArrowsOpIso : From 077d38eee589991c0b64ca315d0166286ec428e2 Mon Sep 17 00:00:00 2001 From: grunweg Date: Wed, 14 Aug 2024 07:23:54 +0000 Subject: [PATCH 253/359] chore: use `assert_not_exists` (#15782) The first location was suggested in #15732; now we can do this. The other locations had comments also suggesting such functionality: they referred to an outdated module name, which we correct. --- Mathlib/Algebra/BigOperators/Group/List.lean | 4 +--- Mathlib/Algebra/Field/Defs.lean | 5 ++--- Mathlib/Data/List/Chain.lean | 3 +-- Mathlib/Data/List/GetD.lean | 4 +--- 4 files changed, 5 insertions(+), 11 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 0251d89f06740..b283972626cff 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -21,9 +21,7 @@ of elements of a list and `List.alternatingProd`, `List.alternatingSum`, their a counterparts. -/ --- Make sure we haven't imported `Data.Nat.Order.Basic` -assert_not_exists OrderedSub -assert_not_exists Ring +assert_not_imported Mathlib.Algebra.Order.Group.Nat variable {ι α β M N P G : Type*} diff --git a/Mathlib/Algebra/Field/Defs.lean b/Mathlib/Algebra/Field/Defs.lean index c29f6e33d0c67..4943b798adf4c 100644 --- a/Mathlib/Algebra/Field/Defs.lean +++ b/Mathlib/Algebra/Field/Defs.lean @@ -43,12 +43,11 @@ a `GroupWithZero` lemma instead. field, division ring, skew field, skew-field, skewfield -/ +assert_not_imported Mathlib.Tactic.Common + -- `NeZero` should not be needed in the basic algebraic hierarchy. assert_not_exists NeZero --- Check that we have not imported `Mathlib.Tactic.Common` yet. -assert_not_exists Mathlib.Tactic.scopedNS - assert_not_exists MonoidHom open Function Set diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index e8e436dc37099..a6c2410414b7d 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -18,8 +18,7 @@ A graph-specialized version is in development and will hopefully be added under sometime soon. -/ --- Make sure we haven't imported `Data.Nat.Order.Basic` -assert_not_exists OrderedSub +assert_not_imported Mathlib.Algebra.Order.Group.Nat universe u v diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index 9b888110c53b9..8e35897c2dd50 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -14,11 +14,9 @@ import Mathlib.Util.AssertExists This file provides theorems for working with the `getD` and `getI` functions. These are used to access an element of a list by numerical index, with a default value as a fallback when the index is out of range. - -/ --- Make sure we haven't imported `Data.Nat.Order.Basic` -assert_not_exists OrderedSub +assert_not_imported Mathlib.Algebra.Order.Group.Nat namespace List From a7e0ad18e89d6413ba748138620e0586f54595fe Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 14 Aug 2024 09:17:18 +0000 Subject: [PATCH 254/359] feat(CategoryTheory/Monoidal): Hopf monoids (#13317) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Defines Hopf monoids in a braided category. Proves two facts (antipode is an antihomomorphism, morphisms intertwine antipode) that are not yet proved for unbundled Hopf algebras. Once we prove `Hopf_ (ModuleCat R)` is equivalent to `HopfAlgebraCat` (having defined that!), we will get those two facts for free. We should also check later that `Hopf_ C` in a cartesian monoidal category is equivalent `GroupObject_ C` (being defined independently elsewhere). - [x] depends on: #13313 - [x] depends on: #13312 (probably doesn't actually depend, but I don't want to sort out the git history now) - [x] depends on: #13315 Co-authored-by: Scott Morrison Co-authored-by: Joël Riou --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Monoidal/Bimon_.lean | 69 ++- Mathlib/CategoryTheory/Monoidal/Comon_.lean | 6 +- Mathlib/CategoryTheory/Monoidal/Hopf_.lean | 466 ++++++++++++++++++++ Mathlib/CategoryTheory/Monoidal/Mon_.lean | 3 +- Mathlib/RingTheory/HopfAlgebra.lean | 3 + 6 files changed, 542 insertions(+), 6 deletions(-) create mode 100644 Mathlib/CategoryTheory/Monoidal/Hopf_.lean diff --git a/Mathlib.lean b/Mathlib.lean index 7b8acb9dff65c..0abdbf82cdafa 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1676,6 +1676,7 @@ import Mathlib.CategoryTheory.Monoidal.Free.Coherence import Mathlib.CategoryTheory.Monoidal.Functor import Mathlib.CategoryTheory.Monoidal.FunctorCategory import Mathlib.CategoryTheory.Monoidal.Functorial +import Mathlib.CategoryTheory.Monoidal.Hopf_ import Mathlib.CategoryTheory.Monoidal.Internal.FunctorCategory import Mathlib.CategoryTheory.Monoidal.Internal.Limits import Mathlib.CategoryTheory.Monoidal.Internal.Module diff --git a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean index 9d5818ef51531..109f528c80123 100644 --- a/Mathlib/CategoryTheory/Monoidal/Bimon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Bimon_.lean @@ -14,8 +14,6 @@ as comonoid objects in the category of monoid objects in `C`. We verify that this is equivalent to the monoid objects in the category of comonoid objects. ## TODO -* Define Hopf monoids, which in a cartesian monoidal category are exactly group objects, - and use this to define group schemes. * Construct the category of modules, and show that it is monoidal with a monoidal forgetful functor to `C`. * Some form of Tannaka reconstruction: @@ -112,4 +110,71 @@ def equivMon_Comon_ : Bimon_ C ≌ Mon_ (Comon_ C) where unitIso := NatIso.ofComponents (fun _ => Comon_.mkIso (Mon_.mkIso (Iso.refl _))) counitIso := NatIso.ofComponents (fun _ => Mon_.mkIso (Comon_.mkIso (Iso.refl _))) +/-! # The trivial bimonoid -/ + +/-- The trivial bimonoid object. -/ +@[simps!] +def trivial : Bimon_ C := Comon_.trivial (Mon_ C) + +/-- The bimonoid morphism from the trivial bimonoid to any bimonoid. -/ +@[simps] +def trivial_to (A : Bimon_ C) : trivial C ⟶ A := + { hom := (default : Mon_.trivial C ⟶ A.X), } + +/-- The bimonoid morphism from any bimonoid to the trivial bimonoid. -/ +@[simps!] +def to_trivial (A : Bimon_ C) : A ⟶ trivial C := + (default : @Quiver.Hom (Comon_ (Mon_ C)) _ A (Comon_.trivial (Mon_ C))) + +/-! # Additional lemmas -/ + +variable {C} + +@[reassoc] +theorem one_comul (M : Bimon_ C) : + M.X.one ≫ M.comul.hom = (λ_ _).inv ≫ (M.X.one ⊗ M.X.one) := by + simp + +@[reassoc] +theorem mul_counit (M : Bimon_ C) : + M.X.mul ≫ M.counit.hom = (M.counit.hom ⊗ M.counit.hom) ≫ (λ_ _).hom := by + simp + +/-- Compatibility of the monoid and comonoid structures, in terms of morphisms in `C`. -/ +@[reassoc (attr := simp)] theorem compatibility (M : Bimon_ C) : + (M.comul.hom ⊗ M.comul.hom) ≫ + (α_ _ _ (M.X.X ⊗ M.X.X)).hom ≫ M.X.X ◁ (α_ _ _ _).inv ≫ + M.X.X ◁ (β_ M.X.X M.X.X).hom ▷ M.X.X ≫ + M.X.X ◁ (α_ _ _ _).hom ≫ (α_ _ _ _).inv ≫ + (M.X.mul ⊗ M.X.mul) = + M.X.mul ≫ M.comul.hom := by + have := (Mon_.Hom.mul_hom M.comul).symm + simpa [-Mon_.Hom.mul_hom, tensor_μ] using this + +@[reassoc (attr := simp)] theorem comul_counit_hom (M : Bimon_ C) : + M.comul.hom ≫ (_ ◁ M.counit.hom) = (ρ_ _).inv := by + simpa [- Comon_.comul_counit] using congr_arg Mon_.Hom.hom M.comul_counit + +@[reassoc (attr := simp)] theorem counit_comul_hom (M : Bimon_ C) : + M.comul.hom ≫ (M.counit.hom ▷ _) = (λ_ _).inv := by + simpa [- Comon_.counit_comul] using congr_arg Mon_.Hom.hom M.counit_comul + +@[reassoc (attr := simp)] theorem comul_assoc_hom (M : Bimon_ C) : + M.comul.hom ≫ (M.X.X ◁ M.comul.hom) = + M.comul.hom ≫ (M.comul.hom ▷ M.X.X) ≫ (α_ M.X.X M.X.X M.X.X).hom := by + simpa [- Comon_.comul_assoc] using congr_arg Mon_.Hom.hom M.comul_assoc + +@[reassoc] theorem comul_assoc_flip_hom (M : Bimon_ C) : + M.comul.hom ≫ (M.comul.hom ▷ M.X.X) = + M.comul.hom ≫ (M.X.X ◁ M.comul.hom) ≫ (α_ M.X.X M.X.X M.X.X).inv := by + simp + +@[reassoc] theorem hom_comul_hom {M N : Bimon_ C} (f : M ⟶ N) : + f.hom.hom ≫ N.comul.hom = M.comul.hom ≫ (f.hom.hom ⊗ f.hom.hom) := by + simpa [- Comon_.Hom.hom_comul] using congr_arg Mon_.Hom.hom f.hom_comul + +@[reassoc] theorem hom_counit_hom {M N : Bimon_ C} (f : M ⟶ N) : + f.hom.hom ≫ N.counit.hom = M.counit.hom := by + simpa [- Comon_.Hom.hom_counit] using congr_arg Mon_.Hom.hom f.hom_counit + end Bimon_ diff --git a/Mathlib/CategoryTheory/Monoidal/Comon_.lean b/Mathlib/CategoryTheory/Monoidal/Comon_.lean index 7da6f2e2289d8..b58290d9ba93a 100644 --- a/Mathlib/CategoryTheory/Monoidal/Comon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Comon_.lean @@ -79,7 +79,7 @@ theorem comul_counit_hom {Z : C} (f : M.X ⟶ Z) : M.comul ≫ (f ⊗ M.counit) @[reassoc] theorem comul_assoc_flip : M.comul ≫ (M.comul ▷ M.X) = M.comul ≫ (M.X ◁ M.comul) ≫ (α_ M.X M.X M.X).inv := by - simp [← comul_assoc] + simp /-- A morphism of comonoid objects. -/ @[ext] @@ -156,11 +156,11 @@ def mkIso {M N : Comon_ C} (f : M.X ≅ N.X) (f_counit : f.hom ≫ N.counit = M. slice_rhs 1 2 => rw [f_comul] simp } +@[simps] instance uniqueHomToTrivial (A : Comon_ C) : Unique (A ⟶ trivial C) where default := { hom := A.counit - hom_counit := by dsimp; simp - hom_comul := by dsimp; simp [A.comul_counit, unitors_inv_equal] } + hom_comul := by simp [A.comul_counit, unitors_inv_equal] } uniq f := by ext; simp rw [← Category.comp_id f.hom] diff --git a/Mathlib/CategoryTheory/Monoidal/Hopf_.lean b/Mathlib/CategoryTheory/Monoidal/Hopf_.lean new file mode 100644 index 0000000000000..12419c19a3575 --- /dev/null +++ b/Mathlib/CategoryTheory/Monoidal/Hopf_.lean @@ -0,0 +1,466 @@ +/- +Copyright (c) 2024 Lean FRO LLC. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kim Morrison +-/ +import Mathlib.CategoryTheory.Monoidal.Bimon_ +import Mathlib.CategoryTheory.Monoidal.Conv + +/-! +# The category of Hopf monoids in a braided monoidal category. + + +## TODO + +* Show that in a cartesian monoidal category Hopf monoids are exactly group objects. +* Show that `Hopf_ (ModuleCat R) ≌ HopfAlgebraCat R`. +-/ + +noncomputable section + +universe v₁ v₂ u₁ u₂ u + +open CategoryTheory MonoidalCategory + +variable (C : Type u₁) [Category.{v₁} C] [MonoidalCategory.{v₁} C] [BraidedCategory C] + +/-- +A Hopf monoid in a braided category `C` is a bimonoid object in `C` equipped with an antipode. +-/ +structure Hopf_ where + /-- The underlying bimonoid of a Hopf monoid. -/ + X : Bimon_ C + /-- The antipode is an endomorphism of the underlying object of the Hopf monoid. -/ + antipode : X.X.X ⟶ X.X.X + antipode_left : X.comul.hom ≫ (antipode ▷ X.X.X) ≫ X.X.mul = X.counit.hom ≫ X.X.one + antipode_right : X.comul.hom ≫ (X.X.X ◁ antipode) ≫ X.X.mul = X.counit.hom ≫ X.X.one + +attribute [reassoc (attr := simp)] Hopf_.antipode_left Hopf_.antipode_right + +namespace Hopf_ + +/-- +Morphisms of Hopf monoids are just morphisms of the underlying bimonoids. +In fact they automatically intertwine the antipodes, proved below. +-/ +instance : Category (Hopf_ C) := inferInstanceAs <| Category (InducedCategory (Bimon_ C) Hopf_.X) + +variable {C} + +/-- Morphisms of Hopf monoids intertwine the antipodes. -/ +theorem hom_antipode {A B : Hopf_ C} (f : A ⟶ B) : + f.hom.hom ≫ B.antipode = A.antipode ≫ f.hom.hom := by + -- We show these elements are equal by exhibiting an element in the convolution algebra + -- between `A` (as a comonoid) and `B` (as a monoid), + -- such that the LHS is a left inverse, and the RHS is a right inverse. + apply left_inv_eq_right_inv + (M := Conv ((Bimon_.toComon_ C).obj A.X) B.X.X) + (a := f.hom.hom) + · erw [Conv.mul_eq, Conv.one_eq] + simp only [Bimon_.toComon__obj_X, Bimon_.toComon__obj_comul, comp_whiskerRight, Category.assoc, + Bimon_.toComon__obj_counit] + slice_lhs 3 4 => + rw [← whisker_exchange] + slice_lhs 2 3 => + rw [← tensorHom_def] + slice_lhs 1 2 => + rw [← Bimon_.hom_comul_hom f] + slice_lhs 2 4 => + rw [B.antipode_left] + slice_lhs 1 2 => + rw [Bimon_.hom_counit_hom f] + · erw [Conv.mul_eq, Conv.one_eq] + simp only [Bimon_.toComon__obj_X, Bimon_.toComon__obj_comul, MonoidalCategory.whiskerLeft_comp, + Category.assoc, Bimon_.toComon__obj_counit] + slice_lhs 2 3 => + rw [← whisker_exchange] + slice_lhs 3 4 => + rw [← tensorHom_def] + slice_lhs 3 4 => + rw [← f.hom.mul_hom] + slice_lhs 1 3 => + rw [A.antipode_right] + slice_lhs 2 3 => + rw [f.hom.one_hom] + +@[reassoc (attr := simp)] +theorem one_antipode (A : Hopf_ C) : A.X.X.one ≫ A.antipode = A.X.X.one := by + have := (rfl : A.X.X.one ≫ A.X.comul.hom ≫ (A.antipode ▷ A.X.X.X) ≫ A.X.X.mul = _) + conv at this => + rhs + rw [A.antipode_left] + rw [A.X.one_comul_assoc, tensorHom_def, Category.assoc, whisker_exchange_assoc] at this + simpa [unitors_inv_equal] + +@[reassoc (attr := simp)] +theorem antipode_counit (A : Hopf_ C) : A.antipode ≫ A.X.counit.hom = A.X.counit.hom := by + have := (rfl : A.X.comul.hom ≫ (A.antipode ▷ A.X.X.X) ≫ A.X.X.mul ≫ A.X.counit.hom = _) + conv at this => + rhs + rw [A.antipode_left_assoc] + rw [A.X.mul_counit, tensorHom_def', Category.assoc, ← whisker_exchange_assoc] at this + simpa [unitors_equal] + +/-! +## The antipode is an antihomomorphism with respect to both the monoid and comonoid structures. +-/ + +theorem antipode_comul₁ (A : Hopf_ C) : + A.X.comul.hom ≫ + A.antipode ▷ A.X.X.X ≫ + A.X.comul.hom ▷ A.X.X.X ≫ + (α_ A.X.X.X A.X.X.X A.X.X.X).hom ≫ + A.X.X.X ◁ A.X.X.X ◁ A.X.comul.hom ≫ + A.X.X.X ◁ (α_ A.X.X.X A.X.X.X A.X.X.X).inv ≫ + A.X.X.X ◁ (β_ A.X.X.X A.X.X.X).hom ▷ A.X.X.X ≫ + A.X.X.X ◁ (α_ A.X.X.X A.X.X.X A.X.X.X).hom ≫ + (α_ A.X.X.X A.X.X.X (A.X.X.X ⊗ A.X.X.X)).inv ≫ + (A.X.X.mul ⊗ A.X.X.mul) = + A.X.counit.hom ≫ (λ_ (𝟙_ C)).inv ≫ (A.X.X.one ⊗ A.X.X.one) := by + dsimp + slice_lhs 3 5 => + rw [← associator_naturality_right, ← Category.assoc, ← tensorHom_def] + slice_lhs 3 9 => + erw [Bimon_.compatibility] + slice_lhs 1 3 => + erw [A.antipode_left] + simp + +/-- +Auxiliary calculation for `antipode_comul`. +This calculation calls for some ASCII art out of This Week's Finds. + + | | + n n + | \ / | + | / | + | / \ | + | | S S + | | \ / + | | / + | | / \ + \ / \ / + v v + \ / + v + | + +We move the left antipode up through the crossing, +the right antipode down through the crossing, +the right multiplication down across the strand, +reassociate the comultiplications, +then use `antipode_right` then `antipode_left` to simplify. +-/ +theorem antipode_comul₂ (A : Hopf_ C) : + A.X.comul.hom ≫ + A.X.comul.hom ▷ A.X.X.X ≫ + (α_ A.X.X.X A.X.X.X A.X.X.X).hom ≫ + A.X.X.X ◁ A.X.X.X ◁ A.X.comul.hom ≫ + A.X.X.X ◁ A.X.X.X ◁ (β_ A.X.X.X A.X.X.X).hom ≫ + A.X.X.X ◁ A.X.X.X ◁ (A.antipode ⊗ A.antipode) ≫ + A.X.X.X ◁ (α_ A.X.X.X A.X.X.X A.X.X.X).inv ≫ + A.X.X.X ◁ (β_ A.X.X.X A.X.X.X).hom ▷ A.X.X.X ≫ + A.X.X.X ◁ (α_ A.X.X.X A.X.X.X A.X.X.X).hom ≫ + (α_ A.X.X.X A.X.X.X (A.X.X.X ⊗ A.X.X.X)).inv ≫ + (A.X.X.mul ⊗ A.X.X.mul) = + A.X.counit.hom ≫ (λ_ (𝟙_ C)).inv ≫ (A.X.X.one ⊗ A.X.X.one) := by + -- We should write a version of `slice_lhs` that zooms through whiskerings. + slice_lhs 6 6 => + simp only [tensorHom_def', MonoidalCategory.whiskerLeft_comp] + slice_lhs 7 8 => + rw [← MonoidalCategory.whiskerLeft_comp, associator_inv_naturality_middle, + MonoidalCategory.whiskerLeft_comp] + slice_lhs 8 9 => + rw [← MonoidalCategory.whiskerLeft_comp, ← comp_whiskerRight, + BraidedCategory.braiding_naturality_right, + comp_whiskerRight, MonoidalCategory.whiskerLeft_comp] + slice_lhs 9 10 => + rw [← MonoidalCategory.whiskerLeft_comp, + associator_naturality_left, + MonoidalCategory.whiskerLeft_comp] + slice_lhs 5 6 => + rw [← MonoidalCategory.whiskerLeft_comp, ← MonoidalCategory.whiskerLeft_comp, + ← BraidedCategory.braiding_naturality_left, + MonoidalCategory.whiskerLeft_comp, MonoidalCategory.whiskerLeft_comp] + slice_lhs 11 12 => + rw [tensorHom_def', ← Category.assoc, ← associator_inv_naturality_right] + slice_lhs 10 11 => + rw [← MonoidalCategory.whiskerLeft_comp, ← whisker_exchange, + MonoidalCategory.whiskerLeft_comp] + slice_lhs 6 10 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [← BraidedCategory.hexagon_reverse_assoc, Iso.inv_hom_id_assoc, + ← BraidedCategory.braiding_naturality_left] + simp only [MonoidalCategory.whiskerLeft_comp] + rw [Bimon_.comul_assoc_flip_hom_assoc, Iso.inv_hom_id_assoc] + slice_lhs 2 3 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [Bimon_.comul_assoc_hom] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 3 7 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [← associator_naturality_middle_assoc, Iso.hom_inv_id_assoc] + simp only [← comp_whiskerRight] + rw [antipode_right] + simp only [comp_whiskerRight] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 2 3 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [Bimon_.counit_comul_hom] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 3 4 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [BraidedCategory.braiding_naturality_left] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 4 5 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [whisker_exchange] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 5 7 => + rw [associator_inv_naturality_right_assoc, whisker_exchange] + simp only [Mon_.monMonoidalStruct_tensorObj_X, Mon_.tensorUnit_X, braiding_tensorUnit_left, + MonoidalCategory.whiskerLeft_comp, whiskerLeft_rightUnitor_inv, + MonoidalCategory.whiskerRight_id, whiskerLeft_rightUnitor, Category.assoc, Iso.hom_inv_id_assoc, + Iso.inv_hom_id_assoc, whiskerLeft_inv_hom_assoc, antipode_right_assoc] + rw [rightUnitor_inv_naturality_assoc, tensorHom_def] + coherence + +theorem antipode_comul (A : Hopf_ C) : + A.antipode ≫ A.X.comul.hom = A.X.comul.hom ≫ (β_ _ _).hom ≫ (A.antipode ⊗ A.antipode) := by + -- Again, it is a "left inverse equals right inverse" argument in the convolution monoid. + apply left_inv_eq_right_inv + (M := Conv ((Bimon_.toComon_ C).obj A.X) (A.X.X ⊗ A.X.X)) + (a := A.X.comul.hom) + · erw [Conv.mul_eq, Conv.one_eq] + simp only [Bimon_.toComon__obj_X, Mon_.monMonoidalStruct_tensorObj_X, Bimon_.toComon__obj_comul, + comp_whiskerRight, tensor_whiskerLeft, Mon_.tensorObj_mul, Category.assoc, + Bimon_.toComon__obj_counit, Mon_.tensorObj_one] + simp only [tensor_μ] + simp only [Category.assoc, Iso.inv_hom_id_assoc] + exact antipode_comul₁ A + · erw [Conv.mul_eq, Conv.one_eq] + simp only [Bimon_.toComon__obj_X, Mon_.monMonoidalStruct_tensorObj_X, Bimon_.toComon__obj_comul, + MonoidalCategory.whiskerLeft_comp, tensor_whiskerLeft, Category.assoc, Iso.inv_hom_id_assoc, + Mon_.tensorObj_mul, Bimon_.toComon__obj_counit, Mon_.tensorObj_one] + simp only [tensor_μ] + simp only [Category.assoc, Iso.inv_hom_id_assoc] + exact antipode_comul₂ A + +theorem mul_antipode₁ (A : Hopf_ C) : + (A.X.comul.hom ⊗ A.X.comul.hom) ≫ + (α_ A.X.X.X A.X.X.X (A.X.X.X ⊗ A.X.X.X)).hom ≫ + A.X.X.X ◁ (α_ A.X.X.X A.X.X.X A.X.X.X).inv ≫ + A.X.X.X ◁ (β_ A.X.X.X A.X.X.X).hom ▷ A.X.X.X ≫ + (α_ A.X.X.X (A.X.X.X ⊗ A.X.X.X) A.X.X.X).inv ≫ + (α_ A.X.X.X A.X.X.X A.X.X.X).inv ▷ A.X.X.X ≫ + A.X.X.mul ▷ A.X.X.X ▷ A.X.X.X ≫ + A.antipode ▷ A.X.X.X ▷ A.X.X.X ≫ + (α_ A.X.X.X A.X.X.X A.X.X.X).hom ≫ + A.X.X.X ◁ A.X.X.mul ≫ + A.X.X.mul = + (A.X.counit.hom ⊗ A.X.counit.hom) ≫ (λ_ (𝟙_ C)).hom ≫ A.X.X.one := by + slice_lhs 8 9 => + rw [associator_naturality_left] + slice_lhs 9 10 => + rw [← whisker_exchange] + slice_lhs 7 8 => + rw [associator_naturality_left] + slice_lhs 8 9 => + rw [← tensorHom_def] + simp only [Mon_.monMonoidalStruct_tensorObj_X, Category.assoc, pentagon_inv_inv_hom_hom_inv_assoc, + Mon_.tensorUnit_X] + slice_lhs 1 7 => + erw [Bimon_.compatibility] + slice_lhs 2 4 => + rw [antipode_left] + simp + + +/-- +Auxiliary calculation for `mul_antipode`. + + | + n + / \ + | n + | / \ + | S S + | \ / + n / + / \ / \ + | / | + \ / \ / + v v + | | + +We move the leftmost multiplication up, so we can reassociate. +We then move the rightmost comultiplication under the strand, +and simplify using `antipode_right`. +-/ +theorem mul_antipode₂ (A : Hopf_ C) : + (A.X.comul.hom ⊗ A.X.comul.hom) ≫ + (α_ A.X.X.X A.X.X.X (A.X.X.X ⊗ A.X.X.X)).hom ≫ + A.X.X.X ◁ (α_ A.X.X.X A.X.X.X A.X.X.X).inv ≫ + A.X.X.X ◁ (β_ A.X.X.X A.X.X.X).hom ▷ A.X.X.X ≫ + (α_ A.X.X.X (A.X.X.X ⊗ A.X.X.X) A.X.X.X).inv ≫ + (α_ A.X.X.X A.X.X.X A.X.X.X).inv ▷ A.X.X.X ≫ + A.X.X.mul ▷ A.X.X.X ▷ A.X.X.X ≫ + (α_ A.X.X.X A.X.X.X A.X.X.X).hom ≫ + A.X.X.X ◁ (β_ A.X.X.X A.X.X.X).hom ≫ + A.X.X.X ◁ (A.antipode ⊗ A.antipode) ≫ + A.X.X.X ◁ A.X.X.mul ≫ A.X.X.mul = + (A.X.counit.hom ⊗ A.X.counit.hom) ≫ (λ_ (𝟙_ C)).hom ≫ A.X.X.one := by + slice_lhs 7 8 => + rw [associator_naturality_left] + slice_lhs 8 9 => + rw [← whisker_exchange] + slice_lhs 9 10 => + rw [← whisker_exchange] + slice_lhs 11 12 => + rw [Mon_.mul_assoc_flip] + slice_lhs 10 11 => + rw [associator_inv_naturality_left] + slice_lhs 11 12 => + simp only [← comp_whiskerRight] + rw [Mon_.mul_assoc] + simp only [comp_whiskerRight] + rw [tensorHom_def] + rw [tensor_whiskerLeft] + rw [pentagon_inv_inv_hom_hom_inv_assoc] + slice_lhs 7 8 => + rw [Iso.inv_hom_id] + rw [Category.id_comp] + slice_lhs 5 7 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [← BraidedCategory.hexagon_forward] + simp only [MonoidalCategory.whiskerLeft_comp] + simp only [Mon_.monMonoidalStruct_tensorObj_X, tensor_whiskerLeft, + MonoidalCategory.whiskerLeft_comp, Category.assoc, + whiskerLeft_inv_hom, Category.comp_id, whiskerLeft_hom_inv_assoc, Iso.inv_hom_id_assoc, + pentagon_inv_inv_hom_inv_inv, whisker_assoc, Mon_.mul_assoc, Mon_.tensorUnit_X] + slice_lhs 4 5 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [Iso.inv_hom_id] + simp only [MonoidalCategory.whiskerLeft_comp] + rw [MonoidalCategory.whiskerLeft_id, Category.id_comp] + slice_lhs 3 4 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [BraidedCategory.braiding_naturality_right] + simp only [MonoidalCategory.whiskerLeft_comp] + rw [tensorHom_def'] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 5 6 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [← associator_naturality_right] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 4 5 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [← whisker_exchange] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 5 9 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [associator_inv_naturality_middle_assoc, Iso.hom_inv_id_assoc] + simp only [← comp_whiskerRight] + rw [antipode_right] + simp only [comp_whiskerRight] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 6 7 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [A.X.X.one_mul] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 3 4 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [← BraidedCategory.braiding_naturality_left] + simp only [MonoidalCategory.whiskerLeft_comp] + slice_lhs 4 5 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [← BraidedCategory.braiding_naturality_right] + simp only [MonoidalCategory.whiskerLeft_comp] + rw [← associator_naturality_middle_assoc] + simp only [Mon_.tensorUnit_X, braiding_tensorUnit_right, MonoidalCategory.whiskerLeft_comp] + slice_lhs 6 7 => + simp only [← MonoidalCategory.whiskerLeft_comp] + rw [Iso.inv_hom_id] + simp only [MonoidalCategory.whiskerLeft_comp] + simp only [MonoidalCategory.whiskerLeft_id, Category.id_comp] + slice_lhs 5 6 => + rw [whiskerLeft_rightUnitor, Category.assoc, ← rightUnitor_naturality] + rw [associator_inv_naturality_right_assoc, Iso.hom_inv_id_assoc] + slice_lhs 3 4 => + rw [whisker_exchange] + slice_lhs 1 3 => + simp only [← comp_whiskerRight] + rw [antipode_right] + simp only [comp_whiskerRight] + slice_lhs 2 3 => + rw [← whisker_exchange] + slice_lhs 1 2 => + dsimp + rw [← tensorHom_def] + slice_lhs 2 3 => + rw [rightUnitor_naturality] + simp only [Mon_.tensorUnit_X] + coherence + +theorem mul_antipode (A : Hopf_ C) : + A.X.X.mul ≫ A.antipode = (A.antipode ⊗ A.antipode) ≫ (β_ _ _).hom ≫ A.X.X.mul := by + -- Again, it is a "left inverse equals right inverse" argument in the convolution monoid. + apply left_inv_eq_right_inv + (M := Conv (((Bimon_.toComon_ C).obj A.X) ⊗ ((Bimon_.toComon_ C).obj A.X)) A.X.X) + (a := A.X.X.mul) + · -- Unfold the algebra structure in the convolution monoid, + -- then `simp?, simp only [tensor_μ], simp?`. + erw [Conv.mul_eq, Conv.one_eq] + simp only [Monoidal.transportStruct_tensorObj, Equivalence.symm_functor, + Comon_.Comon_EquivMon_OpOp_inverse, Equivalence.symm_inverse, + Comon_.Comon_EquivMon_OpOp_functor, Comon_.Comon_ToMon_OpOp_obj, Comon_.Mon_OpOpToComon__obj, + unop_tensorObj, Comon_.Mon_OpOpToComon_obj'_X, Mon_.monMonoidalStruct_tensorObj_X, + Comon_.Comon_ToMon_OpOp_obj'_X, Bimon_.toComon__obj_X, Comon_.Mon_OpOpToComon_obj'_comul, + Mon_.tensorObj_mul, Comon_.Comon_ToMon_OpOp_obj'_mul, Bimon_.toComon__obj_comul, unop_comp, + unop_tensorHom, Quiver.Hom.unop_op, whiskerRight_tensor, comp_whiskerRight, Category.assoc, + Comon_.Mon_OpOpToComon_obj'_counit, Mon_.tensorObj_one, Comon_.Comon_ToMon_OpOp_obj'_one, + Bimon_.toComon__obj_counit, unop_tensorUnit, unop_inv_leftUnitor] + simp only [tensor_μ] + simp only [unop_comp, unop_tensorObj, unop_inv_associator, unop_whiskerLeft, + unop_hom_associator, unop_whiskerRight, unop_hom_braiding, Category.assoc, + pentagon_hom_inv_inv_inv_inv_assoc] + exact mul_antipode₁ A + · erw [Conv.mul_eq, Conv.one_eq] + simp only [Monoidal.transportStruct_tensorObj, Equivalence.symm_functor, + Comon_.Comon_EquivMon_OpOp_inverse, Equivalence.symm_inverse, + Comon_.Comon_EquivMon_OpOp_functor, Comon_.Comon_ToMon_OpOp_obj, Comon_.Mon_OpOpToComon__obj, + unop_tensorObj, Comon_.Mon_OpOpToComon_obj'_X, Mon_.monMonoidalStruct_tensorObj_X, + Comon_.Comon_ToMon_OpOp_obj'_X, Bimon_.toComon__obj_X, Comon_.Mon_OpOpToComon_obj'_comul, + Mon_.tensorObj_mul, Comon_.Comon_ToMon_OpOp_obj'_mul, Bimon_.toComon__obj_comul, unop_comp, + unop_tensorHom, Quiver.Hom.unop_op, whiskerRight_tensor, + BraidedCategory.braiding_naturality_assoc, MonoidalCategory.whiskerLeft_comp, Category.assoc, + Comon_.Mon_OpOpToComon_obj'_counit, Mon_.tensorObj_one, Comon_.Comon_ToMon_OpOp_obj'_one, + Bimon_.toComon__obj_counit, unop_tensorUnit, unop_inv_leftUnitor] + simp only [tensor_μ] + simp only [unop_comp, unop_tensorObj, unop_inv_associator, unop_whiskerLeft, + unop_hom_associator, unop_whiskerRight, unop_hom_braiding, Category.assoc, + pentagon_hom_inv_inv_inv_inv_assoc] + exact mul_antipode₂ A + +/-- +In a commutative Hopf algebra, the antipode squares to the identity. +-/ +theorem antipode_antipode (A : Hopf_ C) (comm : (β_ _ _).hom ≫ A.X.X.mul = A.X.X.mul) : + A.antipode ≫ A.antipode = 𝟙 A.X.X.X := by + -- Again, it is a "left inverse equals right inverse" argument in the convolution monoid. + apply left_inv_eq_right_inv + (M := Conv ((Bimon_.toComon_ C).obj A.X) A.X.X) + (a := A.antipode) + · -- Unfold the algebra structure in the convolution monoid, + -- then `simp?`. + erw [Conv.mul_eq, Conv.one_eq] + simp only [Bimon_.toComon__obj_X, Bimon_.toComon__obj_comul, comp_whiskerRight, Category.assoc, + Bimon_.toComon__obj_counit] + rw [← comm, ← tensorHom_def_assoc, ← mul_antipode] + simp + · erw [Conv.mul_eq, Conv.one_eq] + simp + +end Hopf_ + +end diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index c19a846886021..0279c6c4d67d2 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -154,10 +154,11 @@ def mkIso {M N : Mon_ C} (f : M.X ≅ N.X) (one_f : M.one ≫ f.hom = N.one := b slice_rhs 2 3 => rw [mul_f] simp } +@[simps] instance uniqueHomFromTrivial (A : Mon_ C) : Unique (trivial C ⟶ A) where default := { hom := A.one - mul_hom := by dsimp; simp [A.one_mul, unitors_equal] } + mul_hom := by simp [A.one_mul, unitors_equal] } uniq f := by ext simp only [trivial_X] diff --git a/Mathlib/RingTheory/HopfAlgebra.lean b/Mathlib/RingTheory/HopfAlgebra.lean index 20f1a7ca2575f..7356afd1a4d6f 100644 --- a/Mathlib/RingTheory/HopfAlgebra.lean +++ b/Mathlib/RingTheory/HopfAlgebra.lean @@ -28,6 +28,9 @@ agree then the antipodes must also agree). * If `A` is commutative then `antipode` is necessarily a bijection and its square is the identity. +(Note that all three facts have been proved for Hopf bimonoids in an arbitrary braided category, +so we could deduce the facts here from an equivalence `HopfAlgebraCat R ≌ Hopf_ (ModuleCat R)`.) + ## References * From 063d5f4f0e48ce46e0d7efe91a34f791127f72ea Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 14 Aug 2024 09:17:19 +0000 Subject: [PATCH 255/359] fix: add `all` flag to import summary script (#15583) Summary of the PR * Add the `all` flag to the import diff script: running ```bash ./scripts/import_trans_difference.sh all ``` will not limit the output to ~~200 modified files~~ ~130k characters, but will print the whole import diff. * Increase the size cut-off for printing the full output of the `import-diff script` to roughly half the size allowed in GitHub comments. * Add a comment that you can run the script locally. The script switches over between commits and creates new files to store the information that it gathers from the various steps and can be intrusive: the automation is designed more to be ran in CI, with a fresh clone, than locally with possibly uncommitted changes, so I am not too sure that this is a good idea. #15584 shows the new script in action: the printed comment there exceeds the previous limits of the script. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/There.20are.201617.20files.20with.20changed.20transitive.20imports.20.2E.2E.2E) --- scripts/import_trans_difference.sh | 37 +++++++++++++++++++++++------- 1 file changed, 29 insertions(+), 8 deletions(-) diff --git a/scripts/import_trans_difference.sh b/scripts/import_trans_difference.sh index f97a0094cb50f..1fd4102c594d9 100755 --- a/scripts/import_trans_difference.sh +++ b/scripts/import_trans_difference.sh @@ -1,10 +1,13 @@ #!/usr/bin/env bash : <<'BASH_MODULE_DOCS' -`scripts/import_trans_difference.sh ` outputs a full diff of the -change of transitive imports in all the files between `` and ``. +`scripts/import_trans_difference.sh ` outputs a full diff +of the change of transitive imports in all the files between `` and ``. -If the commits are not provided, then it uses the current commit as `commit1` and +The optional flag `` must either be `all` or not be passed. +Without `all`, the script only displays the difference if the output does not exceed 200 lines. + +If the commits are not provided, then the script uses the current commit as `commit1` and current `master` as `commit2`. The output is of the form @@ -18,6 +21,14 @@ The output is of the form with collapsible tabs for file entries with at least 3 files. BASH_MODULE_DOCS +# `all=1` is the flag to print all import changes, without cut-off +all=0 +if [ "${1}" == "all" ] +then + all=1 + shift +fi + if [ -n "${1}" ] then commit1="${1}" @@ -34,7 +45,13 @@ fi #printf 'commit1: %s\ncommit2: %s\n' "$commit1" "$commit2" -currCommit="$(git rev-parse HEAD)" +currCommit="$(git rev-parse --abbrev-ref HEAD)" +# if we are in a detached head, `currCommit` would be the unhelpful `HEAD` +# in this case, we fetch the commit hash +if [ "${currCommit}" == "HEAD" ] +then + currCommit="$(git rev-parse HEAD)" +fi getTransImports () { python3 scripts/count-trans-deps.py Mathlib | @@ -55,16 +72,20 @@ 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, '{ diff[$1]+=$2 } END { - con=0 + (awk -F, -v all="${all}" -v ghLimit='261752' '{ diff[$1]+=$2 } END { + fileCount=0 + outputLength=0 for(fil in diff) { if(!(diff[fil] == 0)) { - con++ + fileCount++ + outputLength+=length(fil)+4 nums[diff[fil]]++ reds[diff[fil]]=reds[diff[fil]]" `"fil"`" } } - if (200 <= con) { printf("There are %s files with changed transitive imports: this is too many to display!\n", con) } else { + if ((all == 0) && (ghLimit/2 <= outputLength)) { + printf("There are %s files with changed transitive imports taking up over %s characters: this is too many to display!\nYou can run `scripts/import_trans_difference.sh all` locally to see the whole output.", fileCount, outputLength) + } else { for(x in reds) { if (nums[x] <= 2) { printf("|%s|%s|\n", reds[x], x) } else { printf("|
%s files%s
|%s|\n", nums[x], reds[x], x) } From 2c151adef4b0ed500834ece409fd587ac30f8bc1 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Wed, 14 Aug 2024 09:17:20 +0000 Subject: [PATCH 256/359] perf(Algebra.GradedMonoid): scope theorem with weak keys (#15628) The keys for this theorem are `HSMul.hSMul _ _ _ _ _ _` so will be attempted on any term which has an `SMul` in it. --- Mathlib/Algebra/GradedMonoid.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/GradedMonoid.lean b/Mathlib/Algebra/GradedMonoid.lean index 4fb0daab42a52..0179a285f612a 100644 --- a/Mathlib/Algebra/GradedMonoid.lean +++ b/Mathlib/Algebra/GradedMonoid.lean @@ -304,7 +304,7 @@ variable {A} theorem mk_zero_smul {i} (a : A 0) (b : A i) : mk _ (a • b) = mk _ a * mk _ b := Sigma.ext (zero_add _).symm <| eq_rec_heq _ _ -@[simp] +@[scoped simp] theorem GradeZero.smul_eq_mul (a b : A 0) : a • b = a * b := rfl From 35b04e0ac0469c18e4441df7a09133ea563c6eba Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 14 Aug 2024 09:17:21 +0000 Subject: [PATCH 257/359] chore(SetTheory/Ordinal/FixedPoint): Rewrite dumb theorems about zero function (#15755) Two theorems referred to the function `fun x => 0 * x`, which is just the zero function. Also, added applied versions of theorems referring to function equality. --- Mathlib/SetTheory/Ordinal/FixedPoint.lean | 48 +++++++++++++---------- 1 file changed, 27 insertions(+), 21 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index d6d8d3151c634..aa8d8d9c57d1b 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -442,7 +442,7 @@ theorem deriv_eq_derivFamily (f : Ordinal → Ordinal) : deriv f = derivFamily f rfl @[simp] -theorem deriv_zero (f) : deriv f 0 = nfp f 0 := +theorem deriv_zero_right (f) : deriv f 0 = nfp f 0 := derivFamily_zero _ @[simp] @@ -477,6 +477,27 @@ theorem deriv_eq_enumOrd (H : IsNormal f) : deriv f = enumOrd (Function.fixedPoi theorem deriv_eq_id_of_nfp_eq_id {f : Ordinal → Ordinal} (h : nfp f = id) : deriv f = id := (IsNormal.eq_iff_zero_and_succ (deriv_isNormal _) IsNormal.refl).2 <| by simp [h] +@[simp] +theorem nfp_zero : nfp 0 = id := by + rw [← sup_iterate_eq_nfp] + refine funext fun a => (sup_le fun n => ?_).antisymm (le_sup (fun n => 0^[n] a) 0) + induction' n with n _ + · rfl + rw [Function.iterate_succ'] + exact Ordinal.zero_le a + +theorem nfp_zero_left (a) : nfp 0 a = a := by + rw [nfp_zero] + rfl + +@[simp] +theorem deriv_zero : deriv 0 = id := + deriv_eq_id_of_nfp_eq_id nfp_zero + +theorem deriv_zero_left (a) : deriv 0 a = a := by + rw [deriv_zero] + rfl + end /-! ### Fixed points of addition -/ @@ -498,7 +519,7 @@ theorem nfp_add_eq_mul_omega {a b} (hba : b ≤ a * omega) : nfp (a + ·) b = a theorem add_eq_right_iff_mul_omega_le {a b : Ordinal} : a + b = b ↔ a * omega ≤ b := by refine ⟨fun h => ?_, fun h => ?_⟩ - · rw [← nfp_add_zero a, ← deriv_zero] + · rw [← nfp_add_zero a, ← deriv_zero_right] cases' (add_isNormal a).fp_iff_deriv.1 h with c hc rw [← hc] exact (deriv_isNormal _).monotone (Ordinal.zero_le _) @@ -514,7 +535,7 @@ theorem deriv_add_eq_mul_omega_add (a b : Ordinal.{u}) : deriv (a + ·) b = a * revert b rw [← funext_iff, IsNormal.eq_iff_zero_and_succ (deriv_isNormal _) (add_isNormal _)] refine ⟨?_, fun a h => ?_⟩ - · rw [deriv_zero, add_zero] + · rw [deriv_zero_right, add_zero] exact nfp_add_zero a · rw [deriv_succ, h, add_succ] exact nfp_eq_self (add_eq_right_iff_mul_omega_le.2 ((le_add_right _ _).trans (le_succ _))) @@ -542,27 +563,12 @@ theorem nfp_mul_zero (a : Ordinal) : nfp (a * ·) 0 = 0 := by induction' n with n hn; · rfl dsimp only; rwa [iterate_succ_apply, mul_zero] -@[simp] -theorem nfp_zero_mul : nfp (HMul.hMul 0) = id := by - rw [← sup_iterate_eq_nfp] - refine funext fun a => (sup_le fun n => ?_).antisymm (le_sup (fun n => (0 * ·)^[n] a) 0) - induction' n with n _ - · rfl - rw [Function.iterate_succ'] - change 0 * _ ≤ a - rw [zero_mul] - exact Ordinal.zero_le a - -@[simp] -theorem deriv_mul_zero : deriv (HMul.hMul 0) = id := - deriv_eq_id_of_nfp_eq_id nfp_zero_mul - theorem nfp_mul_eq_opow_omega {a b : Ordinal} (hb : 0 < b) (hba : b ≤ (a^omega)) : nfp (a * ·) b = (a^omega.{u}) := by rcases eq_zero_or_pos a with ha | ha · rw [ha, zero_opow omega_ne_zero] at hba ⊢ - rw [Ordinal.le_zero.1 hba, nfp_zero_mul] - rfl + simp_rw [Ordinal.le_zero.1 hba, zero_mul] + exact nfp_zero_left 0 apply le_antisymm · apply nfp_le_fp (mul_isNormal ha).monotone hba rw [← opow_one_add, one_add_omega] @@ -623,7 +629,7 @@ theorem deriv_mul_eq_opow_omega_mul {a : Ordinal.{u}} (ha : 0 < a) (b) : rw [← funext_iff, IsNormal.eq_iff_zero_and_succ (deriv_isNormal _) (mul_isNormal (opow_pos omega ha))] refine ⟨?_, fun c h => ?_⟩ - · dsimp only; rw [deriv_zero, nfp_mul_zero, mul_zero] + · dsimp only; rw [deriv_zero_right, nfp_mul_zero, mul_zero] · rw [deriv_succ, h] exact nfp_mul_opow_omega_add c ha zero_lt_one (one_le_iff_pos.2 (opow_pos _ ha)) From 787495f3353d04aaae4b2c992c7e000400ca9413 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 14 Aug 2024 09:17:23 +0000 Subject: [PATCH 258/359] fix: `assert_not_exists` takes at least one input (#15795) [Reported on Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/assert_not_imported/near/462220460). --- Mathlib/Util/AssertExists.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Util/AssertExists.lean b/Mathlib/Util/AssertExists.lean index 81b63120ee16e..0d8637f9604d7 100644 --- a/Mathlib/Util/AssertExists.lean +++ b/Mathlib/Util/AssertExists.lean @@ -75,7 +75,7 @@ among the transitive imports of the current file. It also checks that each one of `m₁ m₂ ... mₙ` is actually the name of an existing module, just one that is not currently imported! -/ -elab "assert_not_imported " ids:ident* : command => do +elab "assert_not_imported " ids:ident+ : command => do let mods := (← getEnv).allImportedModuleNames for id in ids do if mods.contains id.getId then logWarningAt id m!"the module '{id}' is (transitively) imported" From 3158e96910bd757673e59e9d9ce669ef41a3a9b0 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 14 Aug 2024 10:13:09 +0000 Subject: [PATCH 259/359] chore: ensure Batteries is built in lake test (#15635) --- scripts/test.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/scripts/test.lean b/scripts/test.lean index 57c05c325fd17..f02a621637377 100644 --- a/scripts/test.lean +++ b/scripts/test.lean @@ -14,7 +14,8 @@ When https://github.com/leanprover/lean4/issues/4121 is resolved, this file can be replaced with a line in `lakefile.lean`. -/ def main (args : List String) : IO Unit := do - -- ProofWidgets may not have been built by `lake build` yet, but it is needed by some tests. - _ ← (← spawn { cmd := "lake", args := #["build", "ProofWidgets"] }).wait + -- ProofWidgets and Batteries may not have been completely built by `lake build` yet, + -- but they are needed by some tests. + _ ← (← spawn { cmd := "lake", args := #["build", "ProofWidgets", "Batteries"] }).wait let exitcode ← (← spawn { cmd := "lake", args := #["exe", "batteries/test"] ++ args }).wait exit exitcode.toUInt8 From e2c80d427061275467640113b8b8eefa7aadf728 Mon Sep 17 00:00:00 2001 From: grunweg Date: Wed, 14 Aug 2024 10:13:10 +0000 Subject: [PATCH 260/359] chore(Geometry): remove remaining `open Classical` (#15669) --- Mathlib/Geometry/Euclidean/Basic.lean | 4 +--- Mathlib/Geometry/Euclidean/Circumcenter.lean | 7 ++++--- Mathlib/Geometry/Euclidean/MongePoint.lean | 5 ++--- Mathlib/Geometry/Euclidean/Triangle.lean | 11 +---------- .../Geometry/Manifold/LocalInvariantProperties.lean | 5 +---- Mathlib/Geometry/Manifold/MFDeriv/Defs.lean | 4 +++- Mathlib/Geometry/Manifold/PartitionOfUnity.lean | 8 +++++--- Mathlib/Geometry/Manifold/WhitneyEmbedding.lean | 4 +--- 8 files changed, 18 insertions(+), 30 deletions(-) diff --git a/Mathlib/Geometry/Euclidean/Basic.lean b/Mathlib/Geometry/Euclidean/Basic.lean index 61e9e1a2ede92..3431ffc86bbca 100644 --- a/Mathlib/Geometry/Euclidean/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Basic.lean @@ -44,11 +44,8 @@ theorems that need it. -/ - noncomputable section -open scoped Classical - open RealInnerProductSpace namespace EuclideanGeometry @@ -164,6 +161,7 @@ theorem eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two {s : AffineSubspace intro v hv have hr : Set.range b = {c₂ -ᵥ c₁, p₂ -ᵥ p₁} := by have hu : (Finset.univ : Finset (Fin 2)) = {0, 1} := by decide + classical rw [← Fintype.coe_image_univ, hu] simp [b] rw [← hbs, hr, Submodule.mem_span_insert] at hv diff --git a/Mathlib/Geometry/Euclidean/Circumcenter.lean b/Mathlib/Geometry/Euclidean/Circumcenter.lean index 9961e0140f7fa..6ab81a7fe36a9 100644 --- a/Mathlib/Geometry/Euclidean/Circumcenter.lean +++ b/Mathlib/Geometry/Euclidean/Circumcenter.lean @@ -27,11 +27,8 @@ the circumcenter. -/ - noncomputable section -open scoped Classical - open RealInnerProductSpace namespace EuclideanGeometry @@ -205,6 +202,7 @@ theorem _root_.AffineIndependent.existsUnique_dist_eq {ι : Type*} [hne : Nonemp rw [hi default, hdist] · have i := hne.some let ι2 := { x // x ≠ i } + classical have hc : Fintype.card ι2 = m + 1 := by rw [Fintype.card_of_subtype (Finset.univ.filter fun x => x ≠ i)] · rw [Finset.filter_not] @@ -498,6 +496,7 @@ def pointIndexEmbedding (n : ℕ) : Fin (n + 1) ↪ PointsWithCircumcenterIndex theorem sum_pointsWithCircumcenter {α : Type*} [AddCommMonoid α] {n : ℕ} (f : PointsWithCircumcenterIndex n → α) : ∑ i, f i = (∑ i : Fin (n + 1), f (pointIndex i)) + f circumcenterIndex := by + classical have h : univ = insert circumcenterIndex (univ.map (pointIndexEmbedding n)) := by ext x refine ⟨fun h => ?_, fun _ => mem_univ _⟩ @@ -539,6 +538,7 @@ def pointWeightsWithCircumcenter {n : ℕ} (i : Fin (n + 1)) : PointsWithCircumc @[simp] theorem sum_pointWeightsWithCircumcenter {n : ℕ} (i : Fin (n + 1)) : ∑ j, pointWeightsWithCircumcenter i j = 1 := by + classical convert sum_ite_eq' univ (pointIndex i) (Function.const _ (1 : ℝ)) with j · cases j <;> simp [pointWeightsWithCircumcenter] · simp @@ -599,6 +599,7 @@ def circumcenterWeightsWithCircumcenter (n : ℕ) : PointsWithCircumcenterIndex @[simp] theorem sum_circumcenterWeightsWithCircumcenter (n : ℕ) : ∑ i, circumcenterWeightsWithCircumcenter n i = 1 := by + classical convert sum_ite_eq' univ circumcenterIndex (Function.const _ (1 : ℝ)) with j · cases j <;> simp [circumcenterWeightsWithCircumcenter] · simp diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index 95434096a7f94..6343ce224c3da 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -46,11 +46,8 @@ generalization, the Monge point of a simplex. -/ - noncomputable section -open scoped Classical - open scoped RealInnerProductSpace namespace Affine @@ -364,6 +361,7 @@ theorem finrank_direction_altitude {n : ℕ} (s : Simplex ℝ P (n + 1)) (i : Fi (vectorSpan_mono ℝ (Set.image_subset_range s.points ↑(univ.erase i))) have hc : card (univ.erase i) = n + 1 := by rw [card_erase_of_mem (mem_univ _)]; simp refine add_left_cancel (_root_.trans h ?_) + classical rw [s.independent.finrank_vectorSpan (Fintype.card_fin _), ← Finset.coe_image, s.independent.finrank_vectorSpan_image_finset hc] @@ -652,6 +650,7 @@ theorem exists_of_range_subset_orthocentricSystem {t : Triangle ℝ P} · right have hs := Set.subset_diff_singleton hps h rw [Set.insert_diff_self_of_not_mem ho] at hs + classical refine Set.eq_of_subset_of_card_le hs ?_ rw [Set.card_range_of_injective hpi, Set.card_range_of_injective t.independent.injective] diff --git a/Mathlib/Geometry/Euclidean/Triangle.lean b/Mathlib/Geometry/Euclidean/Triangle.lean index c76fa197f9315..5a28ab76286ee 100644 --- a/Mathlib/Geometry/Euclidean/Triangle.lean +++ b/Mathlib/Geometry/Euclidean/Triangle.lean @@ -33,16 +33,9 @@ unnecessarily. -/ - noncomputable section -open scoped CharZero - -open scoped Classical - -open scoped Real - -open scoped RealInnerProductSpace +open scoped CharZero Real RealInnerProductSpace namespace InnerProductGeometry @@ -250,9 +243,7 @@ This section develops some geometrical definitions and results on (possibly degenerate) triangles in Euclidean affine spaces. -/ - open InnerProductGeometry - open scoped EuclideanGeometry variable {V : Type*} {P : Type*} [NormedAddCommGroup V] [InnerProductSpace ℝ V] [MetricSpace P] diff --git a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean index 13b8450ed5ea7..c464a37a39429 100644 --- a/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean +++ b/Mathlib/Geometry/Manifold/LocalInvariantProperties.lean @@ -42,13 +42,10 @@ coincide on `s`, then `LiftPropWithinAt P g' s x` holds. We can't call it in the one for `LiftPropWithinAt`. -/ - noncomputable section -open scoped Classical -open Manifold Topology - open Set Filter TopologicalSpace +open scoped Manifold Topology variable {H M H' M' X : Type*} variable [TopologicalSpace H] [TopologicalSpace M] [ChartedSpace H M] diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean index 7fd0dbf7af7cc..6a68a762193e5 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean @@ -98,7 +98,7 @@ derivative, manifold noncomputable section -open scoped Classical Topology Manifold +open scoped Topology open Set ChartedSpace section DerivativesDefinitions @@ -296,6 +296,7 @@ def HasMFDerivAt (f : M → M') (x : M) (f' : TangentSpace I x →L[𝕜] Tangen ContinuousAt f x ∧ HasFDerivWithinAt (writtenInExtChartAt I I' x f : E → E') f' (range I) ((extChartAt I x) x) +open Classical in /-- Let `f` be a function between two smooth manifolds. Then `mfderivWithin I I' f s x` is the derivative of `f` at `x` within `s`, as a continuous linear map from the tangent space at `x` to the tangent space at `f x`. -/ @@ -306,6 +307,7 @@ def mfderivWithin (f : M → M') (s : Set M) (x : M) : TangentSpace I x →L[ _) else 0 +open Classical in /-- Let `f` be a function between two smooth manifolds. Then `mfderiv I I' f x` is the derivative of `f` at `x`, as a continuous linear map from the tangent space at `x` to the tangent space at `f x`. -/ diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index 7b50a99e3073a..23d818e090332 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -55,12 +55,10 @@ function from local functions. smooth bump function, partition of unity -/ - universe uι uE uH uM uF open Function Filter FiniteDimensional Set - -open scoped Topology Manifold Classical Filter +open scoped Topology Manifold noncomputable section @@ -466,12 +464,14 @@ theorem toSmoothPartitionOfUnity_apply (i : ι) (x : M) : fs.toSmoothPartitionOfUnity i x = fs i x * ∏ᶠ (j) (_ : WellOrderingRel j i), (1 - fs j x) := rfl +open Classical in theorem toSmoothPartitionOfUnity_eq_mul_prod (i : ι) (x : M) (t : Finset ι) (ht : ∀ j, WellOrderingRel j i → fs j x ≠ 0 → j ∈ t) : fs.toSmoothPartitionOfUnity i x = fs i x * ∏ j ∈ t.filter fun j => WellOrderingRel j i, (1 - fs j x) := fs.toBumpCovering.toPartitionOfUnity_eq_mul_prod i x t ht +open Classical in theorem exists_finset_toSmoothPartitionOfUnity_eventuallyEq (i : ι) (x : M) : ∃ t : Finset ι, fs.toSmoothPartitionOfUnity i =ᶠ[𝓝 x] @@ -553,6 +553,7 @@ namespace SmoothPartitionOfUnity defined as an example for `Inhabited` instance. -/ 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] @@ -739,6 +740,7 @@ theorem exists_msmooth_support_eq_eq_one_iff · have : 0 < f x := lt_of_le_of_ne (f_pos x) (Ne.symm xs) linarith [g_pos x] · have : 0 < g x := by + classical apply lt_of_le_of_ne (g_pos x) (Ne.symm ?_) rw [← mem_support, g_supp] contrapose! xs diff --git a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean index 09569f2558908..e99bb7ea3797f 100644 --- a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean +++ b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean @@ -25,7 +25,6 @@ for sufficiently large `n` there exists a smooth embedding `M → ℝ^n`. partition of unity, smooth bump function, whitney theorem -/ - universe uι uE uH uM variable {ι : Type uι} {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] @@ -33,8 +32,7 @@ variable {ι : Type uι} {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E {M : Type uM} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] open Function Filter FiniteDimensional Set - -open scoped Topology Manifold Classical Filter +open scoped Manifold noncomputable section From d69f8430cdb777b66d2e6351f71ed5b7a99442a4 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Wed, 14 Aug 2024 10:13:12 +0000 Subject: [PATCH 261/359] chore: protect TensorProduct.{map_one, map_mul} (#15754) It's quite common to use the root ones, even with the TensorProduct namespace opened. --- Mathlib/LinearAlgebra/TensorProduct/Basic.lean | 8 ++++---- .../TensorProduct/Graded/Internal.lean | 12 ++++++------ Mathlib/RepresentationTheory/Basic.lean | 4 ++-- Mathlib/RingTheory/Bialgebra/Hom.lean | 2 +- Mathlib/RingTheory/Generators.lean | 16 ++++++++-------- Mathlib/RingTheory/IsTensorProduct.lean | 18 ++++++++---------- .../RingTheory/Kaehler/CotangentComplex.lean | 16 +++++++--------- Mathlib/RingTheory/Presentation.lean | 9 ++++----- Mathlib/RingTheory/Smooth/Kaehler.lean | 12 ++++++------ Mathlib/RingTheory/TensorProduct/Basic.lean | 6 +++--- 10 files changed, 49 insertions(+), 54 deletions(-) diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index 8cc95132b62e1..f10afce554362 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -827,10 +827,10 @@ theorem map_id : map (id : M →ₗ[R] M) (id : N →ₗ[R] N) = .id := by simp only [mk_apply, id_coe, compr₂_apply, _root_.id, map_tmul] @[simp] -theorem map_one : map (1 : M →ₗ[R] M) (1 : N →ₗ[R] N) = 1 := +protected theorem map_one : map (1 : M →ₗ[R] M) (1 : N →ₗ[R] N) = 1 := map_id -theorem map_mul (f₁ f₂ : M →ₗ[R] M) (g₁ g₂ : N →ₗ[R] N) : +protected theorem map_mul (f₁ f₂ : M →ₗ[R] M) (g₁ g₂ : N →ₗ[R] N) : map (f₁ * f₂) (g₁ * g₂) = map f₁ g₁ * map f₂ g₂ := map_comp f₁ f₂ g₁ g₂ @@ -838,8 +838,8 @@ theorem map_mul (f₁ f₂ : M →ₗ[R] M) (g₁ g₂ : N →ₗ[R] N) : protected theorem map_pow (f : M →ₗ[R] M) (g : N →ₗ[R] N) (n : ℕ) : map f g ^ n = map (f ^ n) (g ^ n) := by induction' n with n ih - · simp only [Nat.zero_eq, pow_zero, map_one] - · simp only [pow_succ', ih, map_mul] + · simp only [pow_zero, TensorProduct.map_one] + · simp only [pow_succ', ih, TensorProduct.map_mul] theorem map_add_left (f₁ f₂ : M →ₗ[R] P) (g : N →ₗ[R] Q) : map (f₁ + f₂) g = map f₁ g + map f₂ g := by diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean index b66de04b692fd..d351d025cd8df 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean @@ -307,7 +307,7 @@ def lift (f : A →ₐ[R] C) (g : B →ₐ[R] C) ∘ₗ ((of R 𝒜 ℬ).symm : 𝒜 ᵍ⊗[R] ℬ →ₗ[R] A ⊗[R] B)) (by dsimp [Algebra.TensorProduct.one_def] - simp only [_root_.map_one, mul_one]) + simp only [map_one, mul_one]) (by rw [LinearMap.map_mul_iff] ext a₁ : 3 @@ -320,7 +320,7 @@ def lift (f : A →ₐ[R] C) (g : B →ₐ[R] C) rw [@Units.smul_def _ _ (_) (_), ← Int.cast_smul_eq_zsmul R, map_smul, map_smul, map_smul] rw [Int.cast_smul_eq_zsmul R, ← @Units.smul_def _ _ (_) (_)] rw [of_symm_of, map_tmul, LinearMap.mul'_apply] - simp_rw [AlgHom.toLinearMap_apply, _root_.map_mul] + simp_rw [AlgHom.toLinearMap_apply, map_mul] simp_rw [mul_assoc (f a₁), ← mul_assoc _ _ (g b₂), h_anti_commutes, mul_smul_comm, smul_mul_assoc, smul_smul, Int.units_mul_self, one_smul]) @@ -340,14 +340,14 @@ def liftEquiv : toFun fg := lift 𝒜 ℬ _ _ fg.prop invFun F := ⟨(F.comp (includeLeft 𝒜 ℬ), F.comp (includeRight 𝒜 ℬ)), fun i j a b => by dsimp - rw [← _root_.map_mul, ← _root_.map_mul F, tmul_coe_mul_coe_tmul, one_mul, mul_one, - AlgHom.map_smul_of_tower, tmul_one_mul_one_tmul, smul_smul, Int.units_mul_self, one_smul]⟩ - left_inv fg := by ext <;> (dsimp; simp only [_root_.map_one, mul_one, one_mul]) + rw [← map_mul, ← map_mul F, tmul_coe_mul_coe_tmul, one_mul, mul_one, AlgHom.map_smul_of_tower, + tmul_one_mul_one_tmul, smul_smul, Int.units_mul_self, one_smul]⟩ + left_inv fg := by ext <;> (dsimp; simp only [map_one, mul_one, one_mul]) right_inv F := by apply AlgHom.toLinearMap_injective ext dsimp - rw [← _root_.map_mul, tmul_one_mul_one_tmul] + rw [← map_mul, tmul_one_mul_one_tmul] /-- Two algebra morphism from the graded tensor product agree if their compositions with the left and right inclusions agree. -/ diff --git a/Mathlib/RepresentationTheory/Basic.lean b/Mathlib/RepresentationTheory/Basic.lean index 380582b4bd4ce..6f61040dc4ac8 100644 --- a/Mathlib/RepresentationTheory/Basic.lean +++ b/Mathlib/RepresentationTheory/Basic.lean @@ -371,8 +371,8 @@ tensor product `V ⊗[k] W`. -/ noncomputable def tprod : Representation k G (V ⊗[k] W) where toFun g := TensorProduct.map (ρV g) (ρW g) - map_one' := by simp only [_root_.map_one, TensorProduct.map_one] - map_mul' g h := by simp only [_root_.map_mul, TensorProduct.map_mul] + map_one' := by simp only [map_one, TensorProduct.map_one] + map_mul' g h := by simp only [map_mul, TensorProduct.map_mul] local notation ρV " ⊗ " ρW => tprod ρV ρW diff --git a/Mathlib/RingTheory/Bialgebra/Hom.lean b/Mathlib/RingTheory/Bialgebra/Hom.lean index 78feadf839f29..fe171794c8b66 100644 --- a/Mathlib/RingTheory/Bialgebra/Hom.lean +++ b/Mathlib/RingTheory/Bialgebra/Hom.lean @@ -67,7 +67,7 @@ instance (priority := 100) toAlgHomClass : AlgHomClass F R A B where map_add := map_add map_zero := map_zero commutes := fun c r => by - simp only [Algebra.algebraMap_eq_smul_one, map_smul, _root_.map_one] + simp only [Algebra.algebraMap_eq_smul_one, map_smul, map_one] /-- Turn an element of a type `F` satisfying `BialgHomClass F R A B` into an actual `BialgHom`. This is declared as the default coercion from `F` to `A →ₐc[R] B`. -/ diff --git a/Mathlib/RingTheory/Generators.lean b/Mathlib/RingTheory/Generators.lean index d55ec361f5c74..0aac65c1e9c25 100644 --- a/Mathlib/RingTheory/Generators.lean +++ b/Mathlib/RingTheory/Generators.lean @@ -145,7 +145,7 @@ def localizationAway : Generators R S where letI n : ℕ := (IsLocalization.Away.sec r s).2 C a * X () ^ n aeval_val_σ' s := by - rw [_root_.map_mul, algHom_C, map_pow, aeval_X] + rw [map_mul, algHom_C, map_pow, aeval_X] simp only [← IsLocalization.Away.sec_spec, map_pow, IsLocalization.Away.invSelf] rw [← IsLocalization.mk'_pow, one_pow, ← IsLocalization.mk'_one (M := Submonoid.powers r) S r] rw [← IsLocalization.mk'_pow, one_pow, mul_assoc, ← IsLocalization.mk'_mul] @@ -168,9 +168,9 @@ def comp (Q : Generators S T) (P : Generators R S) : Generators R T where have (x : P.Ring) : aeval (algebraMap S T ∘ P.val) x = algebraMap S T (aeval P.val x) := by rw [map_aeval, aeval_def, coe_eval₂Hom, ← IsScalarTower.algebraMap_eq, Function.comp] conv_rhs => rw [← Q.aeval_val_σ s, ← (Q.σ s).sum_single] - simp only [map_finsupp_sum, _root_.map_mul, aeval_rename, Sum.elim_comp_inr, this, aeval_val_σ, - aeval_monomial, _root_.map_one, Finsupp.prod_mapDomain_index_inj Sum.inl_injective, - Sum.elim_inl, one_mul, single_eq_monomial] + simp only [map_finsupp_sum, map_mul, aeval_rename, Sum.elim_comp_inr, this, aeval_val_σ, + aeval_monomial, map_one, Finsupp.prod_mapDomain_index_inj Sum.inl_injective, Sum.elim_inl, + one_mul, single_eq_monomial] variable (S) in /-- If `R → S → T` is a tower of algebras, a family of generators `R[X] → T` @@ -310,7 +310,7 @@ noncomputable def Hom.comp [IsScalarTower R' R'' S''] [IsScalarTower R' S' S''] induction g.val x using MvPolynomial.induction_on with | h_C r => simp [← IsScalarTower.algebraMap_apply] | h_add x y hx hy => simp only [map_add, hx, hy] - | h_X p i hp => simp only [_root_.map_mul, hp, aeval_X, aeval_val] + | h_X p i hp => simp only [map_mul, hp, aeval_X, aeval_val] @[simp] lemma Hom.comp_id [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] (f : Hom P P') : @@ -334,7 +334,7 @@ lemma Hom.toAlgHom_comp_apply induction x using MvPolynomial.induction_on with | h_C r => simp only [← MvPolynomial.algebraMap_eq, AlgHom.map_algebraMap] | h_add x y hx hy => simp only [map_add, hx, hy] - | h_X p i hp => simp only [_root_.map_mul, hp, toAlgHom_X, comp_val]; rfl + | h_X p i hp => simp only [map_mul, hp, toAlgHom_X, comp_val]; rfl variable {T} [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] @@ -438,7 +438,7 @@ instance {R₁ R₂} [CommRing R₁] [CommRing R₂] [Algebra R₁ S] [Algebra R constructor intros r s m show algebraMap R₂ S (r • s) • m = (algebraMap _ S r) • (algebraMap _ S s) • m - rw [Algebra.smul_def, _root_.map_mul, mul_smul, ← IsScalarTower.algebraMap_apply] + rw [Algebra.smul_def, map_mul, mul_smul, ← IsScalarTower.algebraMap_apply] lemma Cotangent.val_smul''' {R₀} [CommRing R₀] [Algebra R₀ S] (r : R₀) (x : P.Cotangent) : (r • x).val = P.σ (algebraMap R₀ S r) • x.val := rfl @@ -488,7 +488,7 @@ def Cotangent.map (f : Hom P P') : P.Cotangent →ₗ[S] P'.Cotangent where ← algebraMap_apply, algebraMap_smul, val_smul', val_of, ← (Ideal.toCotangent _).map_smul] congr 1 ext1 - simp only [SetLike.val_smul, smul_eq_mul, _root_.map_mul] + simp only [SetLike.val_smul, smul_eq_mul, map_mul] @[simp] lemma Cotangent.map_mk (f : Hom P P') (x) : diff --git a/Mathlib/RingTheory/IsTensorProduct.lean b/Mathlib/RingTheory/IsTensorProduct.lean index 07110fac07464..942b0ab45ddf4 100644 --- a/Mathlib/RingTheory/IsTensorProduct.lean +++ b/Mathlib/RingTheory/IsTensorProduct.lean @@ -213,10 +213,9 @@ noncomputable nonrec def IsBaseChange.equiv : S ⊗[R] M ≃ₗ[S] N := · rw [smul_zero, map_zero, smul_zero] · intro x y -- porting note (#10745): was simp [smul_tmul', Algebra.ofId_apply] - simp only [Algebra.linearMap_apply, lift.tmul, smul_eq_mul, - LinearMap.mul_apply, LinearMap.smul_apply, IsTensorProduct.equiv_apply, - Module.algebraMap_end_apply, _root_.map_mul, smul_tmul', eq_self_iff_true, - LinearMap.coe_restrictScalars, LinearMap.flip_apply] + simp only [Algebra.linearMap_apply, lift.tmul, smul_eq_mul, LinearMap.mul_apply, + LinearMap.smul_apply, IsTensorProduct.equiv_apply, Module.algebraMap_end_apply, map_mul, + smul_tmul', eq_self_iff_true, LinearMap.coe_restrictScalars, LinearMap.flip_apply] · intro x y hx hy rw [map_add, smul_add, map_add, smul_add, hx, hy] } @@ -353,8 +352,7 @@ theorem Algebra.IsPushout.symm (h : Algebra.IsPushout R S R' S') : Algebra.IsPus · intro x y simp only [smul_tmul', smul_eq_mul, TensorProduct.comm_tmul, smul_def, TensorProduct.algebraMap_apply, id.map_eq_id, RingHom.id_apply, TensorProduct.tmul_mul_tmul, - one_mul, h.1.equiv_tmul, AlgHom.toLinearMap_apply, _root_.map_mul, - IsScalarTower.coe_toAlgHom'] + one_mul, h.1.equiv_tmul, AlgHom.toLinearMap_apply, map_mul, IsScalarTower.coe_toAlgHom'] ring · intro x y hx hy rw [map_add, map_add, smul_add, map_add, map_add, hx, hy, smul_add] @@ -399,7 +397,7 @@ noncomputable def Algebra.pushoutDesc [H : Algebra.IsPushout R S R' S'] {A : Typ have : ∀ x, H.out.lift g.toLinearMap (algebraMap R' S' x) = g x := H.out.lift_eq _ refine AlgHom.ofLinearMap ((H.out.lift g.toLinearMap).restrictScalars R) ?_ ?_ · dsimp only [LinearMap.restrictScalars_apply] - rw [← (algebraMap R' S').map_one, this, _root_.map_one] + rw [← (algebraMap R' S').map_one, this, map_one] · intro x y refine H.out.inductionOn x _ ?_ ?_ ?_ ?_ · rw [zero_mul, map_zero, zero_mul] @@ -417,7 +415,7 @@ noncomputable def Algebra.pushoutDesc [H : Algebra.IsPushout R S R' S'] {A : Typ · rw [mul_zero, map_zero, mul_zero] · intro y dsimp - rw [← _root_.map_mul, this, this, _root_.map_mul] + rw [← map_mul, this, this, map_mul] · intro s s' e rw [mul_comm, smul_mul_assoc, LinearMap.map_smul, LinearMap.map_smul, mul_comm, e] change f s * (g x * _) = g x * (f s * _) @@ -435,7 +433,7 @@ theorem Algebra.pushoutDesc_left [H : Algebra.IsPushout R S R' S'] {A : Type*} [ show f (r • s) * a = r • (f s * a) by rw [map_smul, smul_mul_assoc] } haveI : IsScalarTower S A A := { smul_assoc := fun r a b => mul_assoc _ _ _ } rw [Algebra.algebraMap_eq_smul_one, pushoutDesc_apply, map_smul, ← - Algebra.pushoutDesc_apply S' f g H, _root_.map_one] + Algebra.pushoutDesc_apply S' f g H, map_one] exact mul_one (f x) theorem Algebra.lift_algHom_comp_left [Algebra.IsPushout R S R' S'] {A : Type*} [Semiring A] @@ -467,7 +465,7 @@ theorem Algebra.IsPushout.algHom_ext [H : Algebra.IsPushout R S R' S'] {A : Type · simp only [map_zero] · exact AlgHom.congr_fun h₁ · intro s s' e - rw [Algebra.smul_def, _root_.map_mul, _root_.map_mul, e] + rw [Algebra.smul_def, map_mul, map_mul, e] congr 1 exact (AlgHom.congr_fun h₂ s : _) · intro s₁ s₂ e₁ e₂ diff --git a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean index 3f687662317e1..bcdc42a93b266 100644 --- a/Mathlib/RingTheory/Kaehler/CotangentComplex.lean +++ b/Mathlib/RingTheory/Kaehler/CotangentComplex.lean @@ -111,7 +111,7 @@ lemma map_tmul (f : Hom P P') (x y) : lemma repr_map (f : Hom P P') (i j) : P'.cotangentSpaceBasis.repr (CotangentSpace.map f (P.cotangentSpaceBasis i)) j = aeval P'.val (pderiv j (f.val i)) := by - simp only [cotangentSpaceBasis_apply, map_tmul, _root_.map_one, Hom.toAlgHom_X, + simp only [cotangentSpaceBasis_apply, map_tmul, map_one, Hom.toAlgHom_X, cotangentSpaceBasis_repr_one_tmul] universe w'' u'' v'' @@ -129,7 +129,7 @@ lemma map_id : CotangentSpace.map (.id P) = LinearMap.id := by apply P.cotangentSpaceBasis.ext intro i - simp only [cotangentSpaceBasis_apply, map_tmul, _root_.map_one, Hom.toAlgHom_X, Hom.id_val, + simp only [cotangentSpaceBasis_apply, map_tmul, map_one, Hom.toAlgHom_X, Hom.id_val, LinearMap.id_coe, id_eq] lemma map_comp (f : Hom P P') (g : Hom P' P'') : @@ -137,7 +137,7 @@ lemma map_comp (f : Hom P P') (g : Hom P' P'') : (CotangentSpace.map g).restrictScalars S ∘ₗ CotangentSpace.map f := by apply P.cotangentSpaceBasis.ext intro i - simp only [cotangentSpaceBasis_apply, map_tmul, _root_.map_one, Hom.toAlgHom_X, Hom.comp_val, + simp only [cotangentSpaceBasis_apply, map_tmul, map_one, Hom.toAlgHom_X, Hom.comp_val, LinearMap.coe_comp, LinearMap.coe_restrictScalars, Function.comp_apply] rfl @@ -148,8 +148,7 @@ lemma map_comp_apply (f : Hom P P') (g : Hom P' P'') (x) : lemma map_cotangentComplex (f : Hom P P') (x) : CotangentSpace.map f (P.cotangentComplex x) = P'.cotangentComplex (.map f x) := by obtain ⟨x, rfl⟩ := Cotangent.mk_surjective x - rw [cotangentComplex_mk, map_tmul, _root_.map_one, Cotangent.map_mk, - cotangentComplex_mk] + rw [cotangentComplex_mk, map_tmul, map_one, Cotangent.map_mk, cotangentComplex_mk] lemma map_comp_cotangentComplex (f : Hom P P') : CotangentSpace.map f ∘ₗ P.cotangentComplex = @@ -180,7 +179,7 @@ lemma Hom.sub_aux (f g : Hom P P') (x y) : coe_eval₂Hom, ← aeval_def, ker, RingHom.mem_ker, map_sub, algebraMap_toAlgHom, aeval_val_σ, sub_self] convert this using 1 - simp only [_root_.map_mul] + simp only [map_mul] ring /-- @@ -222,8 +221,7 @@ def Hom.sub (f g : Hom P P') : P.CotangentSpace →ₗ[S] P'.Cotangent := by simp only [LinearMap.coe_comp, LinearMap.coe_restrictScalars, Function.comp_apply, Cotangent.val_mk, Cotangent.val_zero, Ideal.toCotangent_eq_zero] erw [LinearMap.codRestrict_apply] - simp only [LinearMap.sub_apply, AlgHom.toLinearMap_apply, _root_.map_one, sub_self, - Submodule.zero_mem] + simp only [LinearMap.sub_apply, AlgHom.toLinearMap_apply, map_one, sub_self, Submodule.zero_mem] · intro x y ext simp only [LinearMap.coe_comp, LinearMap.coe_restrictScalars, Function.comp_apply, @@ -249,7 +247,7 @@ lemma CotangentSpace.map_sub_map (f g : Hom P P') : P'.cotangentComplex.restrictScalars S ∘ₗ (f.sub g) := by apply P.cotangentSpaceBasis.ext intro i - simp only [cotangentSpaceBasis_apply, LinearMap.sub_apply, map_tmul, _root_.map_one, + simp only [cotangentSpaceBasis_apply, LinearMap.sub_apply, map_tmul, map_one, Hom.toAlgHom_X, LinearMap.coe_comp, LinearMap.coe_restrictScalars, Function.comp_apply, Hom.sub_one_tmul, cotangentComplex_mk, Hom.subToKer_apply_coe, map_sub, tmul_sub] diff --git a/Mathlib/RingTheory/Presentation.lean b/Mathlib/RingTheory/Presentation.lean index 3212414414659..5dfe5599b9bc1 100644 --- a/Mathlib/RingTheory/Presentation.lean +++ b/Mathlib/RingTheory/Presentation.lean @@ -200,13 +200,13 @@ private lemma span_range_relation_eq_ker_baseChange : id.map_eq_id, RingHom.id_apply, e] erw [← MvPolynomial.algebraMap_eq, AlgEquiv.commutes] simp only [TensorProduct.algebraMap_apply, id.map_eq_id, RingHom.id_apply, - TensorProduct.map_tmul, AlgHom.coe_id, id_eq, _root_.map_one, algebraMap_eq] + TensorProduct.map_tmul, AlgHom.coe_id, id_eq, map_one, algebraMap_eq] erw [aeval_C] simp | h_add p q hp hq => simp only [map_add, hp, hq] | h_X p i hp => - simp only [_root_.map_mul, algebraTensorAlgEquiv_symm_X, hp, TensorProduct.map_tmul, - _root_.map_one, IsScalarTower.coe_toAlgHom', Generators.algebraMap_apply, aeval_X, e] + simp only [map_mul, algebraTensorAlgEquiv_symm_X, hp, TensorProduct.map_tmul, map_one, + IsScalarTower.coe_toAlgHom', Generators.algebraMap_apply, aeval_X, e] congr erw [aeval_X] rw [Generators.baseChange_val] @@ -216,8 +216,7 @@ private lemma span_range_relation_eq_ker_baseChange : Ideal.map_map, Ideal.map_span, ← Set.range_comp] at H' convert H' simp only [AlgHom.toRingHom_eq_coe, RingHom.coe_comp, RingHom.coe_coe, Function.comp_apply, - TensorProduct.includeRight_apply, TensorProduct.lift_tmul, _root_.map_one, mapAlgHom_apply, - one_mul] + TensorProduct.includeRight_apply, TensorProduct.lift_tmul, map_one, mapAlgHom_apply, one_mul] rfl /-- If `P` is a presentation of `S` over `R` and `T` is an `R`-algebra, we diff --git a/Mathlib/RingTheory/Smooth/Kaehler.lean b/Mathlib/RingTheory/Smooth/Kaehler.lean index a5825c5318f5f..06b7cba8d586b 100644 --- a/Mathlib/RingTheory/Smooth/Kaehler.lean +++ b/Mathlib/RingTheory/Smooth/Kaehler.lean @@ -51,10 +51,10 @@ def derivationOfSectionOfKerSqZero (f : P →ₐ[R] S) (hf' : (RingHom.ker f) ^ map_add' x y := by simp only [map_add, AddSubmonoid.mk_add_mk, Subtype.mk.injEq]; ring map_smul' x y := by ext - simp only [Algebra.smul_def, _root_.map_mul, ← IsScalarTower.algebraMap_apply, - AlgHom.commutes, RingHom.id_apply, Submodule.coe_smul_of_tower] + simp only [Algebra.smul_def, map_mul, ← IsScalarTower.algebraMap_apply, AlgHom.commutes, + RingHom.id_apply, Submodule.coe_smul_of_tower] ring - map_one_eq_zero' := by simp only [LinearMap.coe_mk, AddHom.coe_mk, _root_.map_one, sub_self, + map_one_eq_zero' := by simp only [LinearMap.coe_mk, AddHom.coe_mk, map_one, sub_self, AddSubmonoid.mk_eq_zero] leibniz' a b := by have : (a - g (f a)) * (b - g (f b)) = 0 := by @@ -65,8 +65,8 @@ def derivationOfSectionOfKerSqZero (f : P →ₐ[R] S) (hf' : (RingHom.ker f) ^ ext rw [← sub_eq_zero] conv_rhs => rw [← neg_zero, ← this] - simp only [LinearMap.coe_mk, AddHom.coe_mk, _root_.map_mul, SetLike.mk_smul_mk, smul_eq_mul, - mul_sub, AddSubmonoid.mk_add_mk, sub_mul, neg_sub] + simp only [LinearMap.coe_mk, AddHom.coe_mk, map_mul, SetLike.mk_smul_mk, smul_eq_mul, mul_sub, + AddSubmonoid.mk_add_mk, sub_mul, neg_sub] ring variable (hf' : (RingHom.ker (algebraMap P S)) ^ 2 = ⊥) @@ -80,7 +80,7 @@ lemma isScalarTower_of_section_of_ker_sqZero : intro p s m ext show g (p • s) * m = p * (g s * m) - simp only [Algebra.smul_def, _root_.map_mul, mul_assoc, mul_left_comm _ (g s)] + simp only [Algebra.smul_def, map_mul, mul_assoc, mul_left_comm _ (g s)] congr 1 rw [← sub_eq_zero, ← Ideal.mem_bot, ← hf', pow_two, ← sub_mul] refine Ideal.mul_mem_mul ?_ m.2 diff --git a/Mathlib/RingTheory/TensorProduct/Basic.lean b/Mathlib/RingTheory/TensorProduct/Basic.lean index 78c5b9dc2e6f9..228b5be3855a4 100644 --- a/Mathlib/RingTheory/TensorProduct/Basic.lean +++ b/Mathlib/RingTheory/TensorProduct/Basic.lean @@ -462,7 +462,7 @@ theorem ext ⦃f g : (A ⊗[R] B) →ₐ[S] C⦄ ext a b have := congr_arg₂ HMul.hMul (AlgHom.congr_fun ha a) (AlgHom.congr_fun hb b) dsimp at * - rwa [← _root_.map_mul, ← _root_.map_mul, tmul_mul_tmul, one_mul, mul_one] at this + rwa [← map_mul, ← map_mul, tmul_mul_tmul, one_mul, mul_one] at this theorem ext' {g h : A ⊗[R] B →ₐ[S] C} (H : ∀ a b, g (a ⊗ₜ b) = h (a ⊗ₜ b)) : g = h := ext (AlgHom.ext fun _ => H _ _) (AlgHom.ext fun _ => H _ _) @@ -693,8 +693,8 @@ def lift (f : A →ₐ[S] C) (g : B →ₐ[R] C) (hfg : ∀ x y, Commute (f x) ( map_smul' := fun c g => LinearMap.ext fun x => rfl } LinearMap.flip <| (restr ∘ₗ LinearMap.mul S C ∘ₗ f.toLinearMap).flip ∘ₗ g) (fun a₁ a₂ b₁ b₂ => show f (a₁ * a₂) * g (b₁ * b₂) = f a₁ * g b₁ * (f a₂ * g b₂) by - rw [_root_.map_mul, _root_.map_mul, (hfg a₂ b₁).mul_mul_mul_comm]) - (show f 1 * g 1 = 1 by rw [_root_.map_one, _root_.map_one, one_mul]) + rw [map_mul, map_mul, (hfg a₂ b₁).mul_mul_mul_comm]) + (show f 1 * g 1 = 1 by rw [map_one, map_one, one_mul]) @[simp] theorem lift_tmul (f : A →ₐ[S] C) (g : B →ₐ[R] C) (hfg : ∀ x y, Commute (f x) (g y)) From 85ffadc666e4c68a0aa66420b80867c9272b950b Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 14 Aug 2024 10:13:13 +0000 Subject: [PATCH 262/359] chore: inline `getLinterHash`* (#15771) A copy-paste error that propagated to almost all linters -- [pointed out by Eric](https://github.com/leanprover-community/mathlib4/pull/15724#discussion_r1715138669). Resolved by inlining the corresponding definition. In the case of the `refine'` linter, I also recycled one doc-string and removed a `refine` namespace since every declaration already had `refine` in its name anyway. --- Mathlib/Tactic/Linter/HashCommandLinter.lean | 5 +---- Mathlib/Tactic/Linter/Lint.lean | 21 +++++------------- Mathlib/Tactic/Linter/MinImports.lean | 5 +---- Mathlib/Tactic/Linter/OldObtain.lean | 5 +---- Mathlib/Tactic/Linter/RefineLinter.lean | 23 ++++++++++---------- Mathlib/Tactic/Linter/Style.lean | 5 +---- Mathlib/Tactic/Linter/UnusedTactic.lean | 5 +---- 7 files changed, 21 insertions(+), 48 deletions(-) diff --git a/Mathlib/Tactic/Linter/HashCommandLinter.lean b/Mathlib/Tactic/Linter/HashCommandLinter.lean index 09ceae9db6b3a..f50ba4b9f795c 100644 --- a/Mathlib/Tactic/Linter/HashCommandLinter.lean +++ b/Mathlib/Tactic/Linter/HashCommandLinter.lean @@ -37,9 +37,6 @@ namespace HashCommandLinter open Lean Elab -/-- Gets the value of the `linter.hashCommand` option. -/ -def getLinterHash (o : Options) : Bool := Linter.getLinterValue linter.hashCommand o - open Command in /-- Exactly like `withSetOptionIn`, but recursively discards nested uses of `in`. Intended to be used in the `hashCommand` linter, where we want to enter `set_option` `in` commands. @@ -68,7 +65,7 @@ However, in order to avoid local clutter, when `warningAsError` is `false`, the logs a warning only for the `#`-commands that do not already emit a message. -/ def hashCommandLinter : Linter where run := withSetOptionIn' fun stx => do let mod := (← getMainModule).components - if getLinterHash (← getOptions) && + if Linter.getLinterValue linter.hashCommand (← getOptions) && ((← get).messages.toList.isEmpty || warningAsError.get (← getOptions)) && -- we check that the module is either not in `test` or, is `test.HashCommandLinter` (mod.getD 0 default != `test || (mod == [`test, `HashCommandLinter])) diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index bdd5e856f065d..38bd0a9ffc843 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -79,9 +79,6 @@ namespace DupNamespaceLinter open Lean Parser Elab Command Meta -/-- Gets the value of the `linter.dupNamespace` option. -/ -def getLinterDupNamespace (o : Options) : Bool := Linter.getLinterValue linter.dupNamespace o - /-- `getIds stx` extracts the `declId` nodes from the `Syntax` `stx`. If `stx` is an `alias` or an `export`, then it extracts an `ident`, instead of a `declId`. -/ partial @@ -94,7 +91,7 @@ def getIds : Syntax → Array Syntax @[inherit_doc linter.dupNamespace] def dupNamespace : Linter where run := withSetOptionIn fun stx => do - if getLinterDupNamespace (← getOptions) then + if Linter.getLinterValue linter.dupNamespace (← getOptions) then match getIds stx with | #[id] => let ns := (← getScope).currNamespace @@ -129,9 +126,6 @@ register_option linter.missingEnd : Bool := { namespace MissingEnd -/-- Gets the value of the `linter.missingEnd` option. -/ -def getLinterHash (o : Options) : Bool := Linter.getLinterValue linter.missingEnd o - @[inherit_doc Mathlib.Linter.linter.missingEnd] def missingEndLinter : Linter where run := withSetOptionIn fun stx ↦ do -- Only run this linter at the end of a module. @@ -139,7 +133,8 @@ def missingEndLinter : Linter where run := withSetOptionIn fun stx ↦ do -- TODO: once mathlib's Lean version includes leanprover/lean4#4741, make this configurable unless #[`Mathlib, `test, `Archive, `Counterexamples].contains (← getMainModule).getRoot do return - if getLinterHash (← getOptions) && !(← MonadState.get).messages.hasErrors then + if Linter.getLinterValue linter.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. @@ -202,12 +197,9 @@ def unwanted_cdot (stx : Syntax) : Array Syntax := namespace CDotLinter -/-- Gets the value of the `linter.generic` option. -/ -def getLinterHash (o : Options) : Bool := Linter.getLinterValue linter.cdot o - @[inherit_doc linter.cdot] def cdotLinter : Linter where run := withSetOptionIn fun stx => do - unless getLinterHash (← getOptions) do + unless Linter.getLinterValue linter.cdot (← getOptions) do return if (← MonadState.get).messages.hasErrors then return @@ -229,12 +221,9 @@ register_option linter.longLine : Bool := { namespace LongLine -/-- Gets the value of the `linter.longLine` option. -/ -def getLinterHash (o : Options) : Bool := Linter.getLinterValue linter.longLine o - @[inherit_doc Mathlib.Linter.linter.longLine] def longLineLinter : Linter where run := withSetOptionIn fun stx ↦ do - unless getLinterHash (← getOptions) do + unless Linter.getLinterValue linter.longLine (← getOptions) do return if (← MonadState.get).messages.hasErrors then return diff --git a/Mathlib/Tactic/Linter/MinImports.lean b/Mathlib/Tactic/Linter/MinImports.lean index 999aca9d8c390..129150e7718b2 100644 --- a/Mathlib/Tactic/Linter/MinImports.lean +++ b/Mathlib/Tactic/Linter/MinImports.lean @@ -53,12 +53,9 @@ namespace MinImports open Mathlib.Command.MinImports -/-- Gets the value of the `linter.minImports` option. -/ -def getLinterHash (o : Options) : Bool := Linter.getLinterValue linter.minImports o - @[inherit_doc Mathlib.Linter.linter.minImports] def minImportsLinter : Linter where run := withSetOptionIn fun stx => do - unless linter.minImports.get (← getOptions) do + unless Linter.getLinterValue linter.minImports (← getOptions) do return if (← MonadState.get).messages.hasErrors then return diff --git a/Mathlib/Tactic/Linter/OldObtain.lean b/Mathlib/Tactic/Linter/OldObtain.lean index 49f19a5a47846..25069ff70983b 100644 --- a/Mathlib/Tactic/Linter/OldObtain.lean +++ b/Mathlib/Tactic/Linter/OldObtain.lean @@ -66,12 +66,9 @@ register_option linter.oldObtain : Bool := { descr := "enable the `oldObtain` linter" } -/-- Gets the value of the `linter.oldObtain` option. -/ -def getLinterHash (o : Options) : Bool := Linter.getLinterValue linter.oldObtain o - /-- The `oldObtain` linter: see docstring above -/ def oldObtainLinter : Linter where run := withSetOptionIn fun stx => do - unless getLinterHash (← getOptions) do + unless Linter.getLinterValue linter.oldObtain (← getOptions) do return if (← MonadState.get).messages.hasErrors then return diff --git a/Mathlib/Tactic/Linter/RefineLinter.lean b/Mathlib/Tactic/Linter/RefineLinter.lean index d6cb739dbf88c..7518ae9a1b1d5 100644 --- a/Mathlib/Tactic/Linter/RefineLinter.lean +++ b/Mathlib/Tactic/Linter/RefineLinter.lean @@ -20,9 +20,14 @@ This linter is an incentive to discourage uses of `refine'`, without being a ban open Lean Elab -namespace Mathlib.Linter.refine +namespace Mathlib.Linter -/-- The refine linter emits a warning on usages of `refine'`. -/ +/-- The "refine" linter flags usages of the `refine'` tactic. + +The tactics `refine` and `refine'` are similar, but they handle meta-variables slightly differently. +This means that they are not completely interchangeable, nor can one completely replace the other. +However, `refine` is more readable and (heuristically) tends to be more efficient on average. +-/ register_option linter.refine : Bool := { defValue := true descr := "enable the refine linter" @@ -36,17 +41,9 @@ def getRefine' : Syntax → Array Syntax if kind == ``Lean.Parser.Tactic.refine' then rargs.push stx else rargs | _ => default -/-- The "refine" linter flags usages of the `refine'` tactic. - -The tactics `refine` and `refine'` are similar, but they handle meta-variables slightly differently. -This means that they are not completely interchangeable, nor can one completely replace the other. -However, `refine` is more readable and (heuristically) tends to be more efficient on average. --/ -def getLinterHash (o : Options) : Bool := Linter.getLinterValue linter.refine o - -@[inherit_doc getLinterHash] +@[inherit_doc linter.refine] def refineLinter : Linter where run := withSetOptionIn fun _stx => do - unless getLinterHash (← getOptions) do + unless Linter.getLinterValue linter.refine (← getOptions) do return if (← MonadState.get).messages.hasErrors then return @@ -56,3 +53,5 @@ def refineLinter : Linter where run := withSetOptionIn fun _stx => do please strongly consider using `refine` or `apply` instead." initialize addLinter refineLinter + +end Mathlib.Linter diff --git a/Mathlib/Tactic/Linter/Style.lean b/Mathlib/Tactic/Linter/Style.lean index fbceb4a721f86..b7dcce49b4ae3 100644 --- a/Mathlib/Tactic/Linter/Style.lean +++ b/Mathlib/Tactic/Linter/Style.lean @@ -42,9 +42,6 @@ def parse_set_option : Syntax → Option Name def is_set_option : Syntax → Bool := fun stx ↦ parse_set_option stx matches some _name -/-- Gets the value of the `linter.setOption` option. -/ -def getLinterHash (o : Options) : Bool := Linter.getLinterValue linter.setOption o - /-- The `setOption` linter: this lints any `set_option` command, term or tactic which sets a `pp`, `profiler` or `trace` option. @@ -54,7 +51,7 @@ used in production code. (Some tests will intentionally use one of these options; in this case, simply allow the linter.) -/ def setOptionLinter : Linter where run := withSetOptionIn fun stx => do - unless getLinterHash (← getOptions) do + unless Linter.getLinterValue linter.setOption (← getOptions) do return if (← MonadState.get).messages.hasErrors then return diff --git a/Mathlib/Tactic/Linter/UnusedTactic.lean b/Mathlib/Tactic/Linter/UnusedTactic.lean index 425cdc0023e0e..6b9a97032fa33 100644 --- a/Mathlib/Tactic/Linter/UnusedTactic.lean +++ b/Mathlib/Tactic/Linter/UnusedTactic.lean @@ -196,12 +196,9 @@ partial def eraseUsedTactics : InfoTree → M Unit end -/-- Gets the value of the `linter.unusedTactic` option. -/ -def getLinterHash (o : Options) : Bool := Linter.getLinterValue linter.unusedTactic o - /-- The main entry point to the unused tactic linter. -/ def unusedTacticLinter : Linter where run := withSetOptionIn fun stx => do - unless getLinterHash (← getOptions) && (← getInfoState).enabled do + unless Linter.getLinterValue linter.unusedTactic (← getOptions) && (← getInfoState).enabled do return if (← get).messages.hasErrors then return From ca93110a55bca836eab8b0328be6d175dd9d4154 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 14 Aug 2024 10:13:14 +0000 Subject: [PATCH 263/359] chore: cleanup in List/Infix (#15792) --- Mathlib/Data/List/Infix.lean | 15 +++++---------- 1 file changed, 5 insertions(+), 10 deletions(-) diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index 8d4e9784ae69e..4dcb64c9655da 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -186,7 +186,7 @@ instance decidableInfix [DecidableEq α] : ∀ l₁ l₂ : List α, Decidable (l (@instDecidableOr _ _ (l₁.decidablePrefix (b :: l₂)) (l₁.decidableInfix l₂)) infix_cons_iff.symm -theorem prefix_take_le_iff {L : List (List (Option α))} (hm : m < L.length) : +theorem prefix_take_le_iff {L : List α} (hm : m < L.length) : L.take m <+: L.take n ↔ m ≤ n := by simp only [prefix_iff_eq_take, length_take] induction m generalizing L n with @@ -202,21 +202,16 @@ theorem prefix_take_le_iff {L : List (List (Option α))} (hm : m < L.length) : simp only [length_cons, succ_eq_add_one, Nat.add_lt_add_iff_right] at hm simp [← @IH n ls hm, Nat.min_eq_left, Nat.le_of_lt hm] +@[deprecated cons_prefix_cons (since := "2024-08-14")] theorem cons_prefix_iff : a :: l₁ <+: b :: l₂ ↔ a = b ∧ l₁ <+: l₂ := by - constructor - · rintro ⟨L, hL⟩ - simp only [cons_append] at hL - injection hL with hLLeft hLRight - exact ⟨hLLeft, ⟨L, hLRight⟩⟩ - · rintro ⟨rfl, h⟩ - rwa [prefix_cons_inj] + simp protected theorem IsPrefix.map (h : l₁ <+: l₂) (f : α → β) : l₁.map f <+: l₂.map f := by induction' l₁ with hd tl hl generalizing l₂ · simp only [nil_prefix, map_nil] · cases' l₂ with hd₂ tl₂ · simpa only using eq_nil_of_prefix_nil h - · rw [cons_prefix_iff] at h + · rw [cons_prefix_cons] at h simp only [List.map_cons, h, prefix_cons_inj, hl, map] protected theorem IsPrefix.filterMap (h : l₁ <+: l₂) (f : α → Option β) : @@ -225,7 +220,7 @@ protected theorem IsPrefix.filterMap (h : l₁ <+: l₂) (f : α → Option β) · simp only [nil_prefix, filterMap_nil] · cases' l₂ with hd₂ tl₂ · simpa only using eq_nil_of_prefix_nil h - · rw [cons_prefix_iff] at h + · rw [cons_prefix_cons] at h rw [← @singleton_append _ hd₁ _, ← @singleton_append _ hd₂ _, filterMap_append, filterMap_append, h.left, prefix_append_right_inj] exact hl h.right From c9eabb1358d52f6b7dfd1cabdea4e4fa2632b1f3 Mon Sep 17 00:00:00 2001 From: grunweg Date: Wed, 14 Aug 2024 11:19:47 +0000 Subject: [PATCH 264/359] chore(ConcreteCategory/Bundled): remove outdated comment about Bundled.map (#15221) the definition is not reducible, so apparently this is fine. --- .../CategoryTheory/ConcreteCategory/Bundled.lean | 14 +------------- 1 file changed, 1 insertion(+), 13 deletions(-) diff --git a/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean b/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean index b97cff90ac9d4..8ea7c729576e2 100644 --- a/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean +++ b/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean @@ -51,20 +51,8 @@ instance coeSort : CoeSort (Bundled c) (Type u) := theorem coe_mk (α) (str) : (@Bundled.mk c α str : Type u) = α := rfl -/- -`Bundled.map` is reducible so that, if we define a category - - def Ring : Type (u+1) := induced_category SemiRing (bundled.map @ring.to_semiring) - -instance search is able to "see" that a morphism R ⟶ S in Ring is really -a (semi)ring homomorphism from R.α to S.α, and not merely from -`(Bundled.map @Ring.toSemiring R).α` to `(Bundled.map @Ring.toSemiring S).α`. - -TODO: Once at least one use of this has been ported, check if this still needs to be reducible in -Lean 4. --/ /-- Map over the bundled structure -/ -def map (f : ∀ {α}, c α → d α) (b : Bundled c) : Bundled d := +abbrev map (f : ∀ {α}, c α → d α) (b : Bundled c) : Bundled d := ⟨b, f b.str⟩ end Bundled From b74124c74c234d628473646ae4e5f851b4d64f16 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Wed, 14 Aug 2024 12:25:08 +0000 Subject: [PATCH 265/359] refactor(Condensed): reorganize TopComparison (#15401) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define a general functor from `TopCat` to sheaves on `CompHausLike P` satisfying some properties, and redefine `topCatToCondensedSet` and `topCatToLightCondSet` in terms of that functor. --- Mathlib/Condensed/Light/TopCatAdjunction.lean | 5 +- Mathlib/Condensed/Light/TopComparison.lean | 21 ++---- Mathlib/Condensed/TopCatAdjunction.lean | 4 +- Mathlib/Condensed/TopComparison.lean | 69 ++++++++++++------- 4 files changed, 57 insertions(+), 42 deletions(-) diff --git a/Mathlib/Condensed/Light/TopCatAdjunction.lean b/Mathlib/Condensed/Light/TopCatAdjunction.lean index fbdac8cf24455..2b67bf1e860bb 100644 --- a/Mathlib/Condensed/Light/TopCatAdjunction.lean +++ b/Mathlib/Condensed/Light/TopCatAdjunction.lean @@ -112,8 +112,9 @@ def topCatAdjunctionUnit (X : LightCondSet.{u}) : X ⟶ X.toTopCat.toLightCondSe apply continuous_coinduced_rng } naturality := fun _ _ _ ↦ by ext - simp only [types_comp_apply, ContinuousMap.coe_mk, TopCat.toLightCondSet_val_map, - ContinuousMap.comp_apply, ← FunctorToTypes.map_comp_apply] + simp only [TopCat.toSheafCompHausLike_val_obj, CompHausLike.compHausLikeToTop_obj, + Opposite.op_unop, types_comp_apply, TopCat.toSheafCompHausLike_val_map, + ← FunctorToTypes.map_comp_apply] rfl } /-- The adjunction `lightCondSetToTopCat ⊣ topCatToLightCondSet` -/ diff --git a/Mathlib/Condensed/Light/TopComparison.lean b/Mathlib/Condensed/Light/TopComparison.lean index 6fca286eaff84..ee4dbcaabc7ae 100644 --- a/Mathlib/Condensed/Light/TopComparison.lean +++ b/Mathlib/Condensed/Light/TopComparison.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.Condensed.Light.Explicit +import Mathlib.Condensed.Light.Basic import Mathlib.Condensed.TopComparison /-! @@ -14,26 +14,19 @@ We define the functor `topCatToLightCondSet : TopCat.{u} ⥤ LightCondSet.{u}`. -/ -universe w w' v u +universe u -open CategoryTheory Opposite Limits regularTopology ContinuousMap +open CategoryTheory /-- Associate to a `u`-small topological space the corresponding light condensed set, given by `yonedaPresheaf`. -/ -@[simps! val_obj val_map] -noncomputable def TopCat.toLightCondSet (X : TopCat.{u}) : LightCondSet.{u} := - @LightCondSet.ofSheafLightProfinite (yonedaPresheaf LightProfinite.toTopCat.{u} X) _ (by - apply equalizerCondition_yonedaPresheaf LightProfinite.toTopCat.{u} X - intro Z B π he - rw [LightProfinite.effectiveEpi_iff_surjective] at he - apply QuotientMap.of_surjective_continuous he π.continuous) +noncomputable abbrev TopCat.toLightCondSet (X : TopCat.{u}) : LightCondSet.{u} := + toSheafCompHausLike.{u} _ X (fun _ _ _ ↦ (LightProfinite.effectiveEpi_iff_surjective _).mp) /-- `TopCat.toLightCondSet` yields a functor from `TopCat.{u}` to `LightCondSet.{u}`. -/ -@[simps] -noncomputable def topCatToLightCondSet : TopCat.{u} ⥤ LightCondSet.{u} where - obj X := X.toLightCondSet - map f := ⟨⟨fun _ g ↦ f.comp g, by aesop⟩⟩ +noncomputable abbrev topCatToLightCondSet : TopCat.{u} ⥤ LightCondSet.{u} := + topCatToSheafCompHausLike.{u} _ (fun _ _ _ ↦ (LightProfinite.effectiveEpi_iff_surjective _).mp) diff --git a/Mathlib/Condensed/TopCatAdjunction.lean b/Mathlib/Condensed/TopCatAdjunction.lean index c4f72257d9754..13fedf0cb5e3d 100644 --- a/Mathlib/Condensed/TopCatAdjunction.lean +++ b/Mathlib/Condensed/TopCatAdjunction.lean @@ -113,8 +113,8 @@ def topCatAdjunctionUnit (X : CondensedSet.{u}) : X ⟶ X.toTopCat.toCondensedSe apply continuous_coinduced_rng } naturality := fun _ _ _ ↦ by ext - simp only [TopCat.toCondensedSet_val_obj, CompHausLike.compHausLikeToTop_obj, - Opposite.op_unop, types_comp_apply, TopCat.toCondensedSet_val_map, + simp only [TopCat.toSheafCompHausLike_val_obj, CompHausLike.compHausLikeToTop_obj, + Opposite.op_unop, types_comp_apply, TopCat.toSheafCompHausLike_val_map, ← FunctorToTypes.map_comp_apply] rfl } diff --git a/Mathlib/Condensed/TopComparison.lean b/Mathlib/Condensed/TopComparison.lean index 56eeda95ac14a..290e5f0d3433f 100644 --- a/Mathlib/Condensed/TopComparison.lean +++ b/Mathlib/Condensed/TopComparison.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ import Mathlib.CategoryTheory.Limits.Preserves.Opposites +import Mathlib.CategoryTheory.Sites.Coherent.SheafComparison +import Mathlib.Condensed.Basic import Mathlib.Topology.Category.TopCat.Yoneda -import Mathlib.Condensed.Explicit /-! @@ -24,6 +25,8 @@ universe w w' v u open CategoryTheory Opposite Limits regularTopology ContinuousMap +attribute [local instance] ConcreteCategory.instFunLike + variable {C : Type u} [Category.{v} C] (G : C ⥤ TopCat.{w}) (X : Type w') [TopologicalSpace X] @@ -85,37 +88,55 @@ If `G` preserves finite coproducts (which is the case when `C` is `CompHaus`, ` the extensive topology. -/ noncomputable instance [PreservesFiniteCoproducts G] : - PreservesFiniteProducts (yonedaPresheaf G X) := by - change PreservesFiniteProducts (G.op ⋙ yonedaPresheaf' X) - have h' : PreservesFiniteProducts (yonedaPresheaf' X) := inferInstance - have h : PreservesFiniteProducts G.op := - { preserves := fun J _ => by - apply (config := { allowSynthFailures := true }) preservesLimitsOfShapeOp - exact preservesColimitsOfShapeOfEquiv (Discrete.opposite J).symm _ } - constructor - intro J _ - have := h.1 J - have := h'.1 J - exact compPreservesLimitsOfShape _ _ + PreservesFiniteProducts (yonedaPresheaf G X) := + have := preservesFiniteProductsOp G + ⟨fun _ ↦ compPreservesLimitsOfShape G.op (yonedaPresheaf' X)⟩ + +section + +variable (P : TopCat.{u} → Prop) (X : TopCat.{max u w}) + [CompHausLike.HasExplicitFiniteCoproducts.{0} P] [CompHausLike.HasExplicitPullbacks.{u} P] + (hs : ∀ ⦃X Y : CompHausLike P⦄ (f : X ⟶ Y), EffectiveEpi f → Function.Surjective f) /-- -Associate to a `(u+1)`-small topological space the corresponding condensed set, given by -`yonedaPresheaf`. +The sheaf on `CompHausLike P` of continuous maps to a topological space. -/ @[simps! val_obj val_map] -noncomputable def TopCat.toCondensedSet (X : TopCat.{u+1}) : CondensedSet.{u} := - @CondensedSet.ofSheafCompHaus (yonedaPresheaf.{u, u+1, u, u+1} compHausToTop.{u} X) _ (by +def TopCat.toSheafCompHausLike : + have := CompHausLike.preregular hs + Sheaf (coherentTopology (CompHausLike.{u} P)) (Type (max u w)) where + val := yonedaPresheaf.{u, max u w} (CompHausLike.compHausLikeToTop.{u} P) X + cond := by + have := CompHausLike.preregular hs + rw [Presheaf.isSheaf_iff_preservesFiniteProducts_and_equalizerCondition] + refine ⟨⟨inferInstance⟩, ?_⟩ apply (config := { allowSynthFailures := true }) equalizerCondition_yonedaPresheaf - compHausToTop.{u} X + (CompHausLike.compHausLikeToTop.{u} P) X intro Z B π he - rw [((CompHaus.effectiveEpi_tfae π).out 0 2 :)] at he - apply QuotientMap.of_surjective_continuous he π.continuous ) - + apply QuotientMap.of_surjective_continuous (hs _ he) π.continuous /-- -`TopCat.toCondensedSet` yields a functor from `TopCat.{u+1}` to `CondensedSet.{u}`. +`TopCat.toSheafCompHausLike` yields a functor from `TopCat.{max u w}` to +`Sheaf (coherentTopology (CompHausLike.{u} P)) (Type (max u w))`. -/ @[simps] -noncomputable def topCatToCondensedSet : TopCat.{u+1} ⥤ CondensedSet.{u} where - obj X := X.toCondensedSet +noncomputable def topCatToSheafCompHausLike : + have := CompHausLike.preregular hs + TopCat.{max u w} ⥤ Sheaf (coherentTopology (CompHausLike.{u} P)) (Type (max u w)) where + obj X := X.toSheafCompHausLike P hs map f := ⟨⟨fun _ g ↦ f.comp g, by aesop⟩⟩ + +end + +/-- +Associate to a `(u+1)`-small topological space the corresponding condensed set, given by +`yonedaPresheaf`. +-/ +noncomputable abbrev TopCat.toCondensedSet (X : TopCat.{u+1}) : CondensedSet.{u} := + toSheafCompHausLike.{u+1} _ X (fun _ _ _ ↦ ((CompHaus.effectiveEpi_tfae _).out 0 2).mp) + +/-- +`TopCat.toCondensedSet` yields a functor from `TopCat.{u+1}` to `CondensedSet.{u}`. +-/ +noncomputable abbrev topCatToCondensedSet : TopCat.{u+1} ⥤ CondensedSet.{u} := + topCatToSheafCompHausLike.{u+1} _ (fun _ _ _ ↦ ((CompHaus.effectiveEpi_tfae _).out 0 2).mp) From e6273edf333fbced24d8a6eb8625afd5e18ceb39 Mon Sep 17 00:00:00 2001 From: Lorenzo Luccioli Date: Wed, 14 Aug 2024 12:57:00 +0000 Subject: [PATCH 266/359] feat (Algebra/Order/Group): `mulIndicator_le_mulIndicator'` (#15437) Add `mulIndicator_le_mulIndicator'`, a generalization of `mulIndicator_le_mulIndicator`. --- Mathlib/Algebra/Order/Group/Indicator.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Algebra/Order/Group/Indicator.lean b/Mathlib/Algebra/Order/Group/Indicator.lean index 85070045eaeab..6ae4c61cca65e 100644 --- a/Mathlib/Algebra/Order/Group/Indicator.lean +++ b/Mathlib/Algebra/Order/Group/Indicator.lean @@ -101,6 +101,11 @@ lemma mulIndicator_apply_le_one (h : a ∈ s → f a ≤ 1) : mulIndicator s f a lemma mulIndicator_le_one (h : ∀ a ∈ s, f a ≤ 1) (a : α) : mulIndicator s f a ≤ 1 := mulIndicator_apply_le_one (h a) +@[to_additive] +lemma mulIndicator_le_mulIndicator' (h : a ∈ s → f a ≤ g a) : + mulIndicator s f a ≤ mulIndicator s g a := + mulIndicator_rel_mulIndicator le_rfl h + @[to_additive] lemma mulIndicator_le_mulIndicator (h : f a ≤ g a) : mulIndicator s f a ≤ mulIndicator s g a := mulIndicator_rel_mulIndicator le_rfl fun _ ↦ h From 7565a8d76870d86c0b785a8af9d3cd8b394bd48d Mon Sep 17 00:00:00 2001 From: Seewoo Lee Date: Wed, 14 Aug 2024 13:46:22 +0000 Subject: [PATCH 267/359] feat(RingTheory/Radical): Radical of an element in a unique factorization normalization monoid (#14873) For a unique factorization normalization monoid, define a radical of an element as a product of normalized prime factors (without duplication). Co-authored-by: Jineon Baek Co-authored-by: Johan Commelin --- Mathlib.lean | 1 + Mathlib/RingTheory/Radical.lean | 105 ++++++++++++++++++++++++++++++++ 2 files changed, 106 insertions(+) create mode 100644 Mathlib/RingTheory/Radical.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0abdbf82cdafa..279e430d5c489 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3910,6 +3910,7 @@ import Mathlib.RingTheory.PrincipalIdealDomain import Mathlib.RingTheory.QuotSMulTop import Mathlib.RingTheory.QuotientNilpotent import Mathlib.RingTheory.QuotientNoetherian +import Mathlib.RingTheory.Radical import Mathlib.RingTheory.ReesAlgebra import Mathlib.RingTheory.Regular.IsSMulRegular import Mathlib.RingTheory.Regular.RegularSequence diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean new file mode 100644 index 0000000000000..4dca38df1313c --- /dev/null +++ b/Mathlib/RingTheory/Radical.lean @@ -0,0 +1,105 @@ +/- +Copyright (c) 2024 Jineon Back and Seewoo Lee. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jineon Baek, Seewoo Lee +-/ +import Mathlib.RingTheory.UniqueFactorizationDomain + +/-! +# Radical of an element of a unique factorization normalization monoid + +This file defines a radical of an element `a` of a unique factorization normalization +monoid, which is defined as a product of normalized prime factors of `a` without duplication. +This is different from the radical of an ideal. + +## Main declarations + +- `radical`: The radical of an element `a` in a unique factorization monoid is the product of + its prime factors. +- `radical_eq_of_associated`: If `a` and `b` are associates, i.e. `a * u = b` for some unit `u`, + then `radical a = radical b`. +- `radical_unit_mul`: Multiplying unit does not change the radical. +- `radical_dvd_self`: `radical a` divides `a`. +- `radical_pow`: `radical (a ^ n) = radical a` for any `n ≥ 1` +- `radical_of_prime`: Radical of a prime element is equal to its normalization +- `radical_pow_of_prime`: Radical of a power of prime element is equal to its normalization + +## TODO + +- Make a comparison with `Ideal.radical`. Especially, for principal ideal, + `Ideal.radical (Ideal.span {a}) = Ideal.span {radical a}`. +- Prove `radical (radical a) = radical a`. +- Prove a comparison between `primeFactors` and `Nat.primeFactors`. +-/ + +noncomputable section + +open scoped Classical + +open UniqueFactorizationMonoid + +-- `CancelCommMonoidWithZero` is required by `UniqueFactorizationMonoid` +variable {M : Type*} [CancelCommMonoidWithZero M] [NormalizationMonoid M] + [UniqueFactorizationMonoid M] + +/-- The finite set of prime factors of an element in a unique factorization monoid. -/ +def primeFactors (a : M) : Finset M := + (normalizedFactors a).toFinset + +/-- +Radical of an element `a` in a unique factorization monoid is the product of +the prime factors of `a`. +-/ +def radical (a : M) : M := + (primeFactors a).prod id + +@[simp] +theorem radical_zero_eq : radical (0 : M) = 1 := by + rw [radical, primeFactors, normalizedFactors_zero, Multiset.toFinset_zero, Finset.prod_empty] + +@[simp] +theorem radical_one_eq : radical (1 : M) = 1 := by + rw [radical, primeFactors, normalizedFactors_one, Multiset.toFinset_zero, Finset.prod_empty] + +theorem radical_eq_of_associated {a b : M} (h : Associated a b) : radical a = radical b := by + rcases iff_iff_and_or_not_and_not.mp h.eq_zero_iff with (⟨rfl, rfl⟩ | ⟨ha, hb⟩) + · rfl + · simp_rw [radical, primeFactors] + rw [(associated_iff_normalizedFactors_eq_normalizedFactors ha hb).mp h] + +theorem radical_unit_eq_one {a : M} (h : IsUnit a) : radical a = 1 := + (radical_eq_of_associated (associated_one_iff_isUnit.mpr h)).trans radical_one_eq + +theorem radical_unit_mul {u : Mˣ} {a : M} : radical ((↑u : M) * a) = radical a := + radical_eq_of_associated (associated_unit_mul_left _ _ u.isUnit) + +theorem radical_mul_unit {u : Mˣ} {a : M} : radical (a * (↑u : M)) = radical a := + radical_eq_of_associated (associated_mul_unit_left _ _ u.isUnit) + +theorem primeFactors_pow (a : M) {n : ℕ} (hn : 0 < n) : primeFactors (a ^ n) = primeFactors a := by + simp_rw [primeFactors] + simp only [normalizedFactors_pow] + rw [Multiset.toFinset_nsmul] + exact ne_of_gt hn + +theorem radical_pow (a : M) {n : Nat} (hn : 0 < n) : radical (a ^ n) = radical a := by + simp_rw [radical, primeFactors_pow a hn] + +theorem radical_dvd_self (a : M) : radical a ∣ a := by + by_cases ha : a = 0 + · rw [ha] + apply dvd_zero + · rw [radical, ← Finset.prod_val, ← (normalizedFactors_prod ha).dvd_iff_dvd_right] + apply Multiset.prod_dvd_prod_of_le + rw [primeFactors, Multiset.toFinset_val] + apply Multiset.dedup_le + +theorem radical_of_prime {a : M} (ha : Prime a) : radical a = normalize a := by + rw [radical, primeFactors] + rw [normalizedFactors_irreducible ha.irreducible] + simp only [Multiset.toFinset_singleton, id, Finset.prod_singleton] + +theorem radical_pow_of_prime {a : M} (ha : Prime a) {n : ℕ} (hn : 0 < n) : + radical (a ^ n) = normalize a := by + rw [radical_pow a hn] + exact radical_of_prime ha From a08417b75a79fab791d5b392d3a507c1259aec39 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 14 Aug 2024 13:56:09 +0000 Subject: [PATCH 268/359] doc(Algebra/Quaternion): add docstring for `imI`, `imJ`, `imK` (#15789) --- Mathlib/Algebra/Quaternion.lean | 3 +++ scripts/nolints.json | 3 --- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 97126ffd2f19d..56d78e1cdaf61 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -55,8 +55,11 @@ Implemented as a structure with four fields: `re`, `imI`, `imJ`, and `imK`. -/ structure QuaternionAlgebra (R : Type*) (a b : R) where /-- Real part of a quaternion. -/ re : R + /-- First imaginary part (i) of a quaternion. -/ imI : R + /-- Second imaginary part (j) of a quaternion. -/ imJ : R + /-- Third imaginary part (k) of a quaternion. -/ imK : R @[inherit_doc] diff --git a/scripts/nolints.json b/scripts/nolints.json index 28ccef88aecff..6a74c778131bb 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -321,9 +321,6 @@ ["docBlame", "QPF.repr"], ["docBlame", "QuadraticMap.toFun"], ["docBlame", "Quaternion.«termℍ[_]»"], - ["docBlame", "QuaternionAlgebra.imI"], - ["docBlame", "QuaternionAlgebra.imJ"], - ["docBlame", "QuaternionAlgebra.imK"], ["docBlame", "RCLike.im"], ["docBlame", "RCLike.re"], ["docBlame", "Random.randBool"], From f494f96e7a8b3fa24562ddd52a3e827c924e27d5 Mon Sep 17 00:00:00 2001 From: Gabin Date: Wed, 14 Aug 2024 14:06:30 +0000 Subject: [PATCH 269/359] feat(ModelTheory/FinitelyGenerated): only countably many morphisms from a cg structure to a countable one (#11177) Prove that there are only countably many morphisms from a countably generated structure, to a countable one. Co-authored-by: Aaron Anderson <65780815+awainverse@users.noreply.github.com> Co-authored-by: Gabin Kolly Co-authored-by: Gabin <68381468+QuinnLesquimau@users.noreply.github.com> --- Mathlib/ModelTheory/FinitelyGenerated.lean | 48 +++++ Mathlib/ModelTheory/PartialEquiv.lean | 194 ++++++++++++++++++++- Mathlib/ModelTheory/Substructures.lean | 4 + 3 files changed, 243 insertions(+), 3 deletions(-) diff --git a/Mathlib/ModelTheory/FinitelyGenerated.lean b/Mathlib/ModelTheory/FinitelyGenerated.lean index df79d2ed1b59d..744fe0e4c4f1c 100644 --- a/Mathlib/ModelTheory/FinitelyGenerated.lean +++ b/Mathlib/ModelTheory/FinitelyGenerated.lean @@ -61,6 +61,8 @@ theorem fg_iff_exists_fin_generating_family {N : L.Substructure M} : theorem fg_bot : (⊥ : L.Substructure M).FG := ⟨∅, by rw [Finset.coe_empty, closure_empty]⟩ +instance instInhabited_fg : Inhabited { S : L.Substructure M // S.FG } := ⟨⊥, fg_bot⟩ + theorem fg_closure {s : Set M} (hs : s.Finite) : FG (closure L s) := ⟨hs.toFinset, by rw [hs.coe_toFinset]⟩ @@ -161,6 +163,9 @@ theorem cg_iff_countable [Countable (Σl, L.Functions l)] {s : L.Substructure M} rintro ⟨s, h, rfl⟩ exact h.substructure_closure L +theorem cg_of_countable {s : L.Substructure M} [h : Countable s] : s.CG := + ⟨s, h.to_set, s.closure_eq⟩ + end Substructure open Substructure @@ -197,6 +202,31 @@ theorem FG.map_of_surjective {N : Type*} [L.Structure N] (h : FG L M) (f : M → rw [fg_def, ← hs] exact h.range f +theorem FG.countable_hom (N : Type*) [L.Structure N] [Countable N] (h : FG L M) : + Countable (M →[L] N) := by + let ⟨S, finite_S, closure_S⟩ := fg_iff.1 h + let g : (M →[L] N) → (S → N) := + fun f ↦ f ∘ (↑) + have g_inj : Function.Injective g := by + intro f f' h + apply Hom.eq_of_eqOn_dense closure_S + intro x x_in_S + exact congr_fun h ⟨x, x_in_S⟩ + have : Finite ↑S := (S.finite_coe_iff).2 finite_S + exact Function.Embedding.countable ⟨g, g_inj⟩ + +instance FG.instCountable_hom (N : Type*) [L.Structure N] [Countable N] [h : FG L M] : + Countable (M →[L] N) := + FG.countable_hom N h + +theorem FG.countable_embedding (N : Type*) [L.Structure N] [Countable N] (_ : FG L M) : + Countable (M ↪[L] N) := + Function.Embedding.countable ⟨Embedding.toHom, Embedding.toHom_injective⟩ + +instance Fg.instCountable_embedding (N : Type*) [L.Structure N] + [Countable N] [h : FG L M] : Countable (M ↪[L] N) := + FG.countable_embedding N h + theorem cg_def : CG L M ↔ (⊤ : L.Substructure M).CG := ⟨fun h => h.1, fun h => ⟨h⟩⟩ @@ -217,6 +247,9 @@ theorem CG.map_of_surjective {N : Type*} [L.Structure N] (h : CG L M) (f : M → theorem cg_iff_countable [Countable (Σl, L.Functions l)] : CG L M ↔ Countable M := by rw [cg_def, Substructure.cg_iff_countable, topEquiv.toEquiv.countable_iff] +theorem cg_of_countable [Countable M] : CG L M := by + simp only [cg_def, Substructure.cg_of_countable, topEquiv.toEquiv.countable_iff] + theorem FG.cg (h : FG L M) : CG L M := cg_def.2 (fg_def.1 h).cg @@ -253,6 +286,21 @@ theorem Substructure.cg_iff_structure_cg (S : L.Substructure M) : S.CG ↔ Struc rw [← Hom.range_eq_map, range_subtype] at h exact h +theorem Substructure.countable_fg_substructures_of_countable [Countable M] : + Countable { S : L.Substructure M // S.FG } := by + let g : { S : L.Substructure M // S.FG } → Finset M := + fun S ↦ Exists.choose S.prop + have g_inj : Function.Injective g := by + intro S S' h + apply Subtype.eq + rw [(Exists.choose_spec S.prop).symm, (Exists.choose_spec S'.prop).symm] + exact congr_arg ((closure L) ∘ Finset.toSet) h + exact Function.Embedding.countable ⟨g, g_inj⟩ + +instance Substructure.instCountable_fg_substructures_of_countable [Countable M] : + Countable { S : L.Substructure M // S.FG } := + countable_fg_substructures_of_countable + end Language end FirstOrder diff --git a/Mathlib/ModelTheory/PartialEquiv.lean b/Mathlib/ModelTheory/PartialEquiv.lean index 41ce920aed2f2..3277fb597492c 100644 --- a/Mathlib/ModelTheory/PartialEquiv.lean +++ b/Mathlib/ModelTheory/PartialEquiv.lean @@ -3,16 +3,16 @@ Copyright (c) 2024 Gabin Kolly. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Gabin Kolly -/ -import Mathlib.ModelTheory.Substructures +import Mathlib.ModelTheory.DirectLimit +import Mathlib.Order.Ideal /-! # Partial Isomorphisms This file defines partial isomorphisms between first-order structures. ## Main Definitions -- `FirstOrder.Language.Substructure.PartialEquiv` is defined so that `L.PartialEquiv M N`, annotated +- `FirstOrder.Language.PartialEquiv` is defined so that `L.PartialEquiv M N`, annotated `M ≃ₚ[L] N`, is the type of equivalences between substructures of `M` and `N`. - -/ universe u v w w' @@ -237,6 +237,10 @@ theorem toEquivOfEqTop_toEmbedding {f : M ≃ₚ[L] N} (h_dom : f.dom = ⊤) cases h_cod rfl +theorem dom_fg_iff_cod_fg {N : Type*} [L.Structure N] (f : M ≃ₚ[L] N) : + f.dom.FG ↔ f.cod.FG := by + rw [Substructure.fg_iff_structure_fg, f.toEquiv.fg_iff, Substructure.fg_iff_structure_fg] + end PartialEquiv namespace Embedding @@ -269,6 +273,190 @@ theorem toPartialEquiv_toEmbedding {f : M ≃ₚ[L] N} (h : f.dom = ⊤) : end Embedding +namespace DirectLimit + +open PartialEquiv + +variable {ι : Type*} [Preorder ι] [Nonempty ι] [IsDirected ι (· ≤ ·)] +variable (S : ι →o M ≃ₚ[L] N) + +instance : DirectedSystem (fun i ↦ (S i).dom) (fun _ _ h ↦ + Substructure.inclusion (dom_le_dom (S.monotone h))) where + map_self' := fun _ _ _ ↦ rfl + map_map' := fun _ _ _ ↦ rfl + +instance : DirectedSystem (fun i ↦ (S i).cod) (fun _ _ h ↦ + Substructure.inclusion (cod_le_cod (S.monotone h))) where + map_self' := fun _ _ _ ↦ rfl + map_map' := fun _ _ _ ↦ rfl + +/-- The limit of a directed system of PartialEquivs. -/ +noncomputable def partialEquivLimit : M ≃ₚ[L] N where + dom := iSup (fun i ↦ (S i).dom) + cod := iSup (fun i ↦ (S i).cod) + toEquiv := + (Equiv_iSup { + toFun := (fun i ↦ (S i).cod) + monotone' := monotone_cod.comp S.monotone} + ).comp + ((DirectLimit.equiv_lift L ι (fun i ↦ (S i).dom) + (fun _ _ hij ↦ Substructure.inclusion (dom_le_dom (S.monotone hij))) + (fun i ↦ (S i).cod) + (fun _ _ hij ↦ Substructure.inclusion (cod_le_cod (S.monotone hij))) + (fun i ↦ (S i).toEquiv) + (fun _ _ hij _ ↦ toEquiv_inclusion_apply (S.monotone hij) _) + ).comp + (Equiv_iSup { + toFun := (fun i ↦ (S i).dom) + monotone' := monotone_dom.comp S.monotone}).symm) + +@[simp] +theorem dom_partialEquivLimit : (partialEquivLimit S).dom = iSup (fun x ↦ (S x).dom) := rfl + +@[simp] +theorem cod_partialEquivLimit : (partialEquivLimit S).cod = iSup (fun x ↦ (S x).cod) := rfl + +@[simp] +lemma partialEquivLimit_comp_inclusion {i : ι} : + (partialEquivLimit S).toEquiv.toEmbedding.comp (Substructure.inclusion (le_iSup _ i)) = + (Substructure.inclusion (le_iSup _ i)).comp (S i).toEquiv.toEmbedding := by + simp only [partialEquivLimit, Equiv.comp_toEmbedding, Embedding.comp_assoc] + rw [Equiv_isup_symm_inclusion] + congr + +theorem le_partialEquivLimit : ∀ i, S i ≤ partialEquivLimit S := + fun i => ⟨le_iSup (f := fun i ↦ (S i).dom) _, by + simp only [cod_partialEquivLimit, dom_partialEquivLimit, partialEquivLimit_comp_inclusion, ← + Embedding.comp_assoc, subtype_comp_inclusion]⟩ + +end DirectLimit + +section FGEquiv + +open PartialEquiv Set DirectLimit + +variable (M) (N) (L) + +/-- The type of equivalences between finitely generated substructures. -/ +abbrev FGEquiv := {f : M ≃ₚ[L] N // f.dom.FG} + +/-- Two structures `M` and `N` form an extension pair if the domain of any finitely-generated map +from `M` to `N` can be extended to include any element of `M`. -/ +def IsExtensionPair : Prop := ∀ (f : L.FGEquiv M N) (m : M), ∃ g, m ∈ g.1.dom ∧ f ≤ g + +variable {M N L} + +theorem countable_self_fgequiv_of_countable [Countable M] : + Countable (L.FGEquiv M M) := by + let g : L.FGEquiv M M → + Σ U : { S : L.Substructure M // S.FG }, U.val →[L] M := + fun f ↦ ⟨⟨f.val.dom, f.prop⟩, (subtype _).toHom.comp f.val.toEquiv.toHom⟩ + have g_inj : Function.Injective g := by + intro f f' h + ext + let ⟨⟨dom_f, cod_f, equiv_f⟩, f_fin⟩ := f + cases congr_arg (·.1) h + apply PartialEquiv.ext (by rfl) + simp only [g, Sigma.mk.inj_iff, heq_eq_eq, true_and] at h + exact fun x hx ↦ congr_fun (congr_arg (↑) h) ⟨x, hx⟩ + have : ∀ U : { S : L.Substructure M // S.FG }, Structure.FG L U.val := + fun U ↦ (U.val.fg_iff_structure_fg.1 U.prop) + exact Function.Embedding.countable ⟨g, g_inj⟩ + +instance inhabited_self_FGEquiv : Inhabited (L.FGEquiv M M) := + ⟨⟨⟨⊥, ⊥, Equiv.refl L (⊥ : L.Substructure M)⟩, fg_bot⟩⟩ + +/-- Maps to the symmetric finitely-generated partial equivalence. -/ +@[simps] +def FGEquiv.symm (f : L.FGEquiv M N) : L.FGEquiv N M := ⟨f.1.symm, f.1.dom_fg_iff_cod_fg.1 f.2⟩ + +lemma IsExtensionPair_iff_cod : L.IsExtensionPair M N ↔ + ∀ (f : L.FGEquiv N M) (m : M), ∃ g, m ∈ g.1.cod ∧ f ≤ g := by + refine Iff.intro ?_ ?_ <;> + · intro h f m + obtain ⟨g, h1, h2⟩ := h f.symm m + exact ⟨g.symm, h1, monotone_symm h2⟩ + +namespace IsExtensionPair + +protected alias ⟨cod, _⟩ := IsExtensionPair_iff_cod + +/-- The cofinal set of finite equivalences with a given element in their domain. -/ +def definedAtLeft + (h : L.IsExtensionPair M N) (m : M) : Order.Cofinal (FGEquiv L M N) where + carrier := {f | m ∈ f.val.dom} + mem_gt := fun f => h f m + +/-- The cofinal set of finite equivalences with a given element in their codomain. -/ +def definedAtRight + (h : L.IsExtensionPair N M) (n : N) : Order.Cofinal (FGEquiv L M N) where + carrier := {f | n ∈ f.val.cod} + mem_gt := fun f => h.cod f n + +end IsExtensionPair + +/-- For a countably generated structure `M` and a structure `N`, if any partial equivalence +between finitely generated substructures can be extended to any element in the domain, +then there exists an embedding of `M` in `N`. -/ +theorem embedding_from_cg (M_cg : Structure.CG L M) (g : L.FGEquiv M N) + (H : L.IsExtensionPair M N) : + ∃ f : M ↪[L] N, g ≤ f.toPartialEquiv := by + rcases M_cg with ⟨X, _, X_gen⟩ + have _ : Countable (↑X : Type _) := by simpa only [countable_coe_iff] + have _ : Encodable (↑X : Type _) := Encodable.ofCountable _ + let D : X → Order.Cofinal (FGEquiv L M N) := fun x ↦ H.definedAtLeft x + let S : ℕ →o M ≃ₚ[L] N := + ⟨Subtype.val ∘ (Order.sequenceOfCofinals g D), + (Subtype.mono_coe _).comp (Order.sequenceOfCofinals.monotone _ _)⟩ + let F := DirectLimit.partialEquivLimit S + have _ : X ⊆ F.dom := by + intro x hx + have := Order.sequenceOfCofinals.encode_mem g D ⟨x, hx⟩ + exact dom_le_dom + (le_partialEquivLimit S (Encodable.encode (⟨x, hx⟩ : X) + 1)) this + have isTop : F.dom = ⊤ := by rwa [← top_le_iff, ← X_gen, Substructure.closure_le] + exact ⟨toEmbeddingOfEqTop isTop, + by convert (le_partialEquivLimit S 0); apply Embedding.toPartialEquiv_toEmbedding⟩ + +/-- For two countably generated structure `M` and `N`, if any PartialEquiv +between finitely generated substructures can be extended to any element in the domain and to +any element in the codomain, then there exists an equivalence between `M` and `N`. -/ +theorem equiv_between_cg (M_cg : Structure.CG L M) (N_cg : Structure.CG L N) + (g : L.FGEquiv M N) + (ext_dom : L.IsExtensionPair M N) + (ext_cod : L.IsExtensionPair N M) : + ∃ f : M ≃[L] N, g ≤ f.toEmbedding.toPartialEquiv := by + rcases M_cg with ⟨X, X_count, X_gen⟩ + rcases N_cg with ⟨Y, Y_count, Y_gen⟩ + have _ : Countable (↑X : Type _) := by simpa only [countable_coe_iff] + have _ : Encodable (↑X : Type _) := Encodable.ofCountable _ + have _ : Countable (↑Y : Type _) := by simpa only [countable_coe_iff] + have _ : Encodable (↑Y : Type _) := Encodable.ofCountable _ + let D : Sum X Y → Order.Cofinal (FGEquiv L M N) := fun p ↦ + Sum.recOn p (fun x ↦ ext_dom.definedAtLeft x) (fun y ↦ ext_cod.definedAtRight y) + let S : ℕ →o M ≃ₚ[L] N := + ⟨Subtype.val ∘ (Order.sequenceOfCofinals g D), + (Subtype.mono_coe _).comp (Order.sequenceOfCofinals.monotone _ _)⟩ + let F := @DirectLimit.partialEquivLimit L M N _ _ ℕ _ _ _ S + have _ : X ⊆ F.dom := by + intro x hx + have := Order.sequenceOfCofinals.encode_mem g D (Sum.inl ⟨x, hx⟩) + exact dom_le_dom + (le_partialEquivLimit S (Encodable.encode (Sum.inl (⟨x, hx⟩ : X)) + 1)) this + have _ : Y ⊆ F.cod := by + intro y hy + have := Order.sequenceOfCofinals.encode_mem g D (Sum.inr ⟨y, hy⟩) + exact cod_le_cod + (le_partialEquivLimit S (Encodable.encode (Sum.inr (⟨y, hy⟩ : Y)) + 1)) this + have dom_top : F.dom = ⊤ := by rwa [← top_le_iff, ← X_gen, Substructure.closure_le] + have cod_top : F.cod = ⊤ := by rwa [← top_le_iff, ← Y_gen, Substructure.closure_le] + refine ⟨toEquivOfEqTop dom_top cod_top, ?_⟩ + convert le_partialEquivLimit S 0 + rw [toEquivOfEqTop_toEmbedding] + apply Embedding.toPartialEquiv_toEmbedding + +end FGEquiv + end Language end FirstOrder diff --git a/Mathlib/ModelTheory/Substructures.lean b/Mathlib/ModelTheory/Substructures.lean index b08cea3eb212d..a4ac5cf414f99 100644 --- a/Mathlib/ModelTheory/Substructures.lean +++ b/Mathlib/ModelTheory/Substructures.lean @@ -890,6 +890,10 @@ theorem range_subtype (S : L.Substructure M) : S.subtype.toHom.range = S := by rintro ⟨⟨y, hy⟩, rfl⟩ exact hy +@[simp] +lemma subtype_comp_inclusion {S T : L.Substructure M} (h : S ≤ T) : + T.subtype.comp (inclusion h) = S.subtype := rfl + end Substructure end Language From cc30a7edee164b48cb23785a1282fe16ce0b60f2 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Wed, 14 Aug 2024 14:16:39 +0000 Subject: [PATCH 270/359] chore: protect Measure.add_sub_cancel (#15707) --- Mathlib/MeasureTheory/Function/Jacobian.lean | 2 +- Mathlib/MeasureTheory/Measure/Sub.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index 71aec0c9f8971..2477fc7a63860 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -506,7 +506,7 @@ theorem _root_.ApproximatesLinearOn.norm_fderiv_sub_le {A : E →L[ℝ] E} {δ : exact ⟨a, az, by simp only [ha, add_neg_cancel_left]⟩ have norm_a : ‖a‖ ≤ ‖z‖ + ε := calc - ‖a‖ = ‖z + (a - z)‖ := by simp only [_root_.add_sub_cancel] + ‖a‖ = ‖z + (a - z)‖ := by simp only [add_sub_cancel] _ ≤ ‖z‖ + ‖a - z‖ := norm_add_le _ _ _ ≤ ‖z‖ + ε := add_le_add_left (mem_closedBall_iff_norm.1 az) _ -- use the approximation properties to control `(f' x - A) a`, and then `(f' x - A) z` as `z` is diff --git a/Mathlib/MeasureTheory/Measure/Sub.lean b/Mathlib/MeasureTheory/Measure/Sub.lean index 956465498e954..224f04d912129 100644 --- a/Mathlib/MeasureTheory/Measure/Sub.lean +++ b/Mathlib/MeasureTheory/Measure/Sub.lean @@ -91,7 +91,7 @@ theorem sub_add_cancel_of_le [IsFiniteMeasure ν] (h₁ : ν ≤ μ) : μ - ν + rw [add_apply, sub_apply h_s_meas h₁, tsub_add_cancel_of_le (h₁ s)] @[simp] -lemma add_sub_cancel [IsFiniteMeasure ν] : μ + ν - ν = μ := by +protected lemma add_sub_cancel [IsFiniteMeasure ν] : μ + ν - ν = μ := by ext1 s hs rw [sub_apply hs (Measure.le_add_left (le_refl _)), add_apply, ENNReal.add_sub_cancel_right (measure_ne_top ν s)] From 5e8e4a238e119f36df439c824802a60c6dd2bef9 Mon Sep 17 00:00:00 2001 From: Ralf Stephan Date: Wed, 14 Aug 2024 14:16:40 +0000 Subject: [PATCH 271/359] doc(Tactic): expand `clear*` docstring (#15731) Add details and example to `clear*` docstring. --- Mathlib/Tactic/ClearExcept.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/ClearExcept.lean b/Mathlib/Tactic/ClearExcept.lean index 8f3e829ba3472..804d6950d8e2c 100644 --- a/Mathlib/Tactic/ClearExcept.lean +++ b/Mathlib/Tactic/ClearExcept.lean @@ -16,7 +16,11 @@ open Lean.Meta namespace Lean.Elab.Tactic -/-- Clears all hypotheses it can besides those provided -/ +/-- Clears all hypotheses it can, except those provided after a minus sign. Example: +``` + clear * - h₁ h₂ +``` +-/ syntax (name := clearExcept) "clear " "*" " -" (ppSpace colGt ident)* : tactic elab_rules : tactic From ebcf89d6dbfe926a3be9a319f8025ff53a3aab0f Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Wed, 14 Aug 2024 16:13:54 +0000 Subject: [PATCH 272/359] feat(CategoryTheory): Finality of the forget functor on over and under categories (#15799) --- Mathlib/CategoryTheory/Filtered/Final.lean | 36 ++++++++++++++++++++++ 1 file changed, 36 insertions(+) diff --git a/Mathlib/CategoryTheory/Filtered/Final.lean b/Mathlib/CategoryTheory/Filtered/Final.lean index eb3ef28ce871c..73d630e449fad 100644 --- a/Mathlib/CategoryTheory/Filtered/Final.lean +++ b/Mathlib/CategoryTheory/Filtered/Final.lean @@ -18,6 +18,8 @@ final can be restated. We show: * `final_iff_isFiltered_structuredArrow`: `F` is final if and only if `StructuredArrow d F` is filtered for all `d : D`, which strengthens the usual statement that `F` is final if and only if `StructuredArrow d F` is connected for all `d : D`. +* Under categories of objects of filtered categories are filtered and their forgetful functors + are final. Additionally, we show that if `D` is a filtered category and `F : C ⥤ D` is fully faithful and satisfies the additional condition that for every `d : D` there is an object `c : D` and a morphism @@ -164,6 +166,40 @@ theorem Functor.initial_of_exists_of_isCofiltered_of_fullyFaithful [IsCofiltered obtain ⟨c, ⟨f⟩⟩ := h d.unop exact ⟨op c, ⟨f.op⟩⟩ +/-- Any under category on a filtered or empty category is filtered. +(Note that under categories are always cofiltered since they have an initial object.) -/ +instance IsFiltered.under [IsFilteredOrEmpty C] (c : C) : IsFiltered (Under c) := + isFiltered_structuredArrow_of_isFiltered_of_exists _ + (fun c' => ⟨c', ⟨𝟙 _⟩⟩) + (fun s s' => IsFilteredOrEmpty.cocone_maps s s') c + +/-- Any over category on a cofiltered or empty category is cofiltered. +(Note that over categories are always filtered since they have a terminal object.) -/ +instance IsCofiltered.over [IsCofilteredOrEmpty C] (c : C) : IsCofiltered (Over c) := + isCofiltered_costructuredArrow_of_isCofiltered_of_exists _ + (fun c' => ⟨c', ⟨𝟙 _⟩⟩) + (fun s s' => IsCofilteredOrEmpty.cone_maps s s') c + +/-- The forgetful functor of the under category on any filtered or empty category is final. -/ +instance Under.final_forget [IsFilteredOrEmpty C] (c : C) : Final (Under.forget c) := + final_of_exists_of_isFiltered _ + (fun c' => ⟨mk (IsFiltered.leftToMax c c'), ⟨IsFiltered.rightToMax c c'⟩⟩) + (fun {_} {x} s s' => by + use mk (x.hom ≫ IsFiltered.coeqHom s s') + use homMk (IsFiltered.coeqHom s s') (by simp) + simp only [forget_obj, id_obj, mk_right, const_obj_obj, forget_map, homMk_right] + rw [IsFiltered.coeq_condition]) + +/-- The forgetful functor of the over category on any cofiltered or empty category is initial. -/ +instance Over.initial_forget [IsCofilteredOrEmpty C] (c : C) : Initial (Over.forget c) := + initial_of_exists_of_isCofiltered _ + (fun c' => ⟨mk (IsCofiltered.minToLeft c c'), ⟨IsCofiltered.minToRight c c'⟩⟩) + (fun {_} {x} s s' => by + use mk (IsCofiltered.eqHom s s' ≫ x.hom) + use homMk (IsCofiltered.eqHom s s') (by simp) + simp only [forget_obj, mk_left, forget_map, homMk_left] + rw [IsCofiltered.eq_condition]) + end ArbitraryUniverses section LocallySmall From 6786d49a0b8860b3738881c8193f86865c42e922 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Wed, 14 Aug 2024 16:23:34 +0000 Subject: [PATCH 273/359] chore: turn `induction'` into `induction` (#15767) - [x] turn the (soon to be deprecated) `induction'` into `induction` - [x] golf a few proofs --- Counterexamples/Phillips.lean | 8 +- Mathlib/Algebra/Algebra/Operations.lean | 28 +++---- Mathlib/Algebra/BigOperators/Fin.lean | 37 ++++----- .../Algebra/BigOperators/Group/Finset.lean | 6 +- .../Computation/CorrectnessTerminating.lean | 8 +- .../ContinuedFractions/TerminatedStable.lean | 7 +- Mathlib/Algebra/DirectSum/Internal.lean | 8 +- Mathlib/Algebra/DirectSum/Ring.lean | 7 +- Mathlib/Algebra/Group/Basic.lean | 20 +++-- Mathlib/Algebra/GroupWithZero/Basic.lean | 8 +- Mathlib/Algebra/Lie/Solvable.lean | 20 ++--- Mathlib/Algebra/MvPolynomial/Funext.lean | 8 +- .../Polynomial/Degree/Definitions.lean | 17 +++-- Mathlib/Algebra/Polynomial/Derivative.lean | 76 ++++++++++--------- Mathlib/Algebra/Polynomial/EraseLead.lean | 7 +- Mathlib/Algebra/Polynomial/Expand.lean | 11 +-- Mathlib/Algebra/Polynomial/Induction.lean | 6 +- Mathlib/Algebra/Polynomial/Reverse.lean | 6 +- Mathlib/Algebra/Polynomial/Roots.lean | 7 +- Mathlib/Algebra/Ring/CentroidHom.lean | 17 +++-- Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 12 +-- Mathlib/Algebra/Tropical/Basic.lean | 6 +- Mathlib/Analysis/Analytic/IsolatedZeros.lean | 6 +- Mathlib/Analysis/Asymptotics/Asymptotics.lean | 7 +- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 43 ++++++----- .../Analysis/Calculus/FDeriv/Analytic.lean | 28 ++++--- Mathlib/Analysis/Calculus/FDeriv/Comp.lean | 14 ++-- .../Analysis/Calculus/IteratedDeriv/Defs.lean | 7 +- .../Analysis/Normed/Algebra/Exponential.lean | 6 +- Mathlib/Analysis/Normed/Group/Basic.lean | 6 +- Mathlib/Analysis/Normed/Operator/Banach.lean | 13 ++-- Mathlib/Analysis/ODE/PicardLindelof.lean | 8 +- Mathlib/Analysis/PSeries.lean | 21 ++--- .../Analysis/SpecialFunctions/ExpDeriv.lean | 7 +- .../SpecialFunctions/Gamma/Basic.lean | 7 +- .../Analysis/SpecialFunctions/Log/Basic.lean | 11 +-- Mathlib/CategoryTheory/Extensive.lean | 13 ++-- Mathlib/CategoryTheory/Sites/Sheaf.lean | 6 +- Mathlib/Combinatorics/Hindman.lean | 16 ++-- Mathlib/Computability/Language.lean | 9 ++- Mathlib/Computability/Primrec.lean | 14 ++-- Mathlib/Computability/TMToPartrec.lean | 13 ++-- Mathlib/Data/Vector/Mem.lean | 6 +- Mathlib/GroupTheory/GroupAction/Ring.lean | 12 +-- Mathlib/GroupTheory/Perm/Closure.lean | 15 ++-- Mathlib/GroupTheory/Perm/Fin.lean | 7 +- Mathlib/GroupTheory/Solvable.lean | 32 ++++---- .../LinearAlgebra/ExteriorAlgebra/Basic.lean | 7 +- .../LinearAlgebra/TensorProduct/Basic.lean | 6 +- Mathlib/SetTheory/Surreal/Dyadic.lean | 13 ++-- Mathlib/Topology/Algebra/Module/Basic.lean | 7 +- Mathlib/Topology/TietzeExtension.lean | 7 +- Mathlib/Topology/UrysohnsLemma.lean | 29 ++++--- 53 files changed, 393 insertions(+), 333 deletions(-) diff --git a/Counterexamples/Phillips.lean b/Counterexamples/Phillips.lean index 3e9c76aabe1c0..dbccecbab4e8d 100644 --- a/Counterexamples/Phillips.lean +++ b/Counterexamples/Phillips.lean @@ -293,11 +293,13 @@ theorem exists_discrete_support_nonpos (f : BoundedAdditiveMeasure α) : simp only [s, Function.iterate_succ', Subtype.coe_mk, union_diff_left, Function.comp] have I2 : ∀ n : ℕ, (n : ℝ) * (ε / 2) ≤ f ↑(s n) := by intro n - induction' n with n IH - · simp only [s, BoundedAdditiveMeasure.empty, id, Nat.cast_zero, zero_mul, + induction n with + | zero => + simp only [s, BoundedAdditiveMeasure.empty, id, Nat.cast_zero, zero_mul, Function.iterate_zero, Subtype.coe_mk, Nat.zero_eq] rfl - · have : (s (n + 1)).1 = (s (n + 1)).1 \ (s n).1 ∪ (s n).1 := by + | succ n IH => + have : (s (n + 1)).1 = (s (n + 1)).1 \ (s n).1 ∪ (s n).1 := by simpa only [s, Function.iterate_succ', union_diff_self] using (diff_union_of_subset subset_union_left).symm rw [this, f.additive] diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 668fb28578148..d0bb76295683d 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -422,15 +422,15 @@ protected theorem pow_induction_on_left' {C : ∀ (n : ℕ) (x), x ∈ M ^ n → -- Porting note: swapped argument order to match order of `C` {n : ℕ} {x : A} (hx : x ∈ M ^ n) : C n x hx := by - induction' n with n n_ih generalizing x - · rw [pow_zero] at hx + induction n generalizing x with + | zero => + rw [pow_zero] at hx obtain ⟨r, rfl⟩ := hx exact algebraMap r - revert hx - simp_rw [pow_succ'] - intro hx - exact - Submodule.mul_induction_on' (fun m hm x ih => mem_mul _ hm _ _ _ (n_ih ih)) + | succ n n_ih => + revert hx + simp_rw [pow_succ'] + exact fun hx ↦ Submodule.mul_induction_on' (fun m hm x ih => mem_mul _ hm _ _ _ (n_ih ih)) (fun x hx y hy Cx Cy => add _ _ _ _ _ Cx Cy) hx /-- Dependent version of `Submodule.pow_induction_on_right`. -/ @@ -443,15 +443,15 @@ protected theorem pow_induction_on_right' {C : ∀ (n : ℕ) (x), x ∈ M ^ n ∀ m (hm : m ∈ M), C i.succ (x * m) (mul_mem_mul hx hm)) -- Porting note: swapped argument order to match order of `C` {n : ℕ} {x : A} (hx : x ∈ M ^ n) : C n x hx := by - induction' n with n n_ih generalizing x - · rw [pow_zero] at hx + induction n generalizing x with + | zero => + rw [pow_zero] at hx obtain ⟨r, rfl⟩ := hx exact algebraMap r - revert hx - simp_rw [pow_succ] - intro hx - exact - Submodule.mul_induction_on' (fun m hm x ih => mul_mem _ _ hm (n_ih _) _ ih) + | succ n n_ih => + revert hx + simp_rw [pow_succ] + exact fun hx ↦ Submodule.mul_induction_on' (fun m hm x ih => mul_mem _ _ hm (n_ih _) _ ih) (fun x hx y hy Cx Cy => add _ _ _ _ _ Cx Cy) hx /-- To show a property on elements of `M ^ n` holds, it suffices to show that it holds for scalars, diff --git a/Mathlib/Algebra/BigOperators/Fin.lean b/Mathlib/Algebra/BigOperators/Fin.lean index b544301f19d47..0e4e828ea9f8a 100644 --- a/Mathlib/Algebra/BigOperators/Fin.lean +++ b/Mathlib/Algebra/BigOperators/Fin.lean @@ -264,13 +264,15 @@ end Fin def finFunctionFinEquiv {m n : ℕ} : (Fin n → Fin m) ≃ Fin (m ^ n) := Equiv.ofRightInverseOfCardLE (le_of_eq <| by simp_rw [Fintype.card_fun, Fintype.card_fin]) (fun f => ⟨∑ i, f i * m ^ (i : ℕ), by - induction' n with n ih - · simp - cases m - · exact isEmptyElim (f <| Fin.last _) - simp_rw [Fin.sum_univ_castSucc, Fin.coe_castSucc, Fin.val_last] - refine (Nat.add_lt_add_of_lt_of_le (ih _) <| Nat.mul_le_mul_right _ (Fin.is_le _)).trans_eq ?_ - rw [← one_add_mul (_ : ℕ), add_comm, pow_succ']⟩) + induction n with + | zero => simp + | succ n ih => + cases m + · exact isEmptyElim (f <| Fin.last _) + simp_rw [Fin.sum_univ_castSucc, Fin.coe_castSucc, Fin.val_last] + refine (Nat.add_lt_add_of_lt_of_le (ih _) <| Nat.mul_le_mul_right _ + (Fin.is_le _)).trans_eq ?_ + rw [← one_add_mul (_ : ℕ), add_comm, pow_succ']⟩) (fun a b => ⟨a / m ^ (b : ℕ) % m, by cases' n with n · exact b.elim0 @@ -280,16 +282,17 @@ def finFunctionFinEquiv {m n : ℕ} : (Fin n → Fin m) ≃ Fin (m ^ n) := · exact Nat.mod_lt _ m.succ_pos⟩) fun a => by dsimp - induction' n with n ih - · subsingleton [(finCongr <| pow_zero _).subsingleton] - simp_rw [Fin.forall_iff, Fin.ext_iff] at ih - ext - simp_rw [Fin.sum_univ_succ, Fin.val_zero, Fin.val_succ, pow_zero, Nat.div_one, - mul_one, pow_succ', ← Nat.div_div_eq_div_mul, mul_left_comm _ m, ← mul_sum] - rw [ih _ (Nat.div_lt_of_lt_mul ?_), Nat.mod_add_div] - -- Porting note: replaces `a.is_lt` in the wildcard above. Caused by a refactor of the `npow` - -- instance for `Fin`. - exact a.is_lt.trans_eq (pow_succ' _ _) + induction n with + | zero => subsingleton [(finCongr <| pow_zero _).subsingleton] + | succ n ih => + simp_rw [Fin.forall_iff, Fin.ext_iff] at ih + ext + simp_rw [Fin.sum_univ_succ, Fin.val_zero, Fin.val_succ, pow_zero, Nat.div_one, + mul_one, pow_succ', ← Nat.div_div_eq_div_mul, mul_left_comm _ m, ← mul_sum] + rw [ih _ (Nat.div_lt_of_lt_mul ?_), Nat.mod_add_div] + -- Porting note: replaces `a.is_lt` in the wildcard above. + -- Caused by a refactor of the `npow` instance for `Fin`. + exact a.is_lt.trans_eq (pow_succ' _ _) theorem finFunctionFinEquiv_apply {m n : ℕ} (f : Fin n → Fin m) : (finFunctionFinEquiv f : ℕ) = ∑ i : Fin n, ↑(f i) * m ^ (i : ℕ) := diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index f6842570c540f..af2c8fa63c133 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1426,9 +1426,9 @@ This is a discrete analogue of the fundamental theorem of calculus."] theorem prod_range_induction (f s : ℕ → β) (base : s 0 = 1) (step : ∀ n, s (n + 1) = s n * f n) (n : ℕ) : ∏ k ∈ Finset.range n, f k = s n := by - induction' n with k hk - · rw [Finset.prod_range_zero, base] - · simp only [hk, Finset.prod_range_succ, step, mul_comm] + induction n with + | zero => rw [Finset.prod_range_zero, base] + | succ k hk => simp only [hk, Finset.prod_range_succ, step, mul_comm] /-- A telescoping product along `{0, ..., n - 1}` of a commutative group valued function reduces to the ratio of the last and first factors. -/ diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean b/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean index 735682d954a59..882c79f22ade0 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean @@ -96,8 +96,9 @@ theorem compExactValue_correctness_of_stream_eq_some : ∀ {ifp_n : IntFractPair K}, IntFractPair.stream v n = some ifp_n → v = compExactValue ((of v).contsAux n) ((of v).contsAux <| n + 1) ifp_n.fr := by let g := of v - induction' n with n IH - · intro ifp_zero stream_zero_eq + induction n with + | zero => + intro ifp_zero stream_zero_eq -- Nat.zero have : IntFractPair.of v = ifp_zero := by have : IntFractPair.stream v 0 = some (IntFractPair.of v) := rfl @@ -118,7 +119,8 @@ theorem compExactValue_correctness_of_stream_eq_some : -- Porting note: this and the if_neg rewrite are needed have : (IntFractPair.of v).fr = Int.fract v := rfl rw [this, if_neg fract_ne_zero, Int.floor_add_fract] - · intro ifp_succ_n succ_nth_stream_eq + | succ n IH => + intro ifp_succ_n succ_nth_stream_eq -- Nat.succ obtain ⟨ifp_n, nth_stream_eq, nth_fract_ne_zero, -⟩ : ∃ ifp_n, IntFractPair.stream v n = some ifp_n ∧ diff --git a/Mathlib/Algebra/ContinuedFractions/TerminatedStable.lean b/Mathlib/Algebra/ContinuedFractions/TerminatedStable.lean index 712ff1ebf90e9..f2405467d8aa7 100644 --- a/Mathlib/Algebra/ContinuedFractions/TerminatedStable.lean +++ b/Mathlib/Algebra/ContinuedFractions/TerminatedStable.lean @@ -54,9 +54,10 @@ theorem convs'Aux_stable_step_of_terminated {s : Stream'.Seq <| Pair K} theorem convs'Aux_stable_of_terminated {s : Stream'.Seq <| Pair K} (n_le_m : n ≤ m) (terminatedAt_n : s.TerminatedAt n) : convs'Aux s m = convs'Aux s n := by - induction' n_le_m with m n_le_m IH - · rfl - · refine (convs'Aux_stable_step_of_terminated ?_).trans IH + induction n_le_m with + | refl => rfl + | step n_le_m IH => + refine (convs'Aux_stable_step_of_terminated (?_)).trans IH exact s.terminated_stable n_le_m terminatedAt_n theorem conts_stable_of_terminated (n_le_m : n ≤ m) (terminatedAt_n : g.TerminatedAt n) : diff --git a/Mathlib/Algebra/DirectSum/Internal.lean b/Mathlib/Algebra/DirectSum/Internal.lean index ac5d09c97b8de..4c03bf25b2f72 100644 --- a/Mathlib/Algebra/DirectSum/Internal.lean +++ b/Mathlib/Algebra/DirectSum/Internal.lean @@ -66,10 +66,12 @@ theorem SetLike.algebraMap_mem_graded [Zero ι] [CommSemiring S] [Semiring R] [A theorem SetLike.natCast_mem_graded [Zero ι] [AddMonoidWithOne R] [SetLike σ R] [AddSubmonoidClass σ R] (A : ι → σ) [SetLike.GradedOne A] (n : ℕ) : (n : R) ∈ A 0 := by - induction' n with _ n_ih - · rw [Nat.cast_zero] + induction n with + | zero => + rw [Nat.cast_zero] exact zero_mem (A 0) - · rw [Nat.cast_succ] + | succ _ n_ih => + rw [Nat.cast_succ] exact add_mem n_ih (SetLike.one_mem_graded _) @[deprecated (since := "2024-04-17")] diff --git a/Mathlib/Algebra/DirectSum/Ring.lean b/Mathlib/Algebra/DirectSum/Ring.lean index 05a579900593c..52f4424409e6e 100644 --- a/Mathlib/Algebra/DirectSum/Ring.lean +++ b/Mathlib/Algebra/DirectSum/Ring.lean @@ -265,9 +265,10 @@ instance semiring : Semiring (⨁ i, A i) := theorem ofPow {i} (a : A i) (n : ℕ) : of _ i a ^ n = of _ (n • i) (GradedMonoid.GMonoid.gnpow _ a) := by - induction' n with n n_ih - · exact of_eq_of_gradedMonoid_eq (pow_zero <| GradedMonoid.mk _ a).symm - · rw [pow_succ, n_ih, of_mul_of] + induction n with + | zero => exact of_eq_of_gradedMonoid_eq (pow_zero <| GradedMonoid.mk _ a).symm + | succ n n_ih => + rw [pow_succ, n_ih, of_mul_of] exact of_eq_of_gradedMonoid_eq (pow_succ (GradedMonoid.mk _ a) n).symm theorem ofList_dProd {α} (l : List α) (fι : α → ι) (fA : ∀ a, A (fι a)) : diff --git a/Mathlib/Algebra/Group/Basic.lean b/Mathlib/Algebra/Group/Basic.lean index 9ebc3308c99d3..e6d74154fcb84 100644 --- a/Mathlib/Algebra/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Basic.lean @@ -852,11 +852,13 @@ addition by `g` and `-g` on the left. For additive subgroups generated by more t `AddSubgroup.closure_induction_left`."] lemma zpow_induction_left {g : G} {P : G → Prop} (h_one : P (1 : G)) (h_mul : ∀ a, P a → P (g * a)) (h_inv : ∀ a, P a → P (g⁻¹ * a)) (n : ℤ) : P (g ^ n) := by - induction' n using Int.induction_on with n ih n ih - · rwa [zpow_zero] - · rw [Int.add_comm, zpow_add, zpow_one] + induction n using Int.induction_on with + | hz => rwa [zpow_zero] + | hp n ih => + rw [Int.add_comm, zpow_add, zpow_one] exact h_mul _ ih - · rw [Int.sub_eq_add_neg, Int.add_comm, zpow_add, zpow_neg_one] + | hn n ih => + rw [Int.sub_eq_add_neg, Int.add_comm, zpow_add, zpow_neg_one] exact h_inv _ ih /-- To show a property of all powers of `g` it suffices to show it is closed under multiplication @@ -867,11 +869,13 @@ addition by `g` and `-g` on the right. For additive subgroups generated by more see `AddSubgroup.closure_induction_right`."] lemma zpow_induction_right {g : G} {P : G → Prop} (h_one : P (1 : G)) (h_mul : ∀ a, P a → P (a * g)) (h_inv : ∀ a, P a → P (a * g⁻¹)) (n : ℤ) : P (g ^ n) := by - induction' n using Int.induction_on with n ih n ih - · rwa [zpow_zero] - · rw [zpow_add_one] + induction n using Int.induction_on with + | hz => rwa [zpow_zero] + | hp n ih => + rw [zpow_add_one] exact h_mul _ ih - · rw [zpow_sub_one] + | hn n ih => + rw [zpow_sub_one] exact h_inv _ ih end Group diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index d5a2ad48bcf5f..b21f2ccf96142 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -398,10 +398,10 @@ lemma zpow_sub_one₀ (ha : a ≠ 0) (n : ℤ) : a ^ (n - 1) = a ^ n * a⁻¹ := _ = a ^ n * a⁻¹ := by rw [← zpow_add_one₀ ha, Int.sub_add_cancel] lemma zpow_add₀ (ha : a ≠ 0) (m n : ℤ) : a ^ (m + n) = a ^ m * a ^ n := by - induction' n using Int.induction_on with n ihn n ihn - · simp - · simp only [← Int.add_assoc, zpow_add_one₀ ha, ihn, mul_assoc] - · rw [zpow_sub_one₀ ha, ← mul_assoc, ← ihn, ← zpow_sub_one₀ ha, Int.add_sub_assoc] + induction n using Int.induction_on with + | hz => simp + | hp n ihn => simp only [← Int.add_assoc, zpow_add_one₀ ha, ihn, mul_assoc] + | hn n ihn => rw [zpow_sub_one₀ ha, ← mul_assoc, ← ihn, ← zpow_sub_one₀ ha, Int.add_sub_assoc] lemma zpow_add' {m n : ℤ} (h : a ≠ 0 ∨ m + n ≠ 0 ∨ m = 0 ∧ n = 0) : a ^ (m + n) = a ^ m * a ^ n := by diff --git a/Mathlib/Algebra/Lie/Solvable.lean b/Mathlib/Algebra/Lie/Solvable.lean index e458dfb53a90d..e1132be632edc 100644 --- a/Mathlib/Algebra/Lie/Solvable.lean +++ b/Mathlib/Algebra/Lie/Solvable.lean @@ -73,9 +73,9 @@ variable {R L} local notation "D" => derivedSeriesOfIdeal R L theorem derivedSeriesOfIdeal_add (k l : ℕ) : D (k + l) I = D k (D l I) := by - induction' k with k ih - · rw [Nat.zero_add, derivedSeriesOfIdeal_zero] - · rw [Nat.succ_add k l, derivedSeriesOfIdeal_succ, derivedSeriesOfIdeal_succ, ih] + induction k with + | zero => rw [Nat.zero_add, derivedSeriesOfIdeal_zero] + | succ k ih => rw [Nat.succ_add k l, derivedSeriesOfIdeal_succ, derivedSeriesOfIdeal_succ, ih] @[mono] theorem derivedSeriesOfIdeal_le {I J : LieIdeal R L} {k l : ℕ} (h₁ : I ≤ J) (h₂ : l ≤ k) : @@ -172,12 +172,14 @@ theorem derivedSeries_map_eq (k : ℕ) (h : Function.Surjective f) : theorem derivedSeries_succ_eq_top_iff (n : ℕ) : derivedSeries R L (n + 1) = ⊤ ↔ derivedSeries R L 1 = ⊤ := by simp only [derivedSeries_def] - induction' n with n ih; · simp - rw [derivedSeriesOfIdeal_succ] - refine ⟨fun h ↦ ?_, fun h ↦ by rwa [ih.mpr h]⟩ - rw [← ih, eq_top_iff] - conv_lhs => rw [← h] - exact LieSubmodule.lie_le_right _ _ + induction n with + | zero => simp + | succ n ih => + rw [derivedSeriesOfIdeal_succ] + refine ⟨fun h ↦ ?_, fun h ↦ by rwa [ih.mpr h]⟩ + rw [← ih, eq_top_iff] + conv_lhs => rw [← h] + exact LieSubmodule.lie_le_right _ _ theorem derivedSeries_eq_top (n : ℕ) (h : derivedSeries R L 1 = ⊤) : derivedSeries R L n = ⊤ := by diff --git a/Mathlib/Algebra/MvPolynomial/Funext.lean b/Mathlib/Algebra/MvPolynomial/Funext.lean index 2bff666a84d66..eaf63370de40c 100644 --- a/Mathlib/Algebra/MvPolynomial/Funext.lean +++ b/Mathlib/Algebra/MvPolynomial/Funext.lean @@ -27,11 +27,13 @@ variable {R : Type*} [CommRing R] [IsDomain R] [Infinite R] private theorem funext_fin {n : ℕ} {p : MvPolynomial (Fin n) R} (h : ∀ x : Fin n → R, eval x p = 0) : p = 0 := by - induction' n with n ih - · apply (MvPolynomial.isEmptyRingEquiv R (Fin 0)).injective + induction n with + | zero => + apply (MvPolynomial.isEmptyRingEquiv R (Fin 0)).injective rw [RingEquiv.map_zero] convert h finZeroElim - · apply (finSuccEquiv R n).injective + | succ n ih => + apply (finSuccEquiv R n).injective simp only [map_zero] refine Polynomial.funext fun q => ?_ rw [Polynomial.eval_zero] diff --git a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean index 3a4e2fc353d1c..68a3c4f053f61 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Definitions.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Definitions.lean @@ -945,11 +945,11 @@ theorem natDegree_mul_le_of_le (hp : natDegree p ≤ m) (hg : natDegree q ≤ n) natDegree_mul_le.trans <| add_le_add ‹_› ‹_› theorem natDegree_pow_le {p : R[X]} {n : ℕ} : (p ^ n).natDegree ≤ n * p.natDegree := by - induction' n with i hi - · simp - · rw [pow_succ, Nat.succ_mul] - apply le_trans natDegree_mul_le - exact add_le_add_right hi _ + induction n with + | zero => simp + | succ i hi => + rw [pow_succ, Nat.succ_mul] + apply le_trans natDegree_mul_le (add_le_add_right hi _) theorem natDegree_pow_le_of_le (n : ℕ) (hp : natDegree p ≤ m) : natDegree (p ^ n) ≤ n * m := @@ -958,9 +958,10 @@ theorem natDegree_pow_le_of_le (n : ℕ) (hp : natDegree p ≤ m) : @[simp] theorem coeff_pow_mul_natDegree (p : R[X]) (n : ℕ) : (p ^ n).coeff (n * p.natDegree) = p.leadingCoeff ^ n := by - induction' n with i hi - · simp - · rw [pow_succ, pow_succ, Nat.succ_mul] + induction n with + | zero => simp + | succ i hi => + rw [pow_succ, pow_succ, Nat.succ_mul] by_cases hp1 : p.leadingCoeff ^ i = 0 · rw [hp1, zero_mul] by_cases hp2 : p ^ i = 0 diff --git a/Mathlib/Algebra/Polynomial/Derivative.lean b/Mathlib/Algebra/Polynomial/Derivative.lean index b80dbfa008b96..d61d10d525aa9 100644 --- a/Mathlib/Algebra/Polynomial/Derivative.lean +++ b/Mathlib/Algebra/Polynomial/Derivative.lean @@ -333,44 +333,46 @@ theorem coeff_iterate_derivative {k} (p : R[X]) (m : ℕ) : theorem iterate_derivative_mul {n} (p q : R[X]) : derivative^[n] (p * q) = ∑ k ∈ range n.succ, (n.choose k • (derivative^[n - k] p * derivative^[k] q)) := by - induction' n with n IH - · simp [Finset.range] - calc - derivative^[n + 1] (p * q) = - derivative (∑ k ∈ range n.succ, - n.choose k • (derivative^[n - k] p * derivative^[k] q)) := by - rw [Function.iterate_succ_apply', IH] - _ = (∑ k ∈ range n.succ, - n.choose k • (derivative^[n - k + 1] p * derivative^[k] q)) + - ∑ k ∈ range n.succ, - n.choose k • (derivative^[n - k] p * derivative^[k + 1] q) := by - simp_rw [derivative_sum, derivative_smul, derivative_mul, Function.iterate_succ_apply', - smul_add, sum_add_distrib] - _ = (∑ k ∈ range n.succ, - n.choose k.succ • (derivative^[n - k] p * derivative^[k + 1] q)) + - 1 • (derivative^[n + 1] p * derivative^[0] q) + - ∑ k ∈ range n.succ, n.choose k • (derivative^[n - k] p * derivative^[k + 1] q) := - ?_ - _ = ((∑ k ∈ range n.succ, n.choose k • (derivative^[n - k] p * derivative^[k + 1] q)) + - ∑ k ∈ range n.succ, - n.choose k.succ • (derivative^[n - k] p * derivative^[k + 1] q)) + - 1 • (derivative^[n + 1] p * derivative^[0] q) := by - rw [add_comm, add_assoc] - _ = (∑ i ∈ range n.succ, - (n + 1).choose (i + 1) • (derivative^[n + 1 - (i + 1)] p * derivative^[i + 1] q)) + - 1 • (derivative^[n + 1] p * derivative^[0] q) := by - simp_rw [Nat.choose_succ_succ, Nat.succ_sub_succ, add_smul, sum_add_distrib] - _ = ∑ k ∈ range n.succ.succ, - n.succ.choose k • (derivative^[n.succ - k] p * derivative^[k] q) := by - rw [sum_range_succ' _ n.succ, Nat.choose_zero_right, tsub_zero] - congr - refine (sum_range_succ' _ _).trans (congr_arg₂ (· + ·) ?_ ?_) - · rw [sum_range_succ, Nat.choose_succ_self, zero_smul, add_zero] - refine sum_congr rfl fun k hk => ?_ - rw [mem_range] at hk + induction n with + | zero => + simp [Finset.range] + | succ n IH => + calc + derivative^[n + 1] (p * q) = + derivative (∑ k ∈ range n.succ, + n.choose k • (derivative^[n - k] p * derivative^[k] q)) := by + rw [Function.iterate_succ_apply', IH] + _ = (∑ k ∈ range n.succ, + n.choose k • (derivative^[n - k + 1] p * derivative^[k] q)) + + ∑ k ∈ range n.succ, + n.choose k • (derivative^[n - k] p * derivative^[k + 1] q) := by + simp_rw [derivative_sum, derivative_smul, derivative_mul, Function.iterate_succ_apply', + smul_add, sum_add_distrib] + _ = (∑ k ∈ range n.succ, + n.choose k.succ • (derivative^[n - k] p * derivative^[k + 1] q)) + + 1 • (derivative^[n + 1] p * derivative^[0] q) + + ∑ k ∈ range n.succ, n.choose k • (derivative^[n - k] p * derivative^[k + 1] q) := + ?_ + _ = ((∑ k ∈ range n.succ, n.choose k • (derivative^[n - k] p * derivative^[k + 1] q)) + + ∑ k ∈ range n.succ, + n.choose k.succ • (derivative^[n - k] p * derivative^[k + 1] q)) + + 1 • (derivative^[n + 1] p * derivative^[0] q) := by + rw [add_comm, add_assoc] + _ = (∑ i ∈ range n.succ, + (n + 1).choose (i + 1) • (derivative^[n + 1 - (i + 1)] p * derivative^[i + 1] q)) + + 1 • (derivative^[n + 1] p * derivative^[0] q) := by + simp_rw [Nat.choose_succ_succ, Nat.succ_sub_succ, add_smul, sum_add_distrib] + _ = ∑ k ∈ range n.succ.succ, + n.succ.choose k • (derivative^[n.succ - k] p * derivative^[k] q) := by + rw [sum_range_succ' _ n.succ, Nat.choose_zero_right, tsub_zero] congr - omega - · rw [Nat.choose_zero_right, tsub_zero] + refine (sum_range_succ' _ _).trans (congr_arg₂ (· + ·) ?_ ?_) + · rw [sum_range_succ, Nat.choose_succ_self, zero_smul, add_zero] + refine sum_congr rfl fun k hk => ?_ + rw [mem_range] at hk + congr + omega + · rw [Nat.choose_zero_right, tsub_zero] end Semiring diff --git a/Mathlib/Algebra/Polynomial/EraseLead.lean b/Mathlib/Algebra/Polynomial/EraseLead.lean index 56246cc180cc6..5c50263b65651 100644 --- a/Mathlib/Algebra/Polynomial/EraseLead.lean +++ b/Mathlib/Algebra/Polynomial/EraseLead.lean @@ -339,9 +339,10 @@ theorem card_support_eq {n : ℕ} : ∃ (k : Fin n → ℕ) (x : Fin n → R) (hk : StrictMono k) (hx : ∀ i, x i ≠ 0), f = ∑ i, C (x i) * X ^ k i := by refine ⟨?_, fun ⟨k, x, hk, hx, hf⟩ => hf.symm ▸ card_support_eq' k x hk.injective hx⟩ - induction' n with n hn generalizing f - · exact fun hf => ⟨0, 0, fun x => x.elim0, fun x => x.elim0, card_support_eq_zero.mp hf⟩ - · intro h + induction n generalizing f with + | zero => exact fun hf => ⟨0, 0, fun x => x.elim0, fun x => x.elim0, card_support_eq_zero.mp hf⟩ + | succ n hn => + intro h obtain ⟨k, x, hk, hx, hf⟩ := hn (card_support_eraseLead' h) have H : ¬∃ k : Fin n, Fin.castSucc k = Fin.last n := by rintro ⟨i, hi⟩ diff --git a/Mathlib/Algebra/Polynomial/Expand.lean b/Mathlib/Algebra/Polynomial/Expand.lean index d31b700d20a34..e1db94a64141b 100644 --- a/Mathlib/Algebra/Polynomial/Expand.lean +++ b/Mathlib/Algebra/Polynomial/Expand.lean @@ -235,11 +235,12 @@ theorem expand_char (f : R[X]) : map (frobenius R p) (expand R p f) = f ^ p := b theorem map_expand_pow_char (f : R[X]) (n : ℕ) : map (frobenius R p ^ n) (expand R (p ^ n) f) = f ^ p ^ n := by - induction' n with _ n_ih - · simp [RingHom.one_def] - symm - rw [pow_succ, pow_mul, ← n_ih, ← expand_char, pow_succ', RingHom.mul_def, ← map_map, mul_comm, - expand_mul, ← map_expand] + induction n with + | zero => simp [RingHom.one_def] + | succ _ n_ih => + symm + rw [pow_succ, pow_mul, ← n_ih, ← expand_char, pow_succ', RingHom.mul_def, ← map_map, mul_comm, + expand_mul, ← map_expand] end ExpChar diff --git a/Mathlib/Algebra/Polynomial/Induction.lean b/Mathlib/Algebra/Polynomial/Induction.lean index 31c9429ead2ad..3e3fdb290f72c 100644 --- a/Mathlib/Algebra/Polynomial/Induction.lean +++ b/Mathlib/Algebra/Polynomial/Induction.lean @@ -39,9 +39,9 @@ protected theorem induction_on {M : R[X] → Prop} (p : R[X]) (h_C : ∀ a, M (C (h_monomial : ∀ (n : ℕ) (a : R), M (C a * X ^ n) → M (C a * X ^ (n + 1))) : M p := by have A : ∀ {n : ℕ} {a}, M (C a * X ^ n) := by intro n a - induction' n with n ih - · rw [pow_zero, mul_one]; exact h_C a - · exact h_monomial _ _ ih + induction n with + | zero => rw [pow_zero, mul_one]; exact h_C a + | succ n ih => exact h_monomial _ _ ih have B : ∀ s : Finset ℕ, M (s.sum fun n : ℕ => C (p.coeff n) * X ^ n) := by apply Finset.induction · convert h_C 0 diff --git a/Mathlib/Algebra/Polynomial/Reverse.lean b/Mathlib/Algebra/Polynomial/Reverse.lean index 6b21c7277e1c1..788dcb3ce0de2 100644 --- a/Mathlib/Algebra/Polynomial/Reverse.lean +++ b/Mathlib/Algebra/Polynomial/Reverse.lean @@ -313,9 +313,9 @@ theorem coeff_one_reverse (f : R[X]) : coeff (reverse f) 1 = nextCoeff f := by rw [commute_X p, reverse_mul_X] @[simp] lemma reverse_mul_X_pow (p : R[X]) (n : ℕ) : reverse (p * X ^ n) = reverse p := by - induction' n with n ih - · simp - rw [pow_succ, ← mul_assoc, reverse_mul_X, ih] + induction n with + | zero => simp + | succ n ih => rw [pow_succ, ← mul_assoc, reverse_mul_X, ih] @[simp] lemma reverse_X_pow_mul (p : R[X]) (n : ℕ) : reverse (X ^ n * p) = reverse p := by rw [commute_X_pow p, reverse_mul_X_pow] diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index 06b3d6fd29716..6b7ad04a342d1 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -206,9 +206,10 @@ theorem roots_prod {ι : Type*} (f : ι → R[X]) (s : Finset ι) : @[simp] theorem roots_pow (p : R[X]) (n : ℕ) : (p ^ n).roots = n • p.roots := by - induction' n with n ihn - · rw [pow_zero, roots_one, zero_smul, empty_eq_zero] - · rcases eq_or_ne p 0 with (rfl | hp) + induction n with + | zero => rw [pow_zero, roots_one, zero_smul, empty_eq_zero] + | succ n ihn => + rcases eq_or_ne p 0 with (rfl | hp) · rw [zero_pow n.succ_ne_zero, roots_zero, smul_zero] · rw [pow_succ, roots_mul (mul_ne_zero (pow_ne_zero _ hp) hp), ihn, add_smul, one_smul] diff --git a/Mathlib/Algebra/Ring/CentroidHom.lean b/Mathlib/Algebra/Ring/CentroidHom.lean index 111c5f47148d5..68eb72c1e0691 100644 --- a/Mathlib/Algebra/Ring/CentroidHom.lean +++ b/Mathlib/Algebra/Ring/CentroidHom.lean @@ -270,16 +270,17 @@ instance hasNPowNat : Pow (CentroidHom α) ℕ := ⟨fun f n ↦ { toAddMonoidHom := (f.toEnd ^ n : AddMonoid.End α) map_mul_left' := fun a b ↦ by - induction' n with n ih - · exact rfl - · rw [pow_succ'] + induction n with + | zero => rfl + | succ n ih => + rw [pow_succ'] exact (congr_arg f.toEnd ih).trans (f.map_mul_left' _ _) map_mul_right' := fun a b ↦ by - induction' n with n ih - · exact rfl - · rw [pow_succ'] - exact (congr_arg f.toEnd ih).trans (f.map_mul_right' _ _) - }⟩ + induction n with + | zero => rfl + | succ n ih => + rw [pow_succ'] + exact (congr_arg f.toEnd ih).trans (f.map_mul_right' _ _)}⟩ @[simp, norm_cast] theorem coe_zero : ⇑(0 : CentroidHom α) = 0 := diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index 545f71d03475b..72f996c586990 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -102,9 +102,9 @@ instance (priority := 75) toSemiring {R} [Semiring R] [SetLike S R] [Subsemiring @[simp, norm_cast] theorem coe_pow {R} [Semiring R] [SetLike S R] [SubsemiringClass S R] (x : s) (n : ℕ) : ((x ^ n : s) : R) = (x : R) ^ n := by - induction' n with n ih - · simp - · simp [pow_succ, ih] + induction n with + | zero => simp + | succ n ih => simp [pow_succ, ih] /-- A subsemiring of a `CommSemiring` is a `CommSemiring`. -/ instance toCommSemiring {R} [CommSemiring R] [SetLike S R] [SubsemiringClass S R] : @@ -314,9 +314,9 @@ instance toSemiring {R} [Semiring R] (s : Subsemiring R) : Semiring s := @[simp, norm_cast] theorem coe_pow {R} [Semiring R] (s : Subsemiring R) (x : s) (n : ℕ) : ((x ^ n : s) : R) = (x : R) ^ n := by - induction' n with n ih - · simp - · simp [pow_succ, ih] + induction n with + | zero => simp + | succ n ih => simp [pow_succ, ih] /-- A subsemiring of a `CommSemiring` is a `CommSemiring`. -/ instance toCommSemiring {R} [CommSemiring R] (s : Subsemiring R) : CommSemiring s := diff --git a/Mathlib/Algebra/Tropical/Basic.lean b/Mathlib/Algebra/Tropical/Basic.lean index 7d620bea68f04..5454363d36711 100644 --- a/Mathlib/Algebra/Tropical/Basic.lean +++ b/Mathlib/Algebra/Tropical/Basic.lean @@ -489,9 +489,9 @@ instance : CommSemiring (Tropical R) := @[simp] theorem succ_nsmul {R} [LinearOrder R] [OrderTop R] (x : Tropical R) (n : ℕ) : (n + 1) • x = x := by - induction' n with n IH - · simp - · rw [add_nsmul, IH, one_nsmul, add_self] + induction n with + | zero => simp + | succ n IH => rw [add_nsmul, IH, one_nsmul, add_self] -- TODO: find/create the right classes to make this hold (for enat, ennreal, etc) -- Requires `zero_eq_bot` to be true diff --git a/Mathlib/Analysis/Analytic/IsolatedZeros.lean b/Mathlib/Analysis/Analytic/IsolatedZeros.lean index 45a5d24651d57..720708a18aa23 100644 --- a/Mathlib/Analysis/Analytic/IsolatedZeros.lean +++ b/Mathlib/Analysis/Analytic/IsolatedZeros.lean @@ -77,9 +77,9 @@ theorem has_fpower_series_dslope_fslope (hp : HasFPowerSeriesAt f p z₀) : theorem has_fpower_series_iterate_dslope_fslope (n : ℕ) (hp : HasFPowerSeriesAt f p z₀) : HasFPowerSeriesAt ((swap dslope z₀)^[n] f) (fslope^[n] p) z₀ := by - induction' n with n ih generalizing f p - · exact hp - · simpa using ih (has_fpower_series_dslope_fslope hp) + induction n generalizing f p with + | zero => exact hp + | succ n ih => simpa using ih (has_fpower_series_dslope_fslope hp) theorem iterate_dslope_fslope_ne_zero (hp : HasFPowerSeriesAt f p z₀) (h : p ≠ 0) : (swap dslope z₀)^[p.order] f z₀ ≠ 0 := by diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 945fbda42c7d1..1feaea8e67fc7 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -1393,9 +1393,9 @@ theorem IsBigO.of_pow {f : α → 𝕜} {g : α → R} {n : ℕ} (hn : n ≠ 0) theorem IsLittleO.pow {f : α → R} {g : α → 𝕜} (h : f =o[l] g) {n : ℕ} (hn : 0 < n) : (fun x => f x ^ n) =o[l] fun x => g x ^ n := by obtain ⟨n, rfl⟩ := Nat.exists_eq_succ_of_ne_zero hn.ne'; clear hn - induction' n with n ihn - · simpa only [Nat.zero_eq, ← Nat.one_eq_succ_zero, pow_one] - · convert ihn.mul h <;> simp [pow_succ] + induction n with + | zero => simpa only [Nat.zero_eq, ← Nat.one_eq_succ_zero, pow_one] + | succ n ihn => convert ihn.mul h <;> simp [pow_succ] theorem IsLittleO.of_pow {f : α → 𝕜} {g : α → R} {n : ℕ} (h : (f ^ n) =o[l] (g ^ n)) (hn : n ≠ 0) : f =o[l] g := @@ -1403,7 +1403,6 @@ theorem IsLittleO.of_pow {f : α → 𝕜} {g : α → R} {n : ℕ} (h : (f ^ n) /-! ### Inverse -/ - theorem IsBigOWith.inv_rev {f : α → 𝕜} {g : α → 𝕜'} (h : IsBigOWith c l f g) (h₀ : ∀ᶠ x in l, f x = 0 → g x = 0) : IsBigOWith c l (fun x => (g x)⁻¹) fun x => (f x)⁻¹ := by refine IsBigOWith.of_bound (h.bound.mp (h₀.mono fun x h₀ hle => ?_)) diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index c18cbf87b9e2d..f5ce7711b052d 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -839,11 +839,11 @@ lemma iteratedFDerivWithin_two_apply (f : E → F) {z : E} (hs : UniqueDiffOn theorem Filter.EventuallyEq.iteratedFDerivWithin' (h : f₁ =ᶠ[𝓝[s] x] f) (ht : t ⊆ s) (n : ℕ) : iteratedFDerivWithin 𝕜 n f₁ t =ᶠ[𝓝[s] x] iteratedFDerivWithin 𝕜 n f t := by - induction' n with n ihn - · exact h.mono fun y hy => DFunLike.ext _ _ fun _ => hy - · have : fderivWithin 𝕜 _ t =ᶠ[𝓝[s] x] fderivWithin 𝕜 _ t := ihn.fderivWithin' ht - apply this.mono - intro y hy + induction n with + | zero => exact h.mono fun y hy => DFunLike.ext _ _ fun _ => hy + | succ n ihn => + have : fderivWithin 𝕜 _ t =ᶠ[𝓝[s] x] fderivWithin 𝕜 _ t := ihn.fderivWithin' ht + refine this.mono fun y hy => ?_ simp only [iteratedFDerivWithin_succ_eq_comp_left, hy, (· ∘ ·)] protected theorem Filter.EventuallyEq.iteratedFDerivWithin (h : f₁ =ᶠ[𝓝[s] x] f) (n : ℕ) : @@ -873,9 +873,10 @@ protected theorem Set.EqOn.iteratedFDerivWithin (hs : EqOn f₁ f s) (n : ℕ) : theorem iteratedFDerivWithin_eventually_congr_set' (y : E) (h : s =ᶠ[𝓝[{y}ᶜ] x] t) (n : ℕ) : iteratedFDerivWithin 𝕜 n f s =ᶠ[𝓝 x] iteratedFDerivWithin 𝕜 n f t := by - induction' n with n ihn generalizing x - · rfl - · refine (eventually_nhds_nhdsWithin.2 h).mono fun y hy => ?_ + induction n generalizing x with + | zero => rfl + | succ n ihn => + refine (eventually_nhds_nhdsWithin.2 h).mono fun y hy => ?_ simp only [iteratedFDerivWithin_succ_eq_comp_left, (· ∘ ·)] rw [(ihn hy).fderivWithin_eq_nhds, fderivWithin_congr_set' _ hy] @@ -907,8 +908,7 @@ theorem iteratedFDerivWithin_inter_open {n : ℕ} (hu : IsOpen u) (hx : x ∈ u) @[simp] theorem contDiffOn_zero : ContDiffOn 𝕜 0 f s ↔ ContinuousOn f s := by - refine ⟨fun H => H.continuousOn, fun H => ?_⟩ - intro x hx m hm + refine ⟨fun H => H.continuousOn, fun H => fun x hx m hm ↦ ?_⟩ have : (m : ℕ∞) = 0 := le_antisymm hm bot_le rw [this] refine ⟨insert x s, self_mem_nhdsWithin, ftaylorSeriesWithin 𝕜 f s, ?_⟩ @@ -1448,11 +1448,13 @@ theorem fderiv_iteratedFDeriv {n : ℕ} : simp only [Function.comp_apply, LinearIsometryEquiv.symm_apply_apply] theorem tsupport_iteratedFDeriv_subset (n : ℕ) : tsupport (iteratedFDeriv 𝕜 n f) ⊆ tsupport f := by - induction' n with n IH - · rw [iteratedFDeriv_zero_eq_comp] + induction n with + | zero => + rw [iteratedFDeriv_zero_eq_comp] exact closure_minimal ((support_comp_subset (LinearIsometryEquiv.map_zero _) _).trans subset_closure) isClosed_closure - · rw [iteratedFDeriv_succ_eq_comp_left] + | succ n IH => + rw [iteratedFDeriv_succ_eq_comp_left] exact closure_minimal ((support_comp_subset (LinearIsometryEquiv.map_zero _) _).trans ((support_fderiv_subset 𝕜).trans IH)) isClosed_closure @@ -1470,9 +1472,10 @@ theorem norm_fderiv_iteratedFDeriv {n : ℕ} : theorem iteratedFDerivWithin_univ {n : ℕ} : iteratedFDerivWithin 𝕜 n f univ = iteratedFDeriv 𝕜 n f := by - induction' n with n IH - · ext x; simp - · ext x m + induction n with + | zero => ext x; simp + | succ n IH => + ext x m rw [iteratedFDeriv_succ_apply_left, iteratedFDerivWithin_succ_apply_left, IH, fderivWithin_univ] theorem HasFTaylorSeriesUpTo.eq_iteratedFDeriv @@ -1486,11 +1489,13 @@ theorem HasFTaylorSeriesUpTo.eq_iteratedFDeriv derivative. -/ theorem iteratedFDerivWithin_of_isOpen (n : ℕ) (hs : IsOpen s) : EqOn (iteratedFDerivWithin 𝕜 n f s) (iteratedFDeriv 𝕜 n f) s := by - induction' n with n IH - · intro x _ + induction n with + | zero => + intro x _ ext1 simp only [Nat.zero_eq, iteratedFDerivWithin_zero_apply, iteratedFDeriv_zero_apply] - · intro x hx + | succ n IH => + intro x hx rw [iteratedFDeriv_succ_eq_comp_left, iteratedFDerivWithin_succ_eq_comp_left] dsimp congr 1 diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 8b829d21486fe..982fb6e51c423 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -98,10 +98,12 @@ theorem AnalyticOn.fderiv [CompleteSpace F] (h : AnalyticOn 𝕜 f s) : /-- If a function is analytic on a set `s`, so are its successive Fréchet derivative. -/ theorem AnalyticOn.iteratedFDeriv [CompleteSpace F] (h : AnalyticOn 𝕜 f s) (n : ℕ) : AnalyticOn 𝕜 (iteratedFDeriv 𝕜 n f) s := by - induction' n with n IH - · rw [iteratedFDeriv_zero_eq_comp] + induction n with + | zero => + rw [iteratedFDeriv_zero_eq_comp] exact ((continuousMultilinearCurryFin0 𝕜 E F).symm : F →L[𝕜] E[×0]→L[𝕜] F).comp_analyticOn h - · rw [iteratedFDeriv_succ_eq_comp_left] + | succ n IH => + rw [iteratedFDeriv_succ_eq_comp_left] -- Porting note: for reasons that I do not understand at all, `?g` cannot be inlined. convert ContinuousLinearMap.comp_analyticOn ?g IH.fderiv case g => exact ↑(continuousMultilinearCurryLeftEquiv 𝕜 (fun _ : Fin (n + 1) ↦ E) F) @@ -151,9 +153,9 @@ theorem AnalyticOn.deriv [CompleteSpace F] (h : AnalyticOn 𝕜 f s) : AnalyticO /-- If a function is analytic on a set `s`, so are its successive derivatives. -/ theorem AnalyticOn.iterated_deriv [CompleteSpace F] (h : AnalyticOn 𝕜 f s) (n : ℕ) : AnalyticOn 𝕜 (_root_.deriv^[n] f) s := by - induction' n with n IH - · exact h - · simpa only [Function.iterate_succ', Function.comp_apply] using IH.deriv + induction n with + | zero => exact h + | succ n IH => simpa only [Function.iterate_succ', Function.comp_apply] using IH.deriv end deriv section fderiv @@ -221,10 +223,12 @@ theorem CPolynomialOn.fderiv (h : CPolynomialOn 𝕜 f s) : /-- If a function is polynomial on a set `s`, so are its successive Fréchet derivative. -/ theorem CPolynomialOn.iteratedFDeriv (h : CPolynomialOn 𝕜 f s) (n : ℕ) : CPolynomialOn 𝕜 (iteratedFDeriv 𝕜 n f) s := by - induction' n with n IH - · rw [iteratedFDeriv_zero_eq_comp] + induction n with + | zero => + rw [iteratedFDeriv_zero_eq_comp] exact ((continuousMultilinearCurryFin0 𝕜 E F).symm : F →L[𝕜] E[×0]→L[𝕜] F).comp_cPolynomialOn h - · rw [iteratedFDeriv_succ_eq_comp_left] + | succ n IH => + rw [iteratedFDeriv_succ_eq_comp_left] convert ContinuousLinearMap.comp_cPolynomialOn ?g IH.fderiv case g => exact ↑(continuousMultilinearCurryLeftEquiv 𝕜 (fun _ : Fin (n + 1) ↦ E) F) simp @@ -261,9 +265,9 @@ protected theorem CPolynomialOn.deriv (h : CPolynomialOn 𝕜 f s) : CPolynomial /-- If a function is polynomial on a set `s`, so are its successive derivatives. -/ theorem CPolynomialOn.iterated_deriv (h : CPolynomialOn 𝕜 f s) (n : ℕ) : CPolynomialOn 𝕜 (deriv^[n] f) s := by - induction' n with n IH - · exact h - · simpa only [Function.iterate_succ', Function.comp_apply] using IH.deriv + induction n with + | zero => exact h + | succ n IH => simpa only [Function.iterate_succ', Function.comp_apply] using IH.deriv end deriv diff --git a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean index 3a29e467e58b6..4aea3d3f12abc 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean @@ -191,9 +191,10 @@ variable {x} protected theorem HasFDerivAtFilter.iterate {f : E → E} {f' : E →L[𝕜] E} (hf : HasFDerivAtFilter f f' x L) (hL : Tendsto f L L) (hx : f x = x) (n : ℕ) : HasFDerivAtFilter f^[n] (f' ^ n) x L := by - induction' n with n ihn - · exact hasFDerivAtFilter_id x L - · rw [Function.iterate_succ, pow_succ] + induction n with + | zero => exact hasFDerivAtFilter_id x L + | succ n ihn => + rw [Function.iterate_succ, pow_succ] rw [← hx] at ihn exact ihn.comp x hf hL @@ -218,9 +219,10 @@ protected theorem HasFDerivWithinAt.iterate {f : E → E} {f' : E →L[𝕜] E} protected theorem HasStrictFDerivAt.iterate {f : E → E} {f' : E →L[𝕜] E} (hf : HasStrictFDerivAt f f' x) (hx : f x = x) (n : ℕ) : HasStrictFDerivAt f^[n] (f' ^ n) x := by - induction' n with n ihn - · exact hasStrictFDerivAt_id x - · rw [Function.iterate_succ, pow_succ] + induction n with + | zero => exact hasStrictFDerivAt_id x + | succ n ihn => + rw [Function.iterate_succ, pow_succ] rw [← hx] at ihn exact ihn.comp x hf diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean index 32cf207826227..b662ec13e36fc 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean @@ -174,9 +174,10 @@ theorem iteratedDerivWithin_succ {x : 𝕜} (hxs : UniqueDiffWithinAt 𝕜 s x) iterating `n` times the differentiation operation. -/ theorem iteratedDerivWithin_eq_iterate {x : 𝕜} (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : iteratedDerivWithin n f s x = (fun g : 𝕜 → F => derivWithin g s)^[n] f x := by - induction' n with n IH generalizing x - · simp - · rw [iteratedDerivWithin_succ (hs x hx), Function.iterate_succ'] + induction n generalizing x with + | zero => simp + | succ n IH => + rw [iteratedDerivWithin_succ (hs x hx), Function.iterate_succ'] exact derivWithin_congr (fun y hy => IH hy) (IH hx) /-- The `n+1`-th iterated derivative within a set with unique derivatives can be obtained by diff --git a/Mathlib/Analysis/Normed/Algebra/Exponential.lean b/Mathlib/Analysis/Normed/Algebra/Exponential.lean index f9679a22cb191..f95b688f0ee8b 100644 --- a/Mathlib/Analysis/Normed/Algebra/Exponential.lean +++ b/Mathlib/Analysis/Normed/Algebra/Exponential.lean @@ -482,9 +482,9 @@ theorem exp_sum_of_commute {ι} (s : Finset ι) (f : ι → 𝔸) exact h.of_refl (Finset.mem_insert_self _ _) (Finset.mem_insert_of_mem hi) theorem exp_nsmul (n : ℕ) (x : 𝔸) : exp 𝕂 (n • x) = exp 𝕂 x ^ n := by - induction' n with n ih - · rw [zero_smul, pow_zero, exp_zero] - · rw [succ_nsmul, pow_succ, exp_add_of_commute ((Commute.refl x).smul_left n), ih] + induction n with + | zero => rw [zero_smul, pow_zero, exp_zero] + | succ n ih => rw [succ_nsmul, pow_succ, exp_add_of_commute ((Commute.refl x).smul_left n), ih] variable (𝕂) diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 8c619280197ef..d5392d8c567b0 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -1022,8 +1022,10 @@ theorem preimage_mul_sphere (a b : E) (r : ℝ) : (b * ·) ⁻¹' sphere a r = s @[to_additive norm_nsmul_le] theorem norm_pow_le_mul_norm (n : ℕ) (a : E) : ‖a ^ n‖ ≤ n * ‖a‖ := by - induction' n with n ih; · simp - simpa only [pow_succ, Nat.cast_succ, add_mul, one_mul] using norm_mul_le_of_le ih le_rfl + induction n with + | zero => simp + | succ n ih => + simpa only [pow_succ, Nat.cast_succ, add_mul, one_mul] using norm_mul_le_of_le ih le_rfl @[to_additive nnnorm_nsmul_le] theorem nnnorm_pow_le_mul_norm (n : ℕ) (a : E) : ‖a ^ n‖₊ ≤ n * ‖a‖₊ := by diff --git a/Mathlib/Analysis/Normed/Operator/Banach.lean b/Mathlib/Analysis/Normed/Operator/Banach.lean index 26a210ee9dcea..9a19ff37a6355 100644 --- a/Mathlib/Analysis/Normed/Operator/Banach.lean +++ b/Mathlib/Analysis/Normed/Operator/Banach.lean @@ -177,9 +177,10 @@ theorem exists_preimage_norm_le (surj : Surjective f) : refine ⟨2 * C + 1, by linarith, fun y => ?_⟩ have hnle : ∀ n : ℕ, ‖h^[n] y‖ ≤ (1 / 2) ^ n * ‖y‖ := by intro n - induction' n with n IH - · simp only [one_div, Nat.zero_eq, one_mul, iterate_zero_apply, pow_zero, le_rfl] - · rw [iterate_succ'] + induction n with + | zero => simp only [one_div, Nat.zero_eq, one_mul, iterate_zero_apply, pow_zero, le_rfl] + | succ n IH => + rw [iterate_succ'] apply le_trans (hle _) _ rw [pow_succ', mul_assoc] gcongr @@ -205,9 +206,9 @@ theorem exists_preimage_norm_le (surj : Surjective f) : _ = (2 * C + 1) * ‖y‖ := by ring have fsumeq : ∀ n : ℕ, f (∑ i ∈ Finset.range n, u i) = y - h^[n] y := by intro n - induction' n with n IH - · simp [f.map_zero] - · rw [sum_range_succ, f.map_add, IH, iterate_succ_apply', sub_add] + induction n with + | zero => simp [f.map_zero] + | succ n IH => rw [sum_range_succ, f.map_add, IH, iterate_succ_apply', sub_add] have : Tendsto (fun n => ∑ i ∈ Finset.range n, u i) atTop (𝓝 x) := su.hasSum.tendsto_sum_nat have L₁ : Tendsto (fun n => f (∑ i ∈ Finset.range n, u i)) atTop (𝓝 (f x)) := (f.continuous.tendsto _).comp this diff --git a/Mathlib/Analysis/ODE/PicardLindelof.lean b/Mathlib/Analysis/ODE/PicardLindelof.lean index ed06739eb6415..cb09d82fa6836 100644 --- a/Mathlib/Analysis/ODE/PicardLindelof.lean +++ b/Mathlib/Analysis/ODE/PicardLindelof.lean @@ -265,10 +265,12 @@ theorem dist_next_apply_le_of_le {f₁ f₂ : FunSpace v} {n : ℕ} {d : ℝ} theorem dist_iterate_next_apply_le (f₁ f₂ : FunSpace v) (n : ℕ) (t : Icc v.tMin v.tMax) : dist (next^[n] f₁ t) (next^[n] f₂ t) ≤ (v.L * |t.1 - v.t₀|) ^ n / n ! * dist f₁ f₂ := by - induction' n with n ihn generalizing t - · rw [pow_zero, Nat.factorial_zero, Nat.cast_one, div_one, one_mul] + induction n generalizing t with + | zero => + rw [pow_zero, Nat.factorial_zero, Nat.cast_one, div_one, one_mul] exact dist_apply_le_dist f₁ f₂ t - · rw [iterate_succ_apply', iterate_succ_apply'] + | succ n ihn => + rw [iterate_succ_apply', iterate_succ_apply'] exact dist_next_apply_le_of_le ihn _ theorem dist_iterate_next_le (f₁ f₂ : FunSpace v) (n : ℕ) : diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 8a19619191174..12be1fc3643b6 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -48,16 +48,17 @@ variable {M : Type*} [OrderedAddCommMonoid M] {f : ℕ → M} {u : ℕ → ℕ} theorem le_sum_schlomilch' (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f m) (h_pos : ∀ n, 0 < u n) (hu : Monotone u) (n : ℕ) : (∑ k ∈ Ico (u 0) (u n), f k) ≤ ∑ k ∈ range n, (u (k + 1) - u k) • f (u k) := by - induction' n with n ihn - · simp - suffices (∑ k ∈ Ico (u n) (u (n + 1)), f k) ≤ (u (n + 1) - u n) • f (u n) by - rw [sum_range_succ, ← sum_Ico_consecutive] - · exact add_le_add ihn this - exacts [hu n.zero_le, hu n.le_succ] - have : ∀ k ∈ Ico (u n) (u (n + 1)), f k ≤ f (u n) := fun k hk => - hf (Nat.succ_le_of_lt (h_pos n)) (mem_Ico.mp hk).1 - convert sum_le_sum this - simp [pow_succ, mul_two] + induction n with + | zero => simp + | succ n ihn => + suffices (∑ k ∈ Ico (u n) (u (n + 1)), f k) ≤ (u (n + 1) - u n) • f (u n) by + rw [sum_range_succ, ← sum_Ico_consecutive] + · exact add_le_add ihn this + exacts [hu n.zero_le, hu n.le_succ] + have : ∀ k ∈ Ico (u n) (u (n + 1)), f k ≤ f (u n) := fun k hk => + hf (Nat.succ_le_of_lt (h_pos n)) (mem_Ico.mp hk).1 + convert sum_le_sum this + simp [pow_succ, mul_two] theorem le_sum_condensed' (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f m) (n : ℕ) : (∑ k ∈ Ico 1 (2 ^ n), f k) ≤ ∑ k ∈ range n, 2 ^ k • f (2 ^ k) := by diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index 4079b96dab467..31ff5558fe80a 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -56,9 +56,10 @@ theorem contDiff_exp : ∀ {n}, ContDiff 𝕜 n exp := by -- Porting note: added `@` due to `∀ {n}` weirdness above refine @(contDiff_all_iff_nat.2 fun n => ?_) have : ContDiff ℂ (↑n) exp := by - induction' n with n ihn - · exact contDiff_zero.2 continuous_exp - · rw [contDiff_succ_iff_deriv] + induction n with + | zero => exact contDiff_zero.2 continuous_exp + | succ n ihn => + rw [contDiff_succ_iff_deriv] use differentiable_exp rwa [deriv_exp] exact this.restrict_scalars 𝕜 diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index 10561370fb1a4..bd137a24eea7c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -357,9 +357,10 @@ theorem Gamma_conj (s : ℂ) : Gamma (conj s) = conj (Gamma s) := by suffices ∀ (n : ℕ) (s : ℂ), GammaAux n (conj s) = conj (GammaAux n s) by simp [Gamma, this] intro n - induction' n with n IH - · rw [GammaAux]; exact GammaIntegral_conj - · intro s + induction n with + | zero => rw [GammaAux]; exact GammaIntegral_conj + | succ n IH => + intro s rw [GammaAux] dsimp only rw [div_eq_mul_inv _ s, RingHom.map_mul, conj_inv, ← div_eq_mul_inv] diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index 3593e3aa57d9e..7b1feb85d5ed5 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -252,11 +252,12 @@ theorem log_ne_zero {x : ℝ} : log x ≠ 0 ↔ x ≠ 0 ∧ x ≠ 1 ∧ x ≠ -1 @[simp] theorem log_pow (x : ℝ) (n : ℕ) : log (x ^ n) = n * log x := by - induction' n with n ih - · simp - rcases eq_or_ne x 0 with (rfl | hx) - · simp - rw [pow_succ, log_mul (pow_ne_zero _ hx) hx, ih, Nat.cast_succ, add_mul, one_mul] + induction n with + | zero => simp + | succ n ih => + rcases eq_or_ne x 0 with (rfl | hx) + · simp + · rw [pow_succ, log_mul (pow_ne_zero _ hx) hx, ih, Nat.cast_succ, add_mul, one_mul] @[simp] theorem log_zpow (x : ℝ) (n : ℤ) : log (x ^ n) = n * log x := by diff --git a/Mathlib/CategoryTheory/Extensive.lean b/Mathlib/CategoryTheory/Extensive.lean index be20dc8fc8fae..b4f9a69604b9d 100644 --- a/Mathlib/CategoryTheory/Extensive.lean +++ b/Mathlib/CategoryTheory/Extensive.lean @@ -426,15 +426,12 @@ theorem FinitaryPreExtensive.isUniversal_finiteCoproducts_Fin [FinitaryPreExtens Functor.hext (fun _ ↦ rfl) (by rintro ⟨i⟩ ⟨j⟩ ⟨⟨rfl : i = j⟩⟩; simp [f]) clear_value f subst this - induction' n with n IH - · exact (isVanKampenColimit_of_isEmpty _ hc).isUniversal - · apply IsUniversalColimit.of_iso _ + induction n with + | zero => exact (isVanKampenColimit_of_isEmpty _ hc).isUniversal + | succ n IH => + refine IsUniversalColimit.of_iso (@isUniversalColimit_extendCofan _ _ _ _ _ _ + (IH _ (coproductIsCoproduct _)) (FinitaryPreExtensive.universal' _ (coprodIsCoprod _ _)) ?_) ((extendCofanIsColimit f (coproductIsCoproduct _) (coprodIsCoprod _ _)).uniqueUpToIso hc) - apply @isUniversalColimit_extendCofan _ _ _ _ _ _ _ _ ?_ - · apply IH - exact coproductIsCoproduct _ - · apply FinitaryPreExtensive.universal' - exact coprodIsCoprod _ _ · dsimp infer_instance diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index 2de13cf6759ad..13b38d99b0769 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -462,9 +462,9 @@ instance sheafHomHasNSMul : SMul ℕ (P ⟶ Q) where Sheaf.Hom.mk { app := fun U => n • f.1.app U naturality := fun U V i => by - induction' n with n ih - · simp only [zero_smul, comp_zero, zero_comp, Nat.zero_eq] - · simp only [Nat.succ_eq_add_one, add_smul, ih, one_nsmul, comp_add, + induction n with + | zero => simp only [zero_smul, comp_zero, zero_comp, Nat.zero_eq] + | succ n ih => simp only [Nat.succ_eq_add_one, add_smul, ih, one_nsmul, comp_add, NatTrans.naturality, add_comp] } instance : Zero (P ⟶ Q) where zero := Sheaf.Hom.mk 0 diff --git a/Mathlib/Combinatorics/Hindman.lean b/Mathlib/Combinatorics/Hindman.lean index d509718bb8861..0a41f5051974a 100644 --- a/Mathlib/Combinatorics/Hindman.lean +++ b/Mathlib/Combinatorics/Hindman.lean @@ -205,17 +205,17 @@ theorem exists_FP_of_finite_cover {M} [Semigroup M] [Nonempty M] (s : Set (Set M @[to_additive FS_iter_tail_sub_FS] theorem FP_drop_subset_FP {M} [Semigroup M] (a : Stream' M) (n : ℕ) : FP (a.drop n) ⊆ FP a := by - induction' n with n ih - · rfl - rw [Nat.add_comm, ← Stream'.drop_drop] - exact _root_.trans (FP.tail _) ih + induction n with + | zero => rfl + | succ n ih => + rw [Nat.add_comm, ← Stream'.drop_drop] + exact _root_.trans (FP.tail _) ih @[to_additive] theorem FP.singleton {M} [Semigroup M] (a : Stream' M) (i : ℕ) : a.get i ∈ FP a := by - induction' i with i ih generalizing a - · apply FP.head - · apply FP.tail - apply ih + induction i generalizing a with + | zero => exact FP.head _ + | succ i ih => exact FP.tail _ _ (ih _) @[to_additive] theorem FP.mul_two {M} [Semigroup M] (a : Stream' M) (i j : ℕ) (ij : i < j) : diff --git a/Mathlib/Computability/Language.lean b/Mathlib/Computability/Language.lean index e16b96d54f1eb..c3cfcd963d09b 100644 --- a/Mathlib/Computability/Language.lean +++ b/Mathlib/Computability/Language.lean @@ -261,10 +261,11 @@ instance : KleeneAlgebra (Language α) := mul_kstar_le_self := fun l m h ↦ by rw [kstar_eq_iSup_pow, mul_iSup] refine iSup_le (fun n ↦ ?_) - induction' n with n ih - · simp - rw [pow_succ, ← mul_assoc m (l^n) l] - exact le_trans (le_mul_congr ih le_rfl) h } + induction n with + | zero => simp + | succ n ih => + rw [pow_succ, ← mul_assoc m (l^n) l] + exact le_trans (le_mul_congr ih le_rfl) h } /-- Language `l.reverse` is defined as the set of words from `l` backwards. -/ def reverse (l : Language α) : Language α := { w : List α | w.reverse ∈ l } diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index 28f554c1f9089..55c72b43472b4 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -768,10 +768,11 @@ private theorem list_foldl' {f : α → List β} {g : α → σ} {h : α → σ dsimp only [F] generalize f a = l generalize g a = x - induction' n with n IH generalizing l x - · rfl - simp only [iterate_succ, comp_apply] - cases' l with b l <;> simp [IH] + induction n generalizing l x with + | zero => rfl + | succ n IH => + simp only [iterate_succ, comp_apply] + cases' l with b l <;> simp [IH] private theorem list_cons' : (haveI := prim H; Primrec₂ (@List.cons β)) := letI := prim H @@ -993,8 +994,9 @@ theorem nat_strong_rec (f : α → ℕ → σ) {g : α → List σ → Option σ option_map (hg.comp (fst.comp fst) snd) (to₂ <| list_concat.comp (snd.comp fst) snd))).of_eq fun a n => by - induction' n with n IH; · rfl - simp [IH, H, List.range_succ] + induction n with + | zero => rfl + | succ n IH => simp [IH, H, List.range_succ] theorem listLookup [DecidableEq α] : Primrec₂ (List.lookup : α → List (α × β) → Option β) := (to₂ <| list_rec snd (const none) <| diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 3ae1f1aa58149..b4e1c66e921db 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -359,12 +359,13 @@ theorem exists_code {n} {f : Vector ℕ n →. ℕ} (hf : Nat.Partrec' f) : PFun.mem_fix_iff.2 (Or.inl (by simp [hf, hn])) generalize (n.succ :: v.1 : List ℕ) = w at this ⊢ clear hn - induction' n with n IH - · exact this - refine IH (fun {m} h' => hm (Nat.lt_succ_of_lt h')) - (PFun.mem_fix_iff.2 (Or.inr ⟨_, ?_, this⟩)) - simp only [hf, hm n.lt_succ_self, Part.bind_some, List.headI, eq_self_iff_true, if_false, - Part.mem_some_iff, and_self_iff, List.tail_cons] + induction n with + | zero => exact this + | succ n IH => + refine IH (fun {m} h' => hm (Nat.lt_succ_of_lt h')) + (PFun.mem_fix_iff.2 (Or.inr ⟨_, ?_, this⟩)) + simp only [hf, hm n.lt_succ_self, Part.bind_some, List.headI, eq_self_iff_true, if_false, + Part.mem_some_iff, and_self_iff, List.tail_cons] end Code diff --git a/Mathlib/Data/Vector/Mem.lean b/Mathlib/Data/Vector/Mem.lean index 220b1cac938b6..3eb0fe6678d82 100644 --- a/Mathlib/Data/Vector/Mem.lean +++ b/Mathlib/Data/Vector/Mem.lean @@ -58,9 +58,9 @@ theorem mem_cons_of_mem (v : Vector α n) (ha' : a' ∈ v.toList) : a' ∈ (a :: (Vector.mem_cons_iff a a' v).2 (Or.inr ha') theorem mem_of_mem_tail (v : Vector α n) (ha : a ∈ v.tail.toList) : a ∈ v.toList := by - induction' n with n _ - · exact False.elim (Vector.not_mem_zero a v.tail ha) - · exact (mem_succ_iff a v).2 (Or.inr ha) + induction n with + | zero => exact False.elim (Vector.not_mem_zero a v.tail ha) + | succ n _ => exact (mem_succ_iff a v).2 (Or.inr ha) theorem mem_map_iff (b : β) (v : Vector α n) (f : α → β) : b ∈ (v.map f).toList ↔ ∃ a : α, a ∈ v.toList ∧ f a = b := by diff --git a/Mathlib/GroupTheory/GroupAction/Ring.lean b/Mathlib/GroupTheory/GroupAction/Ring.lean index b993a6401f328..1d669771f20c3 100644 --- a/Mathlib/GroupTheory/GroupAction/Ring.lean +++ b/Mathlib/GroupTheory/GroupAction/Ring.lean @@ -25,17 +25,17 @@ variable {α : Type*} instance NonUnitalNonAssocSemiring.nat_smulCommClass [NonUnitalNonAssocSemiring α] : SMulCommClass ℕ α α where smul_comm n x y := by - induction' n with n ih - · simp [zero_nsmul] - · simp_rw [succ_nsmul, smul_eq_mul, mul_add, ← smul_eq_mul, ih] + induction n with + | zero => simp [zero_nsmul] + | succ n ih => simp_rw [succ_nsmul, smul_eq_mul, mul_add, ← smul_eq_mul, ih] /-- Note that `AddCommMonoid.nat_isScalarTower` requires stronger assumptions on `α`. -/ instance NonUnitalNonAssocSemiring.nat_isScalarTower [NonUnitalNonAssocSemiring α] : IsScalarTower ℕ α α where smul_assoc n x y := by - induction' n with n ih - · simp [zero_nsmul] - · simp_rw [succ_nsmul, ← ih, smul_eq_mul, add_mul] + induction n with + | zero => simp [zero_nsmul] + | succ n ih => simp_rw [succ_nsmul, ← ih, smul_eq_mul, add_mul] /-- Note that `AddMonoid.int_smulCommClass` requires stronger assumptions on `α`. -/ instance NonUnitalNonAssocRing.int_smulCommClass [NonUnitalNonAssocRing α] : diff --git a/Mathlib/GroupTheory/Perm/Closure.lean b/Mathlib/GroupTheory/Perm/Closure.lean index 113024e8329e4..9ee578518cde5 100644 --- a/Mathlib/GroupTheory/Perm/Closure.lean +++ b/Mathlib/GroupTheory/Perm/Closure.lean @@ -47,17 +47,20 @@ theorem closure_cycle_adjacent_swap {σ : Perm α} (h1 : IsCycle σ) (h2 : σ.su have h4 : swap x (σ x) ∈ H := subset_closure (Set.mem_insert_of_mem _ (Set.mem_singleton _)) have step1 : ∀ n : ℕ, swap ((σ ^ n) x) ((σ ^ (n + 1) : Perm α) x) ∈ H := by intro n - induction' n with n ih - · exact subset_closure (Set.mem_insert_of_mem _ (Set.mem_singleton _)) - · convert H.mul_mem (H.mul_mem h3 ih) (H.inv_mem h3) + induction n with + | zero => exact subset_closure (Set.mem_insert_of_mem _ (Set.mem_singleton _)) + | succ n ih => + convert H.mul_mem (H.mul_mem h3 ih) (H.inv_mem h3) simp_rw [mul_swap_eq_swap_mul, mul_inv_cancel_right, pow_succ'] rfl have step2 : ∀ n : ℕ, swap x ((σ ^ n) x) ∈ H := by intro n - induction' n with n ih - · simp only [Nat.zero_eq, pow_zero, coe_one, id_eq, swap_self, Set.mem_singleton_iff] + induction n with + | zero => + simp only [Nat.zero_eq, pow_zero, coe_one, id_eq, swap_self, Set.mem_singleton_iff] convert H.one_mem - · by_cases h5 : x = (σ ^ n) x + | succ n ih => + by_cases h5 : x = (σ ^ n) x · rw [pow_succ', mul_apply, ← h5] exact h4 by_cases h6 : x = (σ ^ (n + 1) : Perm α) x diff --git a/Mathlib/GroupTheory/Perm/Fin.lean b/Mathlib/GroupTheory/Perm/Fin.lean index b5ba110fa6dae..7df1d1ceaefbe 100644 --- a/Mathlib/GroupTheory/Perm/Fin.lean +++ b/Mathlib/GroupTheory/Perm/Fin.lean @@ -92,9 +92,10 @@ theorem finRotate_succ_eq_decomposeFin {n : ℕ} : @[simp] theorem sign_finRotate (n : ℕ) : Perm.sign (finRotate (n + 1)) = (-1) ^ n := by - induction' n with n ih - · simp - · rw [finRotate_succ_eq_decomposeFin] + induction n with + | zero => simp + | succ n ih => + rw [finRotate_succ_eq_decomposeFin] simp [ih, pow_succ] @[simp] diff --git a/Mathlib/GroupTheory/Solvable.lean b/Mathlib/GroupTheory/Solvable.lean index 9082598ea518c..769f40a7f6ef5 100644 --- a/Mathlib/GroupTheory/Solvable.lean +++ b/Mathlib/GroupTheory/Solvable.lean @@ -66,17 +66,17 @@ variable (f) theorem map_derivedSeries_le_derivedSeries (n : ℕ) : (derivedSeries G n).map f ≤ derivedSeries G' n := by - induction' n with n ih - · exact le_top - · simp only [derivedSeries_succ, map_commutator, commutator_mono, ih] + induction n with + | zero => exact le_top + | succ n ih => simp only [derivedSeries_succ, map_commutator, commutator_mono, ih] variable {f} theorem derivedSeries_le_map_derivedSeries (hf : Function.Surjective f) (n : ℕ) : derivedSeries G' n ≤ (derivedSeries G n).map f := by - induction' n with n ih - · exact (map_top_of_surjective f hf).ge - · exact commutator_le_map_commutator ih ih + induction n with + | zero => exact (map_top_of_surjective f hf).ge + | succ n ih => exact commutator_le_map_commutator ih ih theorem map_derivedSeries_eq (hf : Function.Surjective f) (n : ℕ) : (derivedSeries G n).map f = derivedSeries G' n := @@ -152,12 +152,13 @@ section IsSimpleGroup variable [IsSimpleGroup G] theorem IsSimpleGroup.derivedSeries_succ {n : ℕ} : derivedSeries G n.succ = commutator G := by - induction' n with n ih - · exact derivedSeries_one G - rw [_root_.derivedSeries_succ, ih, _root_.commutator] - cases' (commutator_normal (⊤ : Subgroup G) (⊤ : Subgroup G)).eq_bot_or_eq_top with h h - · rw [h, commutator_bot_left] - · rwa [h] + induction n with + | zero => exact derivedSeries_one G + | succ n ih => + rw [_root_.derivedSeries_succ, ih, _root_.commutator] + cases' (commutator_normal (⊤ : Subgroup G) (⊤ : Subgroup G)).eq_bot_or_eq_top with h h + · rw [h, commutator_bot_left] + · rwa [h] theorem IsSimpleGroup.comm_iff_isSolvable : (∀ a b : G, a * b = b * a) ↔ IsSolvable G := ⟨isSolvable_of_comm, fun ⟨⟨n, hn⟩⟩ => by @@ -187,9 +188,10 @@ theorem Equiv.Perm.fin_5_not_solvable : ¬IsSolvable (Equiv.Perm (Fin 5)) := by let z : Equiv.Perm (Fin 5) := ⟨![0, 3, 2, 1, 4], ![0, 3, 2, 1, 4], by decide, by decide⟩ have key : x = z * ⁅x, y * x * y⁻¹⁆ * z⁻¹ := by unfold_let; decide refine not_solvable_of_mem_derivedSeries (show x ≠ 1 by decide) fun n => ?_ - induction' n with n ih - · exact mem_top x - · rw [key, (derivedSeries_normal _ _).mem_comm_iff, inv_mul_cancel_left] + induction n with + | zero => exact mem_top x + | succ n ih => + rw [key, (derivedSeries_normal _ _).mem_comm_iff, inv_mul_cancel_left] exact commutator_mem_commutator ih ((derivedSeries_normal _ _).conj_mem _ ih _) theorem Equiv.Perm.not_solvable (X : Type*) (hX : 5 ≤ Cardinal.mk X) : diff --git a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean index 74f30e36f4cda..2ffd70d8b1bb4 100644 --- a/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/ExteriorAlgebra/Basic.lean @@ -244,9 +244,10 @@ theorem ι_add_mul_swap (x y : M) : ι R x * ι R y + ι R y * ι R x = 0 := theorem ι_mul_prod_list {n : ℕ} (f : Fin n → M) (i : Fin n) : (ι R <| f i) * (List.ofFn fun i => ι R <| f i).prod = 0 := by - induction' n with n hn - · exact i.elim0 - · rw [List.ofFn_succ, List.prod_cons, ← mul_assoc] + induction n with + | zero => exact i.elim0 + | succ n hn => + rw [List.ofFn_succ, List.prod_cons, ← mul_assoc] by_cases h : i = 0 · rw [h, ι_sq_zero, zero_mul] · replace hn := diff --git a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean index f10afce554362..6bbe92a9b4b88 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Basic.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Basic.lean @@ -837,9 +837,9 @@ protected theorem map_mul (f₁ f₂ : M →ₗ[R] M) (g₁ g₂ : N →ₗ[R] N @[simp] protected theorem map_pow (f : M →ₗ[R] M) (g : N →ₗ[R] N) (n : ℕ) : map f g ^ n = map (f ^ n) (g ^ n) := by - induction' n with n ih - · simp only [pow_zero, TensorProduct.map_one] - · simp only [pow_succ', ih, TensorProduct.map_mul] + induction n with + | zero => simp only [pow_zero, TensorProduct.map_one] + | succ n ih => simp only [pow_succ', ih, TensorProduct.map_mul] theorem map_add_left (f₁ f₂ : M →ₗ[R] P) (g : N →ₗ[R] Q) : map (f₁ + f₂) g = map f₁ g + map f₂ g := by diff --git a/Mathlib/SetTheory/Surreal/Dyadic.lean b/Mathlib/SetTheory/Surreal/Dyadic.lean index 16ea03efd65ee..c3b61c82a8071 100644 --- a/Mathlib/SetTheory/Surreal/Dyadic.lean +++ b/Mathlib/SetTheory/Surreal/Dyadic.lean @@ -75,9 +75,10 @@ theorem birthday_half : birthday (powHalf 1) = 2 := by /-- For all natural numbers `n`, the pre-games `powHalf n` are numeric. -/ theorem numeric_powHalf (n) : (powHalf n).Numeric := by - induction' n with n hn - · exact numeric_one - · constructor + induction n with + | zero => exact numeric_one + | succ n hn => + constructor · simpa using hn.moveLeft_lt default · exact ⟨fun _ => numeric_zero, fun _ => hn⟩ @@ -88,9 +89,9 @@ theorem powHalf_succ_le_powHalf (n : ℕ) : powHalf (n + 1) ≤ powHalf n := (powHalf_succ_lt_powHalf n).le theorem powHalf_le_one (n : ℕ) : powHalf n ≤ 1 := by - induction' n with n hn - · exact le_rfl - · exact (powHalf_succ_le_powHalf n).trans hn + induction n with + | zero => exact le_rfl + | succ n hn => exact (powHalf_succ_le_powHalf n).trans hn theorem powHalf_succ_lt_one (n : ℕ) : powHalf (n + 1) < 1 := (powHalf_succ_lt_powHalf n).trans_le <| powHalf_le_one n diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index ca811c376a395..d075322b1f185 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -1315,10 +1315,9 @@ theorem intCast_apply [TopologicalAddGroup M] (z : ℤ) (m : M) : (↑z : M →L theorem smulRight_one_pow [TopologicalSpace R] [TopologicalRing R] (c : R) (n : ℕ) : smulRight (1 : R →L[R] R) c ^ n = smulRight (1 : R →L[R] R) (c ^ n) := by - induction' n with n ihn - · ext - simp - · rw [pow_succ, ihn, mul_def, smulRight_comp, smul_eq_mul, pow_succ'] + induction n with + | zero => ext; simp + | succ n ihn => rw [pow_succ, ihn, mul_def, smulRight_comp, smul_eq_mul, pow_succ'] section diff --git a/Mathlib/Topology/TietzeExtension.lean b/Mathlib/Topology/TietzeExtension.lean index 55ec301b782ca..704a18651fd7e 100644 --- a/Mathlib/Topology/TietzeExtension.lean +++ b/Mathlib/Topology/TietzeExtension.lean @@ -227,9 +227,10 @@ theorem exists_extension_norm_eq_of_closedEmbedding' (f : X →ᵇ ℝ) (e : C(X Function.iterate_succ_apply' _ _ _ have hgf : ∀ n, dist ((g n).compContinuous e) f ≤ (2 / 3) ^ n * ‖f‖ := by intro n - induction' n with n ihn - · simp [g0] - · rw [g_succ n, add_compContinuous, ← dist_sub_right, add_sub_cancel_left, pow_succ', mul_assoc] + induction n with + | zero => simp [g0] + | succ n ihn => + rw [g_succ n, add_compContinuous, ← dist_sub_right, add_sub_cancel_left, pow_succ', mul_assoc] refine (hF_dist _).trans (mul_le_mul_of_nonneg_left ?_ (by norm_num1)) rwa [← dist_eq_norm'] have hg_dist : ∀ n, dist (g n) (g (n + 1)) ≤ 1 / 3 * ‖f‖ * (2 / 3) ^ n := by diff --git a/Mathlib/Topology/UrysohnsLemma.lean b/Mathlib/Topology/UrysohnsLemma.lean index d6e4967be156a..3c801c768f6b1 100644 --- a/Mathlib/Topology/UrysohnsLemma.lean +++ b/Mathlib/Topology/UrysohnsLemma.lean @@ -149,30 +149,35 @@ noncomputable def approx : ℕ → CU P → X → ℝ | n + 1, c, x => midpoint ℝ (approx n c.left x) (approx n c.right x) theorem approx_of_mem_C (c : CU P) (n : ℕ) {x : X} (hx : x ∈ c.C) : c.approx n x = 0 := by - induction' n with n ihn generalizing c - · exact indicator_of_not_mem (fun (hU : x ∈ c.Uᶜ) => hU <| c.subset hx) _ - · simp only [approx] + induction n generalizing c with + | zero => exact indicator_of_not_mem (fun (hU : x ∈ c.Uᶜ) => hU <| c.subset hx) _ + | succ n ihn => + simp only [approx] rw [ihn, ihn, midpoint_self] exacts [c.subset_right_C hx, hx] theorem approx_of_nmem_U (c : CU P) (n : ℕ) {x : X} (hx : x ∉ c.U) : c.approx n x = 1 := by - induction' n with n ihn generalizing c - · rw [← mem_compl_iff] at hx + induction n generalizing c with + | zero => + rw [← mem_compl_iff] at hx exact indicator_of_mem hx _ - · simp only [approx] + | succ n ihn => + simp only [approx] rw [ihn, ihn, midpoint_self] exacts [hx, fun hU => hx <| c.left_U_subset hU] theorem approx_nonneg (c : CU P) (n : ℕ) (x : X) : 0 ≤ c.approx n x := by - induction' n with n ihn generalizing c - · exact indicator_nonneg (fun _ _ => zero_le_one) _ - · simp only [approx, midpoint_eq_smul_add, invOf_eq_inv] + induction n generalizing c with + | zero => exact indicator_nonneg (fun _ _ => zero_le_one) _ + | succ n ihn => + simp only [approx, midpoint_eq_smul_add, invOf_eq_inv] refine mul_nonneg (inv_nonneg.2 zero_le_two) (add_nonneg ?_ ?_) <;> apply ihn theorem approx_le_one (c : CU P) (n : ℕ) (x : X) : c.approx n x ≤ 1 := by - induction' n with n ihn generalizing c - · exact indicator_apply_le' (fun _ => le_rfl) fun _ => zero_le_one - · simp only [approx, midpoint_eq_smul_add, invOf_eq_inv, smul_eq_mul, ← div_eq_inv_mul] + induction n generalizing c with + | zero => exact indicator_apply_le' (fun _ => le_rfl) fun _ => zero_le_one + | succ n ihn => + simp only [approx, midpoint_eq_smul_add, invOf_eq_inv, smul_eq_mul, ← div_eq_inv_mul] have := add_le_add (ihn (left c)) (ihn (right c)) norm_num at this exact Iff.mpr (div_le_one zero_lt_two) this From 508aeb0a850841872c7a54d33b2b629ed87542ac Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Wed, 14 Aug 2024 18:28:24 +0000 Subject: [PATCH 274/359] chore: update Mathlib dependencies 2024-08-14 (#15803) 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 a136169bb5368..fdeb2fb3a7531 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "1d25ec7ec98d6d9fb526c997aa014bcabbad8b72", + "rev": "5e5e54c4028f7b6bd086ebb72e5822794c64c35d", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From e6c08721e554cf434f18df148791cb10e70cc0b1 Mon Sep 17 00:00:00 2001 From: Ralf Stephan Date: Wed, 14 Aug 2024 19:25:54 +0000 Subject: [PATCH 275/359] doc(Tactic): expand `ring` and `ring_nf` docstrings (#15691) Copy some explanations from other comments into the actual `ring_nf` docstring. Add details to `ring` docstring. --- Mathlib/Tactic/Ring/RingNF.lean | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/Ring/RingNF.lean b/Mathlib/Tactic/Ring/RingNF.lean index 04fc8327befe7..1b88fd1b13a71 100644 --- a/Mathlib/Tactic/Ring/RingNF.lean +++ b/Mathlib/Tactic/Ring/RingNF.lean @@ -196,6 +196,11 @@ which rewrites all ring expressions into a normal form. * `recursive`: if true, `ring_nf` will also recurse into atoms * `ring_nf` works as both a tactic and a conv tactic. In tactic mode, `ring_nf at h` can be used to rewrite in a hypothesis. + +This can be used non-terminally to normalize ring expressions in the goal such as +`⊢ P (x + x + x)` ~> `⊢ P (x * 3)`, as well as being able to prove some equations that +`ring` cannot because they involve ring reasoning inside a subterm, such as +`sin (x + y) + sin (y + x) = 2 * sin (x + y)`. -/ elab (name := ringNF) "ring_nf" tk:"!"? cfg:(config ?) loc:(location)? : tactic => do let mut cfg ← elabConfig cfg @@ -238,7 +243,8 @@ elab (name := ring1NF) "ring1_nf" tk:"!"? cfg:(config ?) : tactic => do /-- Tactic for evaluating expressions in *commutative* (semi)rings, allowing for variables in the -exponent. +exponent. If the goal is not appropriate for `ring` (e.g. not an equality) `ring_nf` will be +suggested. * `ring!` will use a more aggressive reducibility setting to determine equality of atoms. * `ring1` fails if the target is not an equality. @@ -248,6 +254,7 @@ For example: example (n : ℕ) (m : ℤ) : 2^(n+1) * m = 2 * 2^n * m := by ring example (a b : ℤ) (n : ℕ) : (a + b)^(n + 2) = (a^2 + b^2 + a * b + b * a) * (a + b)^n := by ring example (x y : ℕ) : x + id y = y + id x := by ring! +example (x : ℕ) (h : x * 2 > 5): x + x > 5 := by ring; assumption -- suggests ring_nf ``` -/ macro (name := ring) "ring" : tactic => From 7a75f1dfdbf1e153e9f09695117f3d67fbac423d Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Wed, 14 Aug 2024 20:11:45 +0000 Subject: [PATCH 276/359] chore: revert #15782 (#15816) This breaks the speedcenter. Let's pull it back until we can get it figured out. --- Mathlib/Algebra/BigOperators/Group/List.lean | 4 +++- Mathlib/Algebra/Field/Defs.lean | 5 +++-- Mathlib/Data/List/Chain.lean | 3 ++- Mathlib/Data/List/GetD.lean | 4 +++- 4 files changed, 11 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index b283972626cff..0251d89f06740 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -21,7 +21,9 @@ of elements of a list and `List.alternatingProd`, `List.alternatingSum`, their a counterparts. -/ -assert_not_imported Mathlib.Algebra.Order.Group.Nat +-- Make sure we haven't imported `Data.Nat.Order.Basic` +assert_not_exists OrderedSub +assert_not_exists Ring variable {ι α β M N P G : Type*} diff --git a/Mathlib/Algebra/Field/Defs.lean b/Mathlib/Algebra/Field/Defs.lean index 4943b798adf4c..c29f6e33d0c67 100644 --- a/Mathlib/Algebra/Field/Defs.lean +++ b/Mathlib/Algebra/Field/Defs.lean @@ -43,11 +43,12 @@ a `GroupWithZero` lemma instead. field, division ring, skew field, skew-field, skewfield -/ -assert_not_imported Mathlib.Tactic.Common - -- `NeZero` should not be needed in the basic algebraic hierarchy. assert_not_exists NeZero +-- Check that we have not imported `Mathlib.Tactic.Common` yet. +assert_not_exists Mathlib.Tactic.scopedNS + assert_not_exists MonoidHom open Function Set diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index a6c2410414b7d..e8e436dc37099 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -18,7 +18,8 @@ A graph-specialized version is in development and will hopefully be added under sometime soon. -/ -assert_not_imported Mathlib.Algebra.Order.Group.Nat +-- Make sure we haven't imported `Data.Nat.Order.Basic` +assert_not_exists OrderedSub universe u v diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index 8e35897c2dd50..9b888110c53b9 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -14,9 +14,11 @@ import Mathlib.Util.AssertExists This file provides theorems for working with the `getD` and `getI` functions. These are used to access an element of a list by numerical index, with a default value as a fallback when the index is out of range. + -/ -assert_not_imported Mathlib.Algebra.Order.Group.Nat +-- Make sure we haven't imported `Data.Nat.Order.Basic` +assert_not_exists OrderedSub namespace List From 0458ef860b78346c2bde5f2a56e45d35af21fda0 Mon Sep 17 00:00:00 2001 From: Nicky Mouha Date: Wed, 14 Aug 2024 22:14:17 +0000 Subject: [PATCH 277/359] chore(Data/Nat/Factorial): replace all `induction'` by `induction` (#15757) --- Mathlib/Data/Nat/Factorial/Basic.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/Nat/Factorial/Basic.lean b/Mathlib/Data/Nat/Factorial/Basic.lean index a72bcc344f845..49a834165d323 100644 --- a/Mathlib/Data/Nat/Factorial/Basic.lean +++ b/Mathlib/Data/Nat/Factorial/Basic.lean @@ -61,9 +61,9 @@ theorem factorial_ne_zero (n : ℕ) : n ! ≠ 0 := ne_of_gt (factorial_pos _) theorem factorial_dvd_factorial {m n} (h : m ≤ n) : m ! ∣ n ! := by - induction' h with n _ ih - · exact Nat.dvd_refl _ - · exact Nat.dvd_trans ih (Nat.dvd_mul_left _ _) + induction h with + | refl => exact Nat.dvd_refl _ + | step _ ih => exact Nat.dvd_trans ih (Nat.dvd_mul_left _ _) theorem dvd_factorial : ∀ {m n}, 0 < m → m ≤ n → m ∣ n ! | succ _, _, _, h => Nat.dvd_trans (Nat.dvd_mul_right _ _) (factorial_dvd_factorial h) @@ -84,9 +84,9 @@ theorem factorial_lt (hn : 0 < n) : n ! < m ! ↔ n < m := by intro k hk rw [factorial_succ, succ_mul, Nat.lt_add_left_iff_pos] exact Nat.mul_pos hk k.factorial_pos - induction' h with k hnk ih generalizing hn - · exact this hn - · exact lt_trans (ih hn) $ this <| lt_trans hn <| lt_of_succ_le hnk + induction h generalizing hn with + | refl => exact this hn + | step hnk ih => exact lt_trans (ih hn) $ this <| lt_trans hn <| lt_of_succ_le hnk @[gcongr] lemma factorial_lt_of_lt {m n : ℕ} (hn : 0 < n) (h : n < m) : n ! < m ! := (factorial_lt hn).mpr h From 08ecf4c7e62daf335fa2f9e42d6ad90574211eb5 Mon Sep 17 00:00:00 2001 From: ndcroos Date: Thu, 15 Aug 2024 00:43:09 +0000 Subject: [PATCH 278/359] chore(Algebra/Order/Floor): relax LinearOrderedSemiring to OrderedSemiring (#14911) --- Mathlib/Algebra/Order/Floor.lean | 28 +++++++++++++--------------- 1 file changed, 13 insertions(+), 15 deletions(-) diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 517130d8d4219..d12b1da461bd2 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -54,7 +54,6 @@ many lemmas. rounding, floor, ceil -/ - open Set variable {F α β : Type*} @@ -116,18 +115,25 @@ notation "⌊" a "⌋₊" => Nat.floor a @[inherit_doc] notation "⌈" a "⌉₊" => Nat.ceil a -end OrderedSemiring - -section LinearOrderedSemiring - -variable [LinearOrderedSemiring α] [FloorSemiring α] {a : α} {n : ℕ} - theorem le_floor_iff (ha : 0 ≤ a) : n ≤ ⌊a⌋₊ ↔ (n : α) ≤ a := FloorSemiring.gc_floor ha theorem le_floor (h : (n : α) ≤ a) : n ≤ ⌊a⌋₊ := (le_floor_iff <| n.cast_nonneg.trans h).2 h +theorem gc_ceil_coe : GaloisConnection (ceil : α → ℕ) (↑) := + FloorSemiring.gc_ceil + +@[simp] +theorem ceil_le : ⌈a⌉₊ ≤ n ↔ a ≤ n := + gc_ceil_coe _ _ + +end OrderedSemiring + +section LinearOrderedSemiring + +variable [LinearOrderedSemiring α] [FloorSemiring α] {a : α} {n : ℕ} + theorem floor_lt (ha : 0 ≤ a) : ⌊a⌋₊ < n ↔ a < n := lt_iff_lt_of_le_iff_le <| le_floor_iff ha @@ -242,14 +248,6 @@ theorem preimage_floor_of_ne_zero {n : ℕ} (hn : n ≠ 0) : /-! #### Ceil -/ - -theorem gc_ceil_coe : GaloisConnection (ceil : α → ℕ) (↑) := - FloorSemiring.gc_ceil - -@[simp] -theorem ceil_le : ⌈a⌉₊ ≤ n ↔ a ≤ n := - gc_ceil_coe _ _ - theorem lt_ceil : n < ⌈a⌉₊ ↔ (n : α) < a := lt_iff_lt_of_le_iff_le ceil_le From 9c8aa0953796d9b2add55fc9cda35a56f6a285d2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 15 Aug 2024 07:49:31 +0000 Subject: [PATCH 279/359] chore: remove a porting note and simplify (#15819) The old proof seems to work again. Also tidies some related proofs. --- Mathlib/Data/Set/Image.lean | 14 +------------- Mathlib/Logic/Equiv/Set.lean | 6 ++---- 2 files changed, 3 insertions(+), 17 deletions(-) diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index c0e717e96f023..09886a1171c8b 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -938,19 +938,7 @@ theorem range_diff_image {f : α → β} (H : Injective f) (s : Set α) : range @[simp] theorem range_inclusion (h : s ⊆ t) : range (inclusion h) = { x : t | (x : α) ∈ s } := by ext ⟨x, hx⟩ - -- Porting note: `simp [inclusion]` doesn't solve goal - apply Iff.intro - · rw [mem_range] - rintro ⟨a, ha⟩ - rw [inclusion, Subtype.mk.injEq] at ha - rw [mem_setOf, Subtype.coe_mk, ← ha] - exact Subtype.coe_prop _ - · rw [mem_setOf, Subtype.coe_mk, mem_range] - intro hx' - use ⟨x, hx'⟩ - trivial - -- simp_rw [inclusion, mem_range, Subtype.mk_eq_mk] - -- rw [SetCoe.exists, Subtype.coe_mk, exists_prop, exists_eq_right, mem_set_of, Subtype.coe_mk] + simp -- When `f` is injective, see also `Equiv.ofInjective`. theorem leftInverse_rangeSplitting (f : α → β) : diff --git a/Mathlib/Logic/Equiv/Set.lean b/Mathlib/Logic/Equiv/Set.lean index 9976bc2e53085..65af5ec56aaec 100644 --- a/Mathlib/Logic/Equiv/Set.lean +++ b/Mathlib/Logic/Equiv/Set.lean @@ -328,14 +328,12 @@ theorem sumDiffSubset_apply_inr {α} {s t : Set α} (h : s ⊆ t) [DecidablePred theorem sumDiffSubset_symm_apply_of_mem {α} {s t : Set α} (h : s ⊆ t) [DecidablePred (· ∈ s)] {x : t} (hx : x.1 ∈ s) : (Equiv.Set.sumDiffSubset h).symm x = Sum.inl ⟨x, hx⟩ := by apply (Equiv.Set.sumDiffSubset h).injective - simp only [apply_symm_apply, sumDiffSubset_apply_inl] - exact Subtype.eq rfl + simp only [apply_symm_apply, sumDiffSubset_apply_inl, Set.inclusion_mk] theorem sumDiffSubset_symm_apply_of_not_mem {α} {s t : Set α} (h : s ⊆ t) [DecidablePred (· ∈ s)] {x : t} (hx : x.1 ∉ s) : (Equiv.Set.sumDiffSubset h).symm x = Sum.inr ⟨x, ⟨x.2, hx⟩⟩ := by apply (Equiv.Set.sumDiffSubset h).injective - simp only [apply_symm_apply, sumDiffSubset_apply_inr] - exact Subtype.eq rfl + simp only [apply_symm_apply, sumDiffSubset_apply_inr, Set.inclusion_mk] /-- If `s` is a set with decidable membership, then the sum of `s ∪ t` and `s ∩ t` is equivalent to `s ⊕ t`. -/ From e9deae49fe08d020b3f26334994fe3fcd571ce8b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 15 Aug 2024 08:47:50 +0000 Subject: [PATCH 280/359] chore(SetTheory/Ordinal/FixedPoint): Make sure simp lemma has a reasonable key (#15721) This is the second most called `simp` lemma and never succeeds (see [Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Infrastructure.20for.20tracking.20frequently.20applied.20simp.20theorems/near/459926416)]. It's also way too specific to be reasonably useful outside of the file itself. Also did some pretty minor cleanup. --- Mathlib/SetTheory/Ordinal/FixedPoint.lean | 30 +++++++++-------------- 1 file changed, 12 insertions(+), 18 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index aa8d8d9c57d1b..9c3cc3e2c26af 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -368,10 +368,9 @@ def nfp (f : Ordinal → Ordinal) : Ordinal → Ordinal := theorem nfp_eq_nfpFamily (f : Ordinal → Ordinal) : nfp f = nfpFamily fun _ : Unit => f := rfl -@[simp] -theorem sup_iterate_eq_nfp (f : Ordinal.{u} → Ordinal.{u}) : - (fun a => sup fun n : ℕ => f^[n] a) = nfp f := by - refine funext fun a => le_antisymm ?_ (sup_le fun l => ?_) +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 => ?_) · rw [sup_le_iff] intro n rw [← List.length_replicate n Unit.unit, ← List.foldr_const f a] @@ -477,18 +476,18 @@ theorem deriv_eq_enumOrd (H : IsNormal f) : deriv f = enumOrd (Function.fixedPoi theorem deriv_eq_id_of_nfp_eq_id {f : Ordinal → Ordinal} (h : nfp f = id) : deriv f = id := (IsNormal.eq_iff_zero_and_succ (deriv_isNormal _) IsNormal.refl).2 <| by simp [h] -@[simp] -theorem nfp_zero : nfp 0 = id := by +theorem nfp_zero_left (a) : nfp 0 a = a := by rw [← sup_iterate_eq_nfp] - refine funext fun a => (sup_le fun n => ?_).antisymm (le_sup (fun n => 0^[n] a) 0) + apply (sup_le fun n => ?_).antisymm (le_sup (fun n => 0^[n] a) 0) induction' n with n _ · rfl - rw [Function.iterate_succ'] - exact Ordinal.zero_le a + · rw [Function.iterate_succ'] + exact Ordinal.zero_le a -theorem nfp_zero_left (a) : nfp 0 a = a := by - rw [nfp_zero] - rfl +@[simp] +theorem nfp_zero : nfp 0 = id := by + ext + exact nfp_zero_left _ @[simp] theorem deriv_zero : deriv 0 = id := @@ -502,7 +501,6 @@ end /-! ### Fixed points of addition -/ - @[simp] theorem nfp_add_zero (a) : nfp (a + ·) 0 = a * omega := by simp_rw [← sup_iterate_eq_nfp, ← sup_mul_nat] @@ -542,14 +540,10 @@ theorem deriv_add_eq_mul_omega_add (a b : Ordinal.{u}) : deriv (a + ·) b = a * /-! ### Fixed points of multiplication -/ --- Porting note: commented out, doesn't seem necessary --- local infixr:0 "^" => @Pow.pow Ordinal Ordinal Ordinal.hasPow - @[simp] theorem nfp_mul_one {a : Ordinal} (ha : 0 < a) : nfp (a * ·) 1 = (a^omega) := by rw [← sup_iterate_eq_nfp, ← sup_opow_nat] - · dsimp - congr + · congr funext n induction' n with n hn · rw [Nat.cast_zero, opow_zero, iterate_zero_apply] From 58d407baf472d2089c93202a3481e99dddf916b4 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 15 Aug 2024 08:47:51 +0000 Subject: [PATCH 281/359] feat: add the `WithCStarModule` type synonym (#15763) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit It is often the case that we want to construct a `CStarModule` instance on a type that is already endowed with a norm, and it has a different norm which makes it into a `CStarModule`. For this reason, we create a type synonym `WithCStarModule` which is endowed with the requisite `CStarModule` instance. We also introduce the scoped notation `C⋆ᵐᵒᵈ` for this type synonym. The common use cases are, when `A` is a C⋆-algebra: + `A` itself over `A` + `E × F` where `E` and `F` are `CStarModule`s over `A` + `Π i, E i` where `E i` is a `CStarModule` over `A` and `i : ι` with `ι` a `Fintype` + `E` where `E` is an `InnerProductSpace` over `ℂ` The `WithCStarModule` synonym is of vital importance, especially because the `CStarModule` class marks `A` as an `outParam`. Indeed, we want to infer `A` from the type of `E`, but, as with modules, a type `E` can be a `CStarModule` over different C⋆-algebras. For example, note that if `A` is a C⋆-algebra, then so is `A × A`, and therefore we may consider both `A` and `A × A` as `CStarModule`s over themselves, respectively. However, we may *also* consider `A × A` as a `CStarModule` over `A`. However, by utilizing the type synonym, these actually correspond to *different types*, namely: + `A` as a `CStarModule` over `A` corresponds to `C⋆ᵐᵒᵈ A` + `A × A` as a `CStarModule` over `A × A` corresponds to `C⋆ᵐᵒᵈ (A × A)` + `A × A` as a `CStarModule` over `A` corresponds to `C⋆ᵐᵒᵈ (C⋆ᵐᵒᵈ A × C⋆ᵐᵒᵈ A)` --- Mathlib.lean | 1 + .../Analysis/CStarAlgebra/Module/Defs.lean | 41 ++- .../Analysis/CStarAlgebra/Module/Synonym.lean | 320 ++++++++++++++++++ 3 files changed, 351 insertions(+), 11 deletions(-) create mode 100644 Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean diff --git a/Mathlib.lean b/Mathlib.lean index 279e430d5c489..18395bdf1b054 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -921,6 +921,7 @@ import Mathlib.Analysis.CStarAlgebra.Exponential import Mathlib.Analysis.CStarAlgebra.GelfandDuality import Mathlib.Analysis.CStarAlgebra.Matrix import Mathlib.Analysis.CStarAlgebra.Module.Defs +import Mathlib.Analysis.CStarAlgebra.Module.Synonym import Mathlib.Analysis.CStarAlgebra.Multiplier import Mathlib.Analysis.CStarAlgebra.Spectrum import Mathlib.Analysis.CStarAlgebra.Unitization diff --git a/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean b/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean index 80dc5201159e3..daf2909a3ed1a 100644 --- a/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean +++ b/Mathlib/Analysis/CStarAlgebra/Module/Defs.lean @@ -10,9 +10,8 @@ import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Order # Hilbert C⋆-modules A Hilbert C⋆-module is a complex module `E` together with a right `A`-module structure, where `A` -is a C⋆-algebra, and with an "inner product" that takes values in `A`. This inner -product satisfies the Cauchy-Schwarz inequality, and induces a norm that makes `E` a normed -vector space. +is a C⋆-algebra, and with an `A`-valued inner product. This inner product satisfies the +Cauchy-Schwarz inequality, and induces a norm that makes `E` a normed vector space over `ℂ`. ## Main declarations @@ -36,6 +35,20 @@ and which would send the type class search algorithm on a chase for `A`), we pro `NormedSpace.Core` structure which enables downstream users of the class to easily register these instances themselves on a particular type. +Although the `Norm` is passed as a parameter, it almost never coincides with the norm on the +underlying type, unless that it is a purpose built type, as with the *standard Hilbert C⋆-module*. +However, with generic types already equipped with a norm, the norm as a Hilbert C⋆-module almost +never coincides with the norm on the underlying type. The two notable exceptions to this are when +we view `A` as a C⋆-module over itself, or when `A := ℂ`. For this reason we will later use the +type synonym `WithCStarModule`. + +As an example of just how different the norm can be, consider `CStarModule`s `E` and `F` over `A`. +One would like to put a `CStarModule` structure on (a type synonym of) `E × F`, where the `A`-valued +inner product is given, for `x y : E × F`, `⟪x, y⟫_A := ⟪x.1, y.1⟫_A + ⟪x.2, y.2⟫_A`. The norm this +induces satisfies `‖x‖ ^ 2 = ‖⟪x.1, y.1⟫ + ⟪x.2, y.2⟫‖`, but this doesn't coincide with *any* +natural norm on `E × F` unless `A := ℂ`, in which case it is `WithLp 2 (E × F)` because `E × F` is +then an `InnerProductSpace` over `ℂ`. + ## References + Erin Wittlich. *Formalizing Hilbert Modules in C⋆-algebras with the Lean Proof Assistant*, @@ -145,13 +158,12 @@ variable {A E : Type*} [NonUnitalNormedRing A] [StarRing A] [PartialOrder A] local notation "⟪" x ", " y "⟫" => inner (𝕜 := A) x y -variable (A) in /-- The norm associated with a Hilbert C⋆-module. It is not registered as a norm, since a type might already have a norm defined on it. -/ -noncomputable def norm : Norm E where - norm x := Real.sqrt ‖⟪x, x⟫‖ +noncomputable def norm (A : Type*) {E : Type*} [Norm A] [Inner A E] : Norm E where + norm x := Real.sqrt ‖⟪x, x⟫_A‖ -lemma inner_self_eq_norm_sq {x : E} : ‖⟪x, x⟫‖ = ‖x‖ ^ 2 := by simp [norm_eq_sqrt_norm_inner_self] +lemma norm_sq_eq {x : E} : ‖x‖ ^ 2 = ‖⟪x, x⟫‖ := by simp [norm_eq_sqrt_norm_inner_self] section include A @@ -170,12 +182,10 @@ protected lemma norm_zero : ‖(0 : E)‖ = 0 := by simp [norm_eq_sqrt_norm_inne lemma norm_zero_iff (x : E) : ‖x‖ = 0 ↔ x = 0 := ⟨fun h => by simpa [norm_eq_sqrt_norm_inner_self, inner_self] using h, - fun h => by simp [norm, h]; rw [CStarModule.norm_zero] ⟩ + fun h => by simp [norm, h, norm_eq_sqrt_norm_inner_self]⟩ end -lemma norm_sq_eq {x : E} : ‖x‖ ^ 2 = ‖⟪x, x⟫‖ := by simp [norm_eq_sqrt_norm_inner_self] - variable [CStarRing A] [StarOrderedRing A] [StarModule ℂ A] [IsScalarTower ℂ A A] [SMulCommClass ℂ A A] @@ -245,13 +255,18 @@ lemma normedSpaceCore [CompleteSpace A] : NormedSpace.Core ℂ E where norm_smul c x := by simp [norm_eq_sqrt_norm_inner_self, norm_smul, ← mul_assoc] norm_triangle x y := CStarModule.norm_triangle x y +/-- This is not listed as an instance because we often want to replace the topology, uniformity +and bornology instead of inheriting them from the norm. -/ +abbrev normedAddCommGroup [CompleteSpace A] : NormedAddCommGroup E := + NormedAddCommGroup.ofCore CStarModule.normedSpaceCore + lemma norm_eq_csSup [CompleteSpace A] (v : E) : ‖v‖ = sSup { ‖⟪w, v⟫_A‖ | (w : E) (_ : ‖w‖ ≤ 1) } := by let instNACG : NormedAddCommGroup E := NormedAddCommGroup.ofCore normedSpaceCore let instNS : NormedSpace ℂ E := .ofCore normedSpaceCore refine Eq.symm <| IsGreatest.csSup_eq ⟨⟨‖v‖⁻¹ • v, ?_, ?_⟩, ?_⟩ · simpa only [norm_smul, norm_inv, norm_norm] using inv_mul_le_one_of_le le_rfl (by positivity) - · simp [norm_smul, CStarModule.inner_self_eq_norm_sq, pow_two, ← mul_assoc] + · simp [norm_smul, ← norm_sq_eq, pow_two, ← mul_assoc] · rintro - ⟨w, hw, rfl⟩ calc _ ≤ ‖w‖ * ‖v‖ := norm_inner_le E _ ≤ 1 * ‖v‖ := by gcongr @@ -261,6 +276,10 @@ end norm section NormedAddCommGroup +/- Note: one generally creates a `CStarModule` instance for a type `E` first before getting the +`NormedAddCommGroup` and `NormedSpace` instances via `CStarModule.normedSpaceCore`, especially by +using `NormedAddCommGroup.ofCoreReplaceAll` and `NormedSpace.ofCore`. See +`Analysis.CStarAlgebra.Module.Constructions` for examples. -/ variable {A E : Type*} [NonUnitalNormedRing A] [StarRing A] [CStarRing A] [PartialOrder A] [StarOrderedRing A] [NormedSpace ℂ A] [SMul Aᵐᵒᵖ E] [CompleteSpace A] [NormedAddCommGroup E] [NormedSpace ℂ E] [StarModule ℂ A] [CStarModule A E] [IsScalarTower ℂ A A] diff --git a/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean b/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean new file mode 100644 index 0000000000000..e98ec70042e6b --- /dev/null +++ b/Mathlib/Analysis/CStarAlgebra/Module/Synonym.lean @@ -0,0 +1,320 @@ +/- +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.RingTheory.Finiteness +import Mathlib.Topology.Bornology.Constructions +import Mathlib.Topology.UniformSpace.Equiv + +/-! # Type synonym for types with a `CStarModule` structure + +It is often the case that we want to construct a `CStarModule` instance on a type that is already +endowed with a norm, but this norm is not the one associated to its `CStarModule` structure. For +this reason, we create a type synonym `WithCStarModule` which is endowed with the requisite +`CStarModule` instance. We also introduce the scoped notation `C⋆ᵐᵒᵈ` for this type synonym. + +The common use cases are, when `A` is a C⋆-algebra: + ++ `A` itself over `A` ++ `E × F` where `E` and `F` are `CStarModule`s over `A` ++ `Π i, E i` where `E i` is a `CStarModule` over `A` and `i : ι` with `ι` a `Fintype` ++ `E` where `E` is an `InnerProductSpace` over `ℂ` + +In this way, the set up is very similar to the `WithLp` type synonym, although there is no way to +reuse `WithLp` because the norms *do not* coincide in general. + +The `WithCStarModule` synonym is of vital importance, especially because the `CStarModule` class +marks `A` as an `outParam`. Indeed, we want to infer `A` from the type of `E`, but, as with modules, +a type `E` can be a `CStarModule` over different C⋆-algebras. For example, note that if `A` is a +C⋆-algebra, then so is `A × A`, and therefore we may consider both `A` and `A × A` as `CStarModule`s +over themselves, respectively. However, we may *also* consider `A × A` as a `CStarModule` over `A`. +However, by utilizing the type synonym, these actually correspond to *different types*, namely: + ++ `A` as a `CStarModule` over `A` corresponds to `C⋆ᵐᵒᵈ A` ++ `A × A` as a `CStarModule` over `A × A` corresponds to `C⋆ᵐᵒᵈ (A × A)` ++ `A × A` as a `CStarModule` over `A` corresponds to `C⋆ᵐᵒᵈ (C⋆ᵐᵒᵈ A × C⋆ᵐᵒᵈ A)` + +## Main definitions + +* `WithCStarModule E`: a copy of `E` to be equipped with a `CStarModule A` structure. Note that + `A` is an `outParam` in the class `CStarModule`, so it can be inferred from the type of `E`. +* `WithCStarModule.equiv E`: the canonical equivalence between `WithCStarModule E` and `E`. +* `WithCStarModule.linearEquiv ℂ A E`: the canonical `ℂ`-module isomorphism between + `WithCStarModule E` and `E`. + +## Implementation notes + +The pattern here is the same one as is used by `Lex` for order structures; it avoids having a +separate synonym for each type, and allows all the structure-copying code to be shared. +-/ + +/-- A type synonym for endowing a given type with a `CStarModule` structure. This has the scoped +notation `C⋆ᵐᵒᵈ`. + +Note: because the C⋆-algebra `A` over which `E` is a `CStarModule` is listed as an `outParam` in +that class, we don't pass it as an unused argument to `WithCStarModule`, unlike the `p` parameter +in `WithLp`, which can vary. -/ +def WithCStarModule (E : Type*) := E + +namespace WithCStarModule + +@[inherit_doc] +scoped notation "C⋆ᵐᵒᵈ" => WithCStarModule + +section Basic + +variable (R R' E : Type*) + +/-- The canonical equivalence between `WithCStarModule E` and `E`. This should always be used to +convert back and forth between the representations. -/ +def equiv : WithCStarModule E ≃ E := Equiv.refl _ + +instance instNontrivial [Nontrivial E] : Nontrivial (WithCStarModule E) := ‹Nontrivial E› +instance instInhabited [Inhabited E] : Inhabited (WithCStarModule E) := ‹Inhabited E› +instance instNonempty [Nonempty E] : Nonempty (WithCStarModule E) := ‹Nonempty E› +instance instUnique [Unique E] : Unique (WithCStarModule E) := ‹Unique E› + +/-! ## `WithCStarModule E` inherits various module-adjacent structures from `E`. -/ + +instance instAddCommGroup [AddCommGroup E] : AddCommGroup (WithCStarModule E) := ‹AddCommGroup E› +instance instSMul {R : Type*} [SMul R E] : SMul R (WithCStarModule E) := ‹SMul R E› +instance instModule {R : Type*} [Semiring R] [AddCommGroup E] [Module R E] : + Module R (WithCStarModule E) := + ‹Module R E› + +instance instIsScalarTower [SMul R R'] [SMul R E] [SMul R' E] + [IsScalarTower R R' E] : IsScalarTower R R' (WithCStarModule E) := + ‹IsScalarTower R R' E› + +instance instSMulCommClass [SMul R E] [SMul R' E] [SMulCommClass R R' E] : + SMulCommClass R R' (WithCStarModule E) := + ‹SMulCommClass R R' E› + +instance instModuleFinite [Semiring R] [AddCommGroup E] [Module R E] [Module.Finite R E] : + Module.Finite R (WithCStarModule E) := + ‹Module.Finite R E› + + +section Equiv + +variable {R E} +variable [SMul R E] (c : R) (x y : WithCStarModule E) (x' y' : E) + +/-! `WithCStarModule.equiv` preserves the module structure. -/ + +section AddCommGroup + +variable [AddCommGroup E] + +@[simp] +theorem equiv_zero : equiv E 0 = 0 := + rfl + +@[simp] +theorem equiv_symm_zero : (equiv E).symm 0 = 0 := + rfl + +@[simp] +theorem equiv_add : equiv E (x + y) = equiv E x + equiv E y := + rfl + +@[simp] +theorem equiv_symm_add : + (equiv E).symm (x' + y') = (equiv E).symm x' + (equiv E).symm y' := + rfl + +@[simp] +theorem equiv_sub : equiv E (x - y) = equiv E x - equiv E y := + rfl + +@[simp] +theorem equiv_symm_sub : + (equiv E).symm (x' - y') = (equiv E).symm x' - (equiv E).symm y' := + rfl + +@[simp] +theorem equiv_neg : equiv E (-x) = -equiv E x := + rfl + +@[simp] +theorem equiv_symm_neg : (equiv E).symm (-x') = -(equiv E).symm x' := + rfl + +end AddCommGroup + +@[simp] +theorem equiv_smul : equiv E (c • x) = c • equiv E x := + rfl + +@[simp] +theorem equiv_symm_smul : (equiv E).symm (c • x') = c • (equiv E).symm x' := + rfl + +end Equiv + +/-- `WithCStarModule.equiv` as a linear equivalence. -/ +@[simps (config := .asFn)] +def linearEquiv [Semiring R] [AddCommGroup E] [Module R E] : C⋆ᵐᵒᵈ E ≃ₗ[R] E := + { LinearEquiv.refl _ _ with + toFun := equiv _ + invFun := (equiv _).symm } + +/-! ## `WithCStarModule E` inherits the uniformity and bornology from `E`. -/ + +variable {E} + +instance [u : UniformSpace E] : UniformSpace (C⋆ᵐᵒᵈ E) := u.comap <| equiv E + +instance [Bornology E] : Bornology (C⋆ᵐᵒᵈ E) := Bornology.induced <| equiv E + +/-- `WithCStarModule.equiv` as a uniform equivalence between `C⋆ᵐᵒᵈ E` and `E`. -/ +def uniformEquiv [UniformSpace E] : C⋆ᵐᵒᵈ E ≃ᵤ E := equiv E |>.toUniformEquivOfUniformInducing ⟨rfl⟩ + +instance [UniformSpace E] [CompleteSpace E] : CompleteSpace (C⋆ᵐᵒᵈ E) := + uniformEquiv.completeSpace_iff.mpr inferInstance + +end Basic + +/-! ## Prod + +Register simplification lemmas for the applications of `WithCStarModule (E × F)` elements, as +the usual lemmas for `Prod` will not trigger. -/ + +section Prod + +variable {R E F : Type*} +variable [SMul R E] [SMul R F] +variable (x y : C⋆ᵐᵒᵈ (E × F)) (c : R) + +section AddCommGroup + +variable [AddCommGroup E] [AddCommGroup F] + +@[simp] +theorem zero_fst : (0 : C⋆ᵐᵒᵈ (E × F)).fst = 0 := + rfl + +@[simp] +theorem zero_snd : (0 : C⋆ᵐᵒᵈ (E × F)).snd = 0 := + rfl + +@[simp] +theorem add_fst : (x + y).fst = x.fst + y.fst := + rfl + +@[simp] +theorem add_snd : (x + y).snd = x.snd + y.snd := + rfl + +@[simp] +theorem sub_fst : (x - y).fst = x.fst - y.fst := + rfl + +@[simp] +theorem sub_snd : (x - y).snd = x.snd - y.snd := + rfl + +@[simp] +theorem neg_fst : (-x).fst = -x.fst := + rfl + +@[simp] +theorem neg_snd : (-x).snd = -x.snd := + rfl + +end AddCommGroup + +@[simp] +theorem smul_fst : (c • x).fst = c • x.fst := + rfl + +@[simp] +theorem smul_snd : (c • x).snd = c • x.snd := + rfl + +/-! Note that the unapplied versions of these lemmas are deliberately omitted, as they break +the use of the type synonym. -/ + +@[simp] +theorem equiv_fst (x : C⋆ᵐᵒᵈ (E × F)) : (equiv (E × F) x).fst = x.fst := + rfl + +@[simp] +theorem equiv_snd (x : C⋆ᵐᵒᵈ (E × F)) : (equiv (E × F) x).snd = x.snd := + rfl + +@[simp] +theorem equiv_symm_fst (x : E × F) : ((equiv (E × F)).symm x).fst = x.fst := + rfl + +@[simp] +theorem equiv_symm_snd (x : E × F) : ((equiv (E × F)).symm x).snd = x.snd := + rfl + +end Prod + +/-! ## Pi + +Register simplification lemmas for the applications of `WithCStarModule (Π i, E i)` elements, as +the usual lemmas for `Pi` will not trigger. + +We also provide a `CoeFun` instance for `WithCStarModule (Π i, E i)`. -/ + +section Pi + +/- The following should not be a `FunLike` instance because then the coercion `⇑` would get +unfolded to `FunLike.coe` instead of `WithCStarModule.equiv`. -/ +instance {ι : Type*} (E : ι → Type*) : CoeFun (C⋆ᵐᵒᵈ (Π i, E i)) (fun _ ↦ Π i, E i) where + coe := WithCStarModule.equiv _ + +@[ext] +protected theorem ext {ι : Type*} {E : ι → Type*} {x y : C⋆ᵐᵒᵈ (Π i, E i)} + (h : ∀ i, x i = y i) : x = y := + funext h + +variable {R ι : Type*} {E : ι → Type*} +variable [∀ i, SMul R (E i)] +variable (c : R) (x y : C⋆ᵐᵒᵈ (Π i, E i)) (i : ι) + +section AddCommGroup + +variable [∀ i, AddCommGroup (E i)] + +@[simp] +theorem zero_apply : (0 : C⋆ᵐᵒᵈ (Π i, E i)) i = 0 := + rfl + +@[simp] +theorem add_apply : (x + y) i = x i + y i := + rfl + +@[simp] +theorem sub_apply : (x - y) i = x i - y i := + rfl + +@[simp] +theorem neg_apply : (-x) i = -x i := + rfl + +end AddCommGroup + +@[simp] +theorem smul_apply : (c • x) i = c • x i := + rfl + +/-! Note that the unapplied versions of these lemmas are deliberately omitted, as they break +the use of the type synonym. -/ + +@[simp] +theorem equiv_pi_apply (i : ι) : equiv _ x i = x i := + rfl + +@[simp] +theorem equiv_symm_pi_apply (x : ∀ i, E i) (i : ι) : + (WithCStarModule.equiv _).symm x i = x i := + rfl + +end Pi + +end WithCStarModule From 7bf7fe19f954552c770efde284976b23259ec5a5 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 15 Aug 2024 08:47:52 +0000 Subject: [PATCH 282/359] chore: tidy various files (#15807) --- .../Grp/EquivalenceGroupAddGroup.lean | 12 ++-- .../Algebra/Category/ModuleCat/Abelian.lean | 2 +- Mathlib/Algebra/Group/Semiconj/Defs.lean | 2 +- Mathlib/Algebra/Group/Semiconj/Units.lean | 2 +- Mathlib/Algebra/Module/Defs.lean | 7 +- Mathlib/Algebra/Order/Archimedean/Hom.lean | 4 +- .../PrimeSpectrum/Basic.lean | 2 +- Mathlib/Analysis/InnerProductSpace/Basic.lean | 8 +-- Mathlib/Data/Finset/Image.lean | 2 +- Mathlib/Data/Int/Interval.lean | 4 +- .../Geometry/RingedSpace/OpenImmersion.lean | 4 +- Mathlib/GroupTheory/OrderOfElement.lean | 9 ++- .../BilinearForm/Properties.lean | 9 ++- .../LinearAlgebra/QuadraticForm/Basic.lean | 2 +- .../MeasureTheory/Integral/PeakFunction.lean | 2 +- Mathlib/ModelTheory/PartialEquiv.lean | 18 ++--- Mathlib/NumberTheory/Padics/PadicNumbers.lean | 67 ++++++++----------- Mathlib/Order/RelIso/Basic.lean | 3 +- .../RingTheory/MvPolynomial/Symmetric.lean | 6 +- Mathlib/RingTheory/Trace/Basic.lean | 25 +++---- Mathlib/Topology/Maps/Basic.lean | 2 +- 21 files changed, 85 insertions(+), 107 deletions(-) diff --git a/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean b/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean index 4c58186eba37a..c0ffb41f10b80 100644 --- a/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean +++ b/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean @@ -27,7 +27,7 @@ private instance (X : CommGrp) : MulOneClass X.α := X.str.toMulOneClass private instance (X : AddGrp) : AddZeroClass X.α := X.str.toAddZeroClass private instance (X : AddCommGrp) : AddZeroClass X.α := X.str.toAddZeroClass -/-- The functor `Group ⥤ AddGroup` by sending `X ↦ Additive X` and `f ↦ f`. +/-- The functor `Grp ⥤ AddGrp` by sending `X ↦ Additive X` and `f ↦ f`. -/ @[simps] def toAddGrp : Grp ⥤ AddGrp where @@ -38,7 +38,7 @@ end Grp namespace CommGrp -/-- The functor `CommGroup ⥤ AddCommGroup` by sending `X ↦ Additive X` and `f ↦ f`. +/-- The functor `CommGrp ⥤ AddCommGrp` by sending `X ↦ Additive X` and `f ↦ f`. -/ @[simps] def toAddCommGrp : CommGrp ⥤ AddCommGrp where @@ -49,7 +49,7 @@ end CommGrp namespace AddGrp -/-- The functor `AddGroup ⥤ Group` by sending `X ↦ Multiplicative Y` and `f ↦ f`. +/-- The functor `AddGrp ⥤ Grp` by sending `X ↦ Multiplicative Y` and `f ↦ f`. -/ @[simps] def toGrp : AddGrp ⥤ Grp where @@ -60,7 +60,7 @@ end AddGrp namespace AddCommGrp -/-- The functor `AddCommGroup ⥤ CommGroup` by sending `X ↦ Multiplicative Y` and `f ↦ f`. +/-- The functor `AddCommGrp ⥤ CommGrp` by sending `X ↦ Multiplicative Y` and `f ↦ f`. -/ @[simps] def toCommGrp : AddCommGrp ⥤ CommGrp where @@ -69,14 +69,14 @@ def toCommGrp : AddCommGrp ⥤ CommGrp where end AddCommGrp -/-- The equivalence of categories between `Group` and `AddGroup` +/-- The equivalence of categories between `Grp` and `AddGrp` -/ def groupAddGroupEquivalence : Grp ≌ AddGrp := CategoryTheory.Equivalence.mk Grp.toAddGrp AddGrp.toGrp (NatIso.ofComponents fun X => MulEquiv.toGrpIso (MulEquiv.multiplicativeAdditive X)) (NatIso.ofComponents fun X => AddEquiv.toAddGrpIso (AddEquiv.additiveMultiplicative X)) -/-- The equivalence of categories between `CommGroup` and `AddCommGroup`. +/-- The equivalence of categories between `CommGrp` and `AddCommGrp`. -/ def commGroupAddCommGroupEquivalence : CommGrp ≌ AddCommGrp := CategoryTheory.Equivalence.mk CommGrp.toAddCommGrp AddCommGrp.toCommGrp diff --git a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean index 274f4ea89cd2d..401fb715425c1 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean @@ -39,7 +39,7 @@ def normalMono (hf : Mono f) : NormalMono f where calc M ≃ₗ[R] f.ker.quotient : (Submodule.quotEquivOfEqBot _ (ker_eq_bot_of_mono _)).symm ... ≃ₗ[R] f.range : LinearMap.quotKerEquivRange f - ... ≃ₗ[R] r.range.mkq.ker : LinearEquiv.ofEq _ _ (Submodule.ker_mkQ _).symm + ... ≃ₗ[R] r.range.mkQ.ker : LinearEquiv.ofEq _ _ (Submodule.ker_mkQ _).symm ``` -/ IsKernel.isoKernel _ _ (kernelIsLimit _) diff --git a/Mathlib/Algebra/Group/Semiconj/Defs.lean b/Mathlib/Algebra/Group/Semiconj/Defs.lean index bae3512e418c7..460a09c83cf47 100644 --- a/Mathlib/Algebra/Group/Semiconj/Defs.lean +++ b/Mathlib/Algebra/Group/Semiconj/Defs.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -Some proofs and docs came from `Algebra/Commute` (c) Neil Strickland +Some proofs and docs came from mathlib3 `src/algebra/commute.lean` (c) Neil Strickland -/ import Mathlib.Algebra.Group.Defs import Mathlib.Init.Logic diff --git a/Mathlib/Algebra/Group/Semiconj/Units.lean b/Mathlib/Algebra/Group/Semiconj/Units.lean index 142307a0a9682..11e0a651d129f 100644 --- a/Mathlib/Algebra/Group/Semiconj/Units.lean +++ b/Mathlib/Algebra/Group/Semiconj/Units.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -Some proofs and docs came from `Algebra/Commute` (c) Neil Strickland +Some proofs and docs came from mathlib3 `src/algebra/commute.lean` (c) Neil Strickland -/ import Mathlib.Algebra.Group.Semiconj.Defs import Mathlib.Algebra.Group.Units diff --git a/Mathlib/Algebra/Module/Defs.lean b/Mathlib/Algebra/Module/Defs.lean index 53aee9653932d..f509b6a113bc5 100644 --- a/Mathlib/Algebra/Module/Defs.lean +++ b/Mathlib/Algebra/Module/Defs.lean @@ -355,9 +355,10 @@ def AddCommMonoid.uniqueNatModule : Unique (Module ℕ M) where uniq P := (Module.ext' P _) fun n => by convert nat_smul_eq_nsmul P n instance AddCommMonoid.nat_isScalarTower : IsScalarTower ℕ R M where - smul_assoc n x y := - Nat.recOn n (by simp only [Nat.zero_eq, zero_smul]) - fun n ih => by simp only [Nat.succ_eq_add_one, add_smul, one_smul, ih] + smul_assoc n x y := by + induction n with + | zero => simp only [zero_smul] + | succ n ih => simp only [add_smul, one_smul, ih] end AddCommMonoid diff --git a/Mathlib/Algebra/Order/Archimedean/Hom.lean b/Mathlib/Algebra/Order/Archimedean/Hom.lean index 92a89b61ee68e..c4e7b11c89739 100644 --- a/Mathlib/Algebra/Order/Archimedean/Hom.lean +++ b/Mathlib/Algebra/Order/Archimedean/Hom.lean @@ -23,9 +23,7 @@ instance OrderRingHom.subsingleton [LinearOrderedField α] [LinearOrderedField ⟨fun f g => by ext x by_contra! h' : f x ≠ g x - wlog h : f x < g x generalizing α β with h₂ - -- Porting note: had to add the `generalizing` as there are random variables - -- `F γ δ` flying around in context. + wlog h : f x < g x with h₂ · exact h₂ g f x (Ne.symm h') (h'.lt_or_lt.resolve_left h) obtain ⟨q, hf, hg⟩ := exists_rat_btwn h rw [← map_ratCast f] at hf diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index a9bd9774ad781..fbd5e0acc99f5 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -622,7 +622,7 @@ lemma vanishingIdeal_irreducibleComponents : rw [irreducibleComponents_eq_maximals_closed, minimalPrimes_eq_minimals, image_antitone_setOf_maximal (fun s t hs _ ↦ (vanishingIdeal_anti_mono_iff hs.1).symm), ← funext (@Set.mem_setOf_eq _ · Ideal.IsPrime), ← vanishingIdeal_isClosed_isIrreducible] - rfl + simp only [Set.mem_image, Set.mem_setOf] lemma zeroLocus_minimalPrimes : zeroLocus ∘ (↑) '' minimalPrimes R = diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index f5aa42ed27d81..51f6fe6c16b82 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -299,10 +299,10 @@ theorem inner_add_add_self (x y : F) : ⟪x + y, x + y⟫ = ⟪x, x⟫ + ⟪x, y theorem inner_sub_sub_self (x y : F) : ⟪x - y, x - y⟫ = ⟪x, x⟫ - ⟪x, y⟫ - ⟪y, x⟫ + ⟪y, y⟫ := by simp only [inner_sub_left, inner_sub_right]; ring -theorem inner_smul_ofReal_left (x y : F) {t : ℝ} : ⟪(t : 𝕜) • x, y⟫ = ⟪x, y⟫ * t := by +theorem inner_smul_ofReal_left (x y : F) {t : ℝ} : ⟪(t : 𝕜) • x, y⟫ = ⟪x, y⟫ * t := by rw [inner_smul_left, conj_ofReal, mul_comm] -theorem inner_smul_ofReal_right (x y : F) {t : ℝ} : ⟪x, (t : 𝕜) • y⟫ = ⟪x, y⟫ * t := by +theorem inner_smul_ofReal_right (x y : F) {t : ℝ} : ⟪x, (t : 𝕜) • y⟫ = ⟪x, y⟫ * t := by rw [inner_smul_right, mul_comm] theorem re_inner_smul_ofReal_smul_self (x : F) {t : ℝ} : @@ -324,7 +324,7 @@ lemma cauchy_schwarz_aux' (x y : F) (t : ℝ) : 0 ≤ normSqF x * t * t + 2 * re [re_inner_smul_ofReal_smul_self, inner_smul_ofReal_left, inner_smul_ofReal_right] _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪y, x⟫ * t + re ⟪y, y⟫ := by rw [mul_comm ⟪x,y⟫ _, RCLike.re_ofReal_mul, mul_comm t _, mul_comm ⟪y,x⟫ _, RCLike.re_ofReal_mul, mul_comm t _] - _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪y, x⟫ * t + normSq y := by exact rfl + _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪y, x⟫ * t + normSq y := by rw [← normSq] _ = normSq x * t * t + re ⟪x, y⟫ * t + re ⟪x, y⟫ * t + normSq y := by rw [inner_re_symm] _ = normSq x * t * t + 2 * re ⟪x, y⟫ * t + normSq y := by ring @@ -359,7 +359,7 @@ theorem inner_mul_inner_self_le (x y : F) : ‖⟪x, y⟫‖ * ‖⟪y, x⟫‖ convert cauchy_schwarz_aux' (𝕜 := 𝕜) (⟪x, y⟫ • x) y (t / ‖⟪x, y⟫‖) using 3 · field_simp rw [← sq, normSq, normSq, inner_smul_right, inner_smul_left, ← mul_assoc _ _ ⟪x, x⟫, - mul_conj] + mul_conj] nth_rw 2 [sq] rw [← ofReal_mul, re_ofReal_mul] ring diff --git a/Mathlib/Data/Finset/Image.lean b/Mathlib/Data/Finset/Image.lean index e0e8779501a94..f659fdf322788 100644 --- a/Mathlib/Data/Finset/Image.lean +++ b/Mathlib/Data/Finset/Image.lean @@ -255,7 +255,7 @@ theorem map_nontrivial : (s.map f).Nontrivial ↔ s.Nontrivial := theorem attach_map_val {s : Finset α} : s.attach.map (Embedding.subtype _) = s := eq_of_veq <| by rw [map_val, attach_val]; exact Multiset.attach_map_val _ -theorem disjoint_range_addLeftEmbedding (a : ℕ) (s : Finset ℕ): +theorem disjoint_range_addLeftEmbedding (a : ℕ) (s : Finset ℕ) : Disjoint (range a) (map (addLeftEmbedding a) s) := by simp_rw [disjoint_left, mem_map, mem_range, addLeftEmbedding_apply] rintro _ h ⟨l, -, rfl⟩ diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index 9788ca3705b51..39cd170eb6188 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -112,9 +112,7 @@ theorem card_Ioo : (Ioo a b).card = (b - a - 1).toNat := (card_map _).trans <| c @[simp] theorem card_uIcc : (uIcc a b).card = (b - a).natAbs + 1 := (card_map _).trans <| - Int.ofNat.inj <| by - -- Porting note (#11215): TODO: Restore `Int.ofNat.inj` and remove the `change` - change ((↑) : ℕ → ℤ) _ = ((↑) : ℕ → ℤ) _ + Nat.cast_inj.mp <| by rw [card_range, sup_eq_max, inf_eq_min, Int.toNat_of_nonneg (sub_nonneg_of_le <| le_add_one min_le_max), Int.ofNat_add, Int.natCast_natAbs, add_comm, add_sub_assoc, max_sub_min_eq_abs, add_comm, Int.ofNat_one] diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index bd741b1f3e58f..02fe76d5d981b 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -27,11 +27,11 @@ Abbreviations are also provided for `SheafedSpace`, `LocallyRingedSpace` and `Sc contained in an open immersion factors though the open immersion. * `AlgebraicGeometry.PresheafedSpace.IsOpenImmersion.toSheafedSpace`: If `f : X ⟶ Y` is an open immersion of presheafed spaces, and `Y` is a sheafed space, then `X` is also a sheafed - space. The morphism as morphisms of sheafed spaces is given by `to_SheafedSpaceHom`. + space. The morphism as morphisms of sheafed spaces is given by `toSheafedSpaceHom`. * `AlgebraicGeometry.PresheafedSpace.IsOpenImmersion.toLocallyRingedSpace`: If `f : X ⟶ Y` is an open immersion of presheafed spaces, and `Y` is a locally ringed space, then `X` is also a locally ringed space. The morphism as morphisms of locally ringed spaces is given by - `to_LocallyRingedSpace_hom`. + `toLocallyRingedSpaceHom`. ## Main results diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index faa4010f0d09f..f50a3b8ac4190 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -966,9 +966,12 @@ section PowIsSubgroup @[to_additive "A nonempty idempotent subset of a finite cancellative add monoid is a submonoid"] def submonoidOfIdempotent {M : Type*} [LeftCancelMonoid M] [Finite M] (S : Set M) (hS1 : S.Nonempty) (hS2 : S * S = S) : Submonoid M := - have pow_mem : ∀ a : M, a ∈ S → ∀ n : ℕ, a ^ (n + 1) ∈ S := fun a ha => - Nat.rec (by rwa [Nat.zero_eq, zero_add, pow_one]) fun n ih => - (congr_arg₂ (· ∈ ·) (pow_succ a (n + 1)).symm hS2).mp (Set.mul_mem_mul ih ha) + have pow_mem (a : M) (ha : a ∈ S) (n : ℕ) : a ^ (n + 1) ∈ S := by + induction n with + | zero => rwa [zero_add, pow_one] + | succ n ih => + rw [← hS2, pow_succ] + exact Set.mul_mem_mul ih ha { carrier := S one_mem' := by obtain ⟨a, ha⟩ := hS1 diff --git a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean index 6d283deb55741..47f62c2cd1587 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Properties.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Properties.lean @@ -411,12 +411,11 @@ noncomputable def dualBasis (B : BilinForm K V) (hB : B.Nondegenerate) (b : Basi theorem dualBasis_repr_apply (B : BilinForm K V) (hB : B.Nondegenerate) (b : Basis ι K V) (x i) : (B.dualBasis hB b).repr x i = B x (b i) := by - rw [dualBasis, Basis.map_repr, LinearEquiv.symm_symm, LinearEquiv.trans_apply, - Basis.dualBasis_repr] #adaptation_note - /-- Before leanprover/lean4#4814, we had a final `toDual_def` in the `rw`, - and no need for the `rfl`. I'm confused! -/ - rfl + /-- Before leanprover/lean4#4814, we did not need the `@` in front of `toDual_def` in the `rw`. + I'm confused! -/ + rw [dualBasis, Basis.map_repr, LinearEquiv.symm_symm, LinearEquiv.trans_apply, + Basis.dualBasis_repr, @toDual_def] theorem apply_dualBasis_left (B : BilinForm K V) (hB : B.Nondegenerate) (b : Basis ι K V) (i j) : B (B.dualBasis hB b i) (b j) = if j = i then 1 else 0 := by diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index f3758444777a0..6476d73156547 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -643,7 +643,7 @@ end QuadraticMap 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 <`associated` +maps to bilinear maps giving this identification is called the `QuadraticMap.associated` quadratic map. -/ diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index af565327ab251..a3e735eec79cc 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -30,7 +30,7 @@ functions are also called approximations of unity, or approximations of identity at `0` and integrable. Note that there are related results about convolution with respect to peak functions in the file -`Analysis.Convolution`, such as `MeasureTheory.convolution_tendsto_right` there. +`Mathlib.Analysis.Convolution`, such as `MeasureTheory.convolution_tendsto_right` there. -/ open Set Filter MeasureTheory MeasureTheory.Measure TopologicalSpace Metric diff --git a/Mathlib/ModelTheory/PartialEquiv.lean b/Mathlib/ModelTheory/PartialEquiv.lean index 3277fb597492c..63599f2a61b3f 100644 --- a/Mathlib/ModelTheory/PartialEquiv.lean +++ b/Mathlib/ModelTheory/PartialEquiv.lean @@ -90,7 +90,8 @@ theorem toEquiv_inclusion {f g : M ≃ₚ[L] N} (h : f ≤ g) : g.toEquiv.toEmbedding.comp (Substructure.inclusion (dom_le_dom h)) = (Substructure.inclusion (cod_le_cod h)).comp f.toEquiv.toEmbedding := by rw [← (subtype _).comp_inj, subtype_toEquiv_inclusion h] - rfl + ext + simp theorem toEquiv_inclusion_apply {f g : M ≃ₚ[L] N} (h : f ≤ g) (x : f.dom) : g.toEquiv (Substructure.inclusion (dom_le_dom h) x) = @@ -98,7 +99,7 @@ theorem toEquiv_inclusion_apply {f g : M ≃ₚ[L] N} (h : f ≤ g) (x : f.dom) apply (subtype _).injective change (subtype _).comp (g.toEquiv.toEmbedding.comp (inclusion _)) x = _ rw [subtype_toEquiv_inclusion h] - rfl + simp theorem le_iff {f g : M ≃ₚ[L] N} : f ≤ g ↔ ∃ dom_le_dom : f.dom ≤ g.dom, @@ -115,7 +116,8 @@ theorem le_trans (f g h : M ≃ₚ[L] N) : f ≤ g → g ≤ h → f ≤ h := by rintro ⟨le_fg, eq_fg⟩ ⟨le_gh, eq_gh⟩ refine ⟨le_fg.trans le_gh, ?_⟩ rw [← eq_fg, ← Embedding.comp_assoc (g := g.toEquiv.toEmbedding), ← eq_gh] - rfl + ext + simp private theorem le_refl (f : M ≃ₚ[L] N) : f ≤ f := ⟨le_rfl, rfl⟩ @@ -152,7 +154,7 @@ theorem ext {f g : M ≃ₚ[L] N} (h_dom : f.dom = g.dom) : (∀ x : M, ∀ h : intro h rcases f with ⟨dom_f, cod_f, equiv_f⟩ cases h_dom - apply le_antisymm <;> (rw [le_def]; use (by rfl); ext ⟨x, hx⟩) + apply le_antisymm <;> (rw [le_def]; use le_rfl; ext ⟨x, hx⟩) · exact (h x hx).symm · exact h x hx @@ -206,8 +208,7 @@ def toEmbedding (f : M ≃ₚ[L] N) : f.dom ↪[L] N := @[simp] theorem toEmbedding_apply {f : M ≃ₚ[L] N} (m : f.dom) : - f.toEmbedding m = f.toEquiv m := by - rcases f with ⟨dom, cod, g⟩ + f.toEmbedding m = f.toEquiv m := rfl /-- Given a partial equivalence which has the whole structure as domain, @@ -268,8 +269,9 @@ theorem toPartialEquiv_toEmbedding {f : M ≃ₚ[L] N} (h : f.dom = ⊤) : rcases f with ⟨_, _, _⟩ cases h apply PartialEquiv.ext - intro _ _ - rfl; rfl + · intro _ _ + rfl + · rfl end Embedding diff --git a/Mathlib/NumberTheory/Padics/PadicNumbers.lean b/Mathlib/NumberTheory/Padics/PadicNumbers.lean index 3297fc932dd26..23431de0d49ec 100644 --- a/Mathlib/NumberTheory/Padics/PadicNumbers.lean +++ b/Mathlib/NumberTheory/Padics/PadicNumbers.lean @@ -37,9 +37,6 @@ We use the same concrete Cauchy sequence construction that is used to construct The extension of the norm on `ℚ` to `ℚ_[p]` is *not* analogous to extending the absolute value to `ℝ` and hence the proof that `ℚ_[p]` is complete is different from the proof that ℝ is complete. -A small special-purpose simplification tactic, `padic_index_simp`, is used to manipulate sequence -indices in the proof that the norm extends. - `padicNormE` is the rational-valued `p`-adic norm on `ℚ_[p]`. To instantiate `ℚ_[p]` as a normed field, we must cast this into an `ℝ`-valued norm. The `ℝ`-valued norm, using notation `‖ ‖` from normed spaces, @@ -267,8 +264,8 @@ theorem norm_mul (f g : PadicSeq p) : (f * g).norm = f.norm * g.norm := by simp only [hf, hg, norm, dif_pos, mul_zero] else by unfold norm - split_ifs with hfg - · exact (mul_not_equiv_zero hf hg hfg).elim + have hfg := mul_not_equiv_zero hf hg + simp only [hfg, hf, hg, dite_false] -- Porting note: originally `padic_index_simp [hfg, hf, hg]` rw [lift_index_left_left hfg, lift_index_left hf, lift_index_right hg] apply padicNorm.mul @@ -277,7 +274,7 @@ theorem eq_zero_iff_equiv_zero (f : PadicSeq p) : mk f = 0 ↔ f ≈ 0 := mk_eq theorem ne_zero_iff_nequiv_zero (f : PadicSeq p) : mk f ≠ 0 ↔ ¬f ≈ 0 := - not_iff_not.2 (eq_zero_iff_equiv_zero _) + eq_zero_iff_equiv_zero _ |>.not theorem norm_const (q : ℚ) : norm (const (padicNorm p) q) = padicNorm p q := if hq : q = 0 then by @@ -320,15 +317,12 @@ private theorem norm_eq_of_equiv_aux {f g : PadicSeq p} (hf : ¬f ≈ 0) (hg : private theorem norm_eq_of_equiv {f g : PadicSeq p} (hf : ¬f ≈ 0) (hg : ¬g ≈ 0) (hfg : f ≈ g) : padicNorm p (f (stationaryPoint hf)) = padicNorm p (g (stationaryPoint hg)) := by by_contra h - cases' - Decidable.em - (padicNorm p (g (stationaryPoint hg)) < padicNorm p (f (stationaryPoint hf))) with - hlt hnlt - · exact norm_eq_of_equiv_aux hf hg hfg h hlt - · apply norm_eq_of_equiv_aux hg hf (Setoid.symm hfg) (Ne.symm h) - apply lt_of_le_of_ne - · apply le_of_not_gt hnlt - · apply h + cases lt_or_le (padicNorm p (g (stationaryPoint hg))) (padicNorm p (f (stationaryPoint hf))) with + | inl hlt => + exact norm_eq_of_equiv_aux hf hg hfg h hlt + | inr hle => + apply norm_eq_of_equiv_aux hg hf (Setoid.symm hfg) (Ne.symm h) + exact lt_of_le_of_ne hle h theorem norm_equiv {f g : PadicSeq p} (hfg : f ≈ g) : f.norm = g.norm := by classical @@ -566,9 +560,11 @@ theorem defn (f : PadicSeq p) {ε : ℚ} (hε : 0 < ε) : exact not_lt_of_ge hge hε unfold PadicSeq.norm at hge; split_ifs at hge apply not_le_of_gt _ hge - cases' _root_.em (N ≤ stationaryPoint hne) with hgen hngen - · apply hN _ hgen _ hi - · have := stationaryPoint_spec hne le_rfl (le_of_not_le hngen) + cases _root_.le_total N (stationaryPoint hne) with + | inl hgen => + exact hN _ hgen _ hi + | inr hngen => + have := stationaryPoint_spec hne le_rfl hngen rw [← this] exact hN _ le_rfl _ hi @@ -915,19 +911,11 @@ instance complete : CauSeq.IsComplete ℚ_[p] norm where theorem padicNormE_lim_le {f : CauSeq ℚ_[p] norm} {a : ℝ} (ha : 0 < a) (hf : ∀ i, ‖f i‖ ≤ a) : ‖f.lim‖ ≤ a := by - -- Porting note: `Setoid.symm` cannot work out which `Setoid` to use, so instead swap the order - -- now, I use a rewrite to swap it later - obtain ⟨N, hN⟩ := (CauSeq.equiv_lim f) _ ha - rw [← sub_add_cancel f.lim (f N)] - refine le_trans (padicNormE.nonarchimedean _ _) ?_ - rw [norm_sub_rev] - exact max_le (le_of_lt (hN _ le_rfl)) (hf _) - -- Porting note: the following nice `calc` block does not work - -- exact calc - -- ‖f.lim‖ = ‖f.lim - f N + f N‖ := sorry - -- ‖f.lim - f N + f N‖ ≤ max ‖f.lim - f N‖ ‖f N‖ := sorry -- (padicNormE.nonarchimedean _ _) - -- max ‖f.lim - f N‖ ‖f N‖ = max ‖f N - f.lim‖ ‖f N‖ := sorry -- by congr; rw [norm_sub_rev] - -- max ‖f N - f.lim‖ ‖f N‖ ≤ a := sorry -- max_le (le_of_lt (hN _ le_rfl)) (hf _) + obtain ⟨N, hN⟩ := Setoid.symm (CauSeq.equiv_lim f) _ ha + calc + ‖f.lim‖ = ‖f.lim - f N + f N‖ := by simp + _ ≤ max ‖f.lim - f N‖ ‖f N‖ := padicNormE.nonarchimedean _ _ + _ ≤ a := max_le (le_of_lt (hN _ le_rfl)) (hf _) open Filter Set @@ -963,9 +951,8 @@ theorem valuation_one : valuation (1 : ℚ_[p]) = 0 := by classical change dite (CauSeq.const (padicNorm p) 1 ≈ _) _ _ = _ have h : ¬CauSeq.const (padicNorm p) 1 ≈ 0 := by - intro H - erw [const_equiv p] at H - exact one_ne_zero H + rw [← const_zero, const_equiv p] + exact one_ne_zero rw [dif_neg h] simp @@ -1025,7 +1012,7 @@ def addValuationDef : ℚ_[p] → WithTop ℤ := @[simp] theorem AddValuation.map_zero : addValuationDef (0 : ℚ_[p]) = ⊤ := by - rw [addValuationDef, if_pos (Eq.refl _)] + rw [addValuationDef, if_pos rfl] @[simp] theorem AddValuation.map_one : addValuationDef (1 : ℚ_[p]) = 0 := by @@ -1035,9 +1022,9 @@ theorem AddValuation.map_mul (x y : ℚ_[p]) : addValuationDef (x * y : ℚ_[p]) = addValuationDef x + addValuationDef y := by simp only [addValuationDef] by_cases hx : x = 0 - · rw [hx, if_pos (Eq.refl _), zero_mul, if_pos (Eq.refl _), WithTop.top_add] + · rw [hx, if_pos rfl, zero_mul, if_pos rfl, WithTop.top_add] · by_cases hy : y = 0 - · rw [hy, if_pos (Eq.refl _), mul_zero, if_pos (Eq.refl _), WithTop.add_top] + · rw [hy, if_pos rfl, mul_zero, if_pos rfl, WithTop.add_top] · rw [if_neg hx, if_neg hy, if_neg (mul_ne_zero hx hy), ← WithTop.coe_add, WithTop.coe_eq_coe, valuation_map_mul hx hy] @@ -1045,13 +1032,13 @@ theorem AddValuation.map_add (x y : ℚ_[p]) : min (addValuationDef x) (addValuationDef y) ≤ addValuationDef (x + y : ℚ_[p]) := by simp only [addValuationDef] by_cases hxy : x + y = 0 - · rw [hxy, if_pos (Eq.refl _)] + · rw [hxy, if_pos rfl] exact le_top · by_cases hx : x = 0 - · rw [hx, if_pos (Eq.refl _), min_eq_right, zero_add] + · rw [hx, if_pos rfl, min_eq_right, zero_add] exact le_top · by_cases hy : y = 0 - · rw [hy, if_pos (Eq.refl _), min_eq_left, add_zero] + · rw [hy, if_pos rfl, min_eq_left, add_zero] exact le_top · rw [if_neg hx, if_neg hy, if_neg hxy, ← WithTop.coe_min, WithTop.coe_le_coe] exact valuation_map_add hxy diff --git a/Mathlib/Order/RelIso/Basic.lean b/Mathlib/Order/RelIso/Basic.lean index fee223a097f83..d319840aac665 100644 --- a/Mathlib/Order/RelIso/Basic.lean +++ b/Mathlib/Order/RelIso/Basic.lean @@ -567,8 +567,7 @@ theorem coe_fn_mk (f : α ≃ β) (o : ∀ ⦃a b⦄, s (f a) (f b) ↔ r a b) : theorem coe_fn_toEquiv (f : r ≃r s) : (f.toEquiv : α → β) = f := rfl -/-- The map `coe_fn : (r ≃r s) → (α → β)` is injective. Lean fails to parse -`Function.Injective (fun e : r ≃r s ↦ (e : α → β))`, so we use a trick to say the same. -/ +/-- The map `DFunLike.coe : (r ≃r s) → (α → β)` is injective. -/ theorem coe_fn_injective : Injective fun f : r ≃r s => (f : α → β) := DFunLike.coe_injective diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean index 49097e138f0bb..691265878f204 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric.lean @@ -367,12 +367,12 @@ variable [DecidableEq σ] [DecidableEq τ] {n : ℕ} /-- The monomial symmetric `MvPolynomial σ R` with exponent set μ. It is the sum over all the monomials in `MvPolynomial σ R` such that the multiset of exponents is equal to the multiset of parts of μ. -/ -def msymm (μ : n.Partition) : MvPolynomial σ R := - ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod +def msymm (μ : n.Partition) : MvPolynomial σ R := + ∑ s : {a : Sym σ n // .ofSym a = μ}, (s.1.1.map X).prod @[simp] theorem msymm_zero : msymm σ R (.indiscrete 0) = 1 := by - rw [msymm, Fintype.sum_subsingleton _ ⟨(Sym.nil : Sym σ 0), by rfl⟩] + rw [msymm, Fintype.sum_subsingleton _ ⟨(Sym.nil : Sym σ 0), rfl⟩] simp @[simp] diff --git a/Mathlib/RingTheory/Trace/Basic.lean b/Mathlib/RingTheory/Trace/Basic.lean index 42e25f8be8ce6..ee8a5a44c08f7 100644 --- a/Mathlib/RingTheory/Trace/Basic.lean +++ b/Mathlib/RingTheory/Trace/Basic.lean @@ -119,13 +119,9 @@ variable (K) theorem trace_eq_trace_adjoin [FiniteDimensional K L] (x : L) : Algebra.trace K L x = finrank K⟮x⟯ L • trace K K⟮x⟯ (AdjoinSimple.gen K x) := by - -- Porting note: `conv` was - -- `conv in x => rw [← IntermediateField.AdjoinSimple.algebraMap_gen K x]` - -- and it was after the first `rw`. - conv => - lhs - rw [← IntermediateField.AdjoinSimple.algebraMap_gen K x] - rw [← trace_trace (S := K⟮x⟯), trace_algebraMap, LinearMap.map_smul_of_tower] + rw [← trace_trace (S := K⟮x⟯)] + conv in x => rw [← IntermediateField.AdjoinSimple.algebraMap_gen K x] + rw [trace_algebraMap, LinearMap.map_smul_of_tower] variable {K} @@ -247,12 +243,9 @@ theorem trace_eq_sum_embeddings [FiniteDimensional K L] [Algebra.IsSeparable K L have hx := Algebra.IsSeparable.isIntegral K x let pb := adjoin.powerBasis hx rw [trace_eq_trace_adjoin K x, Algebra.smul_def, RingHom.map_mul, ← adjoin.powerBasis_gen hx, - trace_eq_sum_embeddings_gen E pb (IsAlgClosed.splits_codomain _)] - -- Porting note: the following `convert` was `exact`, with `← Algebra.smul_def, algebraMap_smul` - -- in the previous `rw`. - · convert (sum_embeddings_eq_finrank_mul L E pb).symm - ext - simp + trace_eq_sum_embeddings_gen E pb (IsAlgClosed.splits_codomain _), ← Algebra.smul_def, + algebraMap_smul] + · exact (sum_embeddings_eq_finrank_mul L E pb).symm · haveI := Algebra.isSeparable_tower_bot_of_isSeparable K K⟮x⟯ L exact Algebra.IsSeparable.isSeparable K _ @@ -261,10 +254,8 @@ theorem trace_eq_sum_automorphisms (x : L) [FiniteDimensional K L] [IsGalois K L apply NoZeroSMulDivisors.algebraMap_injective L (AlgebraicClosure L) rw [_root_.map_sum (algebraMap L (AlgebraicClosure L))] rw [← Fintype.sum_equiv (Normal.algHomEquivAut K (AlgebraicClosure L) L)] - · rw [← trace_eq_sum_embeddings (AlgebraicClosure L)] - · simp only [algebraMap_eq_smul_one] - -- Porting note: `smul_one_smul` was in the `simp only`. - apply smul_one_smul + · rw [← trace_eq_sum_embeddings (AlgebraicClosure L) (x := x)] + simp only [algebraMap_eq_smul_one, smul_one_smul] · intro σ simp only [Normal.algHomEquivAut, AlgHom.restrictNormal', Equiv.coe_fn_mk, AlgEquiv.coe_ofBijective, AlgHom.restrictNormal_commutes, id.map_eq_id, RingHom.id_apply] diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index 3f1b53b0f9042..38f3f36cbcec0 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -383,7 +383,7 @@ protected theorem Inducing.isOpenMap (hi : Inducing f) (ho : IsOpen (range f)) : /-- Preimage of a dense set under an open map is dense. -/ protected theorem Dense.preimage {s : Set Y} (hs : Dense s) (hf : IsOpenMap f) : - Dense (f ⁻¹' s) := fun x ↦ + Dense (f ⁻¹' s) := fun x ↦ hf.preimage_closure_subset_closure_preimage <| hs (f x) end OpenMap From c0aa46273530ae96448bdfc9cd7e739bdfe8d3ad Mon Sep 17 00:00:00 2001 From: Jon Eugster Date: Thu, 15 Aug 2024 08:47:53 +0000 Subject: [PATCH 283/359] feat(mk_all): import Batteries in Mathlib.lean (#15810) Add a special case for Mathlib to add `import Batteries` to `Mathlib.lean` --- Mathlib.lean | 1 + scripts/lint-style.lean | 4 +++- scripts/mk_all.lean | 5 ++++- 3 files changed, 8 insertions(+), 2 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 18395bdf1b054..a6265743f1af0 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1,3 +1,4 @@ +import Batteries import Mathlib.Algebra.AddConstMap.Basic import Mathlib.Algebra.AddConstMap.Equiv import Mathlib.Algebra.AddTorsor diff --git a/scripts/lint-style.lean b/scripts/lint-style.lean index 93a1068c14cad..6367daa796ac2 100644 --- a/scripts/lint-style.lean +++ b/scripts/lint-style.lean @@ -25,7 +25,9 @@ def lintStyleCli (args : Cli.Parsed) : IO UInt32 := do -- Read all module names to lint. let mut allModules := #[] for s in ["Archive.lean", "Counterexamples.lean", "Mathlib.lean"] do - allModules := allModules.append ((← IO.FS.lines s).map (·.stripPrefix "import ")) + -- note: since we manually add "Batteries" to "Mathlib.lean", we remove it here manually + allModules := allModules.append ((← IO.FS.lines s).filter (!·.containsSubstr "Batteries") + |>.map (·.stripPrefix "import ")) let numberErrorFiles ← lintModules allModules mode -- Make sure to return an exit code of at most 125, so this return value can be used further -- in shell scripts. diff --git a/scripts/mk_all.lean b/scripts/mk_all.lean index 607a868e518e4..c33a957e6149d 100644 --- a/scripts/mk_all.lean +++ b/scripts/mk_all.lean @@ -50,7 +50,10 @@ def mkAllCLI (args : Parsed) : IO UInt32 := do let mut updates := 0 for d in libs.reverse do -- reverse to create `Mathlib/Tactic.lean` before `Mathlib.lean` let fileName := addExtension d "lean" - let allFiles ← getAllModulesSorted git d + let mut allFiles ← getAllModulesSorted git d + -- mathlib exception: manually import Batteries in `Mathlib.lean` + if d == "Mathlib" then + allFiles := #["Batteries"] ++ allFiles let fileContent := ("\n".intercalate (allFiles.map ("import " ++ ·)).toList).push '\n' if !(← pathExists fileName) then if check then From b664cf66619a565894072939ea9f2f7451d5f690 Mon Sep 17 00:00:00 2001 From: Seewoo Lee Date: Thu, 15 Aug 2024 08:47:54 +0000 Subject: [PATCH 284/359] Fix typo of author name (#15818) --- Mathlib/RingTheory/Radical.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean index 4dca38df1313c..8de194f20477d 100644 --- a/Mathlib/RingTheory/Radical.lean +++ b/Mathlib/RingTheory/Radical.lean @@ -1,5 +1,5 @@ /- -Copyright (c) 2024 Jineon Back and Seewoo Lee. All rights reserved. +Copyright (c) 2024 Jineon Baek, Seewoo Lee. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jineon Baek, Seewoo Lee -/ From d49e55251a723f1aee9c1f682084f86037e6bf25 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 15 Aug 2024 09:43:31 +0000 Subject: [PATCH 285/359] feat: Add `IsCaratheodory` lemmas (#15265) Add `isCaratheodory_diff`, `isCaratheodory_partialSups` and `isCaratheodory_disjointed`. Generalize `isCaratheodory_iUnion_nat`. --- .../OuterMeasure/Caratheodory.lean | 32 ++++++++++++++++--- 1 file changed, 28 insertions(+), 4 deletions(-) diff --git a/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean b/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean index 706d7e9a7b206..04c80d5a09c16 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Caratheodory.lean @@ -89,6 +89,21 @@ theorem isCaratheodory_inter (h₁ : IsCaratheodory m s₁) (h₂ : IsCaratheodo rw [← isCaratheodory_compl_iff, Set.compl_inter] exact isCaratheodory_union _ (isCaratheodory_compl _ h₁) (isCaratheodory_compl _ h₂) +lemma isCaratheodory_diff (h₁ : IsCaratheodory m s₁) (h₂ : IsCaratheodory m s₂) : + IsCaratheodory m (s₁ \ s₂) := m.isCaratheodory_inter h₁ (m.isCaratheodory_compl h₂) + +lemma isCaratheodory_partialSups {s : ℕ → Set α} (h : ∀ i, m.IsCaratheodory (s i)) (i : ℕ) : + m.IsCaratheodory (partialSups s i) := by + induction i with + | zero => exact h 0 + | succ i hi => exact m.isCaratheodory_union hi (h (i + 1)) + +lemma isCaratheodory_disjointed {s : ℕ → Set α} (h : ∀ i, m.IsCaratheodory (s i)) (i : ℕ) : + m.IsCaratheodory (disjointed s i) := by + induction i with + | zero => exact h 0 + | succ i _ => exact m.isCaratheodory_diff (h (i + 1)) (m.isCaratheodory_partialSups h i) + theorem isCaratheodory_sum {s : ℕ → Set α} (h : ∀ i, IsCaratheodory m (s i)) (hd : Pairwise (Disjoint on s)) {t : Set α} : ∀ {n}, (∑ i ∈ Finset.range n, m (t ∩ s i)) = m (t ∩ ⋃ i < n, s i) @@ -99,13 +114,13 @@ theorem isCaratheodory_sum {s : ℕ → Set α} (h : ∀ i, IsCaratheodory m (s intro a simpa using fun (h₁ : a ∈ s n) i (hi : i < n) h₂ => (hd (ne_of_gt hi)).le_bot ⟨h₁, h₂⟩ -set_option linter.deprecated false in -- not immediately obvious how to replace `iUnion` here. -theorem isCaratheodory_iUnion_nat {s : ℕ → Set α} (h : ∀ i, IsCaratheodory m (s i)) +/-- Use `isCaratheodory_iUnion` instead, which does not require the disjoint assumption. -/ +theorem isCaratheodory_iUnion_of_disjoint {s : ℕ → Set α} (h : ∀ i, IsCaratheodory m (s i)) (hd : Pairwise (Disjoint on s)) : IsCaratheodory m (⋃ i, s i) := by apply (isCaratheodory_iff_le' m).mpr intro t have hp : m (t ∩ ⋃ i, s i) ≤ ⨆ n, m (t ∩ ⋃ i < n, s i) := by - convert m.iUnion fun i => t ∩ s i using 1 + convert measure_iUnion_le (μ := m) fun i => t ∩ s i using 1 · simp [inter_iUnion] · simp [ENNReal.tsum_eq_iSup_nat, isCaratheodory_sum m h hd] refine le_trans (add_le_add_right hp _) ?_ @@ -115,6 +130,15 @@ theorem isCaratheodory_iUnion_nat {s : ℕ → Set α} (h : ∀ i, IsCaratheodor refine m.mono (diff_subset_diff_right ?_) exact iUnion₂_subset fun i _ => subset_iUnion _ i +@[deprecated (since := "2024-07-29")] +alias isCaratheodory_iUnion_nat := isCaratheodory_iUnion_of_disjoint + +lemma isCaratheodory_iUnion {s : ℕ → Set α} (h : ∀ i, m.IsCaratheodory (s i)) : + m.IsCaratheodory (⋃ i, s i) := by + rw [← iUnion_disjointed] + exact m.isCaratheodory_iUnion_of_disjoint (m.isCaratheodory_disjointed h) + (disjoint_disjointed _) + theorem f_iUnion {s : ℕ → Set α} (h : ∀ i, IsCaratheodory m (s i)) (hd : Pairwise (Disjoint on s)) : m (⋃ i, s i) = ∑' i, m (s i) := by refine le_antisymm (measure_iUnion_le s) ?_ @@ -129,7 +153,7 @@ def caratheodoryDynkin : MeasurableSpace.DynkinSystem α where Has := IsCaratheodory m has_empty := isCaratheodory_empty m has_compl s := isCaratheodory_compl m s - has_iUnion_nat f hf hn := by apply isCaratheodory_iUnion_nat m hf f + has_iUnion_nat _ hf hn := by apply isCaratheodory_iUnion m hf /-- Given an outer measure `μ`, the Carathéodory-measurable space is defined such that `s` is measurable if `∀t, μ t = μ (t ∩ s) + μ (t \ s)`. -/ From c66e20ecff7734236437c9298384e2e6a4aa34c0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 15 Aug 2024 09:43:32 +0000 Subject: [PATCH 286/359] feat: add NNReal.isOpen_Ico_zero (#15295) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Similar to the lemma `ENNReal.isOpen_Ico_zero` already in Mathlib. Co-authored-by: Rémy Degenne --- Mathlib/Topology/Instances/NNReal.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Topology/Instances/NNReal.lean b/Mathlib/Topology/Instances/NNReal.lean index e12bda18c19cc..f7759c47f7eea 100644 --- a/Mathlib/Topology/Instances/NNReal.lean +++ b/Mathlib/Topology/Instances/NNReal.lean @@ -75,6 +75,9 @@ instance : ContinuousStar ℝ≥0 where continuous_star := continuous_id section coe +lemma isOpen_Ico_zero {x : NNReal} : IsOpen (Set.Ico 0 x) := + Ico_bot (a := x) ▸ isOpen_Iio + variable {α : Type*} open Filter Finset From 4154233a74762fda3092177618e83103647cc4d7 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Thu, 15 Aug 2024 09:43:33 +0000 Subject: [PATCH 287/359] feat(CategoryTheory): `IsIso.comp_isIso` with explicit arguments (#15678) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Often when one wants to prove that a composition is an isomorphism, one proves that each component is. The instance `IsIso.comp_isIso` can be annoying to because of its instance arguments, this PR adds a lemma `IsIso.comp_isIso'` which is exactly the same as the instance, but with explicit arguments. This avoids having to guess the correct number of underscores in `@IsIso.comp_isIso` or doing `apply ( config := { allowSynthFailures := true } ) IsIso.comp_isIso`  --- Mathlib/AlgebraicGeometry/AffineScheme.lean | 2 +- Mathlib/AlgebraicGeometry/Morphisms/Affine.lean | 2 +- Mathlib/AlgebraicGeometry/Spec.lean | 3 +-- Mathlib/CategoryTheory/Iso.lean | 6 ++++++ Mathlib/CategoryTheory/Sites/Plus.lean | 2 +- Mathlib/CategoryTheory/Sites/Preserves.lean | 2 +- Mathlib/Geometry/RingedSpace/OpenImmersion.lean | 6 +++--- 7 files changed, 14 insertions(+), 9 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index 9e2cd3513685e..fd1b6a8831da4 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -441,7 +441,7 @@ def basicOpenSectionsToAffine : instance basicOpenSectionsToAffine_isIso : IsIso (basicOpenSectionsToAffine hU f) := by delta basicOpenSectionsToAffine - apply (config := { allowSynthFailures := true }) IsIso.comp_isIso + refine IsIso.comp_isIso' ?_ inferInstance apply PresheafedSpace.IsOpenImmersion.isIso_of_subset rw [hU.range_fromSpec] exact RingedSpace.basicOpen_le _ _ diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean index e4979dd055a15..9cf21c35955ab 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean @@ -127,7 +127,7 @@ lemma isAffine_of_isAffineOpen_basicOpen (s : Set Γ(X, ⊤)) exact hs₂ _ i.2 · simp only [Functor.comp_obj, Functor.rightOp_obj, Scheme.Γ_obj, Scheme.Spec_obj, id_eq, eq_mpr_eq_cast, Functor.id_obj, Opens.map_top, morphismRestrict_app] - apply (config := { allowSynthFailures := true }) IsIso.comp_isIso + refine IsIso.comp_isIso' ?_ inferInstance convert isIso_ΓSpec_adjunction_unit_app_basicOpen i.1 using 0 refine congr(IsIso ((ΓSpec.adjunction.unit.app X).app $(?_))) rw [Opens.openEmbedding_obj_top] diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index a18add389dda1..eac3cfe696a30 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -306,8 +306,7 @@ theorem Spec_map_localization_isIso (R : CommRingCat.{u}) (M : Submonoid R) erw [← localRingHom_comp_stalkIso] -- Porting note: replaced `apply (config := { instances := false })`. -- See https://github.com/leanprover/lean4/issues/2273 - refine @IsIso.comp_isIso _ _ _ _ _ _ _ _ (?_) - refine @IsIso.comp_isIso _ _ _ _ _ _ _ (?_) _ + refine IsIso.comp_isIso' inferInstance (IsIso.comp_isIso' ?_ inferInstance) /- I do not know why this is defeq to the goal, but I'm happy to accept that it is. -/ show IsIso (IsLocalization.localizationLocalizationAtPrimeIsoLocalization M diff --git a/Mathlib/CategoryTheory/Iso.lean b/Mathlib/CategoryTheory/Iso.lean index c30a8089b8795..77a86f6d853df 100644 --- a/Mathlib/CategoryTheory/Iso.lean +++ b/Mathlib/CategoryTheory/Iso.lean @@ -344,6 +344,12 @@ because `f.hom` is defeq to `(fun x ↦ x) ≫ f.hom`, triggering a loop. -/ instance (priority := 900) comp_isIso [IsIso f] [IsIso h] : IsIso (f ≫ h) := (asIso f ≪≫ asIso h).isIso_hom +/-- +The composition of isomorphisms is an isomorphism. Here the arguments of type `IsIso` are +explicit, to make this easier to use with the `refine` tactic, for instance. +-/ +lemma comp_isIso' (_ : IsIso f) (_ : IsIso h) : IsIso (f ≫ h) := inferInstance + @[simp] theorem inv_id : inv (𝟙 X) = 𝟙 X := by apply inv_eq_of_hom_inv_id diff --git a/Mathlib/CategoryTheory/Sites/Plus.lean b/Mathlib/CategoryTheory/Sites/Plus.lean index 3dc4ed4575b01..f8755d25759dd 100644 --- a/Mathlib/CategoryTheory/Sites/Plus.lean +++ b/Mathlib/CategoryTheory/Sites/Plus.lean @@ -246,7 +246,7 @@ theorem isIso_toPlus_of_isSheaf (hP : Presheaf.IsSheaf J P) : IsIso (J.toPlus P) rw [Presheaf.isSheaf_iff_multiequalizer] at hP suffices ∀ X, IsIso ((J.toPlus P).app X) from NatIso.isIso_of_isIso_app _ intro X - suffices IsIso (colimit.ι (J.diagram P X.unop) (op ⊤)) from IsIso.comp_isIso + refine IsIso.comp_isIso' inferInstance ?_ suffices ∀ (S T : (J.Cover X.unop)ᵒᵖ) (f : S ⟶ T), IsIso ((J.diagram P X.unop).map f) from isIso_ι_of_isInitial (initialOpOfTerminal isTerminalTop) _ intro S T e diff --git a/Mathlib/CategoryTheory/Sites/Preserves.lean b/Mathlib/CategoryTheory/Sites/Preserves.lean index 1fdb884f4aeb3..3c64ca614f34f 100644 --- a/Mathlib/CategoryTheory/Sites/Preserves.lean +++ b/Mathlib/CategoryTheory/Sites/Preserves.lean @@ -148,7 +148,7 @@ def preservesProductOfIsSheafFor have : HasCoproduct X := ⟨⟨c, hc⟩⟩ refine @PreservesProduct.ofIsoComparison _ _ _ _ F _ (fun x ↦ op (X x)) _ _ ?_ rw [piComparison_fac (hc := hc)] - refine @IsIso.comp_isIso _ _ _ _ _ _ _ inferInstance ?_ + refine IsIso.comp_isIso' inferInstance ?_ rw [isIso_iff_bijective, Function.bijective_iff_existsUnique] rw [Equalizer.Presieve.Arrows.sheaf_condition, Limits.Types.type_equalizer_iff_unique] at hF' exact fun b ↦ hF' b (congr_fun (firstMap_eq_secondMap F hF hI c hd) b) diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 02fe76d5d981b..6862e4b4c144c 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -152,7 +152,7 @@ instance comp {Z : PresheafedSpace C} (g : Y ⟶ Z) [hg : IsOpenImmersion g] : generalize_proofs h dsimp only [AlgebraicGeometry.PresheafedSpace.comp_c_app, unop_op, Functor.op, comp_base, Opens.map_comp_obj] - apply (config := { allowSynthFailures := true }) IsIso.comp_isIso + apply IsIso.comp_isIso' · exact c_iso' g ((opensFunctor f).obj U) (by ext; simp) · apply c_iso' f U ext1 @@ -738,7 +738,7 @@ theorem of_stalk_iso {X Y : SheafedSpace C} (f : X ⟶ Y) (hf : OpenEmbedding f. delta PresheafedSpace.Hom.stalkMap at H haveI H' := TopCat.Presheaf.stalkPushforward.stalkPushforward_iso_of_openEmbedding C hf X.presheaf y - have := @IsIso.comp_isIso _ _ _ _ _ _ _ H (@IsIso.inv_isIso _ _ _ _ _ H') + have := IsIso.comp_isIso' H (@IsIso.inv_isIso _ _ _ _ _ H') rwa [Category.assoc, IsIso.hom_inv_id, Category.comp_id] at this } end OfStalkIso @@ -1086,7 +1086,7 @@ theorem pullback_snd_isIso_of_range_subset (H' : Set.range g.1.base ⊆ Set.rang erw [← PreservesPullback.iso_hom_snd (LocallyRingedSpace.forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) f g] -- Porting note: was `inferInstance` - exact @IsIso.comp_isIso _ _ _ _ _ _ _ _ <| + exact IsIso.comp_isIso' inferInstance <| PresheafedSpace.IsOpenImmersion.pullback_snd_isIso_of_range_subset _ _ H' /-- The universal property of open immersions: From e024d5c04f15a1b917b4063810472b4a1bcab46a Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 15 Aug 2024 10:50:26 +0000 Subject: [PATCH 288/359] chore: do not set release-ci label upon failed PR test (#15775) with https://github.com/leanprover/lean4/pull/5022 the need for this should have disappeared --- scripts/lean-pr-testing-comments.sh | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/scripts/lean-pr-testing-comments.sh b/scripts/lean-pr-testing-comments.sh index df6bcb257f4fb..189a6ea6ac536 100755 --- a/scripts/lean-pr-testing-comments.sh +++ b/scripts/lean-pr-testing-comments.sh @@ -58,14 +58,14 @@ if [[ "$branch_name" =~ ^lean-pr-testing-([0-9]+)$ ]]; then -H "Authorization: Bearer $TOKEN" \ -H "X-GitHub-Api-Version: 2022-11-28" \ https://api.github.com/repos/leanprover/lean4/issues/$pr_number/labels/builds-mathlib - echo "Adding labels breaks-mathlib and release-ci" + echo "Adding label breaks-mathlib" curl -L -s \ -X POST \ -H "Accept: application/vnd.github+json" \ -H "Authorization: Bearer $TOKEN" \ -H "X-GitHub-Api-Version: 2022-11-28" \ https://api.github.com/repos/leanprover/lean4/issues/$pr_number/labels \ - -d '{"labels":["breaks-mathlib", "release-ci"]}' + -d '{"labels":["breaks-mathlib"]}' fi # Use GitHub API to check if a comment already exists From 75cc36e80cb9fe76f894b7688be1e0c792ae55d9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 15 Aug 2024 11:25:40 +0000 Subject: [PATCH 289/359] chore: Golf `prod_involution` (#15832) ... and extract `prod_ite_zero` from the proof of `prod_boole` Co-authored-by: Bhavik Mehta, --- .../Algebra/BigOperators/Group/Finset.lean | 76 +++++++++---------- .../BigOperators/GroupWithZero/Finset.lean | 18 +++-- 2 files changed, 47 insertions(+), 47 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index af2c8fa63c133..eaeaf7b7f6f44 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -1510,45 +1510,41 @@ theorem prod_flip {n : ℕ} (f : ℕ → β) : rw [prod_range_succ', prod_range_succ _ (Nat.succ n)] simp [← ih] -@[to_additive] -theorem prod_involution {s : Finset α} {f : α → β} : - ∀ (g : ∀ a ∈ s, α) (_ : ∀ a ha, f a * f (g a ha) = 1) (_ : ∀ a ha, f a ≠ 1 → g a ha ≠ a) - (g_mem : ∀ a ha, g a ha ∈ s) (_ : ∀ a ha, g (g a ha) (g_mem a ha) = a), - ∏ x ∈ s, f x = 1 := by - haveI := Classical.decEq α; haveI := Classical.decEq β - exact - Finset.strongInductionOn s fun s ih g h g_ne g_mem g_inv => - s.eq_empty_or_nonempty.elim (fun hs => hs.symm ▸ rfl) fun ⟨x, hx⟩ => - have hmem : ∀ y ∈ (s.erase x).erase (g x hx), y ∈ s := fun y hy => - mem_of_mem_erase (mem_of_mem_erase hy) - have g_inj : ∀ {x hx y hy}, g x hx = g y hy → x = y := fun {x hx y hy} h => by - rw [← g_inv x hx, ← g_inv y hy]; simp [h] - have ih' : (∏ y ∈ erase (erase s x) (g x hx), f y) = (1 : β) := - ih ((s.erase x).erase (g x hx)) - ⟨Subset.trans (erase_subset _ _) (erase_subset _ _), fun h => - not_mem_erase (g x hx) (s.erase x) (h (g_mem x hx))⟩ - (fun y hy => g y (hmem y hy)) (fun y hy => h y (hmem y hy)) - (fun y hy => g_ne y (hmem y hy)) - (fun y hy => - mem_erase.2 - ⟨fun h : g y _ = g x hx => by simp [g_inj h] at hy, - mem_erase.2 - ⟨fun h : g y _ = x => by - have : y = g x hx := g_inv y (hmem y hy) ▸ by simp [h] - simp [this] at hy, g_mem y (hmem y hy)⟩⟩) - fun y hy => g_inv y (hmem y hy) - if hx1 : f x = 1 then - ih' ▸ - Eq.symm - (prod_subset hmem fun y hy hy₁ => - have : y = x ∨ y = g x hx := by - simpa [hy, -not_and, mem_erase, not_and_or, or_comm] using hy₁ - this.elim (fun hy => hy.symm ▸ hx1) fun hy => - h x hx ▸ hy ▸ hx1.symm ▸ (one_mul _).symm) - else by - rw [← insert_erase hx, prod_insert (not_mem_erase _ _), ← - insert_erase (mem_erase.2 ⟨g_ne x hx hx1, g_mem x hx⟩), - prod_insert (not_mem_erase _ _), ih', mul_one, h x hx] +/-- The difference with `Finset.prod_ninvolution` is that the involution is allowed to use +membership of the domain of the product, rather than being a non-dependent function. -/ +@[to_additive "The difference with `Finset.sum_ninvolution` is that the involution is allowed to use +membership of the domain of the sum, rather than being a non-dependent function."] +lemma prod_involution (g : ∀ a ∈ s, α) (hg₁ : ∀ a ha, f a * f (g a ha) = 1) + (hg₃ : ∀ a ha, f a ≠ 1 → g a ha ≠ a) + (g_mem : ∀ a ha, g a ha ∈ s) (hg₄ : ∀ a ha, g (g a ha) (g_mem a ha) = a) : + ∏ x ∈ s, f x = 1 := by + classical + induction' s using Finset.strongInduction with s ih + obtain rfl | ⟨x, hx⟩ := s.eq_empty_or_nonempty + · simp + have : {x, g x hx} ⊆ s := by simp [insert_subset_iff, hx, g_mem] + suffices h : ∏ x ∈ s \ {x, g x hx}, f x = 1 by + rw [← prod_sdiff this, h, one_mul] + rcases eq_or_ne (g x hx) x + case inl hx' => simpa [hx'] using hg₃ x hx + case inr hx' => rw [prod_pair hx'.symm, hg₁] + suffices h₃ : ∀ a (ha : a ∈ s \ {x, g x hx}), g a (sdiff_subset ha) ∈ s \ {x, g x hx} by + exact ih (s \ {x, g x hx}) (ssubset_iff.2 ⟨x, by simp [insert_subset_iff, hx]⟩) _ + (by simp [hg₁]) (fun _ _ => hg₃ _ _) h₃ (fun _ _ => hg₄ _ _) + simp only [mem_sdiff, mem_insert, mem_singleton, not_or, g_mem, true_and] + rintro a ⟨ha₁, ha₂, ha₃⟩ + refine ⟨fun h => by simp [← h, hg₄] at ha₃, fun h => ?_⟩ + have : g (g a ha₁) (g_mem _ _) = g (g x hx) (g_mem _ _) := by simp only [h] + exact ha₂ (by simpa [hg₄] using this) + +/-- The difference with `Finset.prod_involution` is that the involution is a non-dependent function, +rather than being allowed to use membership of the domain of the product. -/ +@[to_additive "The difference with `Finset.sum_involution` is that the involution is a non-dependent +function, rather than being allowed to use membership of the domain of the sum."] +lemma prod_ninvolution (g : α → α) (hg₁ : ∀ a, f a * f (g a) = 1) (hg₂ : ∀ a, f a ≠ 1 → g a ≠ a) + (g_mem : ∀ a, g a ∈ s) (hg₃ : ∀ a, g (g a) = a) : ∏ x ∈ s, f x = 1 := + prod_involution (fun i _ => g i) (fun i _ => hg₁ i) (fun _ _ hi => hg₂ _ hi) + (fun i _ => g_mem i) (fun i _ => hg₃ i) /-- The product of the composition of functions `f` and `g`, is the product over `b ∈ s.image g` of `f b` to the power of the cardinality of the fibre of `b`. See also `Finset.prod_image`. -/ @@ -1674,7 +1670,7 @@ theorem prod_erase [DecidableEq α] (s : Finset α) {f : α → β} {a : α} (h rw [sdiff_singleton_eq_erase] at hnx rwa [eq_of_mem_of_not_mem_erase hx hnx] -/-- See also `Finset.prod_boole`. -/ +/-- See also `Finset.prod_ite_zero`. -/ @[to_additive "See also `Finset.sum_boole`."] theorem prod_ite_one (s : Finset α) (p : α → Prop) [DecidablePred p] (h : ∀ i ∈ s, ∀ j ∈ s, p i → p j → i = j) (a : β) : diff --git a/Mathlib/Algebra/BigOperators/GroupWithZero/Finset.lean b/Mathlib/Algebra/BigOperators/GroupWithZero/Finset.lean index e1f1eccb24706..ecc383eb492cf 100644 --- a/Mathlib/Algebra/BigOperators/GroupWithZero/Finset.lean +++ b/Mathlib/Algebra/BigOperators/GroupWithZero/Finset.lean @@ -24,15 +24,16 @@ variable [CommMonoidWithZero M₀] {p : ι → Prop} [DecidablePred p] {f : ι lemma prod_eq_zero (hi : i ∈ s) (h : f i = 0) : ∏ j ∈ s, f j = 0 := by classical rw [← prod_erase_mul _ _ hi, h, mul_zero] -lemma prod_boole : ∏ i ∈ s, (ite (p i) 1 0 : M₀) = ite (∀ i ∈ s, p i) 1 0 := by +lemma prod_ite_zero : + (∏ i ∈ s, if p i then f i else 0) = if ∀ i ∈ s, p i then ∏ i ∈ s, f i else 0 := by split_ifs with h - · apply prod_eq_one - intro i hi - rw [if_pos (h i hi)] + · exact prod_congr rfl fun i hi => by simp [h i hi] · push_neg at h rcases h with ⟨i, hi, hq⟩ - apply prod_eq_zero hi - rw [if_neg hq] + exact prod_eq_zero hi (by simp [hq]) + +lemma prod_boole : ∏ i ∈ s, (ite (p i) 1 0 : M₀) = ite (∀ i ∈ s, p i) 1 0 := by + rw [prod_ite_zero, prod_const_one] lemma support_prod_subset (s : Finset ι) (f : ι → κ → M₀) : support (fun x ↦ ∏ i ∈ s, f i x) ⊆ ⋂ i ∈ s, support (f i) := @@ -57,7 +58,10 @@ lemma support_prod (s : Finset ι) (f : ι → κ → M₀) : end Finset namespace Fintype -variable [Fintype ι] [CommMonoidWithZero M₀] {p : ι → Prop} [DecidablePred p] +variable [Fintype ι] [CommMonoidWithZero M₀] {p : ι → Prop} [DecidablePred p] {f : ι → M₀} + +lemma prod_ite_zero : (∏ i, if p i then f i else 0) = if ∀ i, p i then ∏ i, f i else 0 := by + simp [Finset.prod_ite_zero] lemma prod_boole : ∏ i, (ite (p i) 1 0 : M₀) = ite (∀ i, p i) 1 0 := by simp [Finset.prod_boole] From ebc084f34c951dbf91bacec6cf8d77be3c8e7730 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 15 Aug 2024 12:19:28 +0000 Subject: [PATCH 290/359] chore: cleanup in Mathlib/Init/* (#15188) This moves some content out of `Init`, deprecates some material that was ported from Lean 3 but is used nowhere, and removes some attributes and instances which are duplicated from Lean or Batteries. Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib.lean | 8 +-- Mathlib/Algebra/GroupWithZero/Semiconj.lean | 1 - Mathlib/Analysis/Asymptotics/Asymptotics.lean | 2 +- Mathlib/Data/Fin/Basic.lean | 4 +- Mathlib/Data/Fintype/Basic.lean | 1 + Mathlib/Data/List/Basic.lean | 68 ++----------------- Mathlib/Data/List/Defs.lean | 23 +++++-- Mathlib/Data/List/GetD.lean | 1 - Mathlib/Data/List/Intervals.lean | 6 +- .../Instances.lean => Data/List/Monad.lean} | 9 +-- Mathlib/Data/Nat/Bits.lean | 2 +- Mathlib/Data/Nat/Choose/Multinomial.lean | 4 +- Mathlib/Data/Nat/Digits.lean | 3 +- Mathlib/Data/Prod/Basic.lean | 4 ++ Mathlib/Data/Prod/PProd.lean | 4 ++ Mathlib/Data/Sigma/Basic.lean | 11 +++ Mathlib/Data/Stream/Init.lean | 1 - Mathlib/Deprecated/Combinator.lean | 20 ++++++ Mathlib/FieldTheory/KummerExtension.lean | 2 +- Mathlib/Init/Classical.lean | 41 ----------- Mathlib/Init/Core.lean | 54 --------------- Mathlib/Init/Data/Int/Order.lean | 30 +------- Mathlib/Init/Data/List/Basic.lean | 51 -------------- Mathlib/Init/Data/Nat/Lemmas.lean | 39 ++++------- Mathlib/Init/Data/Sigma/Basic.lean | 29 -------- Mathlib/Init/Data/Sigma/Lex.lean | 61 ----------------- Mathlib/Init/Logic.lean | 14 +--- Mathlib/Init/Quot.lean | 1 - .../AffineSpace/FiniteDimensional.lean | 1 - Mathlib/Logic/Basic.lean | 31 ++++++++- Mathlib/Logic/Equiv/Basic.lean | 1 - Mathlib/NumberTheory/Cyclotomic/Basic.lean | 1 - Mathlib/NumberTheory/Pell.lean | 2 +- Mathlib/Order/Fin/Basic.lean | 1 + Mathlib/Order/RelClasses.lean | 25 +++++++ Mathlib/Order/WellFoundedSet.lean | 5 +- .../Cardinal/SchroederBernstein.lean | 1 - Mathlib/SetTheory/Ordinal/Notation.lean | 2 +- Mathlib/Tactic/Group.lean | 2 +- Mathlib/Testing/SlimCheck/Sampleable.lean | 2 +- test/RewriteSearch/Polynomial.lean | 1 - 41 files changed, 157 insertions(+), 412 deletions(-) rename Mathlib/{Init/Data/List/Instances.lean => Data/List/Monad.lean} (65%) create mode 100644 Mathlib/Deprecated/Combinator.lean delete mode 100644 Mathlib/Init/Classical.lean delete mode 100644 Mathlib/Init/Core.lean delete mode 100644 Mathlib/Init/Data/List/Basic.lean delete mode 100644 Mathlib/Init/Data/Sigma/Basic.lean delete mode 100644 Mathlib/Init/Data/Sigma/Lex.lean diff --git a/Mathlib.lean b/Mathlib.lean index a6265743f1af0..bddf041fc866a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2219,6 +2219,7 @@ import Mathlib.Data.List.Lattice import Mathlib.Data.List.Lemmas import Mathlib.Data.List.Lex import Mathlib.Data.List.MinMax +import Mathlib.Data.List.Monad import Mathlib.Data.List.NatAntidiagonal import Mathlib.Data.List.Nodup import Mathlib.Data.List.NodupEquivFin @@ -2538,6 +2539,7 @@ import Mathlib.Data.ZMod.Parity import Mathlib.Data.ZMod.Quotient import Mathlib.Data.ZMod.Units import Mathlib.Deprecated.Aliases +import Mathlib.Deprecated.Combinator import Mathlib.Deprecated.Group import Mathlib.Deprecated.HashMap import Mathlib.Deprecated.Ring @@ -2792,16 +2794,10 @@ import Mathlib.GroupTheory.Torsion import Mathlib.GroupTheory.Transfer import Mathlib.InformationTheory.Hamming import Mathlib.Init.Algebra.Classes -import Mathlib.Init.Classical -import Mathlib.Init.Core import Mathlib.Init.Data.Int.Order -import Mathlib.Init.Data.List.Basic -import Mathlib.Init.Data.List.Instances import Mathlib.Init.Data.List.Lemmas import Mathlib.Init.Data.Nat.Lemmas import Mathlib.Init.Data.Quot -import Mathlib.Init.Data.Sigma.Basic -import Mathlib.Init.Data.Sigma.Lex import Mathlib.Init.Logic import Mathlib.Init.Order.LinearOrder import Mathlib.Init.Quot diff --git a/Mathlib/Algebra/GroupWithZero/Semiconj.lean b/Mathlib/Algebra/GroupWithZero/Semiconj.lean index f702796fd87ed..f3798456188e4 100644 --- a/Mathlib/Algebra/GroupWithZero/Semiconj.lean +++ b/Mathlib/Algebra/GroupWithZero/Semiconj.lean @@ -5,7 +5,6 @@ Authors: Johan Commelin -/ import Mathlib.Algebra.GroupWithZero.Units.Basic import Mathlib.Algebra.Group.Semiconj.Units -import Mathlib.Init.Classical /-! # Lemmas about semiconjugate elements in a `GroupWithZero`. diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 1feaea8e67fc7..2e4f8556940f8 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -1394,7 +1394,7 @@ theorem IsLittleO.pow {f : α → R} {g : α → 𝕜} (h : f =o[l] g) {n : ℕ} (fun x => f x ^ n) =o[l] fun x => g x ^ n := by obtain ⟨n, rfl⟩ := Nat.exists_eq_succ_of_ne_zero hn.ne'; clear hn induction n with - | zero => simpa only [Nat.zero_eq, ← Nat.one_eq_succ_zero, pow_one] + | zero => simpa only [pow_one] | succ n ihn => convert ihn.mul h <;> simp [pow_succ] theorem IsLittleO.of_pow {f : α → 𝕜} {g : α → R} {n : ℕ} (h : (f ^ n) =o[l] (g ^ n)) (hn : n ≠ 0) : diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index bb29c0f7e1ad9..cd8cf284df159 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -5,7 +5,6 @@ Authors: Robert Y. Lewis, Keeley Hoek -/ import Mathlib.Algebra.NeZero import Mathlib.Data.Nat.Defs -import Mathlib.Init.Data.Nat.Lemmas import Mathlib.Logic.Embedding.Basic import Mathlib.Logic.Equiv.Set import Mathlib.Tactic.Common @@ -346,8 +345,7 @@ instance nontrivial {n : ℕ} : Nontrivial (Fin (n + 2)) where theorem nontrivial_iff_two_le : Nontrivial (Fin n) ↔ 2 ≤ n := by rcases n with (_ | _ | n) <;> - simp [← Nat.one_eq_succ_zero, Fin.nontrivial, not_nontrivial, Nat.succ_le_iff] --- Porting note: here and in the next lemma, had to use `← Nat.one_eq_succ_zero`. + simp [Fin.nontrivial, not_nontrivial, Nat.succ_le_iff] section Monoid diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 08d7bacc8a516..eac49025de7cc 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.Init.Data.Nat.Lemmas /-! # Finite types diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 22a70729a558d..8bce8077ffc81 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -6,8 +6,7 @@ Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, M import Mathlib.Data.Nat.Defs import Mathlib.Data.Option.Basic import Mathlib.Data.List.Defs -import Mathlib.Init.Data.List.Basic -import Mathlib.Init.Data.List.Instances +import Mathlib.Data.List.Monad import Mathlib.Init.Data.List.Lemmas import Mathlib.Logic.Unique import Mathlib.Order.Basic @@ -33,6 +32,12 @@ universe u v w variable {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {l₁ l₂ : List α} +/-- `≤` implies not `>` for lists. -/ +@[deprecated (since := "2024-07-27")] +theorem le_eq_not_gt [LT α] : ∀ l₁ l₂ : List α, (l₁ ≤ l₂) = ¬l₂ < l₁ := fun _ _ => rfl + +@[deprecated (since := "2024-06-07")] alias toArray_data := Array.data_toArray + -- Porting note: Delete this attribute -- attribute [inline] List.head! @@ -90,12 +95,6 @@ theorem _root_.Function.Involutive.exists_mem_and_apply_eq_iff {f : α → α} theorem mem_map_of_involutive {f : α → α} (hf : Involutive f) {a : α} {l : List α} : a ∈ map f l ↔ f a ∈ l := by rw [mem_map, hf.exists_mem_and_apply_eq_iff] -attribute [simp] List.mem_join - -attribute [simp] List.mem_bind - --- Porting note: bExists in Lean3, And in Lean4 - /-! ### length -/ alias ⟨_, length_pos_of_ne_nil⟩ := length_pos @@ -132,13 +131,10 @@ theorem length_eq_three {l : List α} : l.length = 3 ↔ ∃ a b c, l = [a, b, c /-! ### set-theoretic notation of lists -/ --- ADHOC Porting note: instance from Lean3 core instance instSingletonList : Singleton α (List α) := ⟨fun x => [x]⟩ --- ADHOC Porting note: instance from Lean3 core instance [DecidableEq α] : Insert α (List α) := ⟨List.insert⟩ --- ADHOC Porting note: instance from Lean3 core instance [DecidableEq α] : LawfulSingleton α (List α) := { insert_emptyc_eq := fun x => show (if x ∈ ([] : List α) then [] else [x]) = [x] from if_neg (not_mem_nil _) } @@ -185,9 +181,6 @@ theorem exists_mem_cons_iff (p : α → Prop) (a : α) (l : List α) : /-! ### list subset -/ -instance : IsTrans (List α) Subset where - trans := fun _ _ _ => List.Subset.trans - theorem cons_subset_of_subset_of_mem {a : α} {l m : List α} (ainm : a ∈ m) (lsubm : l ⊆ m) : a::l ⊆ m := cons_subset.2 ⟨ainm, lsubm⟩ @@ -196,8 +189,6 @@ theorem append_subset_of_subset_of_subset {l₁ l₂ l : List α} (l₁subl : l l₁ ++ l₂ ⊆ l := fun _ h ↦ (mem_append.1 h).elim (@l₁subl _) (@l₂subl _) --- Porting note: in Batteries - alias ⟨eq_nil_of_subset_nil, _⟩ := subset_nil theorem map_subset_iff {l₁ l₂ : List α} (f : α → β) (h : Injective f) : @@ -211,8 +202,6 @@ theorem map_subset_iff {l₁ l₂ : List α} (f : α → β) (h : Injective f) : theorem append_eq_has_append {L₁ L₂ : List α} : List.append L₁ L₂ = L₁ ++ L₂ := rfl --- Porting note: in Batteries - @[deprecated (since := "2024-03-24")] alias append_eq_cons_iff := append_eq_cons @[deprecated (since := "2024-03-24")] alias cons_eq_append_iff := cons_eq_append @@ -292,9 +281,6 @@ theorem bind_eq_bind {α β} (f : α → List β) (l : List α) : l >>= f = l.bi /-! ### reverse -/ --- Porting note: Do we need this? -attribute [local simp] reverseAux - theorem reverse_cons' (a : α) (l : List α) : reverse (a :: l) = concat (reverse l) a := by simp only [reverse_cons, concat_eq_append] @@ -333,9 +319,6 @@ theorem map_reverseAux (f : α → β) (l₁ l₂ : List α) : /-! ### empty -/ --- Porting note: this does not work as desired --- attribute [simp] List.isEmpty - theorem isEmpty_iff_eq_nil {l : List α} : l.isEmpty ↔ l = [] := by cases l <;> simp [isEmpty] /-! ### dropLast -/ @@ -403,7 +386,6 @@ lemma getLast_filter {p : α → Bool} : /-! ### getLast? -/ --- Porting note: Moved earlier in file, for use in subsequent lemmas. @[simp] theorem getLast?_cons_cons (a b : α) (l : List α) : getLast? (a :: b :: l) = getLast? (b :: l) := rfl @@ -734,7 +716,6 @@ theorem Sublist.of_cons_cons {l₁ l₂ : List α} {a b : α} (h : a :: l₁ <+ @[deprecated (since := "2024-04-07")] theorem sublist_of_cons_sublist_cons {a} (h : a :: l₁ <+ a :: l₂) : l₁ <+ l₂ := h.of_cons_cons -attribute [simp] cons_sublist_cons @[deprecated (since := "2024-04-07")] alias cons_sublist_cons_iff := cons_sublist_cons theorem eq_nil_of_sublist_nil {l : List α} (s : l <+ []) : l = [] := @@ -1622,16 +1603,8 @@ theorem intersperse_cons_cons (a b c : α) (tl : List α) : section SplitAtOn -/- Porting note: the new version of `splitOnP` uses a `Bool`-valued predicate instead of a - `Prop`-valued one. All downstream definitions have been updated to match. -/ - variable (p : α → Bool) (xs ys : List α) (ls : List (List α)) (f : List α → List α) -/- Porting note: this had to be rewritten because of the new implementation of `splitAt`. It's - long in large part because `splitAt.go` (`splitAt`'s auxiliary function) works differently - in the case where n ≥ length l, requiring two separate cases (and two separate inductions). Still, - this can hopefully be golfed. -/ - @[simp] theorem splitAt_eq_take_drop (n : ℕ) (l : List α) : splitAt n l = (take n l, drop n l) := by by_cases h : n < l.length <;> rw [splitAt, go_eq_take_drop] @@ -1674,10 +1647,6 @@ theorem splitOn_nil [DecidableEq α] (a : α) : [].splitOn a = [[]] := theorem splitOnP_nil : [].splitOnP p = [[]] := rfl -/- Porting note: `split_on_p_aux` and `split_on_p_aux'` were used to prove facts about - `split_on_p`. `splitOnP` has a different structure, and we need different facts about - `splitOnP.go`. Theorems involving `split_on_p_aux` have been omitted where possible. -/ - theorem splitOnP.go_ne_nil (xs acc : List α) : splitOnP.go p xs acc ≠ [] := by induction xs generalizing acc <;> simp [go]; split <;> simp [*] @@ -1780,7 +1749,6 @@ theorem splitOn_intercalate [DecidableEq α] (x : α) (hx : ∀ l ∈ ls, x ∉ end SplitAtOn -/- Porting note: new; here tentatively -/ /-! ### modifyLast -/ section ModifyLast @@ -1852,18 +1820,6 @@ section find? variable {p : α → Bool} {l : List α} {a : α} --- @[simp] --- Later porting note (at time of this lemma moving to Batteries): --- removing attribute `nolint simpNF` -attribute [simp 1100] find?_cons_of_pos - --- @[simp] --- Later porting note (at time of this lemma moving to Batteries): --- removing attribute `nolint simpNF` -attribute [simp 1100] find?_cons_of_neg - -attribute [simp] find?_eq_none - @[deprecated (since := "2024-05-05")] alias find?_mem := mem_of_find?_eq_some end find? @@ -2090,9 +2046,6 @@ theorem dropWhile_nthLe_zero_not (l : List α) (hl : 0 < (l.dropWhile p).length) by_cases hp : p hd · simp [hp, IH] · simp [hp, nthLe_cons] --- Porting note: How did the Lean 3 proof work, --- without mentioning nthLe_cons? --- Same question for takeWhile_eq_nil_iff below variable {p} {l : List α} @@ -2470,13 +2423,6 @@ theorem zipRight_eq_zipRight' : zipRight as bs = (zipRight' as bs).fst := by end ZipRight -/-! ### toChunks -/ - --- Porting note: --- The definition of `toChunks` has changed substantially from Lean 3. --- The theorems about `toChunks` are not used anywhere in mathlib, anyways. --- TODO: Prove these theorems for the new definitions. - /-! ### Forall -/ section Forall diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index cc5d8513bf39b..fc6eaa2b034dc 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -29,15 +29,30 @@ variable {α β γ δ ε ζ : Type*} instance [DecidableEq α] : SDiff (List α) := ⟨List.diff⟩ --- mathlib3 `array` is not ported. --- Porting note: see --- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/List.2Ehead/near/313204716 --- for the fooI naming convention. /-- "Inhabited" `get` function: returns `default` instead of `none` in the case that the index is out of bounds. -/ def getI [Inhabited α] (l : List α) (n : Nat) : α := getD l n default +/-- The head of a list, or the default element of the type is the list is `nil`. -/ +def headI [Inhabited α] : List α → α + | [] => default + | (a :: _) => a + +@[simp] theorem headI_nil [Inhabited α] : ([] : List α).headI = default := rfl +@[simp] theorem headI_cons [Inhabited α] {h : α} {t : List α} : (h :: t).headI = h := rfl + +/-- The last element of a list, with the default if list empty -/ +def getLastI [Inhabited α] : List α → α + | [] => default + | [a] => a + | [_, b] => b + | _ :: _ :: l => getLastI l + +/-- List with a single given element. -/ +@[inline, deprecated List.pure (since := "2024-03-24")] +protected def ret {α : Type u} (a : α) : List α := [a] + /-- "Inhabited" `take` function: Take `n` elements from a list `l`. If `l` has less than `n` elements, append `n - length l` elements `default`. -/ def takeI [Inhabited α] (n : Nat) (l : List α) : List α := diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index 9b888110c53b9..3ec5febd08640 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -6,7 +6,6 @@ Mario Carneiro -/ import Mathlib.Data.List.Defs import Mathlib.Data.Option.Basic -import Mathlib.Init.Data.List.Basic import Mathlib.Util.AssertExists /-! # getD and getI diff --git a/Mathlib/Data/List/Intervals.lean b/Mathlib/Data/List/Intervals.lean index cefbcfd118662..39e9340abd44f 100644 --- a/Mathlib/Data/List/Intervals.lean +++ b/Mathlib/Data/List/Intervals.lean @@ -5,8 +5,8 @@ Authors: Scott Morrison -/ import Mathlib.Data.List.Lattice import Mathlib.Data.Bool.Basic -import Mathlib.Init.Data.Nat.Lemmas import Mathlib.Order.Lattice + /-! # Intervals in ℕ @@ -116,9 +116,7 @@ theorem eq_cons {n m : ℕ} (h : n < m) : Ico n m = n :: Ico (n + 1) m := by @[simp] theorem pred_singleton {m : ℕ} (h : 0 < m) : Ico (m - 1) m = [m - 1] := by - dsimp [Ico] - rw [Nat.sub_sub_self (succ_le_of_lt h)] - simp [← Nat.one_eq_succ_zero] + simp [Ico, Nat.sub_sub_self (succ_le_of_lt h)] theorem chain'_succ (n m : ℕ) : Chain' (fun a b => b = succ a) (Ico n m) := by by_cases h : n < m diff --git a/Mathlib/Init/Data/List/Instances.lean b/Mathlib/Data/List/Monad.lean similarity index 65% rename from Mathlib/Init/Data/List/Instances.lean rename to Mathlib/Data/List/Monad.lean index f4ae241d5a94a..0dc10994ff059 100644 --- a/Mathlib/Init/Data/List/Instances.lean +++ b/Mathlib/Data/List/Monad.lean @@ -5,14 +5,7 @@ Authors: Leonardo de Moura -/ /-! -# Note about `Mathlib/Init/` -The files in `Mathlib/Init` are leftovers from the port from Mathlib3. -(They contain content moved from lean3 itself that Mathlib needed but was not moved to lean4.) - -We intend to move all the content of these files out into the main `Mathlib` directory structure. -Contributions assisting with this are appreciated. - -# Decidable and Monad instances for `List` not (yet) in `Batteries` +# Monad instances for `List` -/ universe u v w diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 117d9a800575e..df2a618f055ba 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -6,7 +6,7 @@ Authors: Praneeth Kolichala import Mathlib.Algebra.Group.Basic import Mathlib.Algebra.Group.Nat import Mathlib.Data.Nat.Defs -import Mathlib.Init.Data.List.Basic +import Mathlib.Data.List.Defs import Mathlib.Tactic.Convert import Mathlib.Tactic.GeneralizeProofs import Mathlib.Tactic.Says diff --git a/Mathlib/Data/Nat/Choose/Multinomial.lean b/Mathlib/Data/Nat/Choose/Multinomial.lean index f41552ee482ec..62485e17c0efa 100644 --- a/Mathlib/Data/Nat/Choose/Multinomial.lean +++ b/Mathlib/Data/Nat/Choose/Multinomial.lean @@ -73,9 +73,9 @@ lemma multinomial_insert [DecidableEq α] (ha : a ∉ s) (f : α → ℕ) : @[simp] theorem multinomial_insert_one [DecidableEq α] (h : a ∉ s) (h₁ : f a = 1) : multinomial (insert a s) f = (s.sum f).succ * multinomial s f := by - simp only [multinomial, one_mul, factorial] + simp only [multinomial] rw [Finset.sum_insert h, Finset.prod_insert h, h₁, add_comm, ← succ_eq_add_one, factorial_succ] - simp only [factorial_one, one_mul, Function.comp_apply, factorial, mul_one, ← one_eq_succ_zero] + simp only [factorial, succ_eq_add_one, zero_add, mul_one, one_mul] rw [Nat.mul_div_assoc _ (prod_factorial_dvd_factorial_sum _ _)] theorem multinomial_congr {f g : α → ℕ} (h : ∀ a ∈ s, f a = g a) : diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index fd51916d69605..3465ed5369aa7 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -544,8 +544,7 @@ theorem sub_one_mul_sum_div_pow_eq_sub_sum_digits {p : ℕ} Ico_zero_eq_range, mul_add, mul_add, ih, range_one, sum_singleton, List.drop, ofDigits, mul_zero, add_zero, ← Nat.add_sub_assoc <| sum_le_ofDigits _ <| Nat.le_of_lt h] nth_rw 2 [← one_mul <| ofDigits p tl] - rw [← add_mul, one_eq_succ_zero, Nat.sub_add_cancel <| zero_lt_of_lt h, - Nat.add_sub_add_left] + rw [← add_mul, Nat.sub_add_cancel (one_le_of_lt h), Nat.add_sub_add_left] · simp [ofDigits_one] · simp [lt_one_iff.mp h] cases L diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index c384a20b1de7c..c13f7a9e55f7b 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -20,6 +20,10 @@ variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} namespace Prod +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₂ + @[simp] theorem mk.eta : ∀ {p : α × β}, (p.1, p.2) = p | (_, _) => rfl diff --git a/Mathlib/Data/Prod/PProd.lean b/Mathlib/Data/Prod/PProd.lean index bd9537e4bf646..dbbf70e0f9842 100644 --- a/Mathlib/Data/Prod/PProd.lean +++ b/Mathlib/Data/Prod/PProd.lean @@ -16,6 +16,10 @@ variable {α β γ δ : Sort*} namespace PProd +def mk.injArrow {α : Type*} {β : Type*} {x₁ : α} {y₁ : β} {x₂ : α} {y₂ : β} : + (x₁, y₁) = (x₂, y₂) → ∀ ⦃P : Sort*⦄, (x₁ = x₂ → y₁ = y₂ → P) → P := + fun h₁ _ h₂ ↦ Prod.noConfusion h₁ h₂ + @[simp] theorem mk.eta {p : PProd α β} : PProd.mk p.1 p.2 = p := rfl diff --git a/Mathlib/Data/Sigma/Basic.lean b/Mathlib/Data/Sigma/Basic.lean index 238a6f84ac8a5..54166f10c5b7c 100644 --- a/Mathlib/Data/Sigma/Basic.lean +++ b/Mathlib/Data/Sigma/Basic.lean @@ -61,6 +61,10 @@ theorem mk.inj_iff {a₁ a₂ : α} {b₁ : β a₁} {b₂ : β a₂} : theorem eta : ∀ x : Σa, β a, Sigma.mk x.1 x.2 = x | ⟨_, _⟩ => rfl +protected theorem eq {α : Type*} {β : α → Type*} : ∀ {p₁ p₂ : Σ a, β a} (h₁ : p₁.1 = p₂.1), + (Eq.recOn h₁ p₁.2 : β p₂.1) = p₂.2 → p₁ = p₂ + | ⟨_, _⟩, _, rfl, rfl => rfl + /-- A version of `Iff.mp Sigma.ext_iff` for functions from a nonempty type to a sigma type. -/ theorem _root_.Function.eq_of_sigmaMk_comp {γ : Type*} [Nonempty γ] {a b : α} {f : γ → β a} {g : γ → β b} (h : Sigma.mk a ∘ f = Sigma.mk b ∘ g) : @@ -209,6 +213,8 @@ def elim {γ} (f : ∀ a, β a → γ) (a : PSigma β) : γ := theorem elim_val {γ} (f : ∀ a, β a → γ) (a b) : PSigma.elim f ⟨a, b⟩ = f a b := rfl +@[deprecated (since := "2024-07-27")] alias ex_of_psig := ex_of_PSigma + instance [Inhabited α] [Inhabited (β default)] : Inhabited (PSigma β) := ⟨⟨default, default⟩⟩ @@ -232,6 +238,11 @@ theorem mk.inj_iff {a₁ a₂ : α} {b₁ : β a₁} {b₂ : β a₂} : match a₁, a₂, b₁, b₂, h₁, h₂ with | _, _, _, _, Eq.refl _, HEq.refl _ => rfl +@[deprecated PSigma.ext_iff (since := "2024-07-27")] +protected theorem eq {α : Sort*} {β : α → Sort*} : ∀ {p₁ p₂ : Σ' a, β a} (h₁ : p₁.1 = p₂.1), + (Eq.recOn h₁ p₁.2 : β p₂.1) = p₂.2 → p₁ = p₂ + | ⟨_, _⟩, _, rfl, rfl => rfl + @[simp] theorem «forall» {p : (Σ'a, β a) → Prop} : (∀ x, p x) ↔ ∀ a b, p ⟨a, b⟩ := ⟨fun h a b ↦ h ⟨a, b⟩, fun h ⟨a, b⟩ ↦ h a b⟩ diff --git a/Mathlib/Data/Stream/Init.lean b/Mathlib/Data/Stream/Init.lean index c8fb6c851095e..7f714ab54b4df 100644 --- a/Mathlib/Data/Stream/Init.lean +++ b/Mathlib/Data/Stream/Init.lean @@ -5,7 +5,6 @@ Authors: Leonardo de Moura -/ import Mathlib.Data.Stream.Defs import Mathlib.Logic.Function.Basic -import Mathlib.Init.Data.List.Basic import Mathlib.Data.List.Basic /-! diff --git a/Mathlib/Deprecated/Combinator.lean b/Mathlib/Deprecated/Combinator.lean new file mode 100644 index 0000000000000..a19596bdeb6eb --- /dev/null +++ b/Mathlib/Deprecated/Combinator.lean @@ -0,0 +1,20 @@ +/- +Copyright (c) 2014 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Mario Carneiro +-/ + +/-! +# Deprecated combinators, ported from Lean 3 core. +-/ + +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) + +end Combinator diff --git a/Mathlib/FieldTheory/KummerExtension.lean b/Mathlib/FieldTheory/KummerExtension.lean index c36f3f23e6e98..6b3c4cdc76d6e 100644 --- a/Mathlib/FieldTheory/KummerExtension.lean +++ b/Mathlib/FieldTheory/KummerExtension.lean @@ -63,7 +63,7 @@ lemma root_X_pow_sub_C_ne_zero {n : ℕ} (hn : 1 < n) (a : K) : lemma root_X_pow_sub_C_ne_zero' {n : ℕ} {a : K} (hn : 0 < n) (ha : a ≠ 0) : (AdjoinRoot.root (X ^ n - C a)) ≠ 0 := by obtain (rfl|hn) := (Nat.succ_le_iff.mpr hn).eq_or_lt - · rw [← Nat.one_eq_succ_zero, pow_one] + · rw [pow_one] intro e refine mk_ne_zero_of_natDegree_lt (monic_X_sub_C a) (C_ne_zero.mpr ha) (by simp) ?_ trans AdjoinRoot.mk (X - C a) (X - (X - C a)) diff --git a/Mathlib/Init/Classical.lean b/Mathlib/Init/Classical.lean deleted file mode 100644 index 2fd70039da0cb..0000000000000 --- a/Mathlib/Init/Classical.lean +++ /dev/null @@ -1,41 +0,0 @@ -/- -Copyright (c) 2015 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro --/ -import Mathlib.Tactic.IrreducibleDef - -/-! -# Note about `Mathlib/Init/` -The files in `Mathlib/Init` are leftovers from the port from Mathlib3. -(They contain content moved from lean3 itself that Mathlib needed but was not moved to lean4.) - -We intend to move all the content of these files out into the main `Mathlib` directory structure. -Contributions assisting with this are appreciated. - -### alignments from lean 3 `init.classical` --/ - -namespace Classical - -attribute [local instance] propDecidable -attribute [local instance] decidableInhabited - -alias axiom_of_choice := axiomOfChoice -- TODO: fix in core -alias prop_complete := propComplete -- TODO: fix in core - -@[elab_as_elim] 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 - -theorem cases_on (a : Prop) {p : Prop → Prop} (h1 : p True) (h2 : p False) : p a := - @cases_true_false p h1 h2 a - -theorem cases {p : Prop → Prop} (h1 : p True) (h2 : p False) (a) : p a := cases_on a h1 h2 - -alias by_cases := byCases -alias by_contradiction := byContradiction - -theorem eq_false_or_eq_true (a : Prop) : a = False ∨ a = True := (prop_complete a).symm - -end Classical diff --git a/Mathlib/Init/Core.lean b/Mathlib/Init/Core.lean deleted file mode 100644 index 1627fc02f1afc..0000000000000 --- a/Mathlib/Init/Core.lean +++ /dev/null @@ -1,54 +0,0 @@ -/- -Copyright (c) 2014 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura, Mario Carneiro --/ -import Mathlib.Tactic.Relation.Trans - -/-! -# Note about `Mathlib/Init/` -The files in `Mathlib/Init` are leftovers from the port from Mathlib3. -(They contain content moved from lean3 itself that Mathlib needed but was not moved to lean4.) - -We intend to move all the content of these files out into the main `Mathlib` directory structure. -Contributions assisting with this are appreciated. - -# Notation, basic datatypes and type classes - -This file contains alignments from lean 3 `init.core`. --/ - --- Note: we do not currently auto-align constants. - --- TODO --- attribute [elab_as_elim, subst] Eq.subst - -attribute [trans] Eq.trans -attribute [symm] Eq.symm - -universe u v w - -def Prod.mk.injArrow {α : Type u} {β : Type v} {x₁ : α} {y₁ : β} {x₂ : α} {y₂ : β} : - (x₁, y₁) = (x₂, y₂) → ∀ ⦃P : Sort w⦄, (x₁ = x₂ → y₁ = y₂ → P) → P := - fun h₁ _ h₂ ↦ Prod.noConfusion h₁ h₂ - -def PProd.mk.injArrow {α : Type u} {β : Type v} {x₁ : α} {y₁ : β} {x₂ : α} {y₂ : β} : - (x₁, y₁) = (x₂, y₂) → ∀ ⦃P : Sort w⦄, (x₁ = x₂ → y₁ = y₂ → P) → P := - fun h₁ _ h₂ ↦ Prod.noConfusion h₁ h₂ - --- attribute [pp_using_anonymous_constructor] Sigma PSigma Subtype PProd And - --- Generic 'has'-stripping --- Note: we don't currently strip automatically for various reasons. - -attribute [simp] insert_emptyc_eq - --- Combinator calculus -namespace Combinator - -variable {α : Sort u} {β : Sort v} {γ : Sort w} -def I (a : α) := a -def K (a : α) (_b : β) := a -def S (x : α → β → γ) (y : α → β) (z : α) := x z (y z) - -end Combinator diff --git a/Mathlib/Init/Data/Int/Order.lean b/Mathlib/Init/Data/Int/Order.lean index abe8b3642eaa5..bae8f42395a7a 100644 --- a/Mathlib/Init/Data/Int/Order.lean +++ b/Mathlib/Init/Data/Int/Order.lean @@ -47,32 +47,8 @@ instance instLinearOrder : LinearOrder ℤ where decidableLE := by infer_instance decidableLT := by infer_instance -theorem neg_mul_eq_neg_mul_symm (a b : ℤ) : -a * b = -(a * b) := (Int.neg_mul_eq_neg_mul a b).symm - -theorem mul_neg_eq_neg_mul_symm (a b : ℤ) : a * -b = -(a * b) := (Int.neg_mul_eq_mul_neg a b).symm +@[deprecated (since := "2024-07-27")] alias mul_neg_eq_neg_mul_symm := Int.mul_neg +@[deprecated (since := "2024-07-27")] alias neg_mul_eq_neg_mul_symm := Int.neg_mul protected theorem eq_zero_or_eq_zero_of_mul_eq_zero {a b : ℤ} (h : a * b = 0) : a = 0 ∨ b = 0 := - match lt_trichotomy 0 a with - | Or.inl hlt₁ => - match lt_trichotomy 0 b with - | Or.inl hlt₂ => by - have : 0 < a * b := Int.mul_pos hlt₁ hlt₂ - rw [h] at this - exact absurd this (lt_irrefl _) - | Or.inr (Or.inl heq₂) => Or.inr heq₂.symm - | Or.inr (Or.inr hgt₂) => by - have : 0 > a * b := Int.mul_neg_of_pos_of_neg hlt₁ hgt₂ - rw [h] at this - exact absurd this (lt_irrefl _) - | Or.inr (Or.inl heq₁) => Or.inl heq₁.symm - | Or.inr (Or.inr hgt₁) => - match lt_trichotomy 0 b with - | Or.inl hlt₂ => by - have : 0 > a * b := Int.mul_neg_of_neg_of_pos hgt₁ hlt₂ - rw [h] at this - exact absurd this (lt_irrefl _) - | Or.inr (Or.inl heq₂) => Or.inr heq₂.symm - | Or.inr (Or.inr hgt₂) => by - have : 0 < a * b := Int.mul_pos_of_neg_of_neg hgt₁ hgt₂ - rw [h] at this - exact absurd this (lt_irrefl _) + Int.mul_eq_zero.mp h diff --git a/Mathlib/Init/Data/List/Basic.lean b/Mathlib/Init/Data/List/Basic.lean deleted file mode 100644 index f0b5deabfcfcc..0000000000000 --- a/Mathlib/Init/Data/List/Basic.lean +++ /dev/null @@ -1,51 +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 Mathlib.Data.Nat.Notation -import Batteries.Tactic.Alias - -/-! -# Note about `Mathlib/Init/` -The files in `Mathlib/Init` are leftovers from the port from Mathlib3. -(They contain content moved from lean3 itself that Mathlib needed but was not moved to lean4.) - -We intend to move all the content of these files out into the main `Mathlib` directory structure. -Contributions assisting with this are appreciated. - -## Definitions for `List` not (yet) in `Batteries` --/ - -open Decidable List - -universe u v w -variable {α : Type u} - -namespace List - -/-- The head of a list, or the default element of the type is the list is `nil`. -/ -def headI [Inhabited α] : List α → α - | [] => default - | (a :: _) => a - -@[simp] theorem headI_nil [Inhabited α] : ([] : List α).headI = default := rfl -@[simp] theorem headI_cons [Inhabited α] {h : α} {t : List α} : (h :: t).headI = h := rfl - -/-- The last element of a list, with the default if list empty -/ -def getLastI [Inhabited α] : List α → α - | [] => default - | [a] => a - | [_, b] => b - | _ :: _ :: l => getLastI l - -/-- List with a single given element. -/ -@[inline, deprecated List.pure (since := "2024-03-24")] -protected def ret {α : Type u} (a : α) : List α := [a] - -/-- `≤` implies not `>` for lists. -/ -theorem le_eq_not_gt [LT α] : ∀ l₁ l₂ : List α, (l₁ ≤ l₂) = ¬l₂ < l₁ := fun _ _ => rfl - -@[deprecated (since := "2024-06-07")] alias toArray_data := Array.data_toArray - -end List diff --git a/Mathlib/Init/Data/Nat/Lemmas.lean b/Mathlib/Init/Data/Nat/Lemmas.lean index d580fa8f4aed0..0117ea849a1e2 100644 --- a/Mathlib/Init/Data/Nat/Lemmas.lean +++ b/Mathlib/Init/Data/Nat/Lemmas.lean @@ -24,11 +24,8 @@ namespace Nat /-! multiplication -/ -theorem eq_zero_of_mul_eq_zero : ∀ {n m : ℕ}, n * m = 0 → n = 0 ∨ m = 0 - | 0, m => fun _ => Or.inl rfl - | succ n, m => by - rw [succ_mul]; intro h - exact Or.inr (Nat.eq_zero_of_add_eq_zero_left h) +theorem eq_zero_of_mul_eq_zero {n m : ℕ} (h : n * m = 0) : n = 0 ∨ m = 0 := + mul_eq_zero.mp h /-! properties of inequality -/ @@ -41,11 +38,6 @@ protected def ltByCases {a b : ℕ} {C : Sort u} (h₁ : a < b → C) (h₂ : a -- TODO: there are four variations, depending on which variables we assume to be nonneg -/-! bit0/bit1 properties -/ -section bit - -end bit - /-! successor and predecessor -/ def discriminate {B : Sort u} {n : ℕ} (H1 : n = 0 → B) (H2 : ∀ m, n = succ m → B) : B := by @@ -53,24 +45,23 @@ def discriminate {B : Sort u} {n : ℕ} (H1 : n = 0 → B) (H2 : ∀ m, n = succ | zero => exact H1 rfl | succ n _ => apply H2 _ rfl +-- Unused in Mathlib; +-- if downstream projects find this essential please copy it or remove the deprecation. +@[deprecated (since := "2024-07-27")] theorem one_eq_succ_zero : 1 = succ 0 := rfl -/-! subtraction - -Many lemmas are proven more generally in mathlib `Algebra/Order/Sub` -/ - -/-! min -/ - /-! induction principles -/ - def twoStepInduction {P : ℕ → Sort u} (H1 : P 0) (H2 : P 1) (H3 : ∀ (n : ℕ) (_IH1 : P n) (_IH2 : P (succ n)), P (succ (succ n))) : ∀ a : ℕ, P a | 0 => H1 | 1 => H2 | succ (succ _n) => H3 _ (twoStepInduction H1 H2 H3 _) (twoStepInduction H1 H2 H3 _) +-- Unused in Mathlib; +-- if downstream projects find this essential please copy it or remove the deprecation. +@[deprecated (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 _ @@ -92,16 +83,12 @@ protected theorem case_strong_induction_on {p : Nat → Prop} (a : Nat) (hz : p /-! mod -/ +-- Unused in Mathlib; +-- if downstream projects find this essential please copy it or remove the deprecation. +@[deprecated (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 - by_cases h : x % 2 = 1 - · simp! [*] - · cases mod_two_eq_zero_or_one x <;> simp! [*, Nat.zero_ne_one] - -/-! div -/ - -/-! dvd -/ - - + simp only [cond_eq_if, decide_eq_true_eq] + split <;> omega end Nat diff --git a/Mathlib/Init/Data/Sigma/Basic.lean b/Mathlib/Init/Data/Sigma/Basic.lean deleted file mode 100644 index 47e27c6f72e55..0000000000000 --- a/Mathlib/Init/Data/Sigma/Basic.lean +++ /dev/null @@ -1,29 +0,0 @@ -/- -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 --/ - -/-! -# Note about `Mathlib/Init/` -The files in `Mathlib/Init` are leftovers from the port from Mathlib3. -(They contain content moved from lean3 itself that Mathlib needed but was not moved to lean4.) - -We intend to move all the content of these files out into the main `Mathlib` directory structure. -Contributions assisting with this are appreciated. - -# Lemmas about `Sigma` from Lean 3 core. --/ - -universe u v - -theorem ex_of_psig {α : Sort u} {p : α → Prop} : (Σ' x, p x) → ∃ x, p x - | ⟨x, hx⟩ => ⟨x, hx⟩ - -protected theorem Sigma.eq {α : Type u} {β : α → Type v} : ∀ {p₁ p₂ : Σ a, β a} (h₁ : p₁.1 = p₂.1), - (Eq.recOn h₁ p₁.2 : β p₂.1) = p₂.2 → p₁ = p₂ - | ⟨_, _⟩, _, rfl, rfl => rfl - -protected theorem PSigma.eq {α : Sort u} {β : α → Sort v} : ∀ {p₁ p₂ : Σ' a, β a} (h₁ : p₁.1 = p₂.1), - (Eq.recOn h₁ p₁.2 : β p₂.1) = p₂.2 → p₁ = p₂ - | ⟨_, _⟩, _, rfl, rfl => rfl diff --git a/Mathlib/Init/Data/Sigma/Lex.lean b/Mathlib/Init/Data/Sigma/Lex.lean deleted file mode 100644 index 80a3ad677bc92..0000000000000 --- a/Mathlib/Init/Data/Sigma/Lex.lean +++ /dev/null @@ -1,61 +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 --/ - -/-! -# Note about `Mathlib/Init/` -The files in `Mathlib/Init` are leftovers from the port from Mathlib3. -(They contain content moved from lean3 itself that Mathlib needed but was not moved to lean4.) - -We intend to move all the content of these files out into the main `Mathlib` directory structure. -Contributions assisting with this are appreciated. - -# Well-foundedness of orders of well-founded relations - -Porting note: many declarations that were here in mathlib3 have been moved to `Init.WF` -in Lean 4 core. The ones below are all the exceptions. --/ - -universe u v - -namespace PSigma - -section - -open WellFounded - -variable {α : Sort u} {β : α → Sort v} {r : α → α → Prop} {s : ∀ a : α, β a → β a → Prop} - -/-- The lexicographical order of well-founded relations is well-founded. -/ -theorem lex_wf (ha : WellFounded r) (hb : ∀ x, WellFounded (s x)) : WellFounded (Lex r s) := - WellFounded.intro fun ⟨a, b⟩ => lexAccessible (WellFounded.apply ha a) hb b - -end - -section - -open WellFounded - -variable {α : Sort u} {β : Sort v} {r : α → α → Prop} {s : β → β → Prop} - -theorem revLex_wf (ha : WellFounded r) (hb : WellFounded s) : WellFounded (RevLex r s) := - WellFounded.intro fun ⟨a, b⟩ => revLexAccessible (apply hb b) (WellFounded.apply ha) a - -end - -section - -theorem skipLeft_wf (α : Type u) {β : Type v} {s : β → β → Prop} (hb : WellFounded s) : - WellFounded (SkipLeft α s) := - revLex_wf emptyWf.wf hb - -end - -instance hasWellFounded {α : Type u} {β : α → Type v} [s₁ : WellFoundedRelation α] - [s₂ : ∀ a, WellFoundedRelation (β a)] : WellFoundedRelation (PSigma β) where - rel := Lex s₁.rel fun a => (s₂ a).rel - wf := lex_wf s₁.wf fun a => (s₂ a).wf - -end PSigma diff --git a/Mathlib/Init/Logic.lean b/Mathlib/Init/Logic.lean index 0496c9b280296..f6d926751096d 100644 --- a/Mathlib/Init/Logic.lean +++ b/Mathlib/Init/Logic.lean @@ -65,8 +65,6 @@ theorem heq_prop {P Q : Prop} (p : P) (q : Q) : HEq p q := variable {a b c d : Prop} -/- or -/ - /- xor -/ def Xor' (a b : Prop) := (a ∧ ¬ b) ∨ (b ∧ ¬ a) @@ -77,15 +75,8 @@ attribute [refl] Iff.refl attribute [trans] Iff.trans attribute [symm] Iff.symm --- This is needed for `calc` to work with `iff`. -instance : Trans Iff Iff Iff where - trans := fun p q ↦ p.trans q - alias ⟨not_of_not_not_not, _⟩ := not_not_not --- FIXME --- attribute [congr] not_congr - variable (p) -- FIXME: remove _iff and add _eq for the lean 4 core versions @@ -110,9 +101,6 @@ theorem false_iff_iff : (False ↔ a) ↔ ¬a := iff_of_eq (false_iff _) theorem iff_self_iff (a : Prop) : (a ↔ a) ↔ True := iff_of_eq (iff_self _) --- TODO --- attribute [intro] Exists.intro - /- exists unique -/ def ExistsUnique (p : α → Prop) := ∃ x, p x ∧ ∀ y, p y → y = x @@ -218,7 +206,7 @@ def recOn_false [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} alias by_cases := byCases alias by_contradiction := byContradiction -alias not_not_iff := not_not +@[deprecated (since := "2024-07-27")] alias not_not_iff := not_not end Decidable diff --git a/Mathlib/Init/Quot.lean b/Mathlib/Init/Quot.lean index ba61a53e97488..af01552219ceb 100644 --- a/Mathlib/Init/Quot.lean +++ b/Mathlib/Init/Quot.lean @@ -27,7 +27,6 @@ protected abbrev Quot.recOn' (h : (a b : α) → (p : r a b) → Eq.ndrec (f a) (Quot.sound p) = f b) : motive q := q.rec f h --- expected token /-- Version of `Quot.recOnSubsingleton` tagged with `elab_as_elim` -/ @[elab_as_elim] -- Porting note: this attribute is missing in core diff --git a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean index f0f90fc28bd12..5be189e9b0611 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean @@ -3,7 +3,6 @@ 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.Init.Core import Mathlib.LinearAlgebra.AffineSpace.Basis import Mathlib.LinearAlgebra.FiniteDimensional diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 88fe378a09004..40d2b6b3b4dc5 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -37,7 +37,10 @@ section Miscellany -- And.decidable Or.decidable Decidable.false Xor.decidable Iff.decidable Decidable.true -- Implies.decidable Not.decidable Ne.decidable Bool.decidableEq Decidable.toBool -attribute [simp] cast_eq cast_heq imp_false +attribute [simp] cast_heq + +-- This can be removed once we move to Lean v4.11 +attribute [simp] insert_emptyc_eq /-- An identity function with its main argument implicit. This will be printed as `hidden` even if it is applied to a large term, so it can be used for elision, @@ -747,6 +750,32 @@ def choice_of_byContradiction' {α : Sort*} (contra : ¬(α → False) → α) : lemma choose_eq' (a : α) : @Exists.choose _ (a = ·) ⟨a, rfl⟩ = a := (@choose_spec _ (a = ·) _).symm +alias axiom_of_choice := axiomOfChoice -- TODO: remove? rename in core? +alias by_cases := byCases -- TODO: remove? rename in core? +alias by_contradiction := byContradiction -- TODO: remove? rename in core? + +-- The remaining theorems in this section were ported from Lean 3, +-- but are currently unused in Mathlib, so have been deprecated. +-- If any are being used downstream, please remove the deprecation. + +alias prop_complete := propComplete -- TODO: remove? rename in core? + +@[elab_as_elim, deprecated (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")] +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")] +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")] +theorem cases {p : Prop → Prop} (h1 : p True) (h2 : p False) (a) : p a := cases_on a h1 h2 + end Classical /-- This function has the same type as `Exists.recOn`, and can be used to case on an equality, diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index d87713b0db049..73e695fc742af 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -9,7 +9,6 @@ import Mathlib.Data.Prod.Basic import Mathlib.Data.Sigma.Basic import Mathlib.Data.Subtype import Mathlib.Data.Sum.Basic -import Mathlib.Init.Data.Sigma.Basic import Mathlib.Logic.Equiv.Defs import Mathlib.Logic.Function.Conjugate import Mathlib.Tactic.Coe diff --git a/Mathlib/NumberTheory/Cyclotomic/Basic.lean b/Mathlib/NumberTheory/Cyclotomic/Basic.lean index ed4273a3b101f..891a8f3be52b6 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Basic.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Basic.lean @@ -3,7 +3,6 @@ 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.Init.Core import Mathlib.RingTheory.Polynomial.Cyclotomic.Roots import Mathlib.NumberTheory.NumberField.Basic import Mathlib.FieldTheory.Galois diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index d6a8d14bb8b9e..19128510fa244 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -263,7 +263,7 @@ natural exponents have positive `y`. -/ theorem y_pow_succ_pos {a : Solution₁ d} (hax : 0 < a.x) (hay : 0 < a.y) (n : ℕ) : 0 < (a ^ n.succ).y := by induction' n with n ih - · simp only [Nat.zero_eq, ← Nat.one_eq_succ_zero, hay, pow_one] + · simp only [pow_one, hay] · rw [pow_succ'] exact y_mul_pos hax hay (x_pow_pos hax _) ih diff --git a/Mathlib/Order/Fin/Basic.lean b/Mathlib/Order/Fin/Basic.lean index 5fac9d7e3e64c..3477d9965e7f1 100644 --- a/Mathlib/Order/Fin/Basic.lean +++ b/Mathlib/Order/Fin/Basic.lean @@ -5,6 +5,7 @@ Authors: Robert Y. Lewis, Keeley Hoek -/ import Mathlib.Data.Fin.Basic import Mathlib.Order.Hom.Set +import Mathlib.Init.Data.Nat.Lemmas /-! # `Fin n` forms a bounded linear order diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index 1a1ad02cbedff..accea482ee544 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -244,6 +244,31 @@ lemma WellFounded.prod_lex {ra : α → α → Prop} {rb : β → β → Prop} ( (hb : WellFounded rb) : WellFounded (Prod.Lex ra rb) := (Prod.lex ⟨_, ha⟩ ⟨_, hb⟩).wf +section PSigma + +open PSigma + +/-- The lexicographical order of well-founded relations is well-founded. -/ +theorem WellFounded.psigma_lex + {α : Sort*} {β : α → Sort*} {r : α → α → Prop} {s : ∀ a : α, β a → β a → Prop} + (ha : WellFounded r) (hb : ∀ x, WellFounded (s x)) : WellFounded (Lex r s) := + WellFounded.intro fun ⟨a, b⟩ => lexAccessible (WellFounded.apply ha a) hb b + +theorem WellFounded.psigma_revLex + {α : Sort*} {β : Sort*} {r : α → α → Prop} {s : β → β → Prop} + (ha : WellFounded r) (hb : WellFounded s) : WellFounded (RevLex r s) := + WellFounded.intro fun ⟨a, b⟩ => revLexAccessible (apply hb b) (WellFounded.apply ha) a + +theorem WellFounded.psigma_skipLeft (α : Type u) {β : Type v} {s : β → β → Prop} + (hb : WellFounded s) : WellFounded (SkipLeft α s) := + psigma_revLex emptyWf.wf hb + +@[deprecated (since := "2024-07-24")] alias PSigma.lex_wf := WellFounded.psigma_lex +@[deprecated (since := "2024-07-24")] alias PSigma.revLex_wf := WellFounded.psigma_revLex +@[deprecated (since := "2024-07-24")] alias PSigma.skipLeft_wf := WellFounded.psigma_skipLeft + +end PSigma + namespace IsWellFounded variable (r) [IsWellFounded α r] diff --git a/Mathlib/Order/WellFoundedSet.lean b/Mathlib/Order/WellFoundedSet.lean index fbe6f4b58728d..84f0ca38e5c7b 100644 --- a/Mathlib/Order/WellFoundedSet.lean +++ b/Mathlib/Order/WellFoundedSet.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ -import Mathlib.Init.Data.Sigma.Lex import Mathlib.Data.Prod.Lex import Mathlib.Data.Sigma.Lex import Mathlib.Order.Antichain @@ -828,7 +827,7 @@ we only require it to be well-founded on fibers of `f`. -/ theorem WellFounded.prod_lex_of_wellFoundedOn_fiber (hα : WellFounded (rα on f)) (hβ : ∀ a, (f ⁻¹' {a}).WellFoundedOn (rβ on g)) : WellFounded (Prod.Lex rα rβ on fun c => (f c, g c)) := by - refine ((PSigma.lex_wf (wellFoundedOn_range.2 hα) fun a => hβ a).onFun + refine ((psigma_lex (wellFoundedOn_range.2 hα) fun a => hβ a).onFun (f := fun c => ⟨⟨_, c, rfl⟩, c, rfl⟩)).mono fun c c' h => ?_ obtain h' | h' := Prod.lex_iff.1 h · exact PSigma.Lex.left _ _ h' @@ -853,7 +852,7 @@ require it to be well-founded on fibers of `f`. -/ theorem WellFounded.sigma_lex_of_wellFoundedOn_fiber (hι : WellFounded (rι on f)) (hπ : ∀ i, (f ⁻¹' {i}).WellFoundedOn (rπ i on g i)) : WellFounded (Sigma.Lex rι rπ on fun c => ⟨f c, g (f c) c⟩) := by - refine ((PSigma.lex_wf (wellFoundedOn_range.2 hι) fun a => hπ a).onFun + refine ((psigma_lex (wellFoundedOn_range.2 hι) fun a => hπ a).onFun (f := fun c => ⟨⟨_, c, rfl⟩, c, rfl⟩)).mono fun c c' h => ?_ obtain h' | ⟨h', h''⟩ := Sigma.lex_iff.1 h · exact PSigma.Lex.left _ _ h' diff --git a/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean b/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean index 4e8640d0749be..a15d05c552591 100644 --- a/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean +++ b/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean @@ -3,7 +3,6 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Init.Classical import Mathlib.Order.FixedPoints import Mathlib.Order.Zorn diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 71c72e2dc3afd..9feba54454817 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -892,7 +892,7 @@ theorem repr_opow (o₁ o₂) [NF o₁] [NF o₂] : repr (o₁ ^ o₂) = repr o · cases' e₂ : split' o₂ with b' k cases' nf_repr_split' e₂ with _ r₂ by_cases h : m = 0 - · simp [opow_def, opow, e₁, h, r₁, e₂, r₂, ← Nat.one_eq_succ_zero] + · simp [opow_def, opow, e₁, h, r₁, e₂, r₂] simp only [opow_def, opowAux2, opow, e₁, h, r₁, e₂, r₂, repr, opow_zero, Nat.succPNat_coe, Nat.cast_succ, Nat.cast_zero, _root_.zero_add, mul_one, add_zero, one_opow, npow_eq_pow] diff --git a/Mathlib/Tactic/Group.lean b/Mathlib/Tactic/Group.lean index 1f0943d59fc51..050f79785fc0a 100644 --- a/Mathlib/Tactic/Group.lean +++ b/Mathlib/Tactic/Group.lean @@ -52,7 +52,7 @@ macro_rules [commutatorElement_def, mul_one, one_mul, ← zpow_neg_one, ← zpow_natCast, ← zpow_mul, Int.ofNat_add, Int.ofNat_mul, - Int.mul_neg_eq_neg_mul_symm, Int.neg_mul_eq_neg_mul_symm, neg_neg, + Int.mul_neg, Int.neg_mul, neg_neg, one_zpow, zpow_zero, zpow_one, mul_zpow_neg_one, ← mul_assoc, ← zpow_add, ← zpow_add_one, ← zpow_one_add, zpow_trick, zpow_trick_one, zpow_trick_one', diff --git a/Mathlib/Testing/SlimCheck/Sampleable.lean b/Mathlib/Testing/SlimCheck/Sampleable.lean index 6e6da32e42935..b2c011792955b 100644 --- a/Mathlib/Testing/SlimCheck/Sampleable.lean +++ b/Mathlib/Testing/SlimCheck/Sampleable.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Henrik Böving, Simon Hudon -/ import Mathlib.Algebra.Order.Ring.Int -import Mathlib.Init.Data.List.Instances +import Mathlib.Data.List.Monad import Mathlib.Testing.SlimCheck.Gen /-! diff --git a/test/RewriteSearch/Polynomial.lean b/test/RewriteSearch/Polynomial.lean index 7fb4c512a6d94..2fcec787de1f2 100644 --- a/test/RewriteSearch/Polynomial.lean +++ b/test/RewriteSearch/Polynomial.lean @@ -1,6 +1,5 @@ import Mathlib.Algebra.Polynomial.Eval import Mathlib.Algebra.Polynomial.Inductions -import Mathlib.Init.Core import Mathlib.Tactic.RewriteSearch set_option autoImplicit true From 006e993bcecfe563273d5066c8a4b36f81a6306c Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Thu, 15 Aug 2024 12:19:29 +0000 Subject: [PATCH 291/359] chore(scripts/macos_install.sh): avoid using `brew` completely (#15276) This script is intended as a single step for a beginner user to get started with Lean. As such, we should avoid third-party packages from package managers whose provenance can be dubious. We should also avoid requiring the user to install a lot of software. We have already removed homebrew from the `elan` install. This PR removes it from the VS Code step as well, avoiding the dependency on `brew` completely. We download from the official Microsoft site instead. --- scripts/install_macos.sh | 27 ++++++++++++++------------- 1 file changed, 14 insertions(+), 13 deletions(-) diff --git a/scripts/install_macos.sh b/scripts/install_macos.sh index bb471f1ba4099..ef46de9dc19ff 100755 --- a/scripts/install_macos.sh +++ b/scripts/install_macos.sh @@ -2,23 +2,24 @@ set -exo pipefail -# Install Homebrew -if ! which brew > /dev/null; then - # Install Homebrew - /bin/bash -c "$(curl -fsSL https://raw.githubusercontent.com/Homebrew/install/master/install.sh)" -else - # Update it, in case it has been ages since it's been updated - brew update -fi - +# Install elan using the official script curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh + +# Set the default Lean version to the latest stable release elan toolchain install stable elan default stable -# Install and configure VS Code -if ! which code > /dev/null; then - brew install --cask visual-studio-code -fi +# Install the universal darwin build of VS Code +curl -L https://update.code.visualstudio.com/latest/darwin-universal/stable -o ~/Downloads/VSCode-darwin-universal.zip + +# Unzip the downloaded file to the Applications folder +unzip -o ~/Downloads/VSCode-darwin-universal.zip -d /Applications + +# Add the VS Code binary to the PATH to enable launching from the terminal +cat << EOF >> ~/.zprofile +# Add Visual Studio Code (code) +export PATH="\$PATH:/Applications/Visual Studio Code.app/Contents/Resources/app/bin" +EOF # Install the Lean4 VS Code extension code --install-extension leanprover.lean4 From 74ee7bb7f26508f22fca05060bbc91bfb1b4cd8d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 15 Aug 2024 12:19:30 +0000 Subject: [PATCH 292/359] chore: Deprecate `Real.mul_pos` (#15830) This is a special of the more general `mul_pos` --- Mathlib/Analysis/MeanInequalities.lean | 2 +- .../SpecialFunctions/Log/ENNRealLog.lean | 6 ++--- Mathlib/Data/Real/Basic.lean | 25 +++++++++++-------- .../Measure/LevyProkhorovMetric.lean | 2 +- 4 files changed, 19 insertions(+), 16 deletions(-) diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index 7dc403aec3cca..0f52b6382904c 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -268,7 +268,7 @@ theorem harm_mean_le_geom_mean_weighted (w z : ι → ℝ) (hs : s.Nonempty) (hw have p_pos : 0 < ∏ i in s, (z i)⁻¹ ^ w i := prod_pos fun i hi => rpow_pos_of_pos (inv_pos.2 (hz i hi)) _ have s_pos : 0 < ∑ i in s, w i * (z i)⁻¹ := - sum_pos (fun i hi => Real.mul_pos (hw i hi) (inv_pos.2 (hz i hi))) hs + sum_pos (fun i hi => mul_pos (hw i hi) (inv_pos.2 (hz i hi))) hs norm_num at this rw [← inv_le_inv s_pos p_pos] at this apply le_trans this diff --git a/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLog.lean b/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLog.lean index 18a40a93e64f5..2c0c4d15f5e57 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLog.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/ENNRealLog.lean @@ -163,10 +163,10 @@ theorem log_mul_add {x y : ℝ≥0∞} : log (x * y) = log x + log y := by rcases ENNReal.trichotomy y with (rfl | rfl | y_real) · simp · simp [Ne.symm (ne_of_lt (ENNReal.toReal_pos_iff.1 x_real).1)] - · have xy_real := Real.mul_pos x_real y_real - rw [← ENNReal.toReal_mul] at xy_real - rw_mod_cast [log_pos_real' xy_real, log_pos_real' y_real, ENNReal.toReal_mul] + · rw_mod_cast [log_pos_real', log_pos_real' y_real, ENNReal.toReal_mul] exact Real.log_mul (Ne.symm (ne_of_lt x_real)) (Ne.symm (ne_of_lt y_real)) + rw [toReal_mul] + positivity theorem log_pow {x : ℝ≥0∞} {n : ℕ} : log (x ^ n) = n * log x := by cases' Nat.eq_zero_or_pos n with n_zero n_pos diff --git a/Mathlib/Data/Real/Basic.lean b/Mathlib/Data/Real/Basic.lean index a78b3dc25b856..8de982a5ba82f 100644 --- a/Mathlib/Data/Real/Basic.lean +++ b/Mathlib/Data/Real/Basic.lean @@ -343,22 +343,25 @@ protected theorem zero_lt_one : (0 : ℝ) < 1 := by protected theorem fact_zero_lt_one : Fact ((0 : ℝ) < 1) := ⟨Real.zero_lt_one⟩ +@[deprecated mul_pos (since := "2024-08-15")] protected theorem mul_pos {a b : ℝ} : 0 < a → 0 < b → 0 < a * b := by induction' a using Real.ind_mk with a induction' b using Real.ind_mk with b simpa only [mk_lt, mk_pos, ← mk_mul] using CauSeq.mul_pos -instance : StrictOrderedCommRing ℝ := - { Real.commRing, Real.partialOrder, - Real.semiring with - exists_pair_ne := ⟨0, 1, Real.zero_lt_one.ne⟩ - add_le_add_left := by - simp only [le_iff_eq_or_lt] - rintro a b ⟨rfl, h⟩ - · simp only [lt_self_iff_false, or_false, forall_const] - · exact fun c => Or.inr ((add_lt_add_iff_left c).2 ‹_›) - zero_le_one := le_of_lt Real.zero_lt_one - mul_pos := @Real.mul_pos } +instance instStrictOrderedCommRing : StrictOrderedCommRing ℝ where + __ := Real.commRing + exists_pair_ne := ⟨0, 1, Real.zero_lt_one.ne⟩ + add_le_add_left := by + simp only [le_iff_eq_or_lt] + rintro a b ⟨rfl, h⟩ + · simp only [lt_self_iff_false, or_false, forall_const] + · exact fun c => Or.inr ((add_lt_add_iff_left c).2 ‹_›) + zero_le_one := le_of_lt Real.zero_lt_one + mul_pos a b := by + induction' a using Real.ind_mk with a + induction' b using Real.ind_mk with b + simpa only [mk_lt, mk_pos, ← mk_mul] using CauSeq.mul_pos instance strictOrderedRing : StrictOrderedRing ℝ := inferInstance diff --git a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean index fd452eb706c2f..87d03ba19bd52 100644 --- a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean +++ b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean @@ -412,7 +412,7 @@ lemma LevyProkhorov.continuous_toProbabilityMeasure : have key := (tendsto_integral_meas_thickening_le f (A := Ioc 0 ‖f‖) (by simp) P).comp ε_of_room' have aux : ∀ (z : ℝ), Iio (z + δ/2) ∈ 𝓝 z := fun z ↦ Iio_mem_nhds (by linarith) filter_upwards [key (aux _), ε_of_room <| Iio_mem_nhds <| half_pos <| - Real.mul_pos (inv_pos.mpr norm_f_pos) δ_pos] + mul_pos (inv_pos.mpr norm_f_pos) δ_pos] with n hn hn' simp only [mem_preimage, mem_Iio] at * specialize εs_pos n From 3e38319ac0bfcd1166ca7130f08c1599bd37ab69 Mon Sep 17 00:00:00 2001 From: grunweg Date: Thu, 15 Aug 2024 12:19:31 +0000 Subject: [PATCH 293/359] fix(lint-style.py): do not barf on new copyright exceptions (#15833) --- scripts/lint-style.py | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/lint-style.py b/scripts/lint-style.py index 43e6f34940475..b8b0ed4779a5b 100755 --- a/scripts/lint-style.py +++ b/scripts/lint-style.py @@ -61,7 +61,7 @@ path = ROOT_DIR / filename if errno == "ERR_MOD": exceptions += [(ERR_MOD, path, None)] - elif errno in ["ERR_LIN", "ERR_ADN", "ERR_NUM_LIN"]: + elif errno in ["ERR_COP", "ERR_LIN", "ERR_ADN", "ERR_NUM_LIN"]: pass # maintained by the Lean style linter now else: print(f"Error: unexpected errno in style-exceptions.txt: {errno}") From 277c23ca91cb41878dd5b4b1b5f3b4ddc992becc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 15 Aug 2024 13:31:07 +0000 Subject: [PATCH 294/359] =?UTF-8?q?chore:=20`LinearOrderedSemiring=20?= =?UTF-8?q?=CE=B1=20=E2=86=92=20LinearOrderedCancelAddCommMonoid=20=CE=B1`?= =?UTF-8?q?=20(#15814)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This was accidentally about `LinearOrderedCommSemiring`, probably due to confusion over which operation the `Comm` in the name refers to. --- Mathlib/Algebra/Order/Ring/Defs.lean | 11 +++-------- 1 file changed, 3 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Order/Ring/Defs.lean b/Mathlib/Algebra/Order/Ring/Defs.lean index 61fcecb798dbf..7e1dda2f4617c 100644 --- a/Mathlib/Algebra/Order/Ring/Defs.lean +++ b/Mathlib/Algebra/Order/Ring/Defs.lean @@ -389,16 +389,11 @@ instance (priority := 100) LinearOrderedRing.isDomain : IsDomain α where obtain ha | ha := ha.lt_or_lt exacts [(strictAnti_mul_right ha).injective h, (strictMono_mul_right_of_pos ha).injective h] -end LinearOrderedSemiring - -section LinearOrderedCommSemiring -variable [LinearOrderedCommSemiring α] {a b c d : α} - -- See note [lower instance priority] -instance (priority := 100) LinearOrderedCommSemiring.toLinearOrderedCancelAddCommMonoid : - LinearOrderedCancelAddCommMonoid α where __ := ‹LinearOrderedCommSemiring α› +instance (priority := 100) LinearOrderedSemiring.toLinearOrderedCancelAddCommMonoid : + LinearOrderedCancelAddCommMonoid α where __ := ‹LinearOrderedSemiring α› -end LinearOrderedCommSemiring +end LinearOrderedSemiring section LinearOrderedRing variable [LinearOrderedRing α] {a b c : α} From 665c7ae372ab75794f6651d13bb120c26ab89e95 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 15 Aug 2024 14:22:39 +0000 Subject: [PATCH 295/359] feat(MeasureTheory): method of exhaustion (#14471) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `μ, ν` are two measures with `ν` s-finite, then there exists a set `s` such that `μ` is sigma-finite on `s`, and for all sets `t ⊆ sᶜ`, either `ν t = 0` or `μ t = ∞`. The technique used to obtain this result is what Halmos calls the "method of exhaustion". We use this to prove that if `μ ≪ ν` and `ν` is s-finite, then `μ` is s-finite. This PR also moves the definition of `sigmaFiniteSet` to the new file and refactors it with the new definition `sigmaFiniteSetWRT`. --- Mathlib.lean | 1 + .../Decomposition/Exhaustion.lean | 341 ++++++++++++++++++ .../MeasureTheory/Measure/Typeclasses.lean | 11 + .../MeasureTheory/Measure/WithDensity.lean | 12 + .../Measure/WithDensityFinite.lean | 108 +----- 5 files changed, 373 insertions(+), 100 deletions(-) create mode 100644 Mathlib/MeasureTheory/Decomposition/Exhaustion.lean diff --git a/Mathlib.lean b/Mathlib.lean index bddf041fc866a..c29b4604f3901 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3105,6 +3105,7 @@ import Mathlib.MeasureTheory.Covering.LiminfLimsup import Mathlib.MeasureTheory.Covering.OneDim import Mathlib.MeasureTheory.Covering.Vitali import Mathlib.MeasureTheory.Covering.VitaliFamily +import Mathlib.MeasureTheory.Decomposition.Exhaustion import Mathlib.MeasureTheory.Decomposition.Jordan import Mathlib.MeasureTheory.Decomposition.Lebesgue import Mathlib.MeasureTheory.Decomposition.RadonNikodym diff --git a/Mathlib/MeasureTheory/Decomposition/Exhaustion.lean b/Mathlib/MeasureTheory/Decomposition/Exhaustion.lean new file mode 100644 index 0000000000000..e58004315b7e8 --- /dev/null +++ b/Mathlib/MeasureTheory/Decomposition/Exhaustion.lean @@ -0,0 +1,341 @@ +/- +Copyright (c) 2024 Rémy Degenne. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémy Degenne +-/ +import Mathlib.MeasureTheory.Measure.Typeclasses + +/-! +# Method of exhaustion + +If `μ, ν` are two measures with `ν` s-finite, then there exists a set `s` such that +`μ` is sigma-finite on `s`, and for all sets `t ⊆ sᶜ`, either `ν t = 0` or `μ t = ∞`. + +## Main definitions + +* `MeasureTheory.Measure.sigmaFiniteSetWRT`: if such a set exists, `μ.sigmaFiniteSetWRT ν` is + a measurable set such that `μ.restrict (μ.sigmaFiniteSetWRT ν)` is sigma-finite and + for all sets `t ⊆ (μ.sigmaFiniteSetWRT ν)ᶜ`, either `ν t = 0` or `μ t = ∞`. + If no such set exists (which is only possible if `ν` is not s-finite), we define + `μ.sigmaFiniteSetWRT ν = ∅`. +* `MeasureTheory.Measure.sigmaFiniteSet`: for an s-finite measure `μ`, a measurable set such that + `μ.restrict μ.sigmaFiniteSet` is sigma-finite, and for all sets `s ⊆ μ.sigmaFiniteSetᶜ`, + either `μ s = 0` or `μ s = ∞`. + Defined as `μ.sigmaFiniteSetWRT μ`. + +## Main statements + +* `measure_eq_top_of_subset_compl_sigmaFiniteSetWRT`: for s-finite `ν`, for all sets `s` + in `(sigmaFiniteSetWRT μ ν)ᶜ`, if `ν s ≠ 0` then `μ s = ∞`. +* An instance showing that `μ.restrict (sigmaFiniteSetWRT μ ν)` is sigma-finite. +* `restrict_compl_sigmaFiniteSetWRT`: if `μ ≪ ν` and `ν` is s-finite, then + `μ.restrict (μ.sigmaFiniteSetWRT ν)ᶜ = ∞ • ν.restrict (μ.sigmaFiniteSetWRT ν)ᶜ`. As a consequence, + that restriction is s-finite. + +* An instance showing that `μ.restrict μ.sigmaFiniteSet` is sigma-finite. +* `restrict_compl_sigmaFiniteSet_eq_zero_or_top`: the measure `μ.restrict μ.sigmaFiniteSetᶜ` takes + only two values: 0 and ∞ . +* `measure_compl_sigmaFiniteSet_eq_zero_iff_sigmaFinite`: a measure `μ` is sigma-finite + iff `μ μ.sigmaFiniteSetᶜ = 0`. + +## References + +* [P. R. Halmos, *Measure theory*, 17.3 and 30.11][halmos1950measure] + +-/ + +open scoped ENNReal Topology + +open Filter + +namespace MeasureTheory + +variable {α : Type*} {mα : MeasurableSpace α} {μ ν : Measure α} {s t : Set α} + +open Classical in +/-- A measurable set such that `μ.restrict (μ.sigmaFiniteSetWRT ν)` is sigma-finite and for all +measurable sets `t ⊆ sᶜ`, either `ν t = 0` or `μ t = ∞`. -/ +def Measure.sigmaFiniteSetWRT (μ ν : Measure α) : Set α := + if h : ∃ s : Set α, MeasurableSet s ∧ SigmaFinite (μ.restrict s) + ∧ (∀ t, t ⊆ sᶜ → ν t ≠ 0 → μ t = ∞) + then h.choose + else ∅ + +@[measurability] +lemma measurableSet_sigmaFiniteSetWRT : + MeasurableSet (μ.sigmaFiniteSetWRT ν) := by + rw [Measure.sigmaFiniteSetWRT] + split_ifs with h + · exact h.choose_spec.1 + · exact MeasurableSet.empty + +instance : SigmaFinite (μ.restrict (μ.sigmaFiniteSetWRT ν)) := by + rw [Measure.sigmaFiniteSetWRT] + split_ifs with h + · exact h.choose_spec.2.1 + · rw [Measure.restrict_empty] + infer_instance + +section IsFiniteMeasure + +/-! We prove that the condition in the definition of `sigmaFiniteSetWRT` is true for finite +measures. Since every s-finite measure is absolutely continuous with respect to a finite measure, +the condition will then also be true for s-finite measures. -/ + +/-- Let `C` be the supremum of `ν s` over all measurable sets `s` such that `μ.restrict s` is +sigma-finite. `C` is finite since `ν` is a finite measure. Then there exists a measurable set `t` +with `μ.restrict t` sigma-finite such that `ν t ≥ C - 1/n`. -/ +lemma exists_isSigmaFiniteSet_measure_ge (μ ν : Measure α) [IsFiniteMeasure ν] (n : ℕ) : + ∃ t, MeasurableSet t ∧ SigmaFinite (μ.restrict t) + ∧ (⨆ (s) (_ : MeasurableSet s) (_ : SigmaFinite (μ.restrict s)), ν s) - 1/n ≤ ν t := by + by_cases hC_lt : 1/n < ⨆ (s) (_ : MeasurableSet s) (_ : SigmaFinite (μ.restrict s)), ν s + · have h_lt_top : ⨆ (s) (_ : MeasurableSet s) (_ : SigmaFinite (μ.restrict s)), ν s < ∞ := by + refine (?_ : ⨆ (s) (_ : MeasurableSet s) + (_ : SigmaFinite (μ.restrict s)), ν s ≤ ν Set.univ).trans_lt (measure_lt_top _ _) + refine iSup_le (fun s ↦ ?_) + exact iSup_le (fun _ ↦ iSup_le (fun _ ↦ measure_mono (Set.subset_univ s))) + obtain ⟨t, ht⟩ := exists_lt_of_lt_ciSup + (ENNReal.sub_lt_self h_lt_top.ne (ne_zero_of_lt hC_lt) (by simp) : + (⨆ (s) (_ : MeasurableSet s) (_ : SigmaFinite (μ.restrict s)), ν s) - 1/n + < ⨆ (s) (_ : MeasurableSet s) (_ : SigmaFinite (μ.restrict s)), ν s) + have ht_meas : MeasurableSet t := by + by_contra h_not_mem + simp only [h_not_mem] at ht + simp at ht + have ht_mem : SigmaFinite (μ.restrict t) := by + by_contra h_not_mem + simp only [h_not_mem] at ht + simp at ht + refine ⟨t, ht_meas, ht_mem, ?_⟩ + simp only [ht_meas, ht_mem, iSup_true] at ht + exact ht.le + · refine ⟨∅, MeasurableSet.empty, by rw [Measure.restrict_empty]; infer_instance, ?_⟩ + rw [tsub_eq_zero_of_le (not_lt.mp hC_lt)] + exact zero_le' + +/-- A measurable set such that `μ.restrict (μ.sigmaFiniteSetGE ν n)` is sigma-finite and +for `C` the supremum of `ν s` over all measurable sets `s` with `μ.restrict s` sigma-finite, +`ν (μ.sigmaFiniteSetGE ν n) ≥ C - 1/n`. -/ +def Measure.sigmaFiniteSetGE (μ ν : Measure α) [IsFiniteMeasure ν] (n : ℕ) : Set α := + (exists_isSigmaFiniteSet_measure_ge μ ν n).choose + +lemma measurableSet_sigmaFiniteSetGE [IsFiniteMeasure ν] (n : ℕ) : + MeasurableSet (μ.sigmaFiniteSetGE ν n) := + (exists_isSigmaFiniteSet_measure_ge μ ν n).choose_spec.1 + +lemma sigmaFinite_restrict_sigmaFiniteSetGE (μ ν : Measure α) [IsFiniteMeasure ν] (n : ℕ) : + SigmaFinite (μ.restrict (μ.sigmaFiniteSetGE ν n)) := + (exists_isSigmaFiniteSet_measure_ge μ ν n).choose_spec.2.1 + +lemma measure_sigmaFiniteSetGE_le (μ ν : Measure α) [IsFiniteMeasure ν] (n : ℕ) : + ν (μ.sigmaFiniteSetGE ν n) + ≤ ⨆ (s) (_ : MeasurableSet s) (_ : SigmaFinite (μ.restrict s)), ν s := by + refine (le_iSup (f := fun s ↦ _) + (sigmaFinite_restrict_sigmaFiniteSetGE μ ν n)).trans ?_ + exact le_iSup₂ (f := fun s _ ↦ ⨆ (_ : SigmaFinite (μ.restrict s)), ν s) (μ.sigmaFiniteSetGE ν n) + (measurableSet_sigmaFiniteSetGE n) + +lemma measure_sigmaFiniteSetGE_ge (μ ν : Measure α) [IsFiniteMeasure ν] (n : ℕ) : + (⨆ (s) (_ : MeasurableSet s) (_ : SigmaFinite (μ.restrict s)), ν s) - 1/n + ≤ ν (μ.sigmaFiniteSetGE ν n) := + (exists_isSigmaFiniteSet_measure_ge μ ν n).choose_spec.2.2 + +lemma tendsto_measure_sigmaFiniteSetGE (μ ν : Measure α) [IsFiniteMeasure ν] : + Tendsto (fun n ↦ ν (μ.sigmaFiniteSetGE ν n)) atTop + (𝓝 (⨆ (s) (_ : MeasurableSet s) (_ : SigmaFinite (μ.restrict s)), ν s)) := by + refine tendsto_of_tendsto_of_tendsto_of_le_of_le ?_ + tendsto_const_nhds (measure_sigmaFiniteSetGE_ge μ ν) (measure_sigmaFiniteSetGE_le μ ν) + nth_rewrite 2 [← tsub_zero (⨆ (s) (_ : MeasurableSet s) (_ : SigmaFinite (μ.restrict s)), ν s)] + refine ENNReal.Tendsto.sub tendsto_const_nhds ?_ (Or.inr ENNReal.zero_ne_top) + simp only [one_div] + exact ENNReal.tendsto_inv_nat_nhds_zero + +/-- A measurable set such that `μ.restrict (μ.sigmaFiniteSetWRT' ν)` is sigma-finite and +`ν (μ.sigmaFiniteSetWRT' ν)` has maximal measure among such sets. -/ +def Measure.sigmaFiniteSetWRT' (μ ν : Measure α) [IsFiniteMeasure ν] : Set α := + ⋃ n, μ.sigmaFiniteSetGE ν n + +lemma measurableSet_sigmaFiniteSetWRT' [IsFiniteMeasure ν] : + MeasurableSet (μ.sigmaFiniteSetWRT' ν) := + MeasurableSet.iUnion measurableSet_sigmaFiniteSetGE + +lemma sigmaFinite_restrict_sigmaFiniteSetWRT' (μ ν : Measure α) [IsFiniteMeasure ν] : + SigmaFinite (μ.restrict (μ.sigmaFiniteSetWRT' ν)) := by + have := sigmaFinite_restrict_sigmaFiniteSetGE μ ν + let f : ℕ × ℕ → Set α := fun p : ℕ × ℕ ↦ (μ.sigmaFiniteSetWRT' ν)ᶜ + ∪ (spanningSets (μ.restrict (μ.sigmaFiniteSetGE ν p.1)) p.2 ∩ (μ.sigmaFiniteSetGE ν p.1)) + suffices (μ.restrict (μ.sigmaFiniteSetWRT' ν)).FiniteSpanningSetsIn (Set.range f) from + this.sigmaFinite + let e : ℕ ≃ ℕ × ℕ := Nat.pairEquiv.symm + refine ⟨fun n ↦ f (e n), fun _ ↦ by simp, fun n ↦ ?_, ?_⟩ + · simp only [Nat.pairEquiv_symm_apply, gt_iff_lt, measure_union_lt_top_iff, f, e] + rw [Measure.restrict_apply' measurableSet_sigmaFiniteSetWRT', Set.compl_inter_self, + Measure.restrict_apply' measurableSet_sigmaFiniteSetWRT'] + simp only [measure_empty, ENNReal.zero_lt_top, true_and] + refine (measure_mono Set.inter_subset_left).trans_lt ?_ + rw [← Measure.restrict_apply' (measurableSet_sigmaFiniteSetGE _)] + exact measure_spanningSets_lt_top _ _ + · simp only [Nat.pairEquiv_symm_apply, f, e] + rw [← Set.union_iUnion] + suffices ⋃ n, (spanningSets (μ.restrict (μ.sigmaFiniteSetGE ν (Nat.unpair n).1)) n.unpair.2 + ∩ μ.sigmaFiniteSetGE ν n.unpair.1) = μ.sigmaFiniteSetWRT' ν by + rw [this, Set.compl_union_self] + calc ⋃ n, (spanningSets (μ.restrict (μ.sigmaFiniteSetGE ν (Nat.unpair n).1)) n.unpair.2 + ∩ μ.sigmaFiniteSetGE ν n.unpair.1) + = ⋃ n, ⋃ m, (spanningSets (μ.restrict (μ.sigmaFiniteSetGE ν n)) m + ∩ μ.sigmaFiniteSetGE ν n) := + Set.iUnion_unpair (fun n m ↦ spanningSets (μ.restrict (μ.sigmaFiniteSetGE ν n)) m + ∩ μ.sigmaFiniteSetGE ν n) + _ = ⋃ n, μ.sigmaFiniteSetGE ν n := by + refine Set.iUnion_congr (fun n ↦ ?_) + rw [← Set.iUnion_inter, iUnion_spanningSets, Set.univ_inter] + _ = μ.sigmaFiniteSetWRT' ν := rfl + +/-- `μ.sigmaFiniteSetWRT' ν` has maximal `ν`-measure among all measurable sets `s` with sigma-finite +`μ.restrict s`. -/ +lemma measure_sigmaFiniteSetWRT' (μ ν : Measure α) [IsFiniteMeasure ν] : + ν (μ.sigmaFiniteSetWRT' ν) + = ⨆ (s) (_ : MeasurableSet s) (_ : SigmaFinite (μ.restrict s)), ν s := by + apply le_antisymm + · refine (le_iSup (f := fun _ ↦ _) + (sigmaFinite_restrict_sigmaFiniteSetWRT' μ ν)).trans ?_ + exact le_iSup₂ (f := fun s _ ↦ ⨆ (_ : SigmaFinite (μ.restrict s)), ν s) (μ.sigmaFiniteSetWRT' ν) + measurableSet_sigmaFiniteSetWRT' + · exact le_of_tendsto' (tendsto_measure_sigmaFiniteSetGE μ ν) + (fun _ ↦ measure_mono (Set.subset_iUnion _ _)) + +/-- Auxiliary lemma for `measure_eq_top_of_subset_compl_sigmaFiniteSetWRT'`. -/ +lemma measure_eq_top_of_subset_compl_sigmaFiniteSetWRT'_of_measurableSet [IsFiniteMeasure ν] + (hs : MeasurableSet s) (hs_subset : s ⊆ (μ.sigmaFiniteSetWRT' ν)ᶜ) (hνs : ν s ≠ 0) : + μ s = ∞ := by + suffices ¬ SigmaFinite (μ.restrict s) by + by_contra h + have h_lt_top : Fact (μ s < ∞) := ⟨Ne.lt_top h⟩ + exact this inferInstance + intro hsσ + have h_lt : ν (μ.sigmaFiniteSetWRT' ν) < ν (μ.sigmaFiniteSetWRT' ν ∪ s) := by + rw [measure_union _ hs] + · exact ENNReal.lt_add_right (measure_ne_top _ _) hνs + · exact disjoint_compl_right.mono_right hs_subset + have h_le : ν (μ.sigmaFiniteSetWRT' ν ∪ s) ≤ ν (μ.sigmaFiniteSetWRT' ν) := by + conv_rhs => rw [measure_sigmaFiniteSetWRT'] + refine (le_iSup + (f := fun (_ : SigmaFinite (μ.restrict (μ.sigmaFiniteSetWRT' ν ∪ s))) ↦ _) ?_).trans ?_ + · have := sigmaFinite_restrict_sigmaFiniteSetWRT' μ ν + infer_instance + · exact le_iSup₂ (f := fun s _ ↦ ⨆ (_ : SigmaFinite (μ.restrict _)), ν s) + (μ.sigmaFiniteSetWRT' ν ∪ s) (measurableSet_sigmaFiniteSetWRT'.union hs) + exact h_lt.not_le h_le + +/-- For all sets `s` in `(μ.sigmaFiniteSetWRT ν)ᶜ`, if `ν s ≠ 0` then `μ s = ∞`. -/ +lemma measure_eq_top_of_subset_compl_sigmaFiniteSetWRT' [IsFiniteMeasure ν] + (hs_subset : s ⊆ (μ.sigmaFiniteSetWRT' ν)ᶜ) (hνs : ν s ≠ 0) : + μ s = ∞ := by + rw [measure_eq_iInf] + simp_rw [iInf_eq_top] + suffices ∀ t, t ⊆ (μ.sigmaFiniteSetWRT' ν)ᶜ → s ⊆ t → MeasurableSet t → μ t = ∞ by + intro t hts ht + suffices μ (t ∩ (μ.sigmaFiniteSetWRT' ν)ᶜ) = ∞ from + measure_mono_top Set.inter_subset_left this + have hs_subset_t : s ⊆ t ∩ (μ.sigmaFiniteSetWRT' ν)ᶜ := Set.subset_inter hts hs_subset + exact this (t ∩ (μ.sigmaFiniteSetWRT' ν)ᶜ) Set.inter_subset_right hs_subset_t + (ht.inter measurableSet_sigmaFiniteSetWRT'.compl) + intro t ht_subset hst ht + refine measure_eq_top_of_subset_compl_sigmaFiniteSetWRT'_of_measurableSet ht ht_subset ?_ + exact fun hνt ↦ hνs (measure_mono_null hst hνt) + +end IsFiniteMeasure + +section SFinite + +/-- For all sets `s` in `(μ.sigmaFiniteSetWRT ν)ᶜ`, if `ν s ≠ 0` then `μ s = ∞`. -/ +lemma measure_eq_top_of_subset_compl_sigmaFiniteSetWRT [SFinite ν] + (hs_subset : s ⊆ (μ.sigmaFiniteSetWRT ν)ᶜ) (hνs : ν s ≠ 0) : + μ s = ∞ := by + have ⟨ν', hν', hνν'⟩ := exists_absolutelyContinuous_isFiniteMeasure ν + have h : ∃ s : Set α, MeasurableSet s ∧ SigmaFinite (μ.restrict s) + ∧ (∀ t (_ : t ⊆ sᶜ) (_ : ν t ≠ 0), μ t = ∞) := by + refine ⟨μ.sigmaFiniteSetWRT' ν', measurableSet_sigmaFiniteSetWRT', + sigmaFinite_restrict_sigmaFiniteSetWRT' _ _, + fun t ht_subset hνt ↦ measure_eq_top_of_subset_compl_sigmaFiniteSetWRT' ht_subset ?_⟩ + exact fun hν't ↦ hνt (hνν' hν't) + rw [Measure.sigmaFiniteSetWRT, dif_pos h] at hs_subset + exact h.choose_spec.2.2 s hs_subset hνs + +lemma restrict_compl_sigmaFiniteSetWRT [SFinite ν] (hμν : μ ≪ ν) : + μ.restrict (μ.sigmaFiniteSetWRT ν)ᶜ = ∞ • ν.restrict (μ.sigmaFiniteSetWRT ν)ᶜ := by + ext s + rw [Measure.restrict_apply' measurableSet_sigmaFiniteSetWRT.compl, + Measure.smul_apply, smul_eq_mul, + Measure.restrict_apply' measurableSet_sigmaFiniteSetWRT.compl] + by_cases hνs : ν (s ∩ (μ.sigmaFiniteSetWRT ν)ᶜ) = 0 + · rw [hνs, mul_zero] + exact hμν hνs + · rw [ENNReal.top_mul hνs, measure_eq_top_of_subset_compl_sigmaFiniteSetWRT + Set.inter_subset_right hνs] + +end SFinite + +@[simp] +lemma measure_compl_sigmaFiniteSetWRT (hμν : μ ≪ ν) [SigmaFinite μ] [SFinite ν] : + ν (μ.sigmaFiniteSetWRT ν)ᶜ = 0 := by + have h : ν (μ.sigmaFiniteSetWRT ν)ᶜ ≠ 0 → μ (μ.sigmaFiniteSetWRT ν)ᶜ = ∞ := + measure_eq_top_of_subset_compl_sigmaFiniteSetWRT subset_rfl + by_contra h0 + refine ENNReal.top_ne_zero ?_ + rw [← h h0, ← Measure.iSup_restrict_spanningSets] + simp_rw [Measure.restrict_apply' (measurable_spanningSets μ _), ENNReal.iSup_eq_zero] + intro i + by_contra h_ne_zero + have h_zero_top := measure_eq_top_of_subset_compl_sigmaFiniteSetWRT + (Set.inter_subset_left : (μ.sigmaFiniteSetWRT ν)ᶜ ∩ spanningSets μ i ⊆ _) ?_ + swap; · exact fun h ↦ h_ne_zero (hμν h) + refine absurd h_zero_top (ne_of_lt ?_) + exact (measure_mono Set.inter_subset_right).trans_lt (measure_spanningSets_lt_top μ i) + +section SigmaFiniteSet + +/-- A measurable set such that `μ.restrict μ.sigmaFiniteSet` is sigma-finite, + and for all measurable sets `s ⊆ μ.sigmaFiniteSetᶜ`, either `μ s = 0` or `μ s = ∞`. -/ +def Measure.sigmaFiniteSet (μ : Measure α) : Set α := μ.sigmaFiniteSetWRT μ + +@[measurability] +lemma measurableSet_sigmaFiniteSet : MeasurableSet μ.sigmaFiniteSet := + measurableSet_sigmaFiniteSetWRT + +lemma measure_eq_zero_or_top_of_subset_compl_sigmaFiniteSet [SFinite μ] + (ht_subset : t ⊆ μ.sigmaFiniteSetᶜ) : + μ t = 0 ∨ μ t = ∞ := by + by_cases h0 : μ t = 0 + · exact Or.inl h0 + · exact Or.inr <| measure_eq_top_of_subset_compl_sigmaFiniteSetWRT ht_subset h0 + +/-- The measure `μ.restrict μ.sigmaFiniteSetᶜ` takes only two values: 0 and ∞ . -/ +lemma restrict_compl_sigmaFiniteSet_eq_zero_or_top (μ : Measure α) [SFinite μ] (s : Set α) : + μ.restrict μ.sigmaFiniteSetᶜ s = 0 ∨ μ.restrict μ.sigmaFiniteSetᶜ s = ∞ := by + rw [Measure.restrict_apply' measurableSet_sigmaFiniteSet.compl] + exact measure_eq_zero_or_top_of_subset_compl_sigmaFiniteSet Set.inter_subset_right + +/-- The restriction of an s-finite measure `μ` to `μ.sigmaFiniteSet` is sigma-finite. -/ +instance : SigmaFinite (μ.restrict μ.sigmaFiniteSet) := by + rw [Measure.sigmaFiniteSet] + infer_instance + +lemma sigmaFinite_of_measure_compl_sigmaFiniteSet_eq_zero (h : μ μ.sigmaFiniteSetᶜ = 0) : + SigmaFinite μ := by + rw [← Measure.restrict_add_restrict_compl (μ := μ) (measurableSet_sigmaFiniteSet (μ := μ)), + Measure.restrict_eq_zero.mpr h, add_zero] + infer_instance + +@[simp] +lemma measure_compl_sigmaFiniteSet (μ : Measure α) [SigmaFinite μ] : μ μ.sigmaFiniteSetᶜ = 0 := + measure_compl_sigmaFiniteSetWRT Measure.AbsolutelyContinuous.rfl + +/-- An s-finite measure `μ` is sigma-finite iff `μ μ.sigmaFiniteSetᶜ = 0`. -/ +lemma measure_compl_sigmaFiniteSet_eq_zero_iff_sigmaFinite (μ : Measure α) : + μ μ.sigmaFiniteSetᶜ = 0 ↔ SigmaFinite μ := + ⟨sigmaFinite_of_measure_compl_sigmaFiniteSet_eq_zero, fun _ ↦ measure_compl_sigmaFiniteSet μ⟩ + +end SigmaFiniteSet + +end MeasureTheory diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index 0777f34ae5b0a..b76803bde0962 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -1036,6 +1036,17 @@ instance SMul.sigmaFinite {μ : Measure α} [SigmaFinite μ] (c : ℝ≥0) : exact ENNReal.mul_lt_top ENNReal.coe_ne_top (measure_spanningSets_lt_top μ i).ne spanning := iUnion_spanningSets μ }⟩ +instance [SigmaFinite (μ.restrict s)] [SigmaFinite (μ.restrict t)] : + SigmaFinite (μ.restrict (s ∪ t)) := sigmaFinite_of_le _ (restrict_union_le _ _) + +instance [h : SigmaFinite (μ.restrict s)] : SigmaFinite (μ.restrict (s ∩ t)) := by + convert sigmaFinite_of_le _ (restrict_mono_ae (ae_of_all _ Set.inter_subset_left)) + exact h + +instance [h : SigmaFinite (μ.restrict t)] : SigmaFinite (μ.restrict (s ∩ t)) := by + convert sigmaFinite_of_le _ (restrict_mono_ae (ae_of_all _ Set.inter_subset_right)) + exact h + theorem SigmaFinite.of_map (μ : Measure α) {f : α → β} (hf : AEMeasurable f μ) (h : SigmaFinite (μ.map f)) : SigmaFinite μ := ⟨⟨⟨fun n => f ⁻¹' spanningSets (μ.map f) n, fun _ => trivial, fun n => by diff --git a/Mathlib/MeasureTheory/Measure/WithDensity.lean b/Mathlib/MeasureTheory/Measure/WithDensity.lean index 4697f0f82f956..2822135c290aa 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensity.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensity.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Johannes Hölzl -/ +import Mathlib.MeasureTheory.Decomposition.Exhaustion import Mathlib.MeasureTheory.Integral.Lebesgue /-! @@ -649,6 +650,17 @@ lemma sFinite_withDensity_of_measurable (μ : Measure α) [SFinite μ] SFinite (μ.withDensity f) := inferInstance +instance [SFinite μ] (c : ℝ≥0∞) : SFinite (c • μ) := by + rw [← withDensity_const] + infer_instance + +/-- If `μ ≪ ν` and `ν` is s-finite, then `μ` is s-finite. -/ +theorem sFinite_of_absolutelyContinuous {ν : Measure α} [SFinite ν] (hμν : μ ≪ ν) : + SFinite μ := by + rw [← Measure.restrict_add_restrict_compl (μ := μ) measurableSet_sigmaFiniteSetWRT, + restrict_compl_sigmaFiniteSetWRT hμν] + infer_instance + end SFinite variable [TopologicalSpace α] [OpensMeasurableSpace α] [IsLocallyFiniteMeasure μ] diff --git a/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean b/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean index 5452770ee04c0..955c03a3c1454 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean @@ -28,8 +28,6 @@ In these definitions and the results below, `μ` is an s-finite measure (`SFinit If `μ ≠ 0`, this is a probability measure. * `MeasureTheory.Measure.densityToFinite`: a measurable function such that `μ = μ.toFinite.withDensity μ.densityToFinite`. -* `MeasureTheory.Measure.sigmaFiniteSet`: a measurable set such that `μ.restrict μ.sigmaFiniteSet` - is sigma-finite, and for all sets `s ⊆ μ.sigmaFiniteSetᶜ`, either `μ s = 0` or `μ s = ∞`. ## Main statements @@ -37,12 +35,6 @@ In these definitions and the results below, `μ` is an s-finite measure (`SFinit * `toFinite_absolutelyContinuous`: `μ.toFinite ≪ μ`. * `withDensity_densitytoFinite`: `μ.toFinite.withDensity μ.densityToFinite = μ`. -* An instance showing that `μ.restrict μ.sigmaFiniteSet` is sigma-finite. -* `restrict_compl_sigmaFiniteSet_eq_zero_or_top`: the measure `μ.restrict μ.sigmaFiniteSetᶜ` takes - only two values: 0 and ∞ . -* `measure_compl_sigmaFiniteSet_eq_zero_iff_sigmaFinite`: an s-finite measure `μ` is sigma-finite - iff `μ μ.sigmaFiniteSetᶜ = 0`. - -/ open scoped ENNReal @@ -180,100 +172,16 @@ lemma densityToFinite_ae_ne_top (μ : Measure α) [SigmaFinite μ] : ∀ᵐ x ∂μ, μ.densityToFinite x ≠ ∞ := (densityToFinite_ae_lt_top μ).mono (fun _ hx ↦ hx.ne) -section SigmaFiniteSet - -variable {s t : Set α} - -/-- A measurable set such that `μ.restrict μ.sigmaFiniteSet` is sigma-finite, - and for all measurable sets `s ⊆ μ.sigmaFiniteSetᶜ`, either `μ s = 0` or `μ s = ∞`. -/ -def Measure.sigmaFiniteSet (μ : Measure α) [SFinite μ] : Set α := {x | μ.densityToFinite x ≠ ∞} - -lemma measurableSet_sigmaFiniteSet (μ : Measure α) [SFinite μ] : - MeasurableSet μ.sigmaFiniteSet := - (measurable_densityToFinite μ (measurableSet_singleton ∞)).compl - lemma restrict_compl_sigmaFiniteSet [SFinite μ] : μ.restrict μ.sigmaFiniteSetᶜ = ∞ • μ.toFinite.restrict μ.sigmaFiniteSetᶜ := by + rw [Measure.sigmaFiniteSet, + restrict_compl_sigmaFiniteSetWRT (Measure.AbsolutelyContinuous.refl μ)] ext t ht - have : μ.restrict μ.sigmaFiniteSetᶜ - = (μ.toFinite.withDensity μ.densityToFinite).restrict μ.sigmaFiniteSetᶜ := by - rw [withDensity_densitytoFinite μ] - simp only [Measure.coe_smul, Pi.smul_apply, smul_eq_mul] - rw [this, Measure.restrict_apply ht, - withDensity_apply _ (ht.inter (measurableSet_sigmaFiniteSet μ).compl), - Measure.restrict_apply ht] - calc ∫⁻ a in t ∩ μ.sigmaFiniteSetᶜ, μ.densityToFinite a ∂μ.toFinite - _ = ∫⁻ _ in t ∩ μ.sigmaFiniteSetᶜ, ∞ ∂μ.toFinite := by - refine setLIntegral_congr_fun (ht.inter (measurableSet_sigmaFiniteSet μ).compl) - (ae_of_all _ (fun x hx ↦ ?_)) - simpa [Measure.sigmaFiniteSet] using ((Set.inter_subset_right) hx) - _ = ∞ * μ.toFinite (t ∩ μ.sigmaFiniteSetᶜ) := by simp - -/-- The measure `μ.restrict μ.sigmaFiniteSetᶜ` takes only two values: 0 and ∞ . -/ -lemma restrict_compl_sigmaFiniteSet_eq_zero_or_top (μ : Measure α) [SFinite μ] (s : Set α) : - μ.restrict μ.sigmaFiniteSetᶜ s = 0 ∨ μ.restrict μ.sigmaFiniteSetᶜ s = ∞ := by - rw [restrict_compl_sigmaFiniteSet] - simp only [Measure.coe_smul, Pi.smul_apply, smul_eq_mul, mul_eq_zero, ENNReal.top_ne_zero, - false_or] - rw [Measure.restrict_apply' (measurableSet_sigmaFiniteSet μ).compl] - by_cases h_zero : μ.toFinite (s ∩ μ.sigmaFiniteSetᶜ) = 0 - · exact Or.inl h_zero - · exact Or.inr (ENNReal.top_mul h_zero) - -lemma measure_eq_zero_or_top_of_subset_compl_sigmaFiniteSet [SFinite μ] - (ht_subset : t ⊆ μ.sigmaFiniteSetᶜ) : - μ t = 0 ∨ μ t = ∞ := by - have : μ t = μ.restrict μ.sigmaFiniteSetᶜ t := by - rw [Measure.restrict_apply' (measurableSet_sigmaFiniteSet μ).compl, - Set.inter_eq_left.mpr ht_subset] - rw [this] - exact restrict_compl_sigmaFiniteSet_eq_zero_or_top μ t - -lemma toFinite_withDensity_restrict_sigmaFiniteSet (μ : Measure α) [SFinite μ] : - (μ.toFinite.withDensity - (fun x ↦ if μ.densityToFinite x ≠ ∞ then μ.densityToFinite x else 1)).restrict - μ.sigmaFiniteSet - = μ.restrict μ.sigmaFiniteSet := by - have : μ.restrict μ.sigmaFiniteSet - = (μ.toFinite.withDensity μ.densityToFinite).restrict μ.sigmaFiniteSet := by - rw [withDensity_densitytoFinite μ] - rw [this] - ext s hs - rw [Measure.restrict_apply hs, Measure.restrict_apply hs, - withDensity_apply _ (hs.inter (measurableSet_sigmaFiniteSet μ)), - withDensity_apply _ (hs.inter (measurableSet_sigmaFiniteSet μ))] - refine setLIntegral_congr_fun (hs.inter (measurableSet_sigmaFiniteSet μ)) - (ae_of_all _ (fun x hx ↦ Eq.symm ?_)) - simp only [Measure.sigmaFiniteSet, Set.mem_inter_iff, Set.mem_compl_iff, Set.mem_setOf_eq, - ne_eq] at hx - rw [if_pos hx.2] - -/-- The restriction of an s-finite measure `μ` to `μ.sigmaFiniteSet` is sigma-finite. -/ -instance (μ : Measure α) [SFinite μ] : SigmaFinite (μ.restrict μ.sigmaFiniteSet) := by - rw [← toFinite_withDensity_restrict_sigmaFiniteSet] - have : SigmaFinite (μ.toFinite.withDensity - (fun x ↦ if μ.densityToFinite x ≠ ∞ then μ.densityToFinite x else 1)) := by - refine SigmaFinite.withDensity_of_ne_top ?_ - · refine ae_of_all _ (fun x ↦ ?_) - split_ifs with h <;> simp [h] - exact Restrict.sigmaFinite _ _ - -lemma sigmaFinite_of_measure_compl_sigmaFiniteSet_eq_zero [SFinite μ] - (h : μ μ.sigmaFiniteSetᶜ = 0) : - SigmaFinite μ := by - rw [← Measure.restrict_add_restrict_compl (μ := μ) (measurableSet_sigmaFiniteSet μ), - Measure.restrict_eq_zero.mpr h, add_zero] - infer_instance - -@[simp] -lemma measure_compl_sigmaFiniteSet (μ : Measure α) [SigmaFinite μ] : μ μ.sigmaFiniteSetᶜ = 0 := - densityToFinite_ae_ne_top μ - -/-- An s-finite measure `μ` is sigma-finite iff `μ μ.sigmaFiniteSetᶜ = 0`. -/ -lemma measure_compl_sigmaFiniteSet_eq_zero_iff_sigmaFinite (μ : Measure α) [SFinite μ] : - μ μ.sigmaFiniteSetᶜ = 0 ↔ SigmaFinite μ := - ⟨sigmaFinite_of_measure_compl_sigmaFiniteSet_eq_zero, fun _ ↦ measure_compl_sigmaFiniteSet μ⟩ - -end SigmaFiniteSet + simp only [Measure.smul_apply, smul_eq_mul] + rw [Measure.restrict_apply ht, Measure.restrict_apply ht] + by_cases hμt : μ (t ∩ (μ.sigmaFiniteSetWRT μ)ᶜ) = 0 + · rw [hμt, toFinite_absolutelyContinuous μ hμt] + · rw [ENNReal.top_mul hμt, ENNReal.top_mul] + exact fun h ↦ hμt (absolutelyContinuous_toFinite μ h) end MeasureTheory From 6ac635f534640dbebd165c702118a89d4e98282f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 15 Aug 2024 14:22:40 +0000 Subject: [PATCH 296/359] feat: Add partialSups lemmas (#15291) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Rémy Degenne --- Mathlib/MeasureTheory/SetSemiring.lean | 11 +++++++++++ Mathlib/Order/PartialSups.lean | 18 +++++++++++++++--- 2 files changed, 26 insertions(+), 3 deletions(-) diff --git a/Mathlib/MeasureTheory/SetSemiring.lean b/Mathlib/MeasureTheory/SetSemiring.lean index 4b0455aed5556..c17719e996a4f 100644 --- a/Mathlib/MeasureTheory/SetSemiring.lean +++ b/Mathlib/MeasureTheory/SetSemiring.lean @@ -320,6 +320,17 @@ lemma biInter_mem {ι : Type*} (hC : IsSetRing C) {s : ι → Set α} refine hC.inter_mem hs.1 ?_ exact h (fun n hnS ↦ hs.2 n hnS) +lemma partialSups_mem (hC : IsSetRing C) {s : ℕ → Set α} (hs : ∀ n, s n ∈ C) (n : ℕ) : + partialSups s n ∈ C := by + rw [partialSups_eq_biUnion_range] + exact hC.biUnion_mem _ (fun n _ ↦ hs n) + +lemma disjointed_mem (hC : IsSetRing C) {s : ℕ → Set α} (hs : ∀ n, s n ∈ C) (n : ℕ) : + disjointed s n ∈ C := by + cases n with + | zero => exact hs 0 + | succ n => exact hC.diff_mem (hs n.succ) (hC.partialSups_mem hs n) + end IsSetRing end MeasureTheory diff --git a/Mathlib/Order/PartialSups.lean b/Mathlib/Order/PartialSups.lean index 188df49ffbced..cfb8d50271716 100644 --- a/Mathlib/Order/PartialSups.lean +++ b/Mathlib/Order/PartialSups.lean @@ -94,6 +94,9 @@ theorem Monotone.partialSups_eq {f : ℕ → α} (hf : Monotone f) : (partialSup theorem partialSups_mono : Monotone (partialSups : (ℕ → α) → ℕ →o α) := fun _f _g h _n ↦ partialSups_le_iff.2 fun k hk ↦ (h k).trans (le_partialSups_of_le _ hk) +lemma partialSups_monotone (f : ℕ → α) : Monotone (partialSups f) := + fun n _ hnm ↦ partialSups_le f n _ (fun _ hm'n ↦ le_partialSups_of_le _ (hm'n.trans hnm)) + /-- `partialSups` forms a Galois insertion with the coercion from monotone functions to functions. -/ def partialSups.gi : GaloisInsertion (partialSups : (ℕ → α) → ℕ →o α) (↑) where @@ -157,11 +160,20 @@ end ConditionallyCompleteLattice section CompleteLattice -variable [CompleteLattice α] - -theorem partialSups_eq_biSup (f : ℕ → α) (n : ℕ) : partialSups f n = ⨆ i ≤ n, f i := by +theorem partialSups_eq_biSup [CompleteLattice α] (f : ℕ → α) (n : ℕ) : + partialSups f n = ⨆ i ≤ n, f i := by simpa only [iSup_subtype] using partialSups_eq_ciSup_Iic f n +lemma partialSups_eq_sUnion_image [DecidableEq (Set α)] (s : ℕ → Set α) (n : ℕ) : + partialSups s n = ⋃₀ ↑((Finset.range (n + 1)).image s) := by + ext; simp [partialSups_eq_biSup, Nat.lt_succ_iff] + +lemma partialSups_eq_biUnion_range (s : ℕ → Set α) (n : ℕ) : + partialSups s n = ⋃ i ∈ Finset.range (n + 1), s i := by + ext; simp [partialSups_eq_biSup, Nat.lt_succ] + +variable [CompleteLattice α] + -- Porting note (#10618): simp can prove this @[simp] theorem iSup_partialSups_eq (f : ℕ → α) : ⨆ n, partialSups f n = ⨆ n, f n := ciSup_partialSups_eq <| OrderTop.bddAbove _ From 318082b0bccc3abd9d654496f7b60267f277d5fd Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 15 Aug 2024 14:22:42 +0000 Subject: [PATCH 297/359] fix: remove check for `olean`s in `assert_not_imported` (#15839) See [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/assert_not_imported) for a discussion of the issues arising from looking for `olean`s, rather than `lean`s. As a test, this PR also incorporates some uses of `assert_not_imported` to verify that CI and `!bench` work! --- Mathlib/Algebra/BigOperators/Group/List.lean | 4 +--- Mathlib/Algebra/Field/Defs.lean | 5 ++--- Mathlib/Data/List/Chain.lean | 3 +-- Mathlib/Data/List/GetD.lean | 4 +--- Mathlib/Util/AssertExists.lean | 6 ++---- test/AssertExists.lean | 6 +----- 6 files changed, 8 insertions(+), 20 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 0251d89f06740..b283972626cff 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -21,9 +21,7 @@ of elements of a list and `List.alternatingProd`, `List.alternatingSum`, their a counterparts. -/ --- Make sure we haven't imported `Data.Nat.Order.Basic` -assert_not_exists OrderedSub -assert_not_exists Ring +assert_not_imported Mathlib.Algebra.Order.Group.Nat variable {ι α β M N P G : Type*} diff --git a/Mathlib/Algebra/Field/Defs.lean b/Mathlib/Algebra/Field/Defs.lean index c29f6e33d0c67..4943b798adf4c 100644 --- a/Mathlib/Algebra/Field/Defs.lean +++ b/Mathlib/Algebra/Field/Defs.lean @@ -43,12 +43,11 @@ a `GroupWithZero` lemma instead. field, division ring, skew field, skew-field, skewfield -/ +assert_not_imported Mathlib.Tactic.Common + -- `NeZero` should not be needed in the basic algebraic hierarchy. assert_not_exists NeZero --- Check that we have not imported `Mathlib.Tactic.Common` yet. -assert_not_exists Mathlib.Tactic.scopedNS - assert_not_exists MonoidHom open Function Set diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index e8e436dc37099..a6c2410414b7d 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -18,8 +18,7 @@ A graph-specialized version is in development and will hopefully be added under sometime soon. -/ --- Make sure we haven't imported `Data.Nat.Order.Basic` -assert_not_exists OrderedSub +assert_not_imported Mathlib.Algebra.Order.Group.Nat universe u v diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index 3ec5febd08640..d006f872973fa 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -13,11 +13,9 @@ import Mathlib.Util.AssertExists This file provides theorems for working with the `getD` and `getI` functions. These are used to access an element of a list by numerical index, with a default value as a fallback when the index is out of range. - -/ --- Make sure we haven't imported `Data.Nat.Order.Basic` -assert_not_exists OrderedSub +assert_not_imported Mathlib.Algebra.Order.Group.Nat namespace List diff --git a/Mathlib/Util/AssertExists.lean b/Mathlib/Util/AssertExists.lean index 0d8637f9604d7..8197e8a99b225 100644 --- a/Mathlib/Util/AssertExists.lean +++ b/Mathlib/Util/AssertExists.lean @@ -72,12 +72,10 @@ elab "assert_not_exists " n:ident : command => do /-- `assert_not_imported m₁ m₂ ... mₙ` checks that each one of the modules `m₁ m₂ ... mₙ` is not among the transitive imports of the current file. -It also checks that each one of `m₁ m₂ ... mₙ` is actually the name of an existing module, just -one that is not currently imported! +The command does not currently check whether the modules `m₁ m₂ ... mₙ` actually exist. -/ +-- TODO: make sure that each one of `m₁ m₂ ... mₙ` is the name of an actually existing module! elab "assert_not_imported " ids:ident+ : command => do let mods := (← getEnv).allImportedModuleNames for id in ids do if mods.contains id.getId then logWarningAt id m!"the module '{id}' is (transitively) imported" - if let none ← (← searchPathRef.get).findModuleWithExt "olean" id.getId then - logErrorAt id m!"the module '{id}' does not exist" diff --git a/test/AssertExists.lean b/test/AssertExists.lean index 8e40f4c5aaf7c..72dcbbb7a2de9 100644 --- a/test/AssertExists.lean +++ b/test/AssertExists.lean @@ -1,10 +1,6 @@ import Mathlib.Util.AssertExists -/-- -warning: the module 'Lean.Elab.Command' is (transitively) imported ---- -error: the module 'I_do_not_exist' does not exist --/ +/-- warning: the module 'Lean.Elab.Command' is (transitively) imported -/ #guard_msgs in assert_not_imported Mathlib.Tactic.Common From 17196af28d4d13eeb0c36f959557805fd1e445b5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 15 Aug 2024 15:26:59 +0000 Subject: [PATCH 298/359] feat: symmetric monoidal structure on graded objects (#7389) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this PR, we construct the braiding for the monoidal structure on graded objects (#14457). We show it is symmetric if the original category is symmetric. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + .../GradedObject/Bifunctor.lean | 1 - .../CategoryTheory/GradedObject/Braiding.lean | 180 ++++++++++++++++++ 3 files changed, 181 insertions(+), 1 deletion(-) create mode 100644 Mathlib/CategoryTheory/GradedObject/Braiding.lean diff --git a/Mathlib.lean b/Mathlib.lean index c29b4604f3901..57bac5b7799cf 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1484,6 +1484,7 @@ import Mathlib.CategoryTheory.GlueData import Mathlib.CategoryTheory.GradedObject import Mathlib.CategoryTheory.GradedObject.Associator import Mathlib.CategoryTheory.GradedObject.Bifunctor +import Mathlib.CategoryTheory.GradedObject.Braiding import Mathlib.CategoryTheory.GradedObject.Monoidal import Mathlib.CategoryTheory.GradedObject.Single import Mathlib.CategoryTheory.GradedObject.Trifunctor diff --git a/Mathlib/CategoryTheory/GradedObject/Bifunctor.lean b/Mathlib/CategoryTheory/GradedObject/Bifunctor.lean index 7d642ffec1ca6..735f21eb95ec0 100644 --- a/Mathlib/CategoryTheory/GradedObject/Bifunctor.lean +++ b/Mathlib/CategoryTheory/GradedObject/Bifunctor.lean @@ -46,7 +46,6 @@ variable {I J K : Type*} (p : I × J → K) /-- Given a bifunctor `F : C₁ ⥤ C₂ ⥤ C₃`, graded objects `X : GradedObject I C₁` and `Y : GradedObject J C₂` and a map `p : I × J → K`, this is the `K`-graded object sending `k` to the coproduct of `(F.obj (X i)).obj (Y j)` for `p ⟨i, j⟩ = k`. -/ -@[simp] noncomputable def mapBifunctorMapObj (X : GradedObject I C₁) (Y : GradedObject J C₂) [HasMap (((mapBifunctor F I J).obj X).obj Y) p] : GradedObject K C₃ := (((mapBifunctor F I J).obj X).obj Y).mapObj p diff --git a/Mathlib/CategoryTheory/GradedObject/Braiding.lean b/Mathlib/CategoryTheory/GradedObject/Braiding.lean new file mode 100644 index 0000000000000..303a59d11668e --- /dev/null +++ b/Mathlib/CategoryTheory/GradedObject/Braiding.lean @@ -0,0 +1,180 @@ +/- +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.GradedObject.Monoidal +import Mathlib.CategoryTheory.Monoidal.Braided.Basic +/-! +# The braided and symmetric category structures on graded objects + +In this file, we construct the braiding +`GradedObject.Monoidal.braiding : tensorObj X Y ≅ tensorObj Y X` +for two objects `X` and `Y` in `GradedObject I C`, when `I` is a commutative +additive monoid (and suitable coproducts exist in a braided category `C`). + +When `C` is a braided category and suitable assumptions are made, we obtain the braided category +structure on `GradedObject I C` and show that it is symmetric if `C` is symmetric. + +-/ + +namespace CategoryTheory + +open Category Limits + +variable {I : Type*} [AddCommMonoid I] {C : Type*} [Category C] [MonoidalCategory C] + +namespace GradedObject + +namespace Monoidal + +variable (X Y Z : GradedObject I C) + +section Braided + +variable [BraidedCategory C] + +/-- The braiding `tensorObj X Y ≅ tensorObj Y X` when `X` and `Y` are graded objects +indexed by a commutative additive monoid. -/ +noncomputable def braiding [HasTensor X Y] [HasTensor Y X] : tensorObj X Y ≅ tensorObj Y X where + hom k := tensorObjDesc (fun i j hij => (β_ _ _).hom ≫ + ιTensorObj Y X j i k (by simpa only [add_comm j i] using hij)) + inv k := tensorObjDesc (fun i j hij => (β_ _ _).inv ≫ + ιTensorObj X Y j i k (by simpa only [add_comm j i] using hij)) + +variable {Y Z} in +lemma braiding_naturality_right [HasTensor X Y] [HasTensor Y X] [HasTensor X Z] [HasTensor Z X] + (f : Y ⟶ Z) : + whiskerLeft X f ≫ (braiding X Z).hom = (braiding X Y).hom ≫ whiskerRight f X := by + dsimp [braiding] + aesop_cat + +variable {X Y} in +lemma braiding_naturality_left [HasTensor Y Z] [HasTensor Z Y] [HasTensor X Z] [HasTensor Z X] + (f : X ⟶ Y) : + whiskerRight f Z ≫ (braiding Y Z).hom = (braiding X Z).hom ≫ whiskerLeft Z f := by + dsimp [braiding] + aesop_cat + +lemma hexagon_forward [HasTensor X Y] [HasTensor Y X] [HasTensor Y Z] + [HasTensor Z X] [HasTensor X Z] + [HasTensor (tensorObj X Y) Z] [HasTensor X (tensorObj Y Z)] + [HasTensor (tensorObj Y Z) X] [HasTensor Y (tensorObj Z X)] + [HasTensor (tensorObj Y X) Z] [HasTensor Y (tensorObj X Z)] + [HasGoodTensor₁₂Tensor X Y Z] [HasGoodTensorTensor₂₃ X Y Z] + [HasGoodTensor₁₂Tensor Y Z X] [HasGoodTensorTensor₂₃ Y Z X] + [HasGoodTensor₁₂Tensor Y X Z] [HasGoodTensorTensor₂₃ Y X Z] : + (associator X Y Z).hom ≫ (braiding X (tensorObj Y Z)).hom ≫ (associator Y Z X).hom = + whiskerRight (braiding X Y).hom Z ≫ (associator Y X Z).hom ≫ + whiskerLeft Y (braiding X Z).hom := by + ext k i₁ i₂ i₃ h + dsimp [braiding] + conv_lhs => rw [ιTensorObj₃'_associator_hom_assoc, ιTensorObj₃_eq X Y Z i₁ i₂ i₃ k h _ rfl, + assoc, ι_tensorObjDesc_assoc, assoc, ← MonoidalCategory.id_tensorHom, + BraidedCategory.braiding_naturality_assoc, + BraidedCategory.braiding_tensor_right, assoc, assoc, assoc, assoc, Iso.hom_inv_id_assoc, + MonoidalCategory.tensorHom_id, + ← ιTensorObj₃'_eq_assoc Y Z X i₂ i₃ i₁ k (by rw [add_comm _ i₁, ← add_assoc, h]) _ rfl, + ιTensorObj₃'_associator_hom, Iso.inv_hom_id_assoc] + conv_rhs => rw [ιTensorObj₃'_eq X Y Z i₁ i₂ i₃ k h _ rfl, assoc, ι_tensorHom_assoc, + ← MonoidalCategory.tensorHom_id, + ← MonoidalCategory.tensor_comp_assoc, id_comp, ι_tensorObjDesc, + categoryOfGradedObjects_id, MonoidalCategory.comp_tensor_id, assoc, + MonoidalCategory.tensorHom_id, MonoidalCategory.tensorHom_id, + ← ιTensorObj₃'_eq_assoc Y X Z i₂ i₁ i₃ k + (by rw [add_comm i₂ i₁, h]) (i₁ + i₂) (add_comm i₂ i₁), + ιTensorObj₃'_associator_hom_assoc, + ιTensorObj₃_eq Y X Z i₂ i₁ i₃ k (by rw [add_comm i₂ i₁, h]) _ rfl, assoc, + ι_tensorHom, categoryOfGradedObjects_id, ← MonoidalCategory.tensorHom_id, + ← MonoidalCategory.id_tensorHom, + ← MonoidalCategory.id_tensor_comp_assoc, + ι_tensorObjDesc, MonoidalCategory.id_tensor_comp, assoc, + ← MonoidalCategory.id_tensor_comp_assoc, MonoidalCategory.tensorHom_id, + MonoidalCategory.id_tensorHom, MonoidalCategory.whiskerLeft_comp, assoc, + ← ιTensorObj₃_eq Y Z X i₂ i₃ i₁ k (by rw [add_comm _ i₁, ← add_assoc, h]) + (i₁ + i₃) (add_comm _ _ )] + +lemma hexagon_reverse [HasTensor X Y] [HasTensor Y Z] [HasTensor Z X] + [HasTensor Z Y] [HasTensor X Z] + [HasTensor (tensorObj X Y) Z] [HasTensor X (tensorObj Y Z)] + [HasTensor Z (tensorObj X Y)] [HasTensor (tensorObj Z X) Y] + [HasTensor X (tensorObj Z Y)] [HasTensor (tensorObj X Z) Y] + [HasGoodTensor₁₂Tensor X Y Z] [HasGoodTensorTensor₂₃ X Y Z] + [HasGoodTensor₁₂Tensor Z X Y] [HasGoodTensorTensor₂₃ Z X Y] + [HasGoodTensor₁₂Tensor X Z Y] [HasGoodTensorTensor₂₃ X Z Y]: + (associator X Y Z).inv ≫ (braiding (tensorObj X Y) Z).hom ≫ (associator Z X Y).inv = + whiskerLeft X (braiding Y Z).hom ≫ (associator X Z Y).inv ≫ + whiskerRight (braiding X Z).hom Y := by + ext k i₁ i₂ i₃ h + dsimp [braiding] + conv_lhs => rw [ιTensorObj₃_associator_inv_assoc, ιTensorObj₃'_eq X Y Z i₁ i₂ i₃ k h _ rfl, assoc, + ι_tensorObjDesc_assoc, assoc, ← MonoidalCategory.tensorHom_id, + BraidedCategory.braiding_naturality_assoc, + BraidedCategory.braiding_tensor_left, assoc, assoc, assoc, assoc, Iso.inv_hom_id_assoc, + MonoidalCategory.id_tensorHom, + ← ιTensorObj₃_eq_assoc Z X Y i₃ i₁ i₂ k (by rw [add_assoc, add_comm i₃, h]) _ rfl, + ιTensorObj₃_associator_inv, Iso.hom_inv_id_assoc] + conv_rhs => rw [ιTensorObj₃_eq X Y Z i₁ i₂ i₃ k h _ rfl, assoc, ι_tensorHom_assoc, + ← MonoidalCategory.id_tensorHom, + ← MonoidalCategory.tensor_comp_assoc, id_comp, ι_tensorObjDesc, + categoryOfGradedObjects_id, MonoidalCategory.id_tensor_comp, assoc, + MonoidalCategory.id_tensorHom, MonoidalCategory.id_tensorHom, + ← ιTensorObj₃_eq_assoc X Z Y i₁ i₃ i₂ k + (by rw [add_assoc, add_comm i₃, ← add_assoc, h]) (i₂ + i₃) (add_comm _ _), + ιTensorObj₃_associator_inv_assoc, + ιTensorObj₃'_eq X Z Y i₁ i₃ i₂ k (by rw [add_assoc, add_comm i₃, ← add_assoc, h]) _ rfl, + assoc, ι_tensorHom, categoryOfGradedObjects_id, ← MonoidalCategory.tensorHom_id, + ← MonoidalCategory.comp_tensor_id_assoc, + ι_tensorObjDesc, MonoidalCategory.comp_tensor_id, assoc, + MonoidalCategory.tensorHom_id, MonoidalCategory.tensorHom_id, + ← ιTensorObj₃'_eq Z X Y i₃ i₁ i₂ k (by rw [add_assoc, add_comm i₃, h]) + (i₁ + i₃) (add_comm _ _)] + +end Braided + +@[reassoc (attr := simp)] +lemma symmetry [SymmetricCategory C] [HasTensor X Y] [HasTensor Y X] : + (braiding X Y).hom ≫ (braiding Y X).hom = 𝟙 _ := by + dsimp [braiding] + aesop_cat + +end Monoidal + +section Instances + +variable + [∀ (X₁ X₂ : GradedObject I C), HasTensor X₁ X₂] + [∀ (X₁ X₂ X₃ : GradedObject I C), HasGoodTensor₁₂Tensor X₁ X₂ X₃] + [∀ (X₁ X₂ X₃ : GradedObject I C), HasGoodTensorTensor₂₃ X₁ X₂ X₃] + [DecidableEq I] [HasInitial C] + [∀ X₁, PreservesColimit (Functor.empty.{0} C) + ((MonoidalCategory.curriedTensor C).obj X₁)] + [∀ X₂, PreservesColimit (Functor.empty.{0} C) + ((MonoidalCategory.curriedTensor C).flip.obj X₂)] + [∀ (X₁ X₂ X₃ X₄ : GradedObject I C), HasTensor₄ObjExt X₁ X₂ X₃ X₄] + +noncomputable instance braidedCategory [BraidedCategory C] : + BraidedCategory (GradedObject I C) where + braiding X Y := Monoidal.braiding X Y + braiding_naturality_left _ _:= Monoidal.braiding_naturality_left _ _ + braiding_naturality_right _ _ _ _ := Monoidal.braiding_naturality_right _ _ + hexagon_forward _ _ _ := Monoidal.hexagon_forward _ _ _ + hexagon_reverse _ _ _ := Monoidal.hexagon_reverse _ _ _ + +noncomputable instance symmetricCategory [SymmetricCategory C] : + SymmetricCategory (GradedObject I C) where + symmetry _ _ := Monoidal.symmetry _ _ + +/-! +The braided/symmetric monoidal category structure on `GradedObject ℕ C` can +be inferred from the assumptions `[HasFiniteCoproducts C]`, +`[∀ (X : C), PreservesFiniteCoproducts ((curriedTensor C).obj X)]` and +`[∀ (X : C), PreservesFiniteCoproducts ((curriedTensor C).flip.obj X)]`. +This requires importing `Mathlib.CategoryTheory.Limits.Preserves.Finite`. +-/ + +end Instances + +end GradedObject + +end CategoryTheory From 2d973ca2c5ad1dc812ba1548665ea50394240631 Mon Sep 17 00:00:00 2001 From: grunweg Date: Thu, 15 Aug 2024 15:27:00 +0000 Subject: [PATCH 299/359] chore: add docstrings to some tactic items (#15751) Remove a couple of nolint entries. The `cases'` docstring has essentially been copied from the module docstring; the other entries are new. --- Mathlib/Tactic/Basic.lean | 5 ++++ Mathlib/Tactic/Cases.lean | 48 ++++++++++++++++++++++++++++++++++++-- Mathlib/Tactic/Rename.lean | 2 ++ scripts/nolints.json | 5 ---- 4 files changed, 53 insertions(+), 7 deletions(-) diff --git a/Mathlib/Tactic/Basic.lean b/Mathlib/Tactic/Basic.lean index 7ebc0c52d04c7..45d7eec835f00 100644 --- a/Mathlib/Tactic/Basic.lean +++ b/Mathlib/Tactic/Basic.lean @@ -14,6 +14,7 @@ import Mathlib.Tactic.Linter.OldObtain # Basic tactics and utilities for tactic writing This file defines some basic utilities for tactic writing, and also +- a dummy `variables` macro (which warns that the Lean 4 name is `variable`) - the `introv` tactic, which allows the user to automatically introduce the variables of a theorem and explicitly name the non-dependent hypotheses, - an `assumption` macro, calling the `assumption` tactic on all goals @@ -25,8 +26,12 @@ changing them into regular hypotheses). namespace Mathlib.Tactic open Lean Parser.Tactic Elab Command Elab.Tactic Meta +/-- Syntax for the `variables` command: this command is just a stub, +and merely warns that it has been renamed to `variable` in Lean 4. -/ syntax (name := «variables») "variables" (ppSpace bracketedBinder)* : command +/-- The `variables` command: this is just a stub, +and merely warns that it has been renamed to `variable` in Lean 4. -/ @[command_elab «variables»] def elabVariables : CommandElab | `(variables%$pos $binders*) => do logWarningAt pos "'variables' has been replaced by 'variable' in lean 4" diff --git a/Mathlib/Tactic/Cases.lean b/Mathlib/Tactic/Cases.lean index 91ee82d214cb9..3f949e67bc8ef 100644 --- a/Mathlib/Tactic/Cases.lean +++ b/Mathlib/Tactic/Cases.lean @@ -8,10 +8,9 @@ import Batteries.Tactic.OpenPrivate import Mathlib.Lean.Expr.Basic /-! - # Backward compatible implementation of lean 3 `cases` tactic -This tactic is similar to the `cases` tactic in lean 4 core, but the syntax for giving +This tactic is similar to the `cases` tactic in Lean 4 core, but the syntax for giving names is different: ``` @@ -68,6 +67,29 @@ def ElimApp.evalNames (elimInfo : ElimInfo) (alts : Array ElimApp.Alt) (withArg pure subgoals open private getElimNameInfo generalizeTargets generalizeVars from Lean.Elab.Tactic.Induction +/-- The `induction'` tactic is similar to the `induction` tactic in Lean 4 core, +but with slightly different syntax (such as, no requirement to name the constructors). + +``` +open Nat + +example (n : ℕ) : 0 < factorial n := by + induction' n with n ih + · rw [factorial_zero] + simp + · rw [factorial_succ] + apply mul_pos (succ_pos n) ih + +example (n : ℕ) : 0 < factorial n := by + induction n + case zero => + rw [factorial_zero] + simp + case succ n ih => + rw [factorial_succ] + apply mul_pos (succ_pos n) ih +``` + -/ elab (name := induction') "induction' " tgts:(Parser.Tactic.casesTarget,+) usingArg:((" using " ident)?) withArg:((" with" (ppSpace colGt binderIdent)+)?) @@ -101,6 +123,28 @@ elab (name := induction') "induction' " tgts:(Parser.Tactic.casesTarget,+) (generalized := fvarIds) (toClear := targetFVarIds) (toTag := toTag) setGoals <| (subgoals ++ result.others).toList ++ gs +/-- The `cases'` tactic is similar to the `cases` tactic in Lean 4 core, but the syntax for giving +names is different: + +``` +example (h : p ∨ q) : q ∨ p := by + cases h with + | inl hp => exact Or.inr hp + | inr hq => exact Or.inl hq + +example (h : p ∨ q) : q ∨ p := by + cases' h with hp hq + · exact Or.inr hp + · exact Or.inl hq + +example (h : p ∨ q) : q ∨ p := by + rcases h with hp | hq + · exact Or.inr hp + · exact Or.inl hq +``` + +Prefer `cases` or `rcases` when possible, because these tactics promote structured proofs. +-/ elab (name := cases') "cases' " tgts:(Parser.Tactic.casesTarget,+) usingArg:((" using " ident)?) withArg:((" with" (ppSpace colGt binderIdent)+)?) : tactic => do let (targets, toTag) ← elabCasesTargets tgts.1.getSepArgs diff --git a/Mathlib/Tactic/Rename.lean b/Mathlib/Tactic/Rename.lean index f082fe62abf99..5ad80c9d3a68a 100644 --- a/Mathlib/Tactic/Rename.lean +++ b/Mathlib/Tactic/Rename.lean @@ -14,6 +14,8 @@ namespace Mathlib.Tactic open Lean Elab.Tactic Meta +/-- A parser for a single rename to perform in the `rename` tactic: +these should have the form `h => hnew` (describing a rename of a hypothesis `h` to `hnew`). -/ syntax renameArg := term " => " ident /-- `rename' h => hnew` renames the hypothesis named `h` to `hnew`. diff --git a/scripts/nolints.json b/scripts/nolints.json index 6a74c778131bb..c8efbe16155ea 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -528,12 +528,8 @@ ["docBlame", "Mathlib.Notation3.expandFoldl"], ["docBlame", "Mathlib.Notation3.expandFoldr"], ["docBlame", "Mathlib.Notation3.prettyPrintOpt"], - ["docBlame", "Mathlib.Tactic.cases'"], ["docBlame", "Mathlib.Tactic.elabConfig"], - ["docBlame", "Mathlib.Tactic.elabVariables"], ["docBlame", "Mathlib.Tactic.evalIntrov"], - ["docBlame", "Mathlib.Tactic.induction'"], - ["docBlame", "Mathlib.Tactic.renameArg"], ["docBlame", "Mathlib.Tactic.setArgsRest"], ["docBlame", "Mathlib.Tactic.setTactic"], ["docBlame", "Mathlib.Tactic.tacticHave_"], @@ -543,7 +539,6 @@ ["docBlame", "Mathlib.Tactic.tacticSuffices_"], ["docBlame", "Mathlib.Tactic.tacticTransitivity___"], ["docBlame", "Mathlib.Tactic.usingArg"], - ["docBlame", "Mathlib.Tactic.variables"], ["docBlame", "Mathlib.Tactic.withArgs"], ["docBlame", "Mathlib.WhatsNew.diffExtension"], ["docBlame", "Mathlib.WhatsNew.whatsNew"], From 7307bae5ca20422e7d589c4bd419fb99c5fa5853 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 15 Aug 2024 15:27:02 +0000 Subject: [PATCH 300/359] chore: create Mathlib.Init (#15811) This PR creates the `Mathlib.Init` file. This file is intended to be imported by virtually all `Mathlib` files. This PR also imports `Mathlib.Init` in all `Mathlib` files that import no `Mathlib` file. In the long run, some very basic files, some tactics and some linters will be imported in `Mathlib.Init`, so that the commands defined there would be available as widely as possible in `Mathlib`. There is also the file `scripts/init_creation.sh` that contains the instructions to regenerate the `Mathlib/Init.lean` file. --- Mathlib.lean | 1 + Mathlib/Algebra/HierarchyDesign.lean | 1 + Mathlib/CategoryTheory/Category/Init.lean | 1 + .../ConcreteCategory/Bundled.lean | 1 + Mathlib/Combinatorics/SimpleGraph/Init.lean | 1 + Mathlib/Control/Combinators.lean | 1 + Mathlib/Control/ULift.lean | 1 + Mathlib/Data/Array/Defs.lean | 1 + Mathlib/Data/Array/ExtractLemmas.lean | 1 + Mathlib/Data/ByteArray.lean | 1 + Mathlib/Data/Finset/Attr.lean | 1 + Mathlib/Data/Int/Align.lean | 1 + Mathlib/Data/Int/Notation.lean | 1 + Mathlib/Data/List/Monad.lean | 1 + Mathlib/Data/MLList/Dedup.lean | 1 + Mathlib/Data/MLList/IO.lean | 1 + Mathlib/Data/Matroid/Init.lean | 1 + Mathlib/Data/Nat/Notation.lean | 1 + Mathlib/Data/Ordering/Basic.lean | 1 + Mathlib/Data/String/Defs.lean | 1 + Mathlib/Data/Sym/Sym2/Init.lean | 1 + Mathlib/Deprecated/Aliases.lean | 1 + Mathlib/Deprecated/HashMap.lean | 1 + Mathlib/Init.lean | 12 ++++++ Mathlib/Init/Data/Quot.lean | 1 + Mathlib/Init/Quot.lean | 1 + Mathlib/Init/Set.lean | 1 + Mathlib/Lean/Elab/Term.lean | 1 + Mathlib/Lean/EnvExtension.lean | 1 + Mathlib/Lean/Exception.lean | 1 + Mathlib/Lean/Expr/Basic.lean | 1 + Mathlib/Lean/GoalsLocation.lean | 1 + Mathlib/Lean/Json.lean | 1 + Mathlib/Lean/LocalContext.lean | 1 + Mathlib/Lean/Message.lean | 1 + Mathlib/Lean/Meta.lean | 1 + Mathlib/Lean/Meta/Basic.lean | 1 + Mathlib/Lean/Meta/DiscrTree.lean | 1 + Mathlib/Lean/Meta/KAbstractPositions.lean | 1 + Mathlib/Lean/Meta/Simp.lean | 1 + Mathlib/Lean/Name.lean | 1 + Mathlib/Lean/PrettyPrinter/Delaborator.lean | 1 + Mathlib/Lean/Thunk.lean | 1 + Mathlib/Tactic/AdaptationNote.lean | 1 + Mathlib/Tactic/ApplyWith.lean | 1 + Mathlib/Tactic/ArithMult/Init.lean | 1 + Mathlib/Tactic/Attr/Register.lean | 1 + Mathlib/Tactic/Bound/Init.lean | 1 + Mathlib/Tactic/CasesM.lean | 1 + Mathlib/Tactic/Change.lean | 1 + Mathlib/Tactic/Check.lean | 1 + Mathlib/Tactic/Clean.lean | 1 + Mathlib/Tactic/ClearExcept.lean | 1 + Mathlib/Tactic/ClearExclamation.lean | 1 + Mathlib/Tactic/Clear_.lean | 1 + Mathlib/Tactic/Coe.lean | 1 + Mathlib/Tactic/Constructor.lean | 1 + Mathlib/Tactic/Continuity/Init.lean | 1 + Mathlib/Tactic/Conv.lean | 1 + Mathlib/Tactic/Eqns.lean | 1 + Mathlib/Tactic/Eval.lean | 1 + Mathlib/Tactic/ExistsI.lean | 1 + Mathlib/Tactic/Explode/Datatypes.lean | 1 + Mathlib/Tactic/ExtendDoc.lean | 1 + Mathlib/Tactic/ExtractGoal.lean | 1 + Mathlib/Tactic/FailIfNoProgress.lean | 1 + Mathlib/Tactic/Find.lean | 1 + Mathlib/Tactic/FunProp/Decl.lean | 1 + Mathlib/Tactic/FunProp/StateList.lean | 1 + Mathlib/Tactic/FunProp/ToBatteries.lean | 1 + Mathlib/Tactic/GCongr/ForwardAttr.lean | 1 + Mathlib/Tactic/Generalize.lean | 1 + Mathlib/Tactic/GuardGoalNums.lean | 1 + Mathlib/Tactic/GuardHypNums.lean | 1 + Mathlib/Tactic/Have.lean | 1 + Mathlib/Tactic/HaveI.lean | 1 + Mathlib/Tactic/HelpCmd.lean | 1 + Mathlib/Tactic/InferParam.lean | 1 + Mathlib/Tactic/Lemma.lean | 1 + .../Oracle/SimplexAlgorithm/Datatypes.lean | 1 + Mathlib/Tactic/Linter/GlobalAttributeIn.lean | 1 + Mathlib/Tactic/Linter/HashCommandLinter.lean | 1 + Mathlib/Tactic/Linter/HaveLetLinter.lean | 1 + Mathlib/Tactic/Linter/Lint.lean | 1 + Mathlib/Tactic/Linter/OldObtain.lean | 1 + Mathlib/Tactic/Linter/RefineLinter.lean | 1 + Mathlib/Tactic/Linter/Style.lean | 1 + Mathlib/Tactic/Linter/UnusedTactic.lean | 1 + Mathlib/Tactic/Measurability/Init.lean | 1 + Mathlib/Tactic/MinImports.lean | 1 + Mathlib/Tactic/Monotonicity/Attr.lean | 1 + Mathlib/Tactic/NthRewrite.lean | 1 + Mathlib/Tactic/Observe.lean | 1 + Mathlib/Tactic/PPWithUniv.lean | 1 + Mathlib/Tactic/ProjectionNotation.lean | 1 + Mathlib/Tactic/Recall.lean | 1 + Mathlib/Tactic/Recover.lean | 1 + Mathlib/Tactic/ReduceModChar/Ext.lean | 1 + Mathlib/Tactic/Relation/Rfl.lean | 1 + Mathlib/Tactic/Relation/Symm.lean | 1 + Mathlib/Tactic/Rename.lean | 1 + Mathlib/Tactic/Says.lean | 1 + Mathlib/Tactic/Set.lean | 1 + Mathlib/Tactic/SetLike.lean | 1 + Mathlib/Tactic/SimpIntro.lean | 1 + Mathlib/Tactic/SimpRw.lean | 1 + Mathlib/Tactic/Simps/NotationClass.lean | 1 + Mathlib/Tactic/Spread.lean | 1 + Mathlib/Tactic/Substs.lean | 1 + Mathlib/Tactic/SuccessIfFailWithMsg.lean | 1 + Mathlib/Tactic/SudoSetOption.lean | 1 + Mathlib/Tactic/SuppressCompilation.lean | 1 + Mathlib/Tactic/Trace.lean | 1 + Mathlib/Tactic/TryThis.lean | 1 + Mathlib/Tactic/TypeCheck.lean | 1 + Mathlib/Tactic/TypeStar.lean | 1 + Mathlib/Tactic/UnsetOption.lean | 1 + Mathlib/Tactic/Use.lean | 1 + Mathlib/Tactic/Variable.lean | 1 + .../Widget/SelectInsertParamsClass.lean | 1 + Mathlib/Topology/Sheaves/Init.lean | 1 + Mathlib/Util/AddRelatedDecl.lean | 1 + Mathlib/Util/AssertExists.lean | 1 + Mathlib/Util/AssertNoSorry.lean | 1 + Mathlib/Util/AtomM.lean | 1 + Mathlib/Util/CompileInductive.lean | 1 + Mathlib/Util/CountHeartbeats.lean | 1 + Mathlib/Util/Delaborators.lean | 1 + Mathlib/Util/DischargerAsTactic.lean | 1 + Mathlib/Util/Export.lean | 1 + Mathlib/Util/GetAllModules.lean | 1 + Mathlib/Util/IncludeStr.lean | 1 + Mathlib/Util/MemoFix.lean | 1 + Mathlib/Util/Qq.lean | 1 + Mathlib/Util/SleepHeartbeats.lean | 1 + Mathlib/Util/Superscript.lean | 1 + Mathlib/Util/SynthesizeUsing.lean | 1 + Mathlib/Util/Tactic.lean | 1 + Mathlib/Util/TermBeta.lean | 1 + Mathlib/Util/Time.lean | 1 + Mathlib/Util/WhatsNew.lean | 1 + Mathlib/Util/WithWeakNamespace.lean | 1 + scripts/init_creation.sh | 40 +++++++++++++++++++ scripts/nolints-style.txt | 5 +++ scripts/noshake.json | 1 + scripts/style-exceptions.txt | 2 - 146 files changed, 199 insertions(+), 2 deletions(-) create mode 100644 Mathlib/Init.lean create mode 100644 scripts/init_creation.sh diff --git a/Mathlib.lean b/Mathlib.lean index 57bac5b7799cf..51a6ebf92e2b5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2794,6 +2794,7 @@ import Mathlib.GroupTheory.Sylow import Mathlib.GroupTheory.Torsion import Mathlib.GroupTheory.Transfer import Mathlib.InformationTheory.Hamming +import Mathlib.Init import Mathlib.Init.Algebra.Classes import Mathlib.Init.Data.Int.Order import Mathlib.Init.Data.List.Lemmas diff --git a/Mathlib/Algebra/HierarchyDesign.lean b/Mathlib/Algebra/HierarchyDesign.lean index 3101d26f94bdc..862788ae48368 100644 --- a/Mathlib/Algebra/HierarchyDesign.lean +++ b/Mathlib/Algebra/HierarchyDesign.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Eric Wieser -/ +import Mathlib.Init import Batteries.Util.LibraryNote /-! diff --git a/Mathlib/CategoryTheory/Category/Init.lean b/Mathlib/CategoryTheory/Category/Init.lean index bb562ef63fef7..9c42b8ce10881 100644 --- a/Mathlib/CategoryTheory/Category/Init.lean +++ b/Mathlib/CategoryTheory/Category/Init.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ +import Mathlib.Init import Aesop /-! diff --git a/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean b/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean index 8ea7c729576e2..365a1c9c6a877 100644 --- a/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean +++ b/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Johannes Hölzl, Reid Barton, Sean Leather -/ +import Mathlib.Init import Batteries.Tactic.Lint.Misc /-! diff --git a/Mathlib/Combinatorics/SimpleGraph/Init.lean b/Mathlib/Combinatorics/SimpleGraph/Init.lean index efd1b1d9dd8c3..349009999ed43 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Init.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Init.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ +import Mathlib.Init import Aesop /-! diff --git a/Mathlib/Control/Combinators.lean b/Mathlib/Control/Combinators.lean index f414787695265..b0aaad216354b 100644 --- a/Mathlib/Control/Combinators.lean +++ b/Mathlib/Control/Combinators.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura -/ +import Mathlib.Init /-! # Monad combinators, as in Haskell's Control.Monad. -/ diff --git a/Mathlib/Control/ULift.lean b/Mathlib/Control/ULift.lean index a18277cf38b2e..785d0d187c5d7 100644 --- a/Mathlib/Control/ULift.lean +++ b/Mathlib/Control/ULift.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Jannis Limperg -/ +import Mathlib.Init /-! # Monadic instances for `ULift` and `PLift` diff --git a/Mathlib/Data/Array/Defs.lean b/Mathlib/Data/Array/Defs.lean index f3d25eac2c584..d6497446d3a52 100644 --- a/Mathlib/Data/Array/Defs.lean +++ b/Mathlib/Data/Array/Defs.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Floris van Doorn -/ +import Mathlib.Init /-! ## Definitions on Arrays diff --git a/Mathlib/Data/Array/ExtractLemmas.lean b/Mathlib/Data/Array/ExtractLemmas.lean index 53aadf1b86929..bc66fc0660fbc 100644 --- a/Mathlib/Data/Array/ExtractLemmas.lean +++ b/Mathlib/Data/Array/ExtractLemmas.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Jiecheng Zhao. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jiecheng Zhao -/ +import Mathlib.Init /-! # Lemmas about `Array.extract` diff --git a/Mathlib/Data/ByteArray.lean b/Mathlib/Data/ByteArray.lean index 8f5407658504d..012385bb8d7cf 100644 --- a/Mathlib/Data/ByteArray.lean +++ b/Mathlib/Data/ByteArray.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Init /-! # Main result Introduce main properties of `Up` (well-ordered relation for "upwards" induction on `ℕ`) and of diff --git a/Mathlib/Data/Finset/Attr.lean b/Mathlib/Data/Finset/Attr.lean index d122deb225f9a..43bc734336e24 100644 --- a/Mathlib/Data/Finset/Attr.lean +++ b/Mathlib/Data/Finset/Attr.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.Init import Aesop import Qq diff --git a/Mathlib/Data/Int/Align.lean b/Mathlib/Data/Int/Align.lean index 8f45148a5452b..70c7b26e8b80d 100644 --- a/Mathlib/Data/Int/Align.lean +++ b/Mathlib/Data/Int/Align.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ +import Mathlib.Init /-! # Align statements for results about the integers -/ diff --git a/Mathlib/Data/Int/Notation.lean b/Mathlib/Data/Int/Notation.lean index 7db60724ac2a7..57ac150fcbfe1 100644 --- a/Mathlib/Data/Int/Notation.lean +++ b/Mathlib/Data/Int/Notation.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ +import Mathlib.Init /-! # Notation `ℤ` for the integers. -/ diff --git a/Mathlib/Data/List/Monad.lean b/Mathlib/Data/List/Monad.lean index 0dc10994ff059..a89b5f05d5272 100644 --- a/Mathlib/Data/List/Monad.lean +++ b/Mathlib/Data/List/Monad.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ +import Mathlib.Init /-! # Monad instances for `List` -/ diff --git a/Mathlib/Data/MLList/Dedup.lean b/Mathlib/Data/MLList/Dedup.lean index f4579cb51a78d..880fe43d4416e 100644 --- a/Mathlib/Data/MLList/Dedup.lean +++ b/Mathlib/Data/MLList/Dedup.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import Mathlib.Init import Batteries.Data.MLList.Basic import Batteries.Data.HashMap.Basic diff --git a/Mathlib/Data/MLList/IO.lean b/Mathlib/Data/MLList/IO.lean index 38b8d9242d245..a5c8a8ed25243 100644 --- a/Mathlib/Data/MLList/IO.lean +++ b/Mathlib/Data/MLList/IO.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import Mathlib.Init import Batteries.Data.MLList.Basic /-! diff --git a/Mathlib/Data/Matroid/Init.lean b/Mathlib/Data/Matroid/Init.lean index d888905b6f315..0ba1866f02d1e 100644 --- a/Mathlib/Data/Matroid/Init.lean +++ b/Mathlib/Data/Matroid/Init.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Peter Nelson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Peter Nelson -/ +import Mathlib.Init import Aesop /-! diff --git a/Mathlib/Data/Nat/Notation.lean b/Mathlib/Data/Nat/Notation.lean index 729e9b67b0f52..582c017ea9161 100644 --- a/Mathlib/Data/Nat/Notation.lean +++ b/Mathlib/Data/Nat/Notation.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura -/ +import Mathlib.Init /-! # Notation `ℕ` for the natural numbers. -/ diff --git a/Mathlib/Data/Ordering/Basic.lean b/Mathlib/Data/Ordering/Basic.lean index 64950a0331059..47a83a56ad26f 100644 --- a/Mathlib/Data/Ordering/Basic.lean +++ b/Mathlib/Data/Ordering/Basic.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura -/ +import Mathlib.Init /-! # Helper definitions and instances for `Ordering` -/ diff --git a/Mathlib/Data/String/Defs.lean b/Mathlib/Data/String/Defs.lean index 0df9b21527948..a6c9b5e9fd401 100644 --- a/Mathlib/Data/String/Defs.lean +++ b/Mathlib/Data/String/Defs.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Keeley Hoek, Floris van Doorn, Chris Bailey -/ +import Mathlib.Init /-! # Definitions for `String` diff --git a/Mathlib/Data/Sym/Sym2/Init.lean b/Mathlib/Data/Sym/Sym2/Init.lean index 8735d0d7603cd..9c7fdd617d017 100644 --- a/Mathlib/Data/Sym/Sym2/Init.lean +++ b/Mathlib/Data/Sym/Sym2/Init.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ +import Mathlib.Init import Aesop /-! diff --git a/Mathlib/Deprecated/Aliases.lean b/Mathlib/Deprecated/Aliases.lean index 3ccfb1ced2c07..5e63adabb8005 100644 --- a/Mathlib/Deprecated/Aliases.lean +++ b/Mathlib/Deprecated/Aliases.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ +import Mathlib.Init import Batteries.Tactic.Alias import Batteries.Data.String.Basic diff --git a/Mathlib/Deprecated/HashMap.lean b/Mathlib/Deprecated/HashMap.lean index 8fa961c5642c6..5917267c7b004 100644 --- a/Mathlib/Deprecated/HashMap.lean +++ b/Mathlib/Deprecated/HashMap.lean @@ -7,6 +7,7 @@ As `HashMap` has been completely reimplemented in `Batteries`, nothing from the mathlib3 file `data.hash_map` is reflected here. The porting header is just here to mark that no further work on `data.hash_map` is desired. -/ +import Mathlib.Init import Batteries.Data.HashMap.Basic import Batteries.Data.RBMap.Basic diff --git a/Mathlib/Init.lean b/Mathlib/Init.lean new file mode 100644 index 0000000000000..c06095f8af584 --- /dev/null +++ b/Mathlib/Init.lean @@ -0,0 +1,12 @@ +/-! +This is the root file in Mathlib: it is imported by virtually *all* Mathlib files. + +For this reason, the imports of this files are carefully curated. +Any modification involving a change in the imports of this file should be discussed beforehand. + +Here are some general guidelines: +* no bucket imports (e.g. `Batteries`/`Lean`/etc); +* every import needs to have a comment explaining why the import is there; +* strong preference for avoiding files that themselves have imports beyond `Lean`, and + any exception to this rule should by accompanied by a comment explaining the transitive imports. +-/ diff --git a/Mathlib/Init/Data/Quot.lean b/Mathlib/Init/Data/Quot.lean index 29dc0adb33bd0..891d304ebce19 100644 --- a/Mathlib/Init/Data/Quot.lean +++ b/Mathlib/Init/Data/Quot.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ +import Mathlib.Init /-! # Note about `Mathlib/Init/` The files in `Mathlib/Init` are leftovers from the port from Mathlib3. diff --git a/Mathlib/Init/Quot.lean b/Mathlib/Init/Quot.lean index af01552219ceb..5db59383719ff 100644 --- a/Mathlib/Init/Quot.lean +++ b/Mathlib/Init/Quot.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ +import Mathlib.Init /-! # Note about `Mathlib/Init/` The files in `Mathlib/Init` are leftovers from the port from Mathlib3. diff --git a/Mathlib/Init/Set.lean b/Mathlib/Init/Set.lean index 4999dfe196e03..71458589bfde7 100644 --- a/Mathlib/Init/Set.lean +++ b/Mathlib/Init/Set.lean @@ -3,6 +3,7 @@ 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.Init import Batteries.Util.ExtendedBinder /-! diff --git a/Mathlib/Lean/Elab/Term.lean b/Mathlib/Lean/Elab/Term.lean index b5b43079267c6..30aa1819157a2 100644 --- a/Mathlib/Lean/Elab/Term.lean +++ b/Mathlib/Lean/Elab/Term.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Init import Lean.Elab.SyntheticMVars /-! diff --git a/Mathlib/Lean/EnvExtension.lean b/Mathlib/Lean/EnvExtension.lean index dc33d0f890559..6ca1adad43b48 100644 --- a/Mathlib/Lean/EnvExtension.lean +++ b/Mathlib/Lean/EnvExtension.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ +import Mathlib.Init import Lean.ScopedEnvExtension /-! diff --git a/Mathlib/Lean/Exception.lean b/Mathlib/Lean/Exception.lean index 9eb2a70deb8a5..ab87792720158 100644 --- a/Mathlib/Lean/Exception.lean +++ b/Mathlib/Lean/Exception.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Edward Ayers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Edward Ayers -/ +import Mathlib.Init import Lean.Exception /-! diff --git a/Mathlib/Lean/Expr/Basic.lean b/Mathlib/Lean/Expr/Basic.lean index b7a76c6339391..0a3a8fa1db16f 100644 --- a/Mathlib/Lean/Expr/Basic.lean +++ b/Mathlib/Lean/Expr/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Simon Hudon, Scott Morrison, Keeley Hoek, Robert Y. Lewis, Floris van Doorn, Edward Ayers, Arthur Paulino -/ +import Mathlib.Init import Lean.Meta.Tactic.Rewrite import Batteries.Lean.Expr import Batteries.Data.Rat.Basic diff --git a/Mathlib/Lean/GoalsLocation.lean b/Mathlib/Lean/GoalsLocation.lean index 59acb29c1015d..d6f34b2a228fe 100644 --- a/Mathlib/Lean/GoalsLocation.lean +++ b/Mathlib/Lean/GoalsLocation.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Jovan Gerbscheid. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jovan Gerbscheid -/ +import Mathlib.Init import Lean.Meta.Tactic.Util import Lean.SubExpr diff --git a/Mathlib/Lean/Json.lean b/Mathlib/Lean/Json.lean index 93889905d6351..020937e439b01 100644 --- a/Mathlib/Lean/Json.lean +++ b/Mathlib/Lean/Json.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ +import Mathlib.Init import Lean.Data.Json.FromToJson /-! diff --git a/Mathlib/Lean/LocalContext.lean b/Mathlib/Lean/LocalContext.lean index f5ba66bd1e966..5bc891c5cdf0b 100644 --- a/Mathlib/Lean/LocalContext.lean +++ b/Mathlib/Lean/LocalContext.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import Mathlib.Init import Lean.LocalContext /-! diff --git a/Mathlib/Lean/Message.lean b/Mathlib/Lean/Message.lean index abfd4e902ad74..a9b79ad767566 100644 --- a/Mathlib/Lean/Message.lean +++ b/Mathlib/Lean/Message.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ +import Mathlib.Init import Lean.Message /-! diff --git a/Mathlib/Lean/Meta.lean b/Mathlib/Lean/Meta.lean index 852de4f890fc9..a4dd8825343f5 100644 --- a/Mathlib/Lean/Meta.lean +++ b/Mathlib/Lean/Meta.lean @@ -3,6 +3,7 @@ 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.Init import Lean.Elab.Term import Lean.Elab.Tactic.Basic import Lean.Meta.Tactic.Assert diff --git a/Mathlib/Lean/Meta/Basic.lean b/Mathlib/Lean/Meta/Basic.lean index de0b468cfe711..a8066da2f4022 100644 --- a/Mathlib/Lean/Meta/Basic.lean +++ b/Mathlib/Lean/Meta/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import Mathlib.Init import Lean.Meta.AppBuilder import Lean.Meta.Basic diff --git a/Mathlib/Lean/Meta/DiscrTree.lean b/Mathlib/Lean/Meta/DiscrTree.lean index 04468c7d9eba0..c20618862b96c 100644 --- a/Mathlib/Lean/Meta/DiscrTree.lean +++ b/Mathlib/Lean/Meta/DiscrTree.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import Mathlib.Init import Lean.Meta.DiscrTree /-! diff --git a/Mathlib/Lean/Meta/KAbstractPositions.lean b/Mathlib/Lean/Meta/KAbstractPositions.lean index c6e9ecaf3c7bf..2e3f801d542e0 100644 --- a/Mathlib/Lean/Meta/KAbstractPositions.lean +++ b/Mathlib/Lean/Meta/KAbstractPositions.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Jovan Gerbscheid. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jovan Gerbscheid -/ +import Mathlib.Init import Lean.HeadIndex import Lean.Meta.ExprLens import Lean.Meta.Check diff --git a/Mathlib/Lean/Meta/Simp.lean b/Mathlib/Lean/Meta/Simp.lean index 82573b55fb821..19d5c973616f2 100644 --- a/Mathlib/Lean/Meta/Simp.lean +++ b/Mathlib/Lean/Meta/Simp.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Gabriel Ebner, Floris van Doorn -/ +import Mathlib.Init import Lean.Elab.Tactic.Simp /-! diff --git a/Mathlib/Lean/Name.lean b/Mathlib/Lean/Name.lean index 84f0943244ae3..268d785610641 100644 --- a/Mathlib/Lean/Name.lean +++ b/Mathlib/Lean/Name.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import Mathlib.Init import Lean.Meta.Match.MatcherInfo import Lean.Meta.Tactic.Delta import Std.Data.HashMap.Basic diff --git a/Mathlib/Lean/PrettyPrinter/Delaborator.lean b/Mathlib/Lean/PrettyPrinter/Delaborator.lean index 2359d34e4d778..6ee7e6719093f 100644 --- a/Mathlib/Lean/PrettyPrinter/Delaborator.lean +++ b/Mathlib/Lean/PrettyPrinter/Delaborator.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Init import Lean.PrettyPrinter.Delaborator.Basic /-! diff --git a/Mathlib/Lean/Thunk.lean b/Mathlib/Lean/Thunk.lean index 6b67ce032964f..e9d504dff0852 100644 --- a/Mathlib/Lean/Thunk.lean +++ b/Mathlib/Lean/Thunk.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon -/ +import Mathlib.Init import Batteries.Data.Thunk /-! diff --git a/Mathlib/Tactic/AdaptationNote.lean b/Mathlib/Tactic/AdaptationNote.lean index 1f6e49a9cb244..af508f76103bf 100644 --- a/Mathlib/Tactic/AdaptationNote.lean +++ b/Mathlib/Tactic/AdaptationNote.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/ApplyWith.lean b/Mathlib/Tactic/ApplyWith.lean index 80bee54facbee..abd1fb386c292 100644 --- a/Mathlib/Tactic/ApplyWith.lean +++ b/Mathlib/Tactic/ApplyWith.lean @@ -3,6 +3,7 @@ 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.Init import Lean.Elab.Eval import Lean.Elab.Tactic.ElabTerm diff --git a/Mathlib/Tactic/ArithMult/Init.lean b/Mathlib/Tactic/ArithMult/Init.lean index e2a7114273970..40a20040c7bdc 100644 --- a/Mathlib/Tactic/ArithMult/Init.lean +++ b/Mathlib/Tactic/ArithMult/Init.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Arend Mellendijk -/ +import Mathlib.Init import Aesop /-! diff --git a/Mathlib/Tactic/Attr/Register.lean b/Mathlib/Tactic/Attr/Register.lean index 9369b682db7c2..98101a86367d7 100644 --- a/Mathlib/Tactic/Attr/Register.lean +++ b/Mathlib/Tactic/Attr/Register.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import Mathlib.Init import Lean.Meta.Tactic.Simp.SimpTheorems import Lean.Meta.Tactic.Simp.RegisterCommand import Lean.LabelAttribute diff --git a/Mathlib/Tactic/Bound/Init.lean b/Mathlib/Tactic/Bound/Init.lean index 68e5fdae37e04..5ca4a86975928 100644 --- a/Mathlib/Tactic/Bound/Init.lean +++ b/Mathlib/Tactic/Bound/Init.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Geoffrey Irving -/ +import Mathlib.Init import Aesop.Frontend.Command /-! diff --git a/Mathlib/Tactic/CasesM.lean b/Mathlib/Tactic/CasesM.lean index f3113a9fab7e4..c8fce2ba7a258 100644 --- a/Mathlib/Tactic/CasesM.lean +++ b/Mathlib/Tactic/CasesM.lean @@ -3,6 +3,7 @@ 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.Init import Lean.Elab.Tactic.Conv.Pattern /-! diff --git a/Mathlib/Tactic/Change.lean b/Mathlib/Tactic/Change.lean index c961f512b9f9e..235826d756da2 100644 --- a/Mathlib/Tactic/Change.lean +++ b/Mathlib/Tactic/Change.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ +import Mathlib.Init import Lean.Elab.Tactic.ElabTerm import Lean.Meta.Tactic.TryThis /-! diff --git a/Mathlib/Tactic/Check.lean b/Mathlib/Tactic/Check.lean index f4120a1084a1a..2579740e185f8 100644 --- a/Mathlib/Tactic/Check.lean +++ b/Mathlib/Tactic/Check.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Init import Lean.Elab.Tactic.Basic import Lean.PrettyPrinter import Lean.Elab.SyntheticMVars diff --git a/Mathlib/Tactic/Clean.lean b/Mathlib/Tactic/Clean.lean index 3c27d21b3c97e..ae5800de840f4 100644 --- a/Mathlib/Tactic/Clean.lean +++ b/Mathlib/Tactic/Clean.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Michail Karatarakis, Kyle Miller -/ +import Mathlib.Init import Lean.Elab.SyntheticMVars /-! diff --git a/Mathlib/Tactic/ClearExcept.lean b/Mathlib/Tactic/ClearExcept.lean index 804d6950d8e2c..b2791172f5776 100644 --- a/Mathlib/Tactic/ClearExcept.lean +++ b/Mathlib/Tactic/ClearExcept.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Joshua Clune. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joshua Clune -/ +import Mathlib.Init import Lean.Elab.Tactic.ElabTerm /-! diff --git a/Mathlib/Tactic/ClearExclamation.lean b/Mathlib/Tactic/ClearExclamation.lean index f15744e865792..2f34db728bb70 100644 --- a/Mathlib/Tactic/ClearExclamation.lean +++ b/Mathlib/Tactic/ClearExclamation.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Joshua Clune. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joshua Clune -/ +import Mathlib.Init import Lean.Elab.Tactic.ElabTerm /-! # `clear!` tactic -/ diff --git a/Mathlib/Tactic/Clear_.lean b/Mathlib/Tactic/Clear_.lean index d1be75bb99ce9..5d0c8245c8c97 100644 --- a/Mathlib/Tactic/Clear_.lean +++ b/Mathlib/Tactic/Clear_.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Joshua Clune. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joshua Clune -/ +import Mathlib.Init import Lean.Meta.Tactic.Clear import Lean.Elab.Tactic.Basic diff --git a/Mathlib/Tactic/Coe.lean b/Mathlib/Tactic/Coe.lean index e4bc51dc6cff2..88a98f81116a5 100644 --- a/Mathlib/Tactic/Coe.lean +++ b/Mathlib/Tactic/Coe.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner -/ +import Mathlib.Init import Lean.Elab.ElabRules /-! diff --git a/Mathlib/Tactic/Constructor.lean b/Mathlib/Tactic/Constructor.lean index 8d7ab63b5ab4e..6ba7e4aee3325 100644 --- a/Mathlib/Tactic/Constructor.lean +++ b/Mathlib/Tactic/Constructor.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Newell Jensen -/ +import Mathlib.Init import Lean.Elab.SyntheticMVars import Lean.Meta.Tactic.Constructor diff --git a/Mathlib/Tactic/Continuity/Init.lean b/Mathlib/Tactic/Continuity/Init.lean index bf0e7dfe0f52f..e8ff04a4ed852 100644 --- a/Mathlib/Tactic/Continuity/Init.lean +++ b/Mathlib/Tactic/Continuity/Init.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ +import Mathlib.Init import Aesop /-! diff --git a/Mathlib/Tactic/Conv.lean b/Mathlib/Tactic/Conv.lean index 840ffec93b3ce..44dfd2db367fd 100644 --- a/Mathlib/Tactic/Conv.lean +++ b/Mathlib/Tactic/Conv.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner -/ +import Mathlib.Init import Lean.Elab.Tactic.Conv.Basic import Lean.Elab.Command diff --git a/Mathlib/Tactic/Eqns.lean b/Mathlib/Tactic/Eqns.lean index f682526f53380..35ce6fc0c98ce 100644 --- a/Mathlib/Tactic/Eqns.lean +++ b/Mathlib/Tactic/Eqns.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ +import Mathlib.Init import Lean.Meta.Eqns import Batteries.Lean.NameMapAttribute import Lean.Elab.Exception diff --git a/Mathlib/Tactic/Eval.lean b/Mathlib/Tactic/Eval.lean index 8edda73b3590a..e8611d5425ea1 100644 --- a/Mathlib/Tactic/Eval.lean +++ b/Mathlib/Tactic/Eval.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ +import Mathlib.Init import Qq.Macro /-! diff --git a/Mathlib/Tactic/ExistsI.lean b/Mathlib/Tactic/ExistsI.lean index 09df3468c3bb8..17f49821f39a9 100644 --- a/Mathlib/Tactic/ExistsI.lean +++ b/Mathlib/Tactic/ExistsI.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Gabriel Ebner, Moritz Doll -/ +import Mathlib.Init /-! # The `existsi` tactic This file defines the `existsi` tactic: its purpose is to instantiate existential quantifiers. diff --git a/Mathlib/Tactic/Explode/Datatypes.lean b/Mathlib/Tactic/Explode/Datatypes.lean index 0ce6999aa004c..d15f255f8b316 100644 --- a/Mathlib/Tactic/Explode/Datatypes.lean +++ b/Mathlib/Tactic/Explode/Datatypes.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Evgenia Karunus, Kyle Miller -/ +import Mathlib.Init import Lean.Util.Trace /-! diff --git a/Mathlib/Tactic/ExtendDoc.lean b/Mathlib/Tactic/ExtendDoc.lean index ece9c14e0436f..87e11b9cbef8d 100644 --- a/Mathlib/Tactic/ExtendDoc.lean +++ b/Mathlib/Tactic/ExtendDoc.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ +import Mathlib.Init import Lean.Elab.ElabRules import Lean.DocString diff --git a/Mathlib/Tactic/ExtractGoal.lean b/Mathlib/Tactic/ExtractGoal.lean index fd7e32324ebc2..f32afd77cc52d 100644 --- a/Mathlib/Tactic/ExtractGoal.lean +++ b/Mathlib/Tactic/ExtractGoal.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Kyle Miller, Damiano Testa -/ +import Mathlib.Init import Lean.Elab.Tactic.ElabTerm import Lean.Meta.Tactic.Cleanup import Lean.PrettyPrinter diff --git a/Mathlib/Tactic/FailIfNoProgress.lean b/Mathlib/Tactic/FailIfNoProgress.lean index 7544d90f2ec45..d8e4dbedbbdd1 100644 --- a/Mathlib/Tactic/FailIfNoProgress.lean +++ b/Mathlib/Tactic/FailIfNoProgress.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Thomas Murrills. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Murrills -/ +import Mathlib.Init import Lean.Elab.Tactic.Basic import Lean.Meta.Tactic.Util diff --git a/Mathlib/Tactic/Find.lean b/Mathlib/Tactic/Find.lean index e52d15513b559..b95098e141720 100644 --- a/Mathlib/Tactic/Find.lean +++ b/Mathlib/Tactic/Find.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Sebastian Ullrich. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sebastian Ullrich -/ +import Mathlib.Init import Batteries.Util.Cache import Lean.HeadIndex import Lean.Elab.Command diff --git a/Mathlib/Tactic/FunProp/Decl.lean b/Mathlib/Tactic/FunProp/Decl.lean index 8ef6ae4bb0367..3e3e3d147570f 100644 --- a/Mathlib/Tactic/FunProp/Decl.lean +++ b/Mathlib/Tactic/FunProp/Decl.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Tomáš Skřivan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Tomáš Skřivan -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/FunProp/StateList.lean b/Mathlib/Tactic/FunProp/StateList.lean index b1f6431c1794f..2174cf4748411 100644 --- a/Mathlib/Tactic/FunProp/StateList.lean +++ b/Mathlib/Tactic/FunProp/StateList.lean @@ -21,6 +21,7 @@ in the do block are combined. -/ +import Mathlib.Init /-! StateList -/ namespace Mathlib.Meta.FunProp diff --git a/Mathlib/Tactic/FunProp/ToBatteries.lean b/Mathlib/Tactic/FunProp/ToBatteries.lean index 5cf39afb38db6..a673dc50a4f3c 100644 --- a/Mathlib/Tactic/FunProp/ToBatteries.lean +++ b/Mathlib/Tactic/FunProp/ToBatteries.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Tomáš Skřivan. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Tomáš Skřivan -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/GCongr/ForwardAttr.lean b/Mathlib/Tactic/GCongr/ForwardAttr.lean index 88078fb3adb8c..ad1561c067f03 100644 --- a/Mathlib/Tactic/GCongr/ForwardAttr.lean +++ b/Mathlib/Tactic/GCongr/ForwardAttr.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Mario Carneiro, Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Heather Macbeth -/ +import Mathlib.Init import Batteries.Tactic.Basic /-! diff --git a/Mathlib/Tactic/Generalize.lean b/Mathlib/Tactic/Generalize.lean index decfd95b40a4f..bcd97b1c34af3 100644 --- a/Mathlib/Tactic/Generalize.lean +++ b/Mathlib/Tactic/Generalize.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Lean FRO, LLC. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import Mathlib.Init import Lean.Elab.Binders import Lean.Elab.Tactic.ElabTerm import Lean.Meta.Tactic.Generalize diff --git a/Mathlib/Tactic/GuardGoalNums.lean b/Mathlib/Tactic/GuardGoalNums.lean index 882cb6b7bf060..27531fd64874f 100644 --- a/Mathlib/Tactic/GuardGoalNums.lean +++ b/Mathlib/Tactic/GuardGoalNums.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis -/ +import Mathlib.Init import Lean.Elab.Tactic.Basic /-! diff --git a/Mathlib/Tactic/GuardHypNums.lean b/Mathlib/Tactic/GuardHypNums.lean index 60a0cf119966a..ad3202937bbca 100644 --- a/Mathlib/Tactic/GuardHypNums.lean +++ b/Mathlib/Tactic/GuardHypNums.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis -/ +import Mathlib.Init import Lean.Elab.Tactic.Basic /-! diff --git a/Mathlib/Tactic/Have.lean b/Mathlib/Tactic/Have.lean index 582bb2c0559e5..1c96ee457da4f 100644 --- a/Mathlib/Tactic/Have.lean +++ b/Mathlib/Tactic/Have.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Arthur Paulino. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Edward Ayers, Mario Carneiro -/ +import Mathlib.Init import Lean.Elab.Binders import Lean.Elab.SyntheticMVars import Lean.Meta.Tactic.Assert diff --git a/Mathlib/Tactic/HaveI.lean b/Mathlib/Tactic/HaveI.lean index 533c20e8d38ee..e08ca9e7301a4 100644 --- a/Mathlib/Tactic/HaveI.lean +++ b/Mathlib/Tactic/HaveI.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner -/ +import Mathlib.Init /-! # Variants of `haveI`/`letI` for use in do-notation. diff --git a/Mathlib/Tactic/HelpCmd.lean b/Mathlib/Tactic/HelpCmd.lean index c5f5709996c83..a46cf4763e507 100644 --- a/Mathlib/Tactic/HelpCmd.lean +++ b/Mathlib/Tactic/HelpCmd.lean @@ -3,6 +3,7 @@ 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.Init import Lean.Elab.Syntax import Lean.DocString diff --git a/Mathlib/Tactic/InferParam.lean b/Mathlib/Tactic/InferParam.lean index e649d5b1f7f72..568fcb26965c2 100644 --- a/Mathlib/Tactic/InferParam.lean +++ b/Mathlib/Tactic/InferParam.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Mario Carneiro -/ +import Mathlib.Init import Lean.Elab.Tactic.Basic import Lean.Meta.Tactic.Replace diff --git a/Mathlib/Tactic/Lemma.lean b/Mathlib/Tactic/Lemma.lean index 4412258b0c428..901874a9f585a 100644 --- a/Mathlib/Tactic/Lemma.lean +++ b/Mathlib/Tactic/Lemma.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kyle Miller -/ +import Mathlib.Init import Lean.Parser.Command /-! diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean index fc26e112464dd..986bb5ed40031 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Vasily Nesterov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Vasily Nesterov -/ +import Mathlib.Init import Lean.Data.HashMap import Batteries.Data.Rat.Basic diff --git a/Mathlib/Tactic/Linter/GlobalAttributeIn.lean b/Mathlib/Tactic/Linter/GlobalAttributeIn.lean index 578c12d114d36..b18e1defd647f 100644 --- a/Mathlib/Tactic/Linter/GlobalAttributeIn.lean +++ b/Mathlib/Tactic/Linter/GlobalAttributeIn.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Rothgang, Damiano Testa -/ +import Mathlib.Init import Lean.Elab.Command import Lean.Linter.Util diff --git a/Mathlib/Tactic/Linter/HashCommandLinter.lean b/Mathlib/Tactic/Linter/HashCommandLinter.lean index f50ba4b9f795c..76e570391be93 100644 --- a/Mathlib/Tactic/Linter/HashCommandLinter.lean +++ b/Mathlib/Tactic/Linter/HashCommandLinter.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ +import Mathlib.Init import Lean.Elab.Command import Lean.Linter.Util import Batteries.Lean.HashSet diff --git a/Mathlib/Tactic/Linter/HaveLetLinter.lean b/Mathlib/Tactic/Linter/HaveLetLinter.lean index e59b5417ef447..9e93804663ceb 100644 --- a/Mathlib/Tactic/Linter/HaveLetLinter.lean +++ b/Mathlib/Tactic/Linter/HaveLetLinter.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ +import Mathlib.Init import Lean.Elab.Command import Lean.Server.InfoUtils diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index 38bd0a9ffc843..3025eadd9f127 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ +import Mathlib.Init import Lean.Linter.Util import Batteries.Data.String.Matcher import Batteries.Tactic.Lint diff --git a/Mathlib/Tactic/Linter/OldObtain.lean b/Mathlib/Tactic/Linter/OldObtain.lean index 25069ff70983b..1c0a6d5783fb8 100644 --- a/Mathlib/Tactic/Linter/OldObtain.lean +++ b/Mathlib/Tactic/Linter/OldObtain.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Rothgang -/ +import Mathlib.Init import Lean.Elab.Command import Lean.Linter.Util diff --git a/Mathlib/Tactic/Linter/RefineLinter.lean b/Mathlib/Tactic/Linter/RefineLinter.lean index 7518ae9a1b1d5..b97dd76db8525 100644 --- a/Mathlib/Tactic/Linter/RefineLinter.lean +++ b/Mathlib/Tactic/Linter/RefineLinter.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ +import Mathlib.Init import Lean.Elab.Command import Lean.Linter.Util diff --git a/Mathlib/Tactic/Linter/Style.lean b/Mathlib/Tactic/Linter/Style.lean index b7dcce49b4ae3..0064675cef1c7 100644 --- a/Mathlib/Tactic/Linter/Style.lean +++ b/Mathlib/Tactic/Linter/Style.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Rothgang -/ +import Mathlib.Init import Lean.Elab.Command import Lean.Linter.Util diff --git a/Mathlib/Tactic/Linter/UnusedTactic.lean b/Mathlib/Tactic/Linter/UnusedTactic.lean index 6b9a97032fa33..9389441b83460 100644 --- a/Mathlib/Tactic/Linter/UnusedTactic.lean +++ b/Mathlib/Tactic/Linter/UnusedTactic.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ +import Mathlib.Init import Lean.Elab.Command import Lean.Linter.Util import Batteries.Tactic.Unreachable diff --git a/Mathlib/Tactic/Measurability/Init.lean b/Mathlib/Tactic/Measurability/Init.lean index 0821e564490e1..8d4de6f8f88c6 100644 --- a/Mathlib/Tactic/Measurability/Init.lean +++ b/Mathlib/Tactic/Measurability/Init.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Miyahara Kō -/ +import Mathlib.Init import Aesop /-! diff --git a/Mathlib/Tactic/MinImports.lean b/Mathlib/Tactic/MinImports.lean index f1419676d8a07..788d4f9df7257 100644 --- a/Mathlib/Tactic/MinImports.lean +++ b/Mathlib/Tactic/MinImports.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ +import Mathlib.Init import ImportGraph.Imports /-! # `#min_imports in` a command to find minimal imports diff --git a/Mathlib/Tactic/Monotonicity/Attr.lean b/Mathlib/Tactic/Monotonicity/Attr.lean index f42c19899eb0b..101f06b95807c 100644 --- a/Mathlib/Tactic/Monotonicity/Attr.lean +++ b/Mathlib/Tactic/Monotonicity/Attr.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Heather Macbeth -/ +import Mathlib.Init import Lean.LabelAttribute /-! # The @[mono] attribute -/ diff --git a/Mathlib/Tactic/NthRewrite.lean b/Mathlib/Tactic/NthRewrite.lean index eff8ce34cc201..b0b2c62bc73f9 100644 --- a/Mathlib/Tactic/NthRewrite.lean +++ b/Mathlib/Tactic/NthRewrite.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ +import Mathlib.Init import Lean.Elab.Tactic.Rewrite /-! diff --git a/Mathlib/Tactic/Observe.lean b/Mathlib/Tactic/Observe.lean index b6f47514b7d3b..9013fc9f12636 100644 --- a/Mathlib/Tactic/Observe.lean +++ b/Mathlib/Tactic/Observe.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ +import Mathlib.Init import Lean.Meta.Tactic.TryThis import Lean.Elab.Tactic.ElabTerm import Lean.Meta.Tactic.LibrarySearch diff --git a/Mathlib/Tactic/PPWithUniv.lean b/Mathlib/Tactic/PPWithUniv.lean index 24a90c3e1394e..c4afe64ed1fe6 100644 --- a/Mathlib/Tactic/PPWithUniv.lean +++ b/Mathlib/Tactic/PPWithUniv.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/ProjectionNotation.lean b/Mathlib/Tactic/ProjectionNotation.lean index b35c8b3dac9b7..d56a57bf0f447 100644 --- a/Mathlib/Tactic/ProjectionNotation.lean +++ b/Mathlib/Tactic/ProjectionNotation.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/Recall.lean b/Mathlib/Tactic/Recall.lean index 373b55ea8addd..2b242f8b21fd4 100644 --- a/Mathlib/Tactic/Recall.lean +++ b/Mathlib/Tactic/Recall.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Mac Malone. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mac Malone, Kyle Miller -/ +import Mathlib.Init import Lean.Elab.Command import Lean.Elab.DeclUtil diff --git a/Mathlib/Tactic/Recover.lean b/Mathlib/Tactic/Recover.lean index 93ee578786e42..a56f0ddf7eaa1 100644 --- a/Mathlib/Tactic/Recover.lean +++ b/Mathlib/Tactic/Recover.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Siddhartha Gadgil. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner, Siddhartha Gadgil, Jannis Limperg -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/ReduceModChar/Ext.lean b/Mathlib/Tactic/ReduceModChar/Ext.lean index 2ee6b19df63b0..4e53350dd2539 100644 --- a/Mathlib/Tactic/ReduceModChar/Ext.lean +++ b/Mathlib/Tactic/ReduceModChar/Ext.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ +import Mathlib.Init import Lean.Meta.Tactic.Simp.Attr /-! diff --git a/Mathlib/Tactic/Relation/Rfl.lean b/Mathlib/Tactic/Relation/Rfl.lean index 038d0e43f1fc4..e114a5b15bd31 100644 --- a/Mathlib/Tactic/Relation/Rfl.lean +++ b/Mathlib/Tactic/Relation/Rfl.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Newell Jensen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Newell Jensen -/ +import Mathlib.Init import Lean.Meta.Tactic.Rfl /-! diff --git a/Mathlib/Tactic/Relation/Symm.lean b/Mathlib/Tactic/Relation/Symm.lean index 4f2d9e03bb9cb..273ff04b3d667 100644 --- a/Mathlib/Tactic/Relation/Symm.lean +++ b/Mathlib/Tactic/Relation/Symm.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import Mathlib.Init import Lean.Meta.Tactic.Symm /-! diff --git a/Mathlib/Tactic/Rename.lean b/Mathlib/Tactic/Rename.lean index 5ad80c9d3a68a..34a4fa944e7a8 100644 --- a/Mathlib/Tactic/Rename.lean +++ b/Mathlib/Tactic/Rename.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/Says.lean b/Mathlib/Tactic/Says.lean index 92726858cd89e..146b5d3aa230e 100644 --- a/Mathlib/Tactic/Says.lean +++ b/Mathlib/Tactic/Says.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Kim Liesinger. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Liesinger -/ +import Mathlib.Init import Batteries.Data.String.Basic import Lean.Meta.Tactic.TryThis import Batteries.Linter.UnreachableTactic diff --git a/Mathlib/Tactic/Set.lean b/Mathlib/Tactic/Set.lean index ec186f24acd28..7ab63a973e54d 100644 --- a/Mathlib/Tactic/Set.lean +++ b/Mathlib/Tactic/Set.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Ian Benway. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Ian Benway -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/SetLike.lean b/Mathlib/Tactic/SetLike.lean index fb0cd27dc2b04..3f55f3606d7d7 100644 --- a/Mathlib/Tactic/SetLike.lean +++ b/Mathlib/Tactic/SetLike.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ +import Mathlib.Init import Aesop /-! diff --git a/Mathlib/Tactic/SimpIntro.lean b/Mathlib/Tactic/SimpIntro.lean index 61e5eeafcdc63..72af478dd156a 100644 --- a/Mathlib/Tactic/SimpIntro.lean +++ b/Mathlib/Tactic/SimpIntro.lean @@ -3,6 +3,7 @@ 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.Init import Lean /-! # `simp_intro` tactic -/ diff --git a/Mathlib/Tactic/SimpRw.lean b/Mathlib/Tactic/SimpRw.lean index c8ce3d76dcf33..5b2e05e4bcc3e 100644 --- a/Mathlib/Tactic/SimpRw.lean +++ b/Mathlib/Tactic/SimpRw.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Mario Carneiro, Alex J. Best -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/Simps/NotationClass.lean b/Mathlib/Tactic/Simps/NotationClass.lean index 42f2271a61ea7..f1e36a9e926cc 100644 --- a/Mathlib/Tactic/Simps/NotationClass.lean +++ b/Mathlib/Tactic/Simps/NotationClass.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ +import Mathlib.Init import Lean.Elab.Exception import Batteries.Lean.NameMapAttribute import Batteries.Lean.Expr diff --git a/Mathlib/Tactic/Spread.lean b/Mathlib/Tactic/Spread.lean index aeb6043b465e1..0252a2fa396db 100644 --- a/Mathlib/Tactic/Spread.lean +++ b/Mathlib/Tactic/Spread.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner -/ +import Mathlib.Init import Lean.Elab.Binders /-! diff --git a/Mathlib/Tactic/Substs.lean b/Mathlib/Tactic/Substs.lean index 7b0b8bb481d9f..1e9dab5a3e5aa 100644 --- a/Mathlib/Tactic/Substs.lean +++ b/Mathlib/Tactic/Substs.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Evan Lohn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Evan Lohn, Mario Carneiro -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/SuccessIfFailWithMsg.lean b/Mathlib/Tactic/SuccessIfFailWithMsg.lean index 33b4f0709426c..d74fee7623ba5 100644 --- a/Mathlib/Tactic/SuccessIfFailWithMsg.lean +++ b/Mathlib/Tactic/SuccessIfFailWithMsg.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Simon Hudon, Sébastien Gouëzel, Scott Morrison, Thomas Murrills -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/SudoSetOption.lean b/Mathlib/Tactic/SudoSetOption.lean index 7d6ba36b7818b..8457c4741033c 100644 --- a/Mathlib/Tactic/SudoSetOption.lean +++ b/Mathlib/Tactic/SudoSetOption.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner -/ +import Mathlib.Init import Lean.Elab.ElabRules /-! diff --git a/Mathlib/Tactic/SuppressCompilation.lean b/Mathlib/Tactic/SuppressCompilation.lean index 945340816281a..aa50ea360913e 100644 --- a/Mathlib/Tactic/SuppressCompilation.lean +++ b/Mathlib/Tactic/SuppressCompilation.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best, Mac Malone -/ +import Mathlib.Init import Lean.Elab.Declaration import Lean.Elab.Notation diff --git a/Mathlib/Tactic/Trace.lean b/Mathlib/Tactic/Trace.lean index 364412c44ada5..2e5673d8cd799 100644 --- a/Mathlib/Tactic/Trace.lean +++ b/Mathlib/Tactic/Trace.lean @@ -3,6 +3,7 @@ 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.Init import Lean.Elab.Tactic.ElabTerm import Lean.Meta.Eval diff --git a/Mathlib/Tactic/TryThis.lean b/Mathlib/Tactic/TryThis.lean index 44e8c914dbc5c..8a8940474d5ef 100644 --- a/Mathlib/Tactic/TryThis.lean +++ b/Mathlib/Tactic/TryThis.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Tactic/TypeCheck.lean b/Mathlib/Tactic/TypeCheck.lean index b72e3f2079b0f..027d17e403e20 100644 --- a/Mathlib/Tactic/TypeCheck.lean +++ b/Mathlib/Tactic/TypeCheck.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ +import Mathlib.Init import Lean.Elab.Tactic.Basic import Lean.Elab.SyntheticMVars diff --git a/Mathlib/Tactic/TypeStar.lean b/Mathlib/Tactic/TypeStar.lean index d020ece4a2492..35c0d6d1b7fae 100644 --- a/Mathlib/Tactic/TypeStar.lean +++ b/Mathlib/Tactic/TypeStar.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Matthew Robert Ballard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Init import Lean.Elab.Term /-! diff --git a/Mathlib/Tactic/UnsetOption.lean b/Mathlib/Tactic/UnsetOption.lean index a57bfe4ffc56c..032d3f10d0be6 100644 --- a/Mathlib/Tactic/UnsetOption.lean +++ b/Mathlib/Tactic/UnsetOption.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Alex J. Best. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best -/ +import Mathlib.Init import Lean.Parser.Term import Lean.Parser.Do import Lean.Elab.Command diff --git a/Mathlib/Tactic/Use.lean b/Mathlib/Tactic/Use.lean index bb735ceacafb9..8c6bc4f992c9a 100644 --- a/Mathlib/Tactic/Use.lean +++ b/Mathlib/Tactic/Use.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Arthur Paulino. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Gabriel Ebner, Kyle Miller -/ +import Mathlib.Init import Lean.Meta.Tactic.Util import Lean.Elab.Tactic.Basic diff --git a/Mathlib/Tactic/Variable.lean b/Mathlib/Tactic/Variable.lean index d09edae33938b..1b303098b1b52 100644 --- a/Mathlib/Tactic/Variable.lean +++ b/Mathlib/Tactic/Variable.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Init import Lean.Meta.Tactic.TryThis /-! diff --git a/Mathlib/Tactic/Widget/SelectInsertParamsClass.lean b/Mathlib/Tactic/Widget/SelectInsertParamsClass.lean index 127473b69b56c..4c7be09054bee 100644 --- a/Mathlib/Tactic/Widget/SelectInsertParamsClass.lean +++ b/Mathlib/Tactic/Widget/SelectInsertParamsClass.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot -/ +import Mathlib.Init import Lean.Widget.InteractiveGoal import Lean.Elab.Deriving.Basic diff --git a/Mathlib/Topology/Sheaves/Init.lean b/Mathlib/Topology/Sheaves/Init.lean index 61db5393fa545..fad0482e9171f 100644 --- a/Mathlib/Topology/Sheaves/Init.lean +++ b/Mathlib/Topology/Sheaves/Init.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ +import Mathlib.Init import Aesop /-! diff --git a/Mathlib/Util/AddRelatedDecl.lean b/Mathlib/Util/AddRelatedDecl.lean index 22625a2347426..229fb6dd488b1 100644 --- a/Mathlib/Util/AddRelatedDecl.lean +++ b/Mathlib/Util/AddRelatedDecl.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Floris van Doorn -/ +import Mathlib.Init import Lean.Elab.DeclarationRange import Lean.Elab.Term diff --git a/Mathlib/Util/AssertExists.lean b/Mathlib/Util/AssertExists.lean index 8197e8a99b225..40beb0cc91e32 100644 --- a/Mathlib/Util/AssertExists.lean +++ b/Mathlib/Util/AssertExists.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Scott Morrison -/ +import Mathlib.Init import Lean.Elab.Command /-! diff --git a/Mathlib/Util/AssertNoSorry.lean b/Mathlib/Util/AssertNoSorry.lean index e29f256c1db8b..5f320192712f4 100644 --- a/Mathlib/Util/AssertNoSorry.lean +++ b/Mathlib/Util/AssertNoSorry.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Renshaw -/ +import Mathlib.Init import Lean.Util.CollectAxioms import Lean.Elab.Command diff --git a/Mathlib/Util/AtomM.lean b/Mathlib/Util/AtomM.lean index 0d865b9f138e8..3eb7d0ffd1304 100644 --- a/Mathlib/Util/AtomM.lean +++ b/Mathlib/Util/AtomM.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Init import Lean.Meta.Tactic.Simp.Types /-! diff --git a/Mathlib/Util/CompileInductive.lean b/Mathlib/Util/CompileInductive.lean index fb9468d94d35f..421224f1fa26e 100644 --- a/Mathlib/Util/CompileInductive.lean +++ b/Mathlib/Util/CompileInductive.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Parth Shastri. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Parth Shastri, Gabriel Ebner, Mario Carneiro -/ +import Mathlib.Init import Lean.Elab.Command import Lean.Compiler.CSimpAttr import Lean.Util.FoldConsts diff --git a/Mathlib/Util/CountHeartbeats.lean b/Mathlib/Util/CountHeartbeats.lean index 3161930d15e2c..f9890c218568b 100644 --- a/Mathlib/Util/CountHeartbeats.lean +++ b/Mathlib/Util/CountHeartbeats.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import Mathlib.Init import Lean.Util.Heartbeats import Lean.Meta.Tactic.TryThis diff --git a/Mathlib/Util/Delaborators.lean b/Mathlib/Util/Delaborators.lean index a37232c230318..561a55e9df20d 100644 --- a/Mathlib/Util/Delaborators.lean +++ b/Mathlib/Util/Delaborators.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Init import Lean.PrettyPrinter.Delaborator.Builtins /-! # Pi type notation diff --git a/Mathlib/Util/DischargerAsTactic.lean b/Mathlib/Util/DischargerAsTactic.lean index c9e350e1ee0ab..b7662072b5f0a 100644 --- a/Mathlib/Util/DischargerAsTactic.lean +++ b/Mathlib/Util/DischargerAsTactic.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Alex J. Best. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best -/ +import Mathlib.Init import Lean.Elab.Tactic.Basic import Lean.Meta.Tactic.Simp.Rewrite import Batteries.Tactic.Exact diff --git a/Mathlib/Util/Export.lean b/Mathlib/Util/Export.lean index 1aef9a303dcb8..503e077034a09 100644 --- a/Mathlib/Util/Export.lean +++ b/Mathlib/Util/Export.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Init import Lean.CoreM import Lean.Util.FoldConsts diff --git a/Mathlib/Util/GetAllModules.lean b/Mathlib/Util/GetAllModules.lean index 216d205cb93bc..c0e303c750528 100644 --- a/Mathlib/Util/GetAllModules.lean +++ b/Mathlib/Util/GetAllModules.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kim Morrison, Damiano Testa -/ +import Mathlib.Init import Lean.Util.Path /-! diff --git a/Mathlib/Util/IncludeStr.lean b/Mathlib/Util/IncludeStr.lean index cad43a37b0703..3cc088bfc91bb 100644 --- a/Mathlib/Util/IncludeStr.lean +++ b/Mathlib/Util/IncludeStr.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Henrik Böving. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Henrik Böving, Xubai Wang -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Util/MemoFix.lean b/Mathlib/Util/MemoFix.lean index ab8344cbc6c6c..75a5d913d317d 100644 --- a/Mathlib/Util/MemoFix.lean +++ b/Mathlib/Util/MemoFix.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Gabriel Ebner. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner, Edward Ayers -/ +import Mathlib.Init import Lean.Data.HashMap /-! diff --git a/Mathlib/Util/Qq.lean b/Mathlib/Util/Qq.lean index fee01e5b0f6e4..af1dbac4f7ad8 100644 --- a/Mathlib/Util/Qq.lean +++ b/Mathlib/Util/Qq.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Alex J. Best -/ +import Mathlib.Init import Qq /-! diff --git a/Mathlib/Util/SleepHeartbeats.lean b/Mathlib/Util/SleepHeartbeats.lean index 2fdfdeb4238a3..f6cb76c320e8d 100644 --- a/Mathlib/Util/SleepHeartbeats.lean +++ b/Mathlib/Util/SleepHeartbeats.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Alex J. Best. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex J. Best -/ +import Mathlib.Init import Lean.Elab.Tactic.Basic /-! diff --git a/Mathlib/Util/Superscript.lean b/Mathlib/Util/Superscript.lean index 340ca3a3d174d..1883098e6ed5e 100644 --- a/Mathlib/Util/Superscript.lean +++ b/Mathlib/Util/Superscript.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Init import Batteries.Tactic.Lint /-! diff --git a/Mathlib/Util/SynthesizeUsing.lean b/Mathlib/Util/SynthesizeUsing.lean index c850177038553..a9979d22f051f 100644 --- a/Mathlib/Util/SynthesizeUsing.lean +++ b/Mathlib/Util/SynthesizeUsing.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import Mathlib.Init import Lean.Elab.Tactic.Basic import Qq diff --git a/Mathlib/Util/Tactic.lean b/Mathlib/Util/Tactic.lean index b306cfd4dd034..f5deed46e9463 100644 --- a/Mathlib/Util/Tactic.lean +++ b/Mathlib/Util/Tactic.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Arthur Paulino. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Arthur Paulino, Jannis Limperg -/ +import Mathlib.Init import Lean.MetavarContext /-! diff --git a/Mathlib/Util/TermBeta.lean b/Mathlib/Util/TermBeta.lean index 2cd8f8d0ccca8..f0a8b2ce66a13 100644 --- a/Mathlib/Util/TermBeta.lean +++ b/Mathlib/Util/TermBeta.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Kyle Miller. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ +import Mathlib.Init import Lean.Elab.Term /-! `beta%` term elaborator diff --git a/Mathlib/Util/Time.lean b/Mathlib/Util/Time.lean index 87e2df48ef086..26da53c0c13ce 100644 --- a/Mathlib/Util/Time.lean +++ b/Mathlib/Util/Time.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Util/WhatsNew.lean b/Mathlib/Util/WhatsNew.lean index 6219583ff984b..43e4c7034f986 100644 --- a/Mathlib/Util/WhatsNew.lean +++ b/Mathlib/Util/WhatsNew.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Gabriel Ebner -/ +import Mathlib.Init import Lean /-! diff --git a/Mathlib/Util/WithWeakNamespace.lean b/Mathlib/Util/WithWeakNamespace.lean index 51e164d83a3c6..e03f6fa9d1a16 100644 --- a/Mathlib/Util/WithWeakNamespace.lean +++ b/Mathlib/Util/WithWeakNamespace.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Daniel Selsam, Gabriel Ebner -/ +import Mathlib.Init import Lean /-! diff --git a/scripts/init_creation.sh b/scripts/init_creation.sh new file mode 100644 index 0000000000000..f7469af5e39c6 --- /dev/null +++ b/scripts/init_creation.sh @@ -0,0 +1,40 @@ +#!/usr/bin/env bash + + : <<'BASH_MODULE_DOC' + +These are the commands to generate a "root" `Mathlib/Init.lean` file, imported by all the +`Mathlib` files that do not import any `Mathlib` file. + +BASH_MODULE_DOC + +# `mathlibNonImportingFiles` generates the list of `Mathlib` files that do not have Mathlib imports. +# The output of `lake exe graph` are many lines like +# ` "Mathlib..." -> "Mathlib...";` +# using `"` as separator, the 2nd field is an imported module ("source") and +# the 4th field is a module that imports it ("target"). +# so we want the sources that are not targets. +mathlibNonImportingFiles () { +lake exe graph | + awk -F'"' '($4 ~ ".") { sources[$2]; targets[$4] } + END{ + printf "for f in" + for(t in sources) { if(!(t in targets)) { + gsub(/\./, "/", t) + print t".lean" + } } + }' +} + +printf 'Adding `import Mathlib.Init` to all file that import no Mathlib file.\n' + +# The `sed` command appends the line `import Mathlib.Init` after the first +# `-/[linebreaks]*` of each file printed by `mathlibNonImportingFiles`. +sed -i -z 's=-/\n*=&import Mathlib.Init\n=' $(mathlibNonImportingFiles) + +printf 'Creating `Mathlib/Init.lean`.\n' + +# Creates the `Mathlib/Init.lean` files. +echo '-- This is the root file in Mathlib: it is imported by virtually *all* Mathlib files' > Mathlib/Init.lean + +printf $'Don\'t forget to add `Mathlib.Init` to the `ignoreImport` field of `scripts/noshake.json` +This ensures that `import Mathlib.Init` does not trigger a `shake` exception.\n' diff --git a/scripts/nolints-style.txt b/scripts/nolints-style.txt index efee926577dd7..2b0bbbe992a8d 100644 --- a/scripts/nolints-style.txt +++ b/scripts/nolints-style.txt @@ -5,6 +5,11 @@ -- the former could be necessary in the long term. -- In this case, it's a side-effect of making the linter stricter than its Python ancestor. +-- The `Mathlib/Init.lean` files does not have a copyright header +Mathlib/Init.lean : line 2 : ERR_COP : Malformed or missing copyright header: Copyright line is malformed +Mathlib/Init.lean : line 3 : ERR_COP : Malformed or missing copyright header: Second line should be "Released under Apache 2.0 license as described in the file LICENSE." +Mathlib/Init.lean : line 4 : ERR_COP : Malformed or missing copyright header: The third line should describe the file's main authors + -- This file was recognised as import-only by the old heuristic, but not by the new, simpler one. Mathlib/Tactic/Linter.lean : line 2 : ERR_COP : Malformed or missing copyright header: Copyright line is malformed Mathlib/Tactic/Linter.lean : line 3 : ERR_COP : Malformed or missing copyright header: Second line should be "Released under Apache 2.0 license as described in the file LICENSE." diff --git a/scripts/noshake.json b/scripts/noshake.json index 021b8a75f57ef..61f866068e7e1 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -48,6 +48,7 @@ "Mathlib.Data.Sym.Sym2.Init", "Mathlib.Data.Vector.Basic", "Mathlib.Geometry.Manifold.Instances.Real", + "Mathlib.Init", "Mathlib.Init.Align", "Mathlib.LinearAlgebra.AffineSpace.Basic", "Mathlib.Mathport.Attributes", diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index 9b6d457b4a950..cf67e34b72362 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -1,4 +1,3 @@ -Mathlib/SetTheory/Ordinal/Segfault.lean : line 6 : ERR_MOD : Module docstring missing, or too late Mathlib/Algebra/BigOperators/Group/Finset.lean : line 1 : ERR_NUM_LIN : 2400 file contains 2245 lines, try to split it up Mathlib/Algebra/Group/Subgroup/Basic.lean : line 1 : ERR_NUM_LIN : 3000 file contains 2893 lines, try to split it up Mathlib/Algebra/MonoidAlgebra/Basic.lean : line 1 : ERR_NUM_LIN : 2100 file contains 1965 lines, try to split it up @@ -52,7 +51,6 @@ Mathlib/Order/LiminfLimsup.lean : line 1 : ERR_NUM_LIN : 1900 file contains 1791 Mathlib/Order/UpperLower/Basic.lean : line 1 : ERR_NUM_LIN : 1900 file contains 1795 lines, try to split it up Mathlib/RingTheory/UniqueFactorizationDomain.lean : line 1 : ERR_NUM_LIN : 2100 file contains 1982 lines, try to split it up Mathlib/SetTheory/Cardinal/Basic.lean : line 1 : ERR_NUM_LIN : 2200 file contains 2004 lines, try to split it up -Mathlib/SetTheory/Cardinal/Ordinal.lean : line 1 : ERR_NUM_LIN : 1700 file contains 1501 lines, try to split it up Mathlib/SetTheory/Game/PGame.lean : line 1 : ERR_NUM_LIN : 1900 file contains 1753 lines, try to split it up Mathlib/SetTheory/Ordinal/Arithmetic.lean : line 1 : ERR_NUM_LIN : 2400 file contains 2269 lines, try to split it up Mathlib/SetTheory/ZFC/Basic.lean : line 1 : ERR_NUM_LIN : 1700 file contains 1549 lines, try to split it up From 8e91009d4efe96e409a343d1da64daf0c1a78d1f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 15 Aug 2024 15:27:03 +0000 Subject: [PATCH 301/359] chore: Fix statement of `eq_iff_not_lt_of_le` (#15829) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The statement said `(a ≤ b → b = a) ↔ ¬a < b` instead of the intended `a ≤ b → (b = a ↔ ¬a < b)`. While both are true, the latter is the more useful one. Also pull out a `PartialOrder` section. --- Mathlib/Order/Basic.lean | 55 ++++++++++++++++------------------------ 1 file changed, 22 insertions(+), 33 deletions(-) diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 49a46ba6bb698..8f3d90b665353 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -298,76 +298,65 @@ alias LE.le.not_lt := not_lt_of_le theorem ne_of_not_le [Preorder α] {a b : α} (h : ¬a ≤ b) : a ≠ b := fun hab ↦ h (le_of_eq hab) +section PartialOrder +variable [PartialOrder α] {a b : α} + -- See Note [decidable namespace] -protected theorem Decidable.le_iff_eq_or_lt [PartialOrder α] [@DecidableRel α (· ≤ ·)] {a b : α} : - a ≤ b ↔ a = b ∨ a < b := +protected theorem Decidable.le_iff_eq_or_lt [@DecidableRel α (· ≤ ·)] : a ≤ b ↔ a = b ∨ a < b := Decidable.le_iff_lt_or_eq.trans or_comm -theorem le_iff_eq_or_lt [PartialOrder α] {a b : α} : a ≤ b ↔ a = b ∨ a < b := - le_iff_lt_or_eq.trans or_comm +theorem le_iff_eq_or_lt : a ≤ b ↔ a = b ∨ a < b := le_iff_lt_or_eq.trans or_comm -theorem lt_iff_le_and_ne [PartialOrder α] {a b : α} : a < b ↔ a ≤ b ∧ a ≠ b := +theorem lt_iff_le_and_ne : a < b ↔ a ≤ b ∧ a ≠ b := ⟨fun h ↦ ⟨le_of_lt h, ne_of_lt h⟩, fun ⟨h1, h2⟩ ↦ h1.lt_of_ne h2⟩ -theorem eq_iff_not_lt_of_le [PartialOrder α] {x y : α} : x ≤ y → y = x ↔ ¬x < y := by - rw [lt_iff_le_and_ne, not_and, Classical.not_not, eq_comm] +lemma eq_iff_not_lt_of_le (hab : a ≤ b) : a = b ↔ ¬ a < b := by simp [hab, lt_iff_le_and_ne] + +alias LE.le.eq_iff_not_lt := eq_iff_not_lt_of_le -- See Note [decidable namespace] -protected theorem Decidable.eq_iff_le_not_lt [PartialOrder α] [@DecidableRel α (· ≤ ·)] {a b : α} : +protected theorem Decidable.eq_iff_le_not_lt [@DecidableRel α (· ≤ ·)] : a = b ↔ a ≤ b ∧ ¬a < b := ⟨fun h ↦ ⟨h.le, h ▸ lt_irrefl _⟩, fun ⟨h₁, h₂⟩ ↦ h₁.antisymm <| Decidable.by_contradiction fun h₃ ↦ h₂ (h₁.lt_of_not_le h₃)⟩ -theorem eq_iff_le_not_lt [PartialOrder α] {a b : α} : a = b ↔ a ≤ b ∧ ¬a < b := +theorem eq_iff_le_not_lt : a = b ↔ a ≤ b ∧ ¬a < b := haveI := Classical.dec Decidable.eq_iff_le_not_lt -theorem eq_or_lt_of_le [PartialOrder α] {a b : α} (h : a ≤ b) : a = b ∨ a < b := - h.lt_or_eq.symm - -theorem eq_or_gt_of_le [PartialOrder α] {a b : α} (h : a ≤ b) : b = a ∨ a < b := - h.lt_or_eq.symm.imp Eq.symm id - -theorem gt_or_eq_of_le [PartialOrder α] {a b : α} (h : a ≤ b) : a < b ∨ b = a := - (eq_or_gt_of_le h).symm +theorem eq_or_lt_of_le (h : a ≤ b) : a = b ∨ a < b := h.lt_or_eq.symm +theorem eq_or_gt_of_le (h : a ≤ b) : b = a ∨ a < b := h.lt_or_eq.symm.imp Eq.symm id +theorem gt_or_eq_of_le (h : a ≤ b) : a < b ∨ b = a := (eq_or_gt_of_le h).symm alias LE.le.eq_or_lt_dec := Decidable.eq_or_lt_of_le - alias LE.le.eq_or_lt := eq_or_lt_of_le - alias LE.le.eq_or_gt := eq_or_gt_of_le - alias LE.le.gt_or_eq := gt_or_eq_of_le -- Porting note: no `decidable_classical` linter -- attribute [nolint decidable_classical] LE.le.eq_or_lt_dec -theorem eq_of_le_of_not_lt [PartialOrder α] {a b : α} (hab : a ≤ b) (hba : ¬a < b) : a = b := - hab.eq_or_lt.resolve_right hba - -theorem eq_of_ge_of_not_gt [PartialOrder α] {a b : α} (hab : a ≤ b) (hba : ¬a < b) : b = a := - (hab.eq_or_lt.resolve_right hba).symm +theorem eq_of_le_of_not_lt (hab : a ≤ b) (hba : ¬a < b) : a = b := hab.eq_or_lt.resolve_right hba +theorem eq_of_ge_of_not_gt (hab : a ≤ b) (hba : ¬a < b) : b = a := (eq_of_le_of_not_lt hab hba).symm alias LE.le.eq_of_not_lt := eq_of_le_of_not_lt - alias LE.le.eq_of_not_gt := eq_of_ge_of_not_gt -theorem Ne.le_iff_lt [PartialOrder α] {a b : α} (h : a ≠ b) : a ≤ b ↔ a < b := - ⟨fun h' ↦ lt_of_le_of_ne h' h, fun h ↦ h.le⟩ +theorem Ne.le_iff_lt (h : a ≠ b) : a ≤ b ↔ a < b := ⟨fun h' ↦ lt_of_le_of_ne h' h, fun h ↦ h.le⟩ -theorem Ne.not_le_or_not_le [PartialOrder α] {a b : α} (h : a ≠ b) : ¬a ≤ b ∨ ¬b ≤ a := - not_and_or.1 <| le_antisymm_iff.not.1 h +theorem Ne.not_le_or_not_le (h : a ≠ b) : ¬a ≤ b ∨ ¬b ≤ a := not_and_or.1 <| le_antisymm_iff.not.1 h -- See Note [decidable namespace] -protected theorem Decidable.ne_iff_lt_iff_le [PartialOrder α] [DecidableEq α] {a b : α} : - (a ≠ b ↔ a < b) ↔ a ≤ b := +protected theorem Decidable.ne_iff_lt_iff_le [DecidableEq α] : (a ≠ b ↔ a < b) ↔ a ≤ b := ⟨fun h ↦ Decidable.byCases le_of_eq (le_of_lt ∘ h.mp), fun h ↦ ⟨lt_of_le_of_ne h, ne_of_lt⟩⟩ @[simp] -theorem ne_iff_lt_iff_le [PartialOrder α] {a b : α} : (a ≠ b ↔ a < b) ↔ a ≤ b := +theorem ne_iff_lt_iff_le : (a ≠ b ↔ a < b) ↔ a ≤ b := haveI := Classical.dec Decidable.ne_iff_lt_iff_le +end PartialOrder + -- Variant of `min_def` with the branches reversed. theorem min_def' [LinearOrder α] (a b : α) : min a b = if b ≤ a then b else a := by rw [min_def] From 28a935cb927c6f2c8a0e27d90ab79892fa042af5 Mon Sep 17 00:00:00 2001 From: Artie Khovanov Date: Thu, 15 Aug 2024 16:43:06 +0000 Subject: [PATCH 302/359] chore(Algebra/Ring/Semireal/Defs): relax `IsSemireal` typeclass constraint (#15844) Typeclass constraint on instance of `IsSemireal` relaxed from `LinearOrderedField` to `LinearOrderedRing`. This instance requires `LinearOrderedSemiring` and well as negation, so this is now the most general instance possible. --- Mathlib/Algebra/Ring/Semireal/Defs.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Ring/Semireal/Defs.lean b/Mathlib/Algebra/Ring/Semireal/Defs.lean index 7e26be34cfffb..7b16914e15ab1 100644 --- a/Mathlib/Algebra/Ring/Semireal/Defs.lean +++ b/Mathlib/Algebra/Ring/Semireal/Defs.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Florent Schaffhauser -/ import Mathlib.Algebra.Ring.SumsOfSquares -import Mathlib.Algebra.Order.Field.Defs /-! # Semireal rings @@ -42,6 +41,6 @@ class IsSemireal [AddMonoid R] [Mul R] [One R] [Neg R] : Prop where @[deprecated (since := "2024-08-09")] alias isSemireal.neg_one_not_SumSq := IsSemireal.not_isSumSq_neg_one -instance [LinearOrderedField R] : IsSemireal R where +instance [LinearOrderedRing R] : IsSemireal R where non_trivial := zero_ne_one not_isSumSq_neg_one := fun h ↦ (not_le (α := R)).2 neg_one_lt_zero h.nonneg From 6572a1f9efe9584614dfd43e21635258d8b12809 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Thu, 15 Aug 2024 17:24:10 +0000 Subject: [PATCH 303/359] feat(RelSeries): eraseLast_last_rel_last (#15386) from the Carleson project --- Mathlib/Order/RelSeries.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index 6b905496d8037..2c25896fdd070 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -496,6 +496,14 @@ def eraseLast (p : RelSeries r) : RelSeries r where @[simp] lemma last_eraseLast (p : RelSeries r) : p.eraseLast.last = p ⟨p.length.pred, Nat.lt_succ_iff.2 (Nat.pred_le _)⟩ := rfl + +/-- In a non-trivial series `p`, the last element of `p.eraseLast` is related to `p.last` -/ +lemma eraseLast_last_rel_last (p : RelSeries r) (h : p.length ≠ 0) : + r p.eraseLast.last p.last := by + simp only [last, Fin.last, eraseLast_length, eraseLast_toFun] + convert p.step ⟨p.length - 1, by omega⟩ + simp only [Nat.succ_eq_add_one, Fin.succ_mk]; omega + /-- Given two series of the form `a₀ -r→ ... -r→ X` and `X -r→ b ---> ...`, then `a₀ -r→ ... -r→ X -r→ b ...` is another series obtained by combining the given two. @@ -674,6 +682,8 @@ lemma strictMono (x : LTSeries α) : StrictMono x := lemma monotone (x : LTSeries α) : Monotone x := x.strictMono.monotone +lemma head_le_last (x : LTSeries α) : x.head ≤ x.last := + LTSeries.monotone x (Fin.zero_le _) /-- An alternative constructor of `LTSeries` from a strictly monotone function. -/ @[simps] From 25706a5c05bde5439153883b1a06e405be024b76 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 15 Aug 2024 17:24:11 +0000 Subject: [PATCH 304/359] =?UTF-8?q?feat:=20Compute=20derivative=20of=20`fu?= =?UTF-8?q?n=20x=20=E2=86=A6=20f=20(a=20-=20x)`=20(#15828)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ... and similar. Also generalise existing lemmas to not require differentiability assumptions. --- Mathlib/Analysis/Calculus/Deriv/Add.lean | 50 +++++++++++++- Mathlib/Analysis/Calculus/Deriv/Shift.lean | 69 +++++++++++-------- Mathlib/NumberTheory/Harmonic/GammaDeriv.lean | 3 +- 3 files changed, 90 insertions(+), 32 deletions(-) diff --git a/Mathlib/Analysis/Calculus/Deriv/Add.lean b/Mathlib/Analysis/Calculus/Deriv/Add.lean index bc4c7216fb90a..d49ad898b452c 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Add.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Add.lean @@ -120,6 +120,26 @@ theorem deriv_const_add (c : F) : deriv (fun y => c + f y) x = deriv f x := by theorem deriv_const_add' (c : F) : (deriv fun y => c + f y) = deriv f := funext fun _ => deriv_const_add c +lemma differentiableAt_comp_const_add {a b : 𝕜} : + DifferentiableAt 𝕜 (fun x ↦ f (b + x)) a ↔ DifferentiableAt 𝕜 f (b + a) := by + refine ⟨fun H ↦ ?_, fun H ↦ H.comp _ (differentiable_id.const_add _).differentiableAt⟩ + convert DifferentiableAt.comp (b + a) (by simpa) + (differentiable_id.const_add (-b)).differentiableAt + ext + simp + +lemma differentiableAt_comp_add_const {a b : 𝕜} : + DifferentiableAt 𝕜 (fun x ↦ f (x + b)) a ↔ DifferentiableAt 𝕜 f (a + b) := by + simpa [add_comm b] using differentiableAt_comp_const_add (f := f) (b := b) + +lemma differentiableAt_iff_comp_const_add {a b : 𝕜} : + DifferentiableAt 𝕜 f a ↔ DifferentiableAt 𝕜 (fun x ↦ f (b + x)) (-b + a) := by + simp [differentiableAt_comp_const_add] + +lemma differentiableAt_iff_comp_add_const {a b : 𝕜} : + DifferentiableAt 𝕜 f a ↔ DifferentiableAt 𝕜 (fun x ↦ f (x + b)) (a - b) := by + simp [differentiableAt_comp_add_const] + end Add section Sum @@ -237,13 +257,17 @@ theorem not_differentiableAt_abs_zero : ¬ DifferentiableAt ℝ (abs : ℝ → (hasDerivWithinAt_neg _ _).congr_of_mem (fun _ h ↦ abs_of_nonpos h) Set.right_mem_Iic linarith -lemma differentiableAt_comp_neg_iff {a : 𝕜} : - DifferentiableAt 𝕜 f (-a) ↔ DifferentiableAt 𝕜 (fun x ↦ f (-x)) a := by - refine ⟨fun H ↦ H.comp a differentiable_neg.differentiableAt, fun H ↦ ?_⟩ +lemma differentiableAt_comp_neg {a : 𝕜} : + DifferentiableAt 𝕜 (fun x ↦ f (-x)) a ↔ DifferentiableAt 𝕜 f (-a) := by + refine ⟨fun H ↦ ?_, fun H ↦ H.comp a differentiable_neg.differentiableAt⟩ convert ((neg_neg a).symm ▸ H).comp (-a) differentiable_neg.differentiableAt ext simp only [Function.comp_apply, neg_neg] +lemma differentiableAt_iff_comp_neg {a : 𝕜} : + DifferentiableAt 𝕜 f a ↔ DifferentiableAt 𝕜 (fun x ↦ f (-x)) (-a) := by + simp_rw [← differentiableAt_comp_neg, neg_neg] + end Neg2 section Sub @@ -319,4 +343,24 @@ theorem deriv_const_sub (c : F) : deriv (fun y => c - f y) x = -deriv f x := by simp only [← derivWithin_univ, derivWithin_const_sub (uniqueDiffWithinAt_univ : UniqueDiffWithinAt 𝕜 _ _)] +lemma differentiableAt_comp_sub_const {a b : 𝕜} : + DifferentiableAt 𝕜 (fun x ↦ f (x - b)) a ↔ DifferentiableAt 𝕜 f (a - b) := by + simp [sub_eq_add_neg, differentiableAt_comp_add_const] + +lemma differentiableAt_comp_const_sub {a b : 𝕜} : + DifferentiableAt 𝕜 (fun x ↦ f (b - x)) a ↔ DifferentiableAt 𝕜 f (b - a) := by + refine ⟨fun H ↦ ?_, fun H ↦ H.comp a (differentiable_id.const_sub _).differentiableAt⟩ + convert ((sub_sub_cancel _ a).symm ▸ H).comp (b - a) + (differentiable_id.const_sub _).differentiableAt + ext + simp + +lemma differentiableAt_iff_comp_sub_const {a b : 𝕜} : + DifferentiableAt 𝕜 f a ↔ DifferentiableAt 𝕜 (fun x ↦ f (x - b)) (a + b) := by + simp [sub_eq_add_neg, differentiableAt_comp_add_const] + +lemma differentiableAt_iff_comp_const_sub {a b : 𝕜} : + DifferentiableAt 𝕜 f a ↔ DifferentiableAt 𝕜 (fun x ↦ f (b - x)) (b - a) := by + simp [differentiableAt_comp_const_sub] + end Sub diff --git a/Mathlib/Analysis/Calculus/Deriv/Shift.lean b/Mathlib/Analysis/Calculus/Deriv/Shift.lean index 777f322714c8b..ae0d346bba6f0 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Shift.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Shift.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2023 Michael Stoll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Michael Stoll +Authors: Michael Stoll, Yaël Dillies -/ import Mathlib.Analysis.Calculus.Deriv.Add import Mathlib.Analysis.Calculus.Deriv.Comp @@ -9,40 +9,55 @@ import Mathlib.Analysis.Calculus.Deriv.Comp /-! ### Invariance of the derivative under translation -We show that if a function `h` has derivative `h'` at a point `a + x`, then `h (a + ·)` -has derivative `h'` at `x`. Similarly for `x + a`. +We show that if a function `f` has derivative `f'` at a point `a + x`, then `f (a + ·)` +has derivative `f'` at `x`. Similarly for `x + a`. -/ +variable {𝕜 F : Type*} [NontriviallyNormedField 𝕜] [NormedAddCommGroup F] [NormedSpace 𝕜 F] + {f : 𝕜 → F} {f' : F} + +/-- Translation in the domain does not change the derivative. -/ +lemma HasDerivAt.comp_const_add (a x : 𝕜) (hf : HasDerivAt f f' (a + x)) : + HasDerivAt (fun x ↦ f (a + x)) f' x := by + simpa [Function.comp_def] using HasDerivAt.scomp (𝕜 := 𝕜) x hf <| hasDerivAt_id' x |>.const_add a + +/-- Translation in the domain does not change the derivative. -/ +lemma HasDerivAt.comp_add_const (x a : 𝕜) (hf : HasDerivAt f f' (x + a)) : + HasDerivAt (fun x ↦ f (x + a)) f' x := by + simpa [Function.comp_def] using HasDerivAt.scomp (𝕜 := 𝕜) x hf <| hasDerivAt_id' x |>.add_const a + /-- Translation in the domain does not change the derivative. -/ -lemma HasDerivAt.comp_const_add {𝕜 : Type*} [NontriviallyNormedField 𝕜] (a x : 𝕜) {𝕜' : Type*} - [NormedAddCommGroup 𝕜'] [NormedSpace 𝕜 𝕜'] {h : 𝕜 → 𝕜'} {h' : 𝕜'} - (hh : HasDerivAt h h' (a + x)) : - HasDerivAt (fun x ↦ h (a + x)) h' x := by - simpa [Function.comp_def] using HasDerivAt.scomp (𝕜 := 𝕜) x hh <| hasDerivAt_id' x |>.const_add a +lemma HasDerivAt.comp_const_sub (a x : 𝕜) (hf : HasDerivAt f f' (a - x)) : + HasDerivAt (fun x ↦ f (a - x)) (-f') x := by + simpa [Function.comp_def] using HasDerivAt.scomp (𝕜 := 𝕜) x hf <| hasDerivAt_id' x |>.const_sub a /-- Translation in the domain does not change the derivative. -/ -lemma HasDerivAt.comp_add_const {𝕜 : Type*} [NontriviallyNormedField 𝕜] (x a : 𝕜) {𝕜' : Type*} - [NormedAddCommGroup 𝕜'] [NormedSpace 𝕜 𝕜'] {h : 𝕜 → 𝕜'} {h' : 𝕜'} - (hh : HasDerivAt h h' (x + a)) : - HasDerivAt (fun x ↦ h (x + a)) h' x := by - simpa [Function.comp_def] using HasDerivAt.scomp (𝕜 := 𝕜) x hh <| hasDerivAt_id' x |>.add_const a +lemma HasDerivAt.comp_sub_const (x a : 𝕜) (hf : HasDerivAt f f' (x - a)) : + HasDerivAt (fun x ↦ f (x - a)) f' x := by + simpa [Function.comp_def] using HasDerivAt.scomp (𝕜 := 𝕜) x hf <| hasDerivAt_id' x |>.sub_const a + +variable (f) (a x : 𝕜) /-- The derivative of `x ↦ f (-x)` at `a` is the negative of the derivative of `f` at `-a`. -/ -lemma deriv_comp_neg {𝕜 : Type*} [NontriviallyNormedField 𝕜] {F : Type*} [NormedAddCommGroup F] - [NormedSpace 𝕜 F] (f : 𝕜 → F) (a : 𝕜) : deriv (fun x ↦ f (-x)) a = -deriv f (-a) := by - by_cases h : DifferentiableAt 𝕜 f (-a) - · simpa only [deriv_neg, neg_one_smul] using deriv.scomp a h (differentiable_neg _) - · rw [deriv_zero_of_not_differentiableAt (mt differentiableAt_comp_neg_iff.mpr h), - deriv_zero_of_not_differentiableAt h, neg_zero] +lemma deriv_comp_neg : deriv (fun x ↦ f (-x)) x = -deriv f (-x) := by + by_cases f : DifferentiableAt 𝕜 f (-x) + · simpa only [deriv_neg, neg_one_smul] using deriv.scomp _ f (differentiable_neg _) + · rw [deriv_zero_of_not_differentiableAt (differentiableAt_comp_neg.not.2 f), + deriv_zero_of_not_differentiableAt f, neg_zero] /-- Translation in the domain does not change the derivative. -/ -lemma deriv_comp_const_add {𝕜 : Type*} [NontriviallyNormedField 𝕜] (a x : 𝕜) {𝕜' : Type*} - [NormedAddCommGroup 𝕜'] [NormedSpace 𝕜 𝕜'] {h : 𝕜 → 𝕜'} - (hh : DifferentiableAt 𝕜 h (a + x)) : - deriv (fun x ↦ h (a + x)) x = deriv h (a + x) := HasDerivAt.deriv hh.hasDerivAt.comp_const_add +lemma deriv_comp_const_add : deriv (fun x ↦ f (a + x)) x = deriv f (a + x) := by + by_cases hf : DifferentiableAt 𝕜 f (a + x) + · exact HasDerivAt.deriv hf.hasDerivAt.comp_const_add + · rw [deriv_zero_of_not_differentiableAt (differentiableAt_comp_const_add.not.2 hf), + deriv_zero_of_not_differentiableAt hf] /-- Translation in the domain does not change the derivative. -/ -lemma deriv_comp_add_const {𝕜 : Type*} [NontriviallyNormedField 𝕜] (a x : 𝕜) {𝕜' : Type*} - [NormedAddCommGroup 𝕜'] [NormedSpace 𝕜 𝕜'] {h : 𝕜 → 𝕜'} - (hh : DifferentiableAt 𝕜 h (x + a)) : - deriv (fun x ↦ h (x + a)) x = deriv h (x + a) := HasDerivAt.deriv hh.hasDerivAt.comp_add_const +lemma deriv_comp_add_const : deriv (fun x ↦ f (x + a)) x = deriv f (x + a) := by + simpa [add_comm] using deriv_comp_const_add f a x + +lemma deriv_comp_const_sub : deriv (fun x ↦ f (a - x)) x = -deriv f (a - x) := by + simp_rw [sub_eq_add_neg, deriv_comp_neg (f $ a + ·), deriv_comp_const_add] + +lemma deriv_comp_sub_const : deriv (fun x ↦ f (x - a)) x = deriv f (x - a) := by + simp_rw [sub_eq_add_neg, deriv_comp_add_const] diff --git a/Mathlib/NumberTheory/Harmonic/GammaDeriv.lean b/Mathlib/NumberTheory/Harmonic/GammaDeriv.lean index 35d6bf5409607..a66d5d7a47b91 100644 --- a/Mathlib/NumberTheory/Harmonic/GammaDeriv.lean +++ b/Mathlib/NumberTheory/Harmonic/GammaDeriv.lean @@ -46,7 +46,7 @@ lemma deriv_Gamma_nat (n : ℕ) : exact fun m ↦ ne_of_gt (by linarith) -- Express derivative at general `n` in terms of value at `1` using recurrence relation have hder_rec (x : ℝ) (hx : 0 < x) : deriv f (x + 1) = deriv f x + 1 / x := by - rw [← deriv_comp_add_const _ _ (hder <| by positivity), one_div, ← deriv_log, + rw [← deriv_comp_add_const, one_div, ← deriv_log, ← deriv_add (hder <| by positivity) (differentiableAt_log hx.ne')] apply EventuallyEq.deriv_eq filter_upwards [eventually_gt_nhds hx] using h_rec @@ -108,7 +108,6 @@ lemma hasDerivAt_Gamma_one_half : HasDerivAt Gamma (-√π * (γ + 2 * log 2)) ( (by norm_num : 1/2 + 1/2 = (1 : ℝ)), Gamma_one, mul_one, eulerMascheroniConstant_eq_neg_deriv, add_neg_cancel, mul_zero, add_zero] · apply h_diff; norm_num -- s = 1 - · apply h_diff; norm_num -- s = 1/2 · exact ((h_diff (by norm_num)).hasDerivAt.comp_add_const).differentiableAt -- s = 1 _ = (deriv (fun s ↦ Gamma (2 * s) * 2 ^ (1 - 2 * s) * √π) (1 / 2)) + √π * γ := by rw [funext Gamma_mul_Gamma_add_half] From 5b6785cc6960c32620eafe6bca60d2271169d4e0 Mon Sep 17 00:00:00 2001 From: grunweg Date: Thu, 15 Aug 2024 17:44:48 +0000 Subject: [PATCH 305/359] chore: fix a few typos in docstrings (#15834) --- Mathlib/Algebra/Algebra/Subalgebra/Basic.lean | 2 +- Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean | 3 +-- Mathlib/AlgebraicTopology/SingularSet.lean | 2 +- Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean | 2 +- 4 files changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index 81ea05a0347ff..a76d7cd9d500a 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -241,7 +241,7 @@ def toSubmodule : Subalgebra R A ↪o Submodule R A where map_rel_iff' := SetLike.coe_subset_coe.symm.trans SetLike.coe_subset_coe /- TODO: bundle other forgetful maps between algebraic substructures, e.g. - `to_subsemiring` and `to_subring` in this file. -/ + `toSubsemiring` and `toSubring` in this file. -/ @[simp] theorem mem_toSubmodule {x} : x ∈ (toSubmodule S) ↔ x ∈ S := Iff.rfl diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 5785868dc943e..02c18404835ab 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -68,7 +68,7 @@ For a homogeneous element `f` of degree `m` defined by sending `x : Proj| (pbo f)` to `A⁰_f ∩ span {g / 1 | g ∈ x}`. We also denote this map as `ψ`. * `ProjIsoSpecTopComponent.ToSpec.preimage_eq`: for any `a: A`, if `a/f^m` has degree zero, - then the preimage of `sbo a/f^m` under `to_Spec f` is `pbo f ∩ pbo a`. + then the preimage of `sbo a/f^m` under `toSpec f` is `pbo f ∩ pbo a`. If we further assume `m` is positive * `ProjIsoSpecTopComponent.fromSpec`: the continuous map between `Spec.T A⁰_f` and `Proj.T| pbo f` @@ -88,7 +88,6 @@ Finally, * [Robin Hartshorne, *Algebraic Geometry*][Har77]: Chapter II.2 Proposition 2.5 -/ - noncomputable section diff --git a/Mathlib/AlgebraicTopology/SingularSet.lean b/Mathlib/AlgebraicTopology/SingularSet.lean index f7da42f37d939..21364f30fc245 100644 --- a/Mathlib/AlgebraicTopology/SingularSet.lean +++ b/Mathlib/AlgebraicTopology/SingularSet.lean @@ -11,7 +11,7 @@ import Mathlib.Topology.Category.TopCat.Limits.Basic /-! # The singular simplicial set of a topological space and geometric realization of a simplicial set -The *singular simplicial set* `TopCat.to_SSet.obj X` of a topological space `X` +The *singular simplicial set* `TopCat.toSSet.obj X` of a topological space `X` has as `n`-simplices the continuous maps `[n].toTop → X`. Here, `[n].toTop` is the standard topological `n`-simplex, defined as `{ f : Fin (n+1) → ℝ≥0 // ∑ i, f i = 1 }` with its subspace topology. diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean b/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean index b1e78a20544fa..6fd78154fd5a4 100644 --- a/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean +++ b/Mathlib/Topology/Algebra/Nonarchimedean/Bases.lean @@ -327,7 +327,7 @@ theorem nonarchimedean (hB : SubmodulesBasis B) : @NonarchimedeanAddGroup M _ hB library_note "nonarchimedean non instances"/-- The non archimedean subgroup basis lemmas cannot be instances because some instances -(such as `MeasureTheory.AEEqFun.instAddMonoid` or `topological_add_group.to_has_continuous_add`) +(such as `MeasureTheory.AEEqFun.instAddMonoid` or `TopologicalAddGroup.toContinuousAdd`) cause the search for `@TopologicalAddGroup β ?m1 ?m2`, i.e. a search for a topological group where the topology/group structure are unknown. -/ From a2c3bef6080c5ed16bdf8daaf6e723e3da52834c Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 15 Aug 2024 18:22:13 +0000 Subject: [PATCH 306/359] fix: silence `aesop` in `cfc_cont_tac` (#15847) If `aesop` fails to discharge a goal in one branch of `fun_prop`, as implemened in `cfc_cont_tac`, it can emit a warning saying it failed to close the goal. If `fun_prop (disch := aesop)` then *succeeds* in a later branch of the same call, then the proof succeeds, but the warning is still emitted, causing the file to be noisy. Even worse, because this is called as an `autoParam`, there is no syntax in the file on which to emit the warning, so it just gets emitted at the top of the file. This changes the configuration of `aesop` to `warnOnNonterminal := false` when used in `cfc_cont_tac`. Note: this won't affect debugging or proof writing, because in that case the user will need to disable the use of the `autoParam` anyway by writing their own terms (or `fun_prop` calls). --- Mathlib/Tactic/ContinuousFunctionalCalculus.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/ContinuousFunctionalCalculus.lean b/Mathlib/Tactic/ContinuousFunctionalCalculus.lean index d65a4da1a676c..0adbd41b52a2f 100644 --- a/Mathlib/Tactic/ContinuousFunctionalCalculus.lean +++ b/Mathlib/Tactic/ContinuousFunctionalCalculus.lean @@ -26,7 +26,10 @@ macro_rules specifically concerning continuity of the functions involved. -/ syntax (name := cfcContTac) "cfc_cont_tac" : tactic macro_rules - | `(tactic| cfc_cont_tac) => `(tactic| try (first | fun_prop (disch := aesop) | assumption)) + | `(tactic| cfc_cont_tac) => + `(tactic| try (first + | fun_prop (disch := aesop (config := {warnOnNonterminal := false})) + | assumption)) /-- A tactic used to automatically discharge goals relating to the non-unital continuous functional calculus, specifically concerning whether `f 0 = 0`. -/ From 14cabf6d543b390026c85d941b3ad50f81a7a11b Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 15 Aug 2024 19:17:43 +0000 Subject: [PATCH 307/359] CI: add deprecated files to tech debts counters (#15853) Adds information about the `Deprecated` folder to the technical debt report. [Asked on Zulip](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Technical.20Debt.20Counters/near/462615354) --- scripts/technical-debt-metrics.sh | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/scripts/technical-debt-metrics.sh b/scripts/technical-debt-metrics.sh index 4e5f937b7aad9..cd49fd69bfaa1 100755 --- a/scripts/technical-debt-metrics.sh +++ b/scripts/technical-debt-metrics.sh @@ -59,6 +59,11 @@ printf '%s|%s\n' "$(git grep "^open .*Classical" | grep -v " in$" -c)" "bare ope # We print the number of files, not the number of matches --- hence, the nested grep. printf '%s|%s\n' "$(git grep -c 'autoImplicit true' | grep -c -v 'test')" "non-test files with autoImplicit true" +deprecatedFiles="$(git ls-files '**/Deprecated/*.lean' | xargs wc -l | sed 's=^ *==')" + +printf '%s|%s\n' "$(printf '%s' "${deprecatedFiles}" | wc -l)" "\`Deprecated\` files" +printf '%s|%s\n' "$(printf '%s\n' "${deprecatedFiles}" | grep total | sed 's= total==')" 'total LoC in `Deprecated` files' + initFiles="$(git ls-files '**/Init/*.lean' | xargs wc -l | sed 's=^ *==')" printf '%s|%s\n' "$(printf '%s' "${initFiles}" | wc -l)" "\`Init\` files" From b73e0accb1f783c3f3347523a90b7eecba0b150f Mon Sep 17 00:00:00 2001 From: Nick Mertes Date: Thu, 15 Aug 2024 19:17:44 +0000 Subject: [PATCH 308/359] feat(CategoryTheory/Category/RelCat): Make `rel_iso_iff` universe polymorphic. (#15854) Make `rel_iso_iff` universe polymorphic by changing `C := Type` to `C := Type u`. --- Mathlib/CategoryTheory/Category/RelCat.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/CategoryTheory/Category/RelCat.lean b/Mathlib/CategoryTheory/Category/RelCat.lean index b505cec20e646..48f8c610da96b 100644 --- a/Mathlib/CategoryTheory/Category/RelCat.lean +++ b/Mathlib/CategoryTheory/Category/RelCat.lean @@ -27,7 +27,7 @@ universe u -- This file is about Lean 3 declaration "Rel". -/-- A type synonym for `Type`, which carries the category instance for which +/-- A type synonym for `Type u`, which carries the category instance for which morphisms are binary relations. -/ def RelCat := Type u @@ -83,9 +83,9 @@ instance graphFunctor_essSurj : graphFunctor.EssSurj := graphFunctor.essSurj_of_surj Function.surjective_id /-- A relation is an isomorphism in `RelCat` iff it is the image of an isomorphism in -`Type`. -/ +`Type u`. -/ theorem rel_iso_iff {X Y : RelCat} (r : X ⟶ Y) : - IsIso (C := RelCat) r ↔ ∃ f : (Iso (C := Type) X Y), graphFunctor.map f.hom = r := by + IsIso (C := RelCat) r ↔ ∃ f : (Iso (C := Type u) X Y), graphFunctor.map f.hom = r := by constructor · intro h have h1 := congr_fun₂ h.hom_inv_id @@ -93,7 +93,7 @@ theorem rel_iso_iff {X Y : RelCat} (r : X ⟶ Y) : simp only [RelCat.Hom.rel_comp_apply₂, RelCat.Hom.rel_id_apply₂, eq_iff_iff] at h1 h2 obtain ⟨f, hf⟩ := Classical.axiomOfChoice (fun a => (h1 a a).mpr rfl) obtain ⟨g, hg⟩ := Classical.axiomOfChoice (fun a => (h2 a a).mpr rfl) - suffices hif : IsIso (C := Type) f by + suffices hif : IsIso (C := Type u) f by use asIso f ext x y simp only [asIso_hom, graphFunctor_map] From c98ad2f61380a2c6a894fd76473450c05147c7d4 Mon Sep 17 00:00:00 2001 From: Rida Hamadani Date: Thu, 15 Aug 2024 22:11:06 +0000 Subject: [PATCH 309/359] feat(Combinatorics/SimpleGraph): Define diameter of simple graphs and provide basic lemmas related to it (#12058) This defines the diameter of a simple graph, and provides a basic API for it. Co-authored-by: Matt Diamond Co-authored-by: Rida Hamadani <106540880+Rida-Hamadani@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/Combinatorics/SimpleGraph/Diam.lean | 203 ++++++++++++++++++ Mathlib/Data/ENat/Lattice.lean | 9 + .../ConditionallyCompleteLattice/Finset.lean | 6 + 4 files changed, 219 insertions(+) create mode 100644 Mathlib/Combinatorics/SimpleGraph/Diam.lean diff --git a/Mathlib.lean b/Mathlib.lean index 51a6ebf92e2b5..710dcbc9c38a5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1918,6 +1918,7 @@ import Mathlib.Combinatorics.SimpleGraph.Connectivity.WalkCounting import Mathlib.Combinatorics.SimpleGraph.Dart import Mathlib.Combinatorics.SimpleGraph.DegreeSum import Mathlib.Combinatorics.SimpleGraph.Density +import Mathlib.Combinatorics.SimpleGraph.Diam import Mathlib.Combinatorics.SimpleGraph.Ends.Defs import Mathlib.Combinatorics.SimpleGraph.Ends.Properties import Mathlib.Combinatorics.SimpleGraph.Finite diff --git a/Mathlib/Combinatorics/SimpleGraph/Diam.lean b/Mathlib/Combinatorics/SimpleGraph/Diam.lean new file mode 100644 index 0000000000000..42ae812095c70 --- /dev/null +++ b/Mathlib/Combinatorics/SimpleGraph/Diam.lean @@ -0,0 +1,203 @@ +/- +Copyright (c) 2024 Rida Hamadani. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rida Hamadani +-/ +import Mathlib.Combinatorics.SimpleGraph.Metric + +/-! +# Diameter of a simple graph + +This module defines the diameter of a simple graph, which measures the maximum distance between +vertices. + +## Main definitions + +- `SimpleGraph.ediam` is the graph extended diameter. + +- `SimpleGraph.diam` is the graph diameter. + +## Todo + +- Prove that `G.egirth ≤ 2 * G.ediam + 1` and `G.girth ≤ 2 * G.diam + 1` when the diameter is + non-zero. + +-/ + +namespace SimpleGraph +variable {α : Type*} {G G' : SimpleGraph α} + +section ediam + +/-- +The extended diameter is the greatest distance between any two vertices, with the value `⊤` in +case the distances are not bounded above, or the graph is not connected. +-/ +noncomputable def ediam (G : SimpleGraph α) : ℕ∞ := + ⨆ u, ⨆ v, G.edist u v + +lemma ediam_def : G.ediam = ⨆ p : α × α, G.edist p.1 p.2 := by + rw [ediam, iSup_prod] + +lemma edist_le_ediam {u v : α} : G.edist u v ≤ G.ediam := + le_iSup₂ (f := G.edist) u v + +lemma ediam_le_of_edist_le {k : ℕ∞} (h : ∀ u v, G.edist u v ≤ k ) : G.ediam ≤ k := + iSup₂_le h + +lemma ediam_le_iff {k : ℕ∞} : G.ediam ≤ k ↔ ∀ u v, G.edist u v ≤ k := + iSup₂_le_iff + +lemma ediam_eq_top : G.ediam = ⊤ ↔ ∀ b < ⊤, ∃ u v, b < G.edist u v := by + simp only [ediam, iSup_eq_top, lt_iSup_iff] + +lemma ediam_eq_zero_of_subsingleton [Subsingleton α] : G.ediam = 0 := by + rw [ediam_def, ENat.iSup_eq_zero] + simpa [edist_eq_zero_iff, Prod.forall] using subsingleton_iff.mp ‹_› + +lemma nontrivial_of_ediam_ne_zero (h : G.ediam ≠ 0) : Nontrivial α := by + contrapose! h + rw [not_nontrivial_iff_subsingleton] at h + exact ediam_eq_zero_of_subsingleton + +lemma ediam_ne_zero [Nontrivial α] : G.ediam ≠ 0 := by + obtain ⟨u, v, huv⟩ := exists_pair_ne ‹_› + contrapose! huv + simp only [ediam, nonpos_iff_eq_zero, ENat.iSup_eq_zero, edist_eq_zero_iff] at huv + exact huv u v + +lemma subsingleton_of_ediam_eq_zero (h : G.ediam = 0) : Subsingleton α := by + contrapose! h + apply not_subsingleton_iff_nontrivial.mp at h + exact ediam_ne_zero + +lemma ediam_ne_zero_iff_nontrivial : + G.ediam ≠ 0 ↔ Nontrivial α := + ⟨nontrivial_of_ediam_ne_zero, fun _ ↦ ediam_ne_zero⟩ + +@[simp] +lemma ediam_eq_zero_iff_subsingleton : + G.ediam = 0 ↔ Subsingleton α := + ⟨subsingleton_of_ediam_eq_zero, fun _ ↦ ediam_eq_zero_of_subsingleton⟩ + +lemma ediam_eq_top_of_not_connected [Nonempty α] (h : ¬G.Connected) : G.ediam = ⊤ := by + rw [connected_iff_exists_forall_reachable] at h + push_neg at h + obtain ⟨_, hw⟩ := h Classical.ofNonempty + rw [eq_top_iff, ← edist_eq_top_of_not_reachable hw] + exact edist_le_ediam + +lemma ediam_eq_top_of_not_preconnected (h : ¬G.Preconnected) : G.ediam = ⊤ := by + cases isEmpty_or_nonempty α + · exfalso + exact h <| IsEmpty.forall_iff.mpr trivial + · apply ediam_eq_top_of_not_connected + rw [connected_iff] + tauto + +lemma exists_edist_eq_ediam_of_ne_top [Nonempty α] (h : G.ediam ≠ ⊤) : + ∃ u v, G.edist u v = G.ediam := + ENat.exists_eq_iSup₂_of_lt_top h.lt_top + +-- Note: Neither `Finite α` nor `G.ediam ≠ ⊤` implies the other. +lemma exists_edist_eq_ediam_of_finite [Nonempty α] [Finite α] : + ∃ u v, G.edist u v = G.ediam := + Prod.exists'.mp <| ediam_def ▸ exists_eq_ciSup_of_finite + +@[gcongr] +lemma ediam_anti (h : G ≤ G') : G'.ediam ≤ G.ediam := + iSup₂_mono fun _ _ ↦ edist_anti h + +@[simp] +lemma ediam_bot [Nontrivial α] : (⊥ : SimpleGraph α).ediam = ⊤ := + ediam_eq_top_of_not_connected bot_not_connected + +@[simp] +lemma ediam_top [Nontrivial α] : (⊤ : SimpleGraph α).ediam = 1 := by + apply le_antisymm ?_ <| ENat.one_le_iff_pos.mpr <| pos_iff_ne_zero.mpr ediam_ne_zero + apply ediam_def ▸ iSup_le_iff.mpr + intro p + by_cases h : (⊤ : SimpleGraph α).Adj p.1 p.2 + · apply le_of_eq <| edist_eq_one_iff_adj.mpr h + · simp_all + +@[simp] +lemma ediam_eq_one [Nontrivial α] : G.ediam = 1 ↔ G = ⊤ := by + refine ⟨fun h₁ ↦ ?_, fun h ↦ h ▸ ediam_top⟩ + ext u v + refine ⟨fun h ↦ h.ne, fun h₂ ↦ ?_⟩ + apply G.edist_pos_of_ne at h₂ + apply le_of_eq at h₁ + rw [ediam_def, iSup_le_iff] at h₁ + exact edist_eq_one_iff_adj.mp <| le_antisymm (h₁ (u, v)) <| ENat.one_le_iff_pos.mpr h₂ + +end ediam + +section diam + +/-- +The diameter is the greatest distance between any two vertices, with the value `0` in +case the distances are not bounded above, or the graph is not connected. +-/ +noncomputable def diam (G : SimpleGraph α) := + G.ediam.toNat + +lemma diam_def : G.diam = (⨆ p : α × α, G.edist p.1 p.2).toNat := by + rw [diam, ediam_def] + +lemma dist_le_diam (h : G.ediam ≠ ⊤) {u v : α} : G.dist u v ≤ G.diam := + ENat.toNat_le_toNat edist_le_ediam h + +lemma nontrivial_of_diam_ne_zero (h : G.diam ≠ 0) : Nontrivial α := by + apply G.nontrivial_of_ediam_ne_zero + contrapose! h + simp [diam, h] + +lemma diam_eq_zero_of_not_connected (h : ¬G.Connected) : G.diam = 0 := by + cases isEmpty_or_nonempty α + · rw [diam, ediam, ciSup_of_empty, bot_eq_zero']; rfl + · rw [diam, ediam_eq_top_of_not_connected h, ENat.toNat_top] + +lemma diam_eq_zero_of_ediam_eq_top (h : G.ediam = ⊤) : G.diam = 0 := by + rw [diam, h, ENat.toNat_top] + +lemma ediam_ne_top_of_diam_ne_zero (h : G.diam ≠ 0) : G.ediam ≠ ⊤ := + mt diam_eq_zero_of_ediam_eq_top h + +lemma exists_dist_eq_diam [Nonempty α] : + ∃ u v, G.dist u v = G.diam := by + by_cases h : G.diam = 0 + · simp [h] + · obtain ⟨u, v, huv⟩ := exists_edist_eq_ediam_of_ne_top <| ediam_ne_top_of_diam_ne_zero h + use u, v + rw [diam, dist, congrArg ENat.toNat huv] + +@[gcongr] +lemma diam_anti_of_ediam_ne_top (h : G ≤ G') (hn : G.ediam ≠ ⊤) : G'.diam ≤ G.diam := + ENat.toNat_le_toNat (ediam_anti h) hn + +@[simp] +lemma diam_bot : (⊥ : SimpleGraph α).diam = 0 := by + rw [diam, ENat.toNat_eq_zero] + cases subsingleton_or_nontrivial α + · exact Or.inl ediam_eq_zero_of_subsingleton + · exact Or.inr ediam_bot + +@[simp] +lemma diam_top [Nontrivial α] : (⊤ : SimpleGraph α).diam = 1 := by + rw [diam, ediam_top] + rfl + +@[simp] +lemma diam_eq_zero : G.diam = 0 ↔ G.ediam = ⊤ ∨ Subsingleton α := by + rw [diam, ENat.toNat_eq_zero] + aesop + +@[simp] +lemma diam_eq_one [Nontrivial α] : G.diam = 1 ↔ G = ⊤ := by + rw [diam, ENat.toNat_eq_iff (by decide)] + exact ediam_eq_one + +end diam + +end SimpleGraph diff --git a/Mathlib/Data/ENat/Lattice.lean b/Mathlib/Data/ENat/Lattice.lean index 869003b893a3e..25c112487f8a7 100644 --- a/Mathlib/Data/ENat/Lattice.lean +++ b/Mathlib/Data/ENat/Lattice.lean @@ -100,4 +100,13 @@ lemma finite_of_sSup_lt_top (h : sSup s < ⊤) : s.Finite := by 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') +lemma exists_eq_iSup_of_lt_top [Nonempty ι] (h : ⨆ i, f i < ⊤) : + ∃ i, f i = ⨆ i, f i := + sSup_mem_of_Nonempty_of_lt_top h + +lemma exists_eq_iSup₂_of_lt_top {ι₁ ι₂ : Type*} {f : ι₁ → ι₂ → ℕ∞} [Nonempty ι₁] [Nonempty ι₂] + (h : ⨆ i, ⨆ j, f i j < ⊤) : ∃ i j, f i j = ⨆ i, ⨆ j, f i j := by + rw [iSup_prod'] at h ⊢ + exact Prod.exists'.mp (exists_eq_iSup_of_lt_top h) + end ENat diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean b/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean index 7c3aa9e1c8185..18c36b633f629 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Finset.lean @@ -46,6 +46,12 @@ theorem Set.Finite.csSup_lt_iff (hs : s.Finite) (h : s.Nonempty) : sSup s < a theorem Set.Finite.lt_csInf_iff (hs : s.Finite) (h : s.Nonempty) : a < sInf s ↔ ∀ x ∈ s, a < x := @Set.Finite.csSup_lt_iff αᵒᵈ _ _ _ hs h +theorem exists_eq_ciSup_of_finite [Nonempty ι] [Finite ι] {f : ι → α} : ∃ i, f i = ⨆ i, f i := + Nonempty.csSup_mem (range_nonempty f) (finite_range f) + +theorem exists_eq_ciInf_of_finite [Nonempty ι] [Finite ι] {f : ι → α} : ∃ i, f i = ⨅ i, f i := + Nonempty.csInf_mem (range_nonempty f) (finite_range f) + end ConditionallyCompleteLinearOrder /-! From 7a3e6979684c0f6f30bd9b8896b54e472cf153ae Mon Sep 17 00:00:00 2001 From: grunweg Date: Thu, 15 Aug 2024 23:53:44 +0000 Subject: [PATCH 310/359] feat: remove `lint-style.sh` (#15051) In detail, this PR - moves the executable bit check into the .yml files, instead of the shell script - adds a `--fix` flag to `lint-style`, for fixing style errors (when possible) - moves calling `lint-style.py` into `lake exe lint-style`, and - switches the CI workflow to call `lake exe lint-style` instead, and removes `lint-style.sh` Ideally, `lint-style.py` will be completely rewritten in Lean soon, so this calling of an external process is only temporary. Helps with #14539. --- .github/workflows/bors.yml | 15 +++++- .github/workflows/build.yml | 15 +++++- .github/workflows/build.yml.in | 15 +++++- .github/workflows/build_fork.yml | 15 +++++- .github/workflows/lint_and_suggest_pr.yml | 2 +- Mathlib/Tactic/Linter/TextBased.lean | 30 +++++++++--- scripts/lint-style.lean | 9 ++-- scripts/lint-style.sh | 57 ----------------------- scripts/print-style-errors.sh | 2 + 9 files changed, 83 insertions(+), 77 deletions(-) delete mode 100755 scripts/lint-style.sh diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index e3a683620ffa6..c5a421e26f3ef 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -44,6 +44,17 @@ jobs: - name: Look for ignored files uses: credfeto/action-no-ignored-files@v1.1.0 + - name: "Check for Lean files with the executable bit set" + shell: bash + run: | + executable_files="$(find . -name '*.lean' -type f \( -perm -u=x -o -perm -g=x -o -perm -o=x \))" + if [[ -n "$executable_files" ]] + then + echo "ERROR: The following Lean files have the executable bit set." + echo "$executable_files" + exit 1 + fi + - name: install Python if: ${{ 'bors' == 'ubuntu-latest' }} uses: actions/setup-python@v5 @@ -57,9 +68,9 @@ jobs: ./elan-init -y --default-toolchain none echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" - - name: "run style linters: Python ones and their Lean rewrite" + - name: "run style linters" run: | - ./scripts/lint-style.sh + lake exe lint-style - name: Install bibtool if: ${{ 'bors' == 'ubuntu-latest' }} diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 6bceb029713e4..bbeef2b24530e 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -51,6 +51,17 @@ jobs: - name: Look for ignored files uses: credfeto/action-no-ignored-files@v1.1.0 + - name: "Check for Lean files with the executable bit set" + shell: bash + run: | + executable_files="$(find . -name '*.lean' -type f \( -perm -u=x -o -perm -g=x -o -perm -o=x \))" + if [[ -n "$executable_files" ]] + then + echo "ERROR: The following Lean files have the executable bit set." + echo "$executable_files" + exit 1 + fi + - name: install Python if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} uses: actions/setup-python@v5 @@ -64,9 +75,9 @@ jobs: ./elan-init -y --default-toolchain none echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" - - name: "run style linters: Python ones and their Lean rewrite" + - name: "run style linters" run: | - ./scripts/lint-style.sh + lake exe lint-style - name: Install bibtool if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index 02dd04d95f5f5..37ec2538a43ce 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -30,6 +30,17 @@ jobs: - name: Look for ignored files uses: credfeto/action-no-ignored-files@v1.1.0 + - name: "Check for Lean files with the executable bit set" + shell: bash + run: | + executable_files="$(find . -name '*.lean' -type f \( -perm -u=x -o -perm -g=x -o -perm -o=x \))" + if [[ -n "$executable_files" ]] + then + echo "ERROR: The following Lean files have the executable bit set." + echo "$executable_files" + exit 1 + fi + - name: install Python if: ${{ 'STYLE_LINT_RUNNER' == 'ubuntu-latest' }} uses: actions/setup-python@v5 @@ -43,9 +54,9 @@ jobs: ./elan-init -y --default-toolchain none echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" - - name: "run style linters: Python ones and their Lean rewrite" + - name: "run style linters" run: | - ./scripts/lint-style.sh + lake exe lint-style - name: Install bibtool if: ${{ 'STYLE_LINT_RUNNER' == 'ubuntu-latest' }} diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 344c91f9a8747..0a2c6fe778dfc 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -48,6 +48,17 @@ jobs: - name: Look for ignored files uses: credfeto/action-no-ignored-files@v1.1.0 + - name: "Check for Lean files with the executable bit set" + shell: bash + run: | + executable_files="$(find . -name '*.lean' -type f \( -perm -u=x -o -perm -g=x -o -perm -o=x \))" + if [[ -n "$executable_files" ]] + then + echo "ERROR: The following Lean files have the executable bit set." + echo "$executable_files" + exit 1 + fi + - name: install Python if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} uses: actions/setup-python@v5 @@ -61,9 +72,9 @@ jobs: ./elan-init -y --default-toolchain none echo "$HOME/.elan/bin" >> "${GITHUB_PATH}" - - name: "run style linters: Python ones and their Lean rewrite" + - name: "run style linters" run: | - ./scripts/lint-style.sh + lake exe lint-style - name: Install bibtool if: ${{ 'ubuntu-latest' == 'ubuntu-latest' }} diff --git a/.github/workflows/lint_and_suggest_pr.yml b/.github/workflows/lint_and_suggest_pr.yml index 64f06540402cd..52888f7128e32 100644 --- a/.github/workflows/lint_and_suggest_pr.yml +++ b/.github/workflows/lint_and_suggest_pr.yml @@ -31,7 +31,7 @@ jobs: - name: lint continue-on-error: true # allows the following `reviewdog` step to add GitHub suggestions run: | - ./scripts/lint-style.sh --fix + lake exe lint-style --fix - name: suggester / lint-style uses: reviewdog/action-suggester@v1 diff --git a/Mathlib/Tactic/Linter/TextBased.lean b/Mathlib/Tactic/Linter/TextBased.lean index 38204f9292cc3..544f44eb1d63c 100644 --- a/Mathlib/Tactic/Linter/TextBased.lean +++ b/Mathlib/Tactic/Linter/TextBased.lean @@ -11,11 +11,21 @@ import Mathlib.Data.Nat.Notation ## Text-based linters This file defines various mathlib linters which are based on reading the source code only. -In practice, only style linters will have this form. -All of these have been rewritten from the `lint-style.py` script. +In practice, all such linters check for code style issues. -For now, this only contains the linters for the copyright and author headers and large files: -further linters will be ported in subsequent PRs. +For now, this only contains linters checking +- that the copyright header and authors line are correctly formatted +- existence of module docstrings (in the right place) +- for certain disallowed imports +- if the string "adaptation note" is used instead of the command #adaptation_note +- lines are at most 100 characters long (except for URLs) +- files are at most 1500 lines long (unless specifically allowed). + +For historic reasons, some of these checks are still written in a Python script `lint-style.py`: +these are gradually being rewritten in Lean. + +This linter maintains a list of exceptions, for legacy reasons. +Ideally, the length of the list of exceptions tends to 0. An executable running all these linters is defined in `scripts/lint-style.lean`. -/ @@ -392,7 +402,6 @@ def lintFile (path : FilePath) (sizeLimit : Option ℕ) (exceptions : Array Erro IO (Array ErrorContext) := do let lines ← IO.FS.lines path -- We don't need to run any checks on imports-only files. - -- NB. The Python script used to still run a few linters; this is in fact not necessary. if isImportsOnlyFile lines then return #[] let mut errors := #[] @@ -409,8 +418,9 @@ Print formatted errors for all unexpected style violations to standard output; update the list of style exceptions if configured so. Return the number of files which had new style errors. `moduleNames` are all the modules to lint, -`mode` specifies what kind of output this script should produce. -/ -def lintModules (moduleNames : Array String) (mode : OutputSetting) : IO UInt32 := do +`mode` specifies what kind of output this script should produce, +`fix` configures whether fixable errors should be corrected in-place. -/ +def lintModules (moduleNames : Array String) (mode : OutputSetting) (fix : Bool) : IO UInt32 := do -- Read the style exceptions file. -- We also have a `nolints` file with manual exceptions for the linter. let exceptionsFilePath : FilePath := "scripts" / "style-exceptions.txt" @@ -441,6 +451,12 @@ def lintModules (moduleNames : Array String) (mode : OutputSetting) : IO UInt32 numberErrorFiles := numberErrorFiles + 1 match mode with | OutputSetting.print style => + -- Run the remaining python linters. It is easier to just run on all files. + -- If this poses an issue, I can either filter the output + -- or wait until lint-style.py is fully rewritten in Lean. + let args := if fix then #["--fix"] else #[] + let pythonOutput ← IO.Process.run { cmd := "./scripts/print-style-errors.sh", args := args } + if pythonOutput != "" then IO.println pythonOutput formatErrors allUnexpectedErrors style if numberErrorFiles > 0 && mode matches OutputSetting.print _ then IO.println s!"error: found {allUnexpectedErrors.size} new style errors\n\ diff --git a/scripts/lint-style.lean b/scripts/lint-style.lean index 6367daa796ac2..5c7ab0d34064f 100644 --- a/scripts/lint-style.lean +++ b/scripts/lint-style.lean @@ -25,10 +25,10 @@ def lintStyleCli (args : Cli.Parsed) : IO UInt32 := do -- Read all module names to lint. let mut allModules := #[] for s in ["Archive.lean", "Counterexamples.lean", "Mathlib.lean"] do - -- note: since we manually add "Batteries" to "Mathlib.lean", we remove it here manually - allModules := allModules.append ((← IO.FS.lines s).filter (!·.containsSubstr "Batteries") - |>.map (·.stripPrefix "import ")) - let numberErrorFiles ← lintModules allModules mode + allModules := allModules.append ((← IO.FS.lines s).map (·.stripPrefix "import ")) + -- note: since we manually add "Batteries" to "Mathlib.lean", we remove it here manually + allModules := allModules.erase "Batteries" + let numberErrorFiles ← lintModules allModules mode (args.hasFlag "fix") -- Make sure to return an exit code of at most 125, so this return value can be used further -- in shell scripts. return min numberErrorFiles 125 @@ -48,6 +48,7 @@ def lintStyle : Cmd := `[Cli| and tries to not modify exception entries unless necessary. To fully regenerate the list of style exceptions, delete `style-exceptions.txt` and run this script again with this flag." + fix; "Automatically fix the style error, if possible" ] /-- The entry point to the `lake exe lint-style` command. -/ diff --git a/scripts/lint-style.sh b/scripts/lint-style.sh deleted file mode 100755 index 91d550ab65b55..0000000000000 --- a/scripts/lint-style.sh +++ /dev/null @@ -1,57 +0,0 @@ -#!/usr/bin/env bash - -set -exo pipefail - -# # Style linter - -# ## Usage - -# Run this script from the root of mathlib using: -# ./scripts/lint-style.sh - -# ## Purpose - -# The style linter checks for new style issues, -# and maintains a list of exceptions for legacy reasons. -# Ideally, the length of the list of exceptions tends to 0. - -# Examples of issues that are checked for are: -# * existence of copyright header -# * existence of module docstrings (in the right place) -# * line length <= 100 (unless URL) - -# ## Implementation details - -# There are two parts. -# 1. A Python script `scripts/lint-style.py` that lints the contents of a Lean file. -# This script is called below on all Lean files in the repository. -# Exceptions are maintained in `scripts/style-exceptions.txt`. -# (Rewriting these checks in Lean is work in progress.) -# 2. The remainder of this shell script -# contains a lint on the global repository. -# -# TODO: This is adapted from the linter for mathlib3. It should be rewritten in Lean. - -################################################################################ - -# 1. Call the Lean file linter, implemented in Python - -touch scripts/style-exceptions.txt - -git ls-files 'Mathlib/*.lean' | xargs ./scripts/lint-style.py "$@" -git ls-files 'Archive/*.lean' | xargs ./scripts/lint-style.py "$@" -git ls-files 'Counterexamples/*.lean' | xargs ./scripts/lint-style.py "$@" - -# Call the in-progress Lean rewrite of these Python lints. -lake exe lint-style --github - -# 2. Global checks on the mathlib repository - -executable_files="$(find . -name '*.lean' -type f \( -perm -u=x -o -perm -g=x -o -perm -o=x \))" - -if [[ -n "$executable_files" ]] -then - echo "ERROR: The following Lean files have the executable bit set." - echo "$executable_files" - exit 1 -fi diff --git a/scripts/print-style-errors.sh b/scripts/print-style-errors.sh index 1eda2b5b93663..f9a9e1c1b42c8 100755 --- a/scripts/print-style-errors.sh +++ b/scripts/print-style-errors.sh @@ -7,3 +7,5 @@ # use C locale so that sorting is the same on macOS and Linux # see https://unix.stackexchange.com/questions/362728/why-does-gnu-sort-sort-differently-on-my-osx-machine-and-linux-machine find Mathlib -name '*.lean' | xargs ./scripts/lint-style.py | LC_ALL=C sort +find Archive -name '*.lean' | xargs ./scripts/lint-style.py | LC_ALL=C sort +find Counterexamples -name '*.lean' | xargs ./scripts/lint-style.py | LC_ALL=C sort From b3b0e2dc2f0551eee64361e48f3ebc70921a12d7 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 16 Aug 2024 00:11:44 +0000 Subject: [PATCH 311/359] chore: better failures for `linear_combination` (#15791) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The implementation of `linear_combination` is as a macro performing effectively `refine ******; ring`. If the term provided to `refine` fails to elaborate, Lean inserts sorries where needed and then continues on to `ring`. This produces a confusing error message. For example, the problem ```lean example (x y : ℤ) (h1 : x * y + 2 * x = 1) (h2 : x = y) : x * y + 2 * x = 1 := by linear_combination h1 + (0 : ℝ) * h2 ``` produces both the "true" error message ``` application type mismatch Mathlib.Tactic.LinearCombination.c_mul_pf h2 0 argument 0 has type ℝ : Type but is expected to have type ℤ : Type ``` and the "tactic is confused because it persevered when it shouldn't have" error message ``` ring failed, ring expressions not equal x y : ℤ h1 : x * y + 2 * x = 1 h2 : x = y ⊢ -(x * sorryAx ℤ true) + y * sorryAx ℤ true = 0 ``` This PR fixes this by strategically inserting a `Term.withoutErrToSorry`. In examples such as the above, it now stops after the `refine`, so the second error message does not appear. --- Mathlib/Tactic/LinearCombination.lean | 2 +- test/linear_combination.lean | 23 +++++++++++++++++++++-- 2 files changed, 22 insertions(+), 3 deletions(-) diff --git a/Mathlib/Tactic/LinearCombination.lean b/Mathlib/Tactic/LinearCombination.lean index 78472a591957f..7b2fdbd850460 100644 --- a/Mathlib/Tactic/LinearCombination.lean +++ b/Mathlib/Tactic/LinearCombination.lean @@ -138,7 +138,7 @@ def elabLinearCombination | .const c => `(Eq.refl $c) | .proof p => pure p let norm := norm?.getD (Unhygienic.run `(tactic| ring1)) - Tactic.evalTactic <| ← withFreshMacroScope <| + Term.withoutErrToSorry <| Tactic.evalTactic <| ← withFreshMacroScope <| if twoGoals then `(tactic| ( refine eq_trans₃ $p ?a ?b diff --git a/test/linear_combination.lean b/test/linear_combination.lean index efb18bd1c26d2..c9e02f7fa847e 100644 --- a/test/linear_combination.lean +++ b/test/linear_combination.lean @@ -184,13 +184,32 @@ example (x y : ℤ) (h1 : x * y + 2 * x = 1) (h2 : x = y) : x * y = -2 * y + 1 : /-! ### Cases that should fail -/ +/-- +error: ring failed, ring expressions not equal +a : ℚ +ha : a = 1 +⊢ -1 = 0 +-/ +#guard_msgs in +example (a : ℚ) (ha : a = 1) : a = 2 := by linear_combination ha + -- This should fail because the second coefficient has a different type than -- the equations it is being combined with. This was a design choice for the -- sake of simplicity, but the tactic could potentially be modified to allow -- this behavior. +/-- +error: application type mismatch + Mathlib.Tactic.LinearCombination.c_mul_pf h2 0 +argument + 0 +has type + ℝ : Type +but is expected to have type + ℤ : Type +-/ +#guard_msgs in example (x y : ℤ) (h1 : x * y + 2 * x = 1) (h2 : x = y) : x * y + 2 * x = 1 := by - fail_if_success linear_combination h1 + (0 : ℝ) * h2 - linear_combination h1 + linear_combination h1 + (0 : ℝ) * h2 -- This fails because the linear_combination tactic requires the equations -- and coefficients to use a type that fulfills the add_group condition, From 44df4ba148579358194381c01d8389fdb58e4e31 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 16 Aug 2024 00:11:45 +0000 Subject: [PATCH 312/359] feat: add some convenience lemmas for `spectrum` and `inv` (#15855) --- Mathlib/Algebra/Algebra/Spectrum.lean | 21 +++++++++++++++++++++ 1 file changed, 21 insertions(+) diff --git a/Mathlib/Algebra/Algebra/Spectrum.lean b/Mathlib/Algebra/Algebra/Spectrum.lean index 933fc61e08ab6..2919c52895656 100644 --- a/Mathlib/Algebra/Algebra/Spectrum.lean +++ b/Mathlib/Algebra/Algebra/Spectrum.lean @@ -295,6 +295,27 @@ theorem sub_singleton_eq (a : A) (r : R) : σ a - {r} = σ (a - ↑ₐ r) := by end ScalarRing +section ScalarSemifield + +variable {R : Type u} {A : Type v} [Semifield R] [Ring A] [Algebra R A] + +@[simp] +lemma inv₀_mem_iff {r : R} {a : Aˣ} : + r⁻¹ ∈ spectrum R (a : A) ↔ r ∈ spectrum R (↑a⁻¹ : A) := by + obtain (rfl | hr) := eq_or_ne r 0 + · simp [zero_mem_iff] + · lift r to Rˣ using hr.isUnit + simp [inv_mem_iff] + +lemma inv₀_mem_inv_iff {r : R} {a : Aˣ} : + r⁻¹ ∈ spectrum R (↑a⁻¹ : A) ↔ r ∈ spectrum R (a : A) := by + simp + +alias ⟨of_inv₀_mem, inv₀_mem⟩ := inv₀_mem_iff +alias ⟨of_inv₀_mem_inv, inv₀_mem_inv⟩ := inv₀_mem_inv_iff + +end ScalarSemifield + section ScalarField variable {𝕜 : Type u} {A : Type v} From 902c3a54ee10932fdfcfa3c9583bcdf122885e87 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 16 Aug 2024 00:11:46 +0000 Subject: [PATCH 313/359] chore: remove unnecessary references to Nat.zero_eq (#15858) The induction tactics no longer put a Nat.zero into the context. --- Archive/Sensitivity.lean | 2 +- Counterexamples/Phillips.lean | 2 +- .../Computation/CorrectnessTerminating.lean | 2 +- .../ContinuedFractions/Computation/Translations.lean | 2 +- Mathlib/Algebra/CubicDiscriminant.lean | 4 ++-- Mathlib/Algebra/GeomSum.lean | 2 +- Mathlib/Algebra/Homology/Augment.lean | 2 +- Mathlib/Algebra/Homology/LocalCohomology.lean | 2 +- Mathlib/Algebra/Lie/CartanExists.lean | 2 +- Mathlib/Algebra/Lie/Engel.lean | 2 +- Mathlib/Algebra/Lie/Nilpotent.lean | 8 ++++---- Mathlib/Algebra/Lie/Solvable.lean | 4 ++-- Mathlib/Algebra/Module/LinearMap/End.lean | 2 +- Mathlib/Algebra/Order/CauSeq/Basic.lean | 2 +- Mathlib/Algebra/Polynomial/Eval.lean | 2 +- Mathlib/Algebra/Polynomial/FieldDivision.lean | 2 +- Mathlib/Algebra/Polynomial/Laurent.lean | 2 +- Mathlib/Algebra/Polynomial/Roots.lean | 4 ++-- Mathlib/Algebra/Polynomial/Smeval.lean | 4 ++-- Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean | 2 +- .../AlgebraicTopology/DoldKan/HomotopyEquivalence.lean | 2 +- Mathlib/AlgebraicTopology/DoldKan/PInfty.lean | 2 +- Mathlib/AlgebraicTopology/SimplexCategory.lean | 2 +- Mathlib/Analysis/Analytic/Basic.lean | 4 ++-- Mathlib/Analysis/Analytic/Constructions.lean | 2 +- Mathlib/Analysis/Analytic/Meromorphic.lean | 2 +- Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean | 2 +- Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean | 2 +- Mathlib/Analysis/CStarAlgebra/Basic.lean | 2 +- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 4 ++-- Mathlib/Analysis/Calculus/ContDiff/Bounds.lean | 2 +- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 2 +- Mathlib/Analysis/Calculus/Deriv/ZPow.lean | 2 +- Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean | 2 +- Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean | 2 +- Mathlib/Analysis/Fourier/FourierTransformDeriv.lean | 2 +- Mathlib/Analysis/InnerProductSpace/OfNorm.lean | 2 +- Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 2 +- Mathlib/Analysis/Normed/Field/Basic.lean | 4 ++-- Mathlib/Analysis/Normed/Operator/Banach.lean | 2 +- Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean | 4 ++-- Mathlib/Analysis/Subadditive.lean | 2 +- Mathlib/CategoryTheory/ComposableArrows.lean | 2 +- Mathlib/CategoryTheory/Sites/Sheaf.lean | 2 +- Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean | 2 +- Mathlib/Computability/TuringMachine.lean | 8 ++++---- Mathlib/Data/Fin/Basic.lean | 4 ++-- Mathlib/Data/Fin/VecNotation.lean | 2 +- Mathlib/Data/List/Cycle.lean | 6 +++--- Mathlib/Data/Nat/Hyperoperation.lean | 2 +- Mathlib/Data/Nat/Totient.lean | 2 +- Mathlib/Data/Num/Lemmas.lean | 2 +- Mathlib/Data/PNat/Factors.lean | 2 +- Mathlib/Data/Seq/Computation.lean | 2 +- Mathlib/Data/Set/Enumerate.lean | 2 +- Mathlib/Data/Set/MemPartition.lean | 2 +- Mathlib/Data/Set/Pointwise/Finite.lean | 2 +- Mathlib/Data/Vector/Basic.lean | 2 +- Mathlib/Data/ZMod/Basic.lean | 4 ++-- Mathlib/FieldTheory/PerfectClosure.lean | 2 +- Mathlib/GroupTheory/FiniteAbelian.lean | 2 +- Mathlib/GroupTheory/Nilpotent.lean | 4 ++-- Mathlib/GroupTheory/Perm/Closure.lean | 2 +- Mathlib/GroupTheory/Perm/Cycle/Concrete.lean | 2 +- Mathlib/GroupTheory/Perm/Cycle/Factors.lean | 2 +- Mathlib/LinearAlgebra/Eigenspace/Zero.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Transvection.lean | 2 +- Mathlib/MeasureTheory/Covering/Besicovitch.lean | 4 ++-- Mathlib/MeasureTheory/Function/SimpleFunc.lean | 2 +- Mathlib/MeasureTheory/Group/Arithmetic.lean | 4 ++-- Mathlib/MeasureTheory/Group/Measure.lean | 2 +- Mathlib/MeasureTheory/Integral/Pi.lean | 4 ++-- .../MeasureTheory/MeasurableSpace/CountablyGenerated.lean | 2 +- Mathlib/ModelTheory/Semantics.lean | 2 +- Mathlib/NumberTheory/Cyclotomic/Rat.lean | 6 +++--- Mathlib/NumberTheory/DiophantineApproximation.lean | 6 +++--- Mathlib/NumberTheory/LSeries/Deriv.lean | 4 ++-- Mathlib/NumberTheory/LSeries/Linearity.lean | 4 ++-- Mathlib/NumberTheory/Multiplicity.lean | 8 ++++---- Mathlib/NumberTheory/Pell.lean | 2 +- Mathlib/Order/Interval/Finset/Nat.lean | 2 +- Mathlib/Order/RelSeries.lean | 2 +- Mathlib/Order/SuccPred/Basic.lean | 2 +- Mathlib/Order/SuccPred/LinearLocallyFinite.lean | 4 ++-- Mathlib/Probability/Martingale/Convergence.lean | 2 +- Mathlib/Probability/Martingale/Upcrossing.lean | 8 ++++---- Mathlib/RingTheory/AdicCompletion/Basic.lean | 2 +- Mathlib/RingTheory/Binomial.lean | 2 +- .../RingTheory/GradedAlgebra/HomogeneousLocalization.lean | 2 +- Mathlib/RingTheory/HahnSeries/Summable.lean | 2 +- Mathlib/RingTheory/Henselian.lean | 2 +- Mathlib/RingTheory/Polynomial/Pochhammer.lean | 2 +- Mathlib/RingTheory/PowerSeries/Basic.lean | 4 ++-- Mathlib/RingTheory/WittVector/FrobeniusFractionField.lean | 2 +- Mathlib/RingTheory/WittVector/Identities.lean | 2 +- Mathlib/RingTheory/WittVector/IsPoly.lean | 2 +- Mathlib/RingTheory/WittVector/StructurePolynomial.lean | 2 +- Mathlib/RingTheory/WittVector/Verschiebung.lean | 3 +-- Mathlib/SetTheory/Surreal/Dyadic.lean | 2 +- Mathlib/Tactic/Attr/Core.lean | 3 --- Mathlib/Topology/MetricSpace/Closeds.lean | 2 +- 101 files changed, 136 insertions(+), 140 deletions(-) diff --git a/Archive/Sensitivity.lean b/Archive/Sensitivity.lean index ae88d8cc7fc27..af363c85fc5cc 100644 --- a/Archive/Sensitivity.lean +++ b/Archive/Sensitivity.lean @@ -279,7 +279,7 @@ using only the addition of `V`. -/ theorem f_squared : ∀ v : V n, (f n) (f n v) = (n : ℝ) • v := by induction' n with n IH _ <;> intro v - · simp only [Nat.zero_eq, Nat.cast_zero, zero_smul]; rfl + · simp only [Nat.cast_zero, zero_smul]; rfl · cases v; rw [f_succ_apply, f_succ_apply]; simp [IH, add_smul (n : ℝ) 1, add_assoc, V]; abel /-! We now compute the matrix of `f` in the `e` basis (`p` is the line index, diff --git a/Counterexamples/Phillips.lean b/Counterexamples/Phillips.lean index dbccecbab4e8d..d4fc4fc003f7e 100644 --- a/Counterexamples/Phillips.lean +++ b/Counterexamples/Phillips.lean @@ -296,7 +296,7 @@ theorem exists_discrete_support_nonpos (f : BoundedAdditiveMeasure α) : induction n with | zero => simp only [s, BoundedAdditiveMeasure.empty, id, Nat.cast_zero, zero_mul, - Function.iterate_zero, Subtype.coe_mk, Nat.zero_eq] + Function.iterate_zero, Subtype.coe_mk] rfl | succ n IH => have : (s (n + 1)).1 = (s (n + 1)).1 \ (s n).1 ∪ (s n).1 := by diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean b/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean index 882c79f22ade0..2f4ac3f34ae67 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/CorrectnessTerminating.lean @@ -102,7 +102,7 @@ theorem compExactValue_correctness_of_stream_eq_some : -- Nat.zero have : IntFractPair.of v = ifp_zero := by have : IntFractPair.stream v 0 = some (IntFractPair.of v) := rfl - simpa only [Nat.zero_eq, this, Option.some.injEq] using stream_zero_eq + simpa only [this, Option.some.injEq] using stream_zero_eq cases this cases' Decidable.em (Int.fract v = 0) with fract_eq_zero fract_ne_zero -- Int.fract v = 0; we must then have `v = ⌊v⌋` diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean b/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean index 34c8347407af7..771e16a57a2bb 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/Translations.lean @@ -304,7 +304,7 @@ are all equal to `a`. -/ theorem convs'_of_int (a : ℤ) : (of (a : K)).convs' n = a := by induction n with - | zero => simp only [zeroth_conv'_eq_h, of_h_eq_floor, floor_intCast, Nat.zero_eq] + | zero => simp only [zeroth_conv'_eq_h, of_h_eq_floor, floor_intCast] | succ => rw [convs', of_h_eq_floor, floor_intCast, add_right_eq_self] exact convs'Aux_succ_none ((of_s_of_int K a).symm ▸ Stream'.Seq.get?_nil 0) _ diff --git a/Mathlib/Algebra/CubicDiscriminant.lean b/Mathlib/Algebra/CubicDiscriminant.lean index b200ea3ac1f44..7280b1a8b0e49 100644 --- a/Mathlib/Algebra/CubicDiscriminant.lean +++ b/Mathlib/Algebra/CubicDiscriminant.lean @@ -240,9 +240,9 @@ def equiv : Cubic R ≃ { p : R[X] // p.degree ≤ 3 } where invFun f := ⟨coeff f 3, coeff f 2, coeff f 1, coeff f 0⟩ left_inv P := by ext <;> simp only [Subtype.coe_mk, coeffs] right_inv f := by - -- Porting note: Added `simp only [Nat.zero_eq, Nat.succ_eq_add_one] <;> ring_nf` + -- Porting note: Added `simp only [Nat.succ_eq_add_one] <;> ring_nf` -- There's probably a better way to do this. - ext (_ | _ | _ | _ | n) <;> simp only [Nat.zero_eq, Nat.succ_eq_add_one] <;> ring_nf + ext (_ | _ | _ | _ | n) <;> simp only [Nat.succ_eq_add_one] <;> ring_nf <;> try simp only [coeffs] have h3 : 3 < 4 + n := by linarith only rw [coeff_eq_zero h3, diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index bce4d36d459b2..1a833ffa86faf 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -440,7 +440,7 @@ theorem geom_sum_alternating_of_le_neg_one [StrictOrderedRing α] (hx : x + 1 if Even n then (∑ i ∈ range n, x ^ i) ≤ 0 else 1 ≤ ∑ i ∈ range n, x ^ i := by have hx0 : x ≤ 0 := (le_add_of_nonneg_right zero_le_one).trans hx induction n with - | zero => simp only [Nat.zero_eq, range_zero, sum_empty, le_refl, ite_true, even_zero] + | zero => simp only [range_zero, sum_empty, le_refl, ite_true, even_zero] | succ n ih => simp only [Nat.even_add_one, geom_sum_succ] split_ifs at ih with h diff --git a/Mathlib/Algebra/Homology/Augment.lean b/Mathlib/Algebra/Homology/Augment.lean index 4daeaa8a4bcf3..0b02a3d3140cd 100644 --- a/Mathlib/Algebra/Homology/Augment.lean +++ b/Mathlib/Algebra/Homology/Augment.lean @@ -231,7 +231,7 @@ def augment (C : CochainComplex V ℕ) {X : V} (f : X ⟶ C.X 0) (w : f ≫ C.d cases k · exact w · rw [C.shape, comp_zero] - simp only [Nat.zero_eq, ComplexShape.up_Rel, zero_add] + simp only [ComplexShape.up_Rel, zero_add] exact (Nat.one_lt_succ_succ _).ne @[simp] diff --git a/Mathlib/Algebra/Homology/LocalCohomology.lean b/Mathlib/Algebra/Homology/LocalCohomology.lean index 77e73c476b11a..e6bcfa9d338a5 100644 --- a/Mathlib/Algebra/Homology/LocalCohomology.lean +++ b/Mathlib/Algebra/Homology/LocalCohomology.lean @@ -201,7 +201,7 @@ def idealPowersToSelfLERadical (J : Ideal R) : ℕᵒᵖ ⥤ SelfLERadical J := FullSubcategory.lift _ (idealPowersDiagram J) fun k => by change _ ≤ (J ^ unop k).radical cases' unop k with n - · simp [Ideal.radical_top, pow_zero, Ideal.one_eq_top, le_top, Nat.zero_eq] + · simp [Ideal.radical_top, pow_zero, Ideal.one_eq_top, le_top] · simp only [J.radical_pow n.succ_ne_zero, Ideal.le_radical] variable {I J K : Ideal R} diff --git a/Mathlib/Algebra/Lie/CartanExists.lean b/Mathlib/Algebra/Lie/CartanExists.lean index 25735e2642f2f..2a91aa8017dfa 100644 --- a/Mathlib/Algebra/Lie/CartanExists.lean +++ b/Mathlib/Algebra/Lie/CartanExists.lean @@ -320,7 +320,7 @@ lemma engel_isBot_of_isMin (hLK : finrank K L ≤ #K) (U : LieSubalgebra K L) rw [← hn] clear hn induction n with - | zero => simp only [Nat.zero_eq, pow_zero, LinearMap.one_apply] + | zero => simp only [pow_zero, LinearMap.one_apply] | succ n ih => rw [pow_succ', pow_succ', LinearMap.mul_apply, ih]; rfl classical -- Now let `n` be the smallest power such that `⁅v, _⁆ ^ n` kills `z'`. diff --git a/Mathlib/Algebra/Lie/Engel.lean b/Mathlib/Algebra/Lie/Engel.lean index 75a7f7b2c26dc..7dff375e7f62d 100644 --- a/Mathlib/Algebra/Lie/Engel.lean +++ b/Mathlib/Algebra/Lie/Engel.lean @@ -110,7 +110,7 @@ theorem lcs_le_lcs_of_is_nilpotent_span_sup_eq_top {n i j : ℕ} by simpa only [bot_sup_eq, LieIdeal.incl_coe, Submodule.map_zero, hxn] using this n intro l induction' l with l ih - · simp only [Nat.zero_eq, add_zero, LieIdeal.lcs_succ, pow_zero, LinearMap.one_eq_id, + · simp only [add_zero, LieIdeal.lcs_succ, pow_zero, LinearMap.one_eq_id, Submodule.map_id] exact le_sup_of_le_left hIM · simp only [LieIdeal.lcs_succ, i.add_succ l, lie_top_eq_of_span_sup_eq_top hxI, sup_le_iff] diff --git a/Mathlib/Algebra/Lie/Nilpotent.lean b/Mathlib/Algebra/Lie/Nilpotent.lean index ad16dcc5bdd9a..2216a2f250436 100644 --- a/Mathlib/Algebra/Lie/Nilpotent.lean +++ b/Mathlib/Algebra/Lie/Nilpotent.lean @@ -157,7 +157,7 @@ variable [LieModule R L M] theorem iterate_toEnd_mem_lowerCentralSeries (x : L) (m : M) (k : ℕ) : (toEnd R L M x)^[k] m ∈ lowerCentralSeries R L M k := by induction' k with k ih - · simp only [Nat.zero_eq, Function.iterate_zero, lowerCentralSeries_zero, LieSubmodule.mem_top] + · simp only [Function.iterate_zero, lowerCentralSeries_zero, LieSubmodule.mem_top] · simp only [lowerCentralSeries_succ, Function.comp_apply, Function.iterate_succ', toEnd_apply_apply] exact LieSubmodule.lie_mem_lie (LieSubmodule.mem_top x) ih @@ -178,7 +178,7 @@ variable {R L M} theorem map_lowerCentralSeries_le (f : M →ₗ⁅R,L⁆ M₂) : (lowerCentralSeries R L M k).map f ≤ lowerCentralSeries R L M₂ k := by induction' k with k ih - · simp only [Nat.zero_eq, lowerCentralSeries_zero, le_top] + · simp only [lowerCentralSeries_zero, le_top] · simp only [LieModule.lowerCentralSeries_succ, LieSubmodule.map_bracket_eq] exact LieSubmodule.mono_lie_right ⊤ ih @@ -621,7 +621,7 @@ theorem coe_lowerCentralSeries_ideal_quot_eq {I : LieIdeal R L} (k : ℕ) : LieSubmodule.toSubmodule (lowerCentralSeries R L (L ⧸ I) k) = LieSubmodule.toSubmodule (lowerCentralSeries R (L ⧸ I) (L ⧸ I) k) := by induction' k with k ih - · simp only [Nat.zero_eq, LieModule.lowerCentralSeries_zero, LieSubmodule.top_coeSubmodule, + · simp only [LieModule.lowerCentralSeries_zero, LieSubmodule.top_coeSubmodule, LieIdeal.top_coe_lieSubalgebra, LieSubalgebra.top_coe_submodule] · simp only [LieModule.lowerCentralSeries_succ, LieSubmodule.lieIdeal_oper_eq_linear_span] congr @@ -662,7 +662,7 @@ theorem LieAlgebra.non_trivial_center_of_isNilpotent [Nontrivial L] [IsNilpotent theorem LieIdeal.map_lowerCentralSeries_le (k : ℕ) {f : L →ₗ⁅R⁆ L'} : LieIdeal.map f (lowerCentralSeries R L L k) ≤ lowerCentralSeries R L' L' k := by induction' k with k ih - · simp only [Nat.zero_eq, LieModule.lowerCentralSeries_zero, le_top] + · simp only [LieModule.lowerCentralSeries_zero, le_top] · simp only [LieModule.lowerCentralSeries_succ] exact le_trans (LieIdeal.map_bracket_le f) (LieSubmodule.mono_lie le_top ih) diff --git a/Mathlib/Algebra/Lie/Solvable.lean b/Mathlib/Algebra/Lie/Solvable.lean index e1132be632edc..987fd1ca9889f 100644 --- a/Mathlib/Algebra/Lie/Solvable.lean +++ b/Mathlib/Algebra/Lie/Solvable.lean @@ -132,7 +132,7 @@ variable {R L} theorem derivedSeries_eq_derivedSeriesOfIdeal_comap (k : ℕ) : derivedSeries R I k = (derivedSeriesOfIdeal R L k I).comap I.incl := by induction' k with k ih - · simp only [Nat.zero_eq, derivedSeries_def, comap_incl_self, derivedSeriesOfIdeal_zero] + · simp only [derivedSeries_def, comap_incl_self, derivedSeriesOfIdeal_zero] · simp only [derivedSeries_def, derivedSeriesOfIdeal_succ] at ih ⊢; rw [ih] exact comap_bracket_incl_of_le I (derivedSeriesOfIdeal_le_self I k) (derivedSeriesOfIdeal_le_self I k) @@ -157,7 +157,7 @@ theorem derivedSeries_add_eq_bot {k l : ℕ} {I J : LieIdeal R L} (hI : derivedS theorem derivedSeries_map_le (k : ℕ) : (derivedSeries R L' k).map f ≤ derivedSeries R L k := by induction' k with k ih - · simp only [Nat.zero_eq, derivedSeries_def, derivedSeriesOfIdeal_zero, le_top] + · simp only [derivedSeries_def, derivedSeriesOfIdeal_zero, le_top] · simp only [derivedSeries_def, derivedSeriesOfIdeal_succ] at ih ⊢ exact le_trans (map_bracket_le f) (LieSubmodule.mono_lie ih ih) diff --git a/Mathlib/Algebra/Module/LinearMap/End.lean b/Mathlib/Algebra/Module/LinearMap/End.lean index 4d09f38caa016..20e40bb532f29 100644 --- a/Mathlib/Algebra/Module/LinearMap/End.lean +++ b/Mathlib/Algebra/Module/LinearMap/End.lean @@ -139,7 +139,7 @@ theorem commute_pow_left_of_commute {f : M →ₛₗ[σ₁₂] M₂} {g : Module.End R M} {g₂ : Module.End R₂ M₂} (h : g₂.comp f = f.comp g) (k : ℕ) : (g₂ ^ k).comp f = f.comp (g ^ k) := by induction' k with k ih - · simp only [Nat.zero_eq, pow_zero, one_eq_id, id_comp, comp_id] + · simp only [pow_zero, one_eq_id, id_comp, comp_id] · rw [pow_succ', pow_succ', LinearMap.mul_eq_comp, LinearMap.comp_assoc, ih, ← LinearMap.comp_assoc, h, LinearMap.comp_assoc, LinearMap.mul_eq_comp] diff --git a/Mathlib/Algebra/Order/CauSeq/Basic.lean b/Mathlib/Algebra/Order/CauSeq/Basic.lean index abb991002f5a5..aa0df8061099e 100644 --- a/Mathlib/Algebra/Order/CauSeq/Basic.lean +++ b/Mathlib/Algebra/Order/CauSeq/Basic.lean @@ -518,7 +518,7 @@ theorem smul_equiv_smul {G : Type*} [SMul G β] [IsScalarTower G β β] {f1 f2 : theorem pow_equiv_pow {f1 f2 : CauSeq β abv} (hf : f1 ≈ f2) (n : ℕ) : f1 ^ n ≈ f2 ^ n := by induction n with - | zero => simp only [Nat.zero_eq, pow_zero, Setoid.refl] + | zero => simp only [pow_zero, Setoid.refl] | succ n ih => simpa only [pow_succ'] using mul_equiv_mul hf ih end Ring diff --git a/Mathlib/Algebra/Polynomial/Eval.lean b/Mathlib/Algebra/Polynomial/Eval.lean index 47cf7b8a57706..e05c9576c8836 100644 --- a/Mathlib/Algebra/Polynomial/Eval.lean +++ b/Mathlib/Algebra/Polynomial/Eval.lean @@ -917,7 +917,7 @@ theorem eval₂_mul' : theorem eval₂_pow' (n : ℕ) : (p ^ n).eval₂ (algebraMap R S) x = (p.eval₂ (algebraMap R S) x) ^ n := by induction n with - | zero => simp only [Nat.zero_eq, pow_zero, eval₂_one] + | zero => simp only [pow_zero, eval₂_one] | succ n ih => rw [pow_succ, pow_succ, eval₂_mul', ih] @[simp] diff --git a/Mathlib/Algebra/Polynomial/FieldDivision.lean b/Mathlib/Algebra/Polynomial/FieldDivision.lean index 955c34ceffb85..0460864bfceb5 100644 --- a/Mathlib/Algebra/Polynomial/FieldDivision.lean +++ b/Mathlib/Algebra/Polynomial/FieldDivision.lean @@ -95,7 +95,7 @@ theorem lt_rootMultiplicity_of_isRoot_iterate_derivative_of_mem_nonZeroDivisors' clear hroot induction n with | zero => - simp only [Nat.zero_eq, Nat.factorial_zero, Nat.cast_one] + simp only [Nat.factorial_zero, Nat.cast_one] exact Submonoid.one_mem _ | succ n ih => rw [Nat.factorial_succ, Nat.cast_mul, mul_mem_nonZeroDivisors] diff --git a/Mathlib/Algebra/Polynomial/Laurent.lean b/Mathlib/Algebra/Polynomial/Laurent.lean index 99365bbe5de8b..7368130e5cb45 100644 --- a/Mathlib/Algebra/Polynomial/Laurent.lean +++ b/Mathlib/Algebra/Polynomial/Laurent.lean @@ -366,7 +366,7 @@ theorem reduce_to_polynomial_of_mul_T (f : R[T;T⁻¹]) {Q : R[T;T⁻¹] → Pro (Qf : ∀ f : R[X], Q (toLaurent f)) (QT : ∀ f, Q (f * T 1) → Q f) : Q f := by induction' f using LaurentPolynomial.induction_on_mul_T with f n induction n with - | zero => simpa only [Nat.zero_eq, Nat.cast_zero, neg_zero, T_zero, mul_one] using Qf _ + | zero => simpa only [Nat.cast_zero, neg_zero, T_zero, mul_one] using Qf _ | succ n hn => convert QT _ _; simpa using hn section Support diff --git a/Mathlib/Algebra/Polynomial/Roots.lean b/Mathlib/Algebra/Polynomial/Roots.lean index 6b7ad04a342d1..93797531db5dd 100644 --- a/Mathlib/Algebra/Polynomial/Roots.lean +++ b/Mathlib/Algebra/Polynomial/Roots.lean @@ -312,7 +312,7 @@ theorem mul_mem_nthRootsFinset η₁ * η₂ ∈ nthRootsFinset n R := by cases n with | zero => - simp only [Nat.zero_eq, nthRootsFinset_zero, not_mem_empty] at hη₁ + simp only [nthRootsFinset_zero, not_mem_empty] at hη₁ | succ n => rw [mem_nthRootsFinset n.succ_pos] at hη₁ hη₂ ⊢ rw [mul_pow, hη₁, hη₂, one_mul] @@ -322,7 +322,7 @@ theorem ne_zero_of_mem_nthRootsFinset {η : R} (hη : η ∈ nthRootsFinset n R) rintro rfl cases n with | zero => - simp only [Nat.zero_eq, nthRootsFinset_zero, not_mem_empty] at hη + simp only [nthRootsFinset_zero, not_mem_empty] at hη | succ n => rw [mem_nthRootsFinset n.succ_pos, zero_pow n.succ_ne_zero] at hη exact zero_ne_one hη diff --git a/Mathlib/Algebra/Polynomial/Smeval.lean b/Mathlib/Algebra/Polynomial/Smeval.lean index 657e2d1b19d52..2eb2eecb587a6 100644 --- a/Mathlib/Algebra/Polynomial/Smeval.lean +++ b/Mathlib/Algebra/Polynomial/Smeval.lean @@ -108,7 +108,7 @@ theorem smeval_add : (p + q).smeval x = p.smeval x + q.smeval x := by theorem smeval_natCast (n : ℕ) : (n : R[X]).smeval x = n • x ^ 0 := by induction n with - | zero => simp only [smeval_zero, Nat.cast_zero, Nat.zero_eq, zero_smul] + | zero => simp only [smeval_zero, Nat.cast_zero, zero_smul] | succ n ih => rw [n.cast_succ, smeval_add, ih, smeval_one, ← add_nsmul] @[deprecated (since := "2024-04-17")] @@ -210,7 +210,7 @@ theorem smeval_at_zero : p.smeval (0 : S) = (p.coeff 0) • (1 : S) := by simp_all only [smeval_add, coeff_add, add_smul] | h_monomial n a => cases n with - | zero => simp only [Nat.zero_eq, monomial_zero_left, smeval_C, npow_zero, coeff_C_zero] + | zero => simp only [monomial_zero_left, smeval_C, npow_zero, coeff_C_zero] | succ n => rw [coeff_monomial_succ, smeval_monomial, npow_add, npow_one, mul_zero, zero_smul, smul_zero] diff --git a/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean b/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean index 768a35876f499..670d14d17f25a 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean @@ -52,7 +52,7 @@ theorem decomposition_Q (n q : ℕ) : ∑ i ∈ Finset.filter (fun i : Fin (n + 1) => (i : ℕ) < q) Finset.univ, (P i).f (n + 1) ≫ X.δ i.rev.succ ≫ X.σ (Fin.rev i) := by induction' q with q hq - · simp only [Nat.zero_eq, Q_zero, HomologicalComplex.zero_f_apply, Nat.not_lt_zero, + · simp only [Q_zero, HomologicalComplex.zero_f_apply, Nat.not_lt_zero, Finset.filter_False, Finset.sum_empty] · by_cases hqn : q + 1 ≤ n + 1 swap diff --git a/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean b/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean index 08e001f128c59..f90156e668cfe 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean @@ -60,7 +60,7 @@ def homotopyPInftyToId : Homotopy (PInfty : K[X] ⟶ _) (𝟙 _) where comm n := by rcases n with _|n · simpa only [Homotopy.dNext_zero_chainComplex, Homotopy.prevD_chainComplex, - PInfty_f, Nat.zero_eq, P_f_0_eq, zero_add] using (homotopyPToId X 2).comm 0 + PInfty_f, P_f_0_eq, zero_add] using (homotopyPToId X 2).comm 0 · simp only [Homotopy.dNext_succ_chainComplex, Homotopy.prevD_chainComplex, HomologicalComplex.id_f, PInfty_f, ← P_is_eventually_constant (rfl.le : n + 1 ≤ n + 1)] -- Porting note(lean4/2146): remaining proof was diff --git a/Mathlib/AlgebraicTopology/DoldKan/PInfty.lean b/Mathlib/AlgebraicTopology/DoldKan/PInfty.lean index b17a90763626e..3ca7762569ccb 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/PInfty.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/PInfty.lean @@ -34,7 +34,7 @@ variable {C : Type*} [Category C] [Preadditive C] {X : SimplicialObject C} theorem P_is_eventually_constant {q n : ℕ} (hqn : n ≤ q) : ((P (q + 1)).f n : X _[n] ⟶ _) = (P q).f n := by rcases n with (_|n) - · simp only [Nat.zero_eq, P_f_0_eq] + · simp only [P_f_0_eq] · simp only [P_succ, add_right_eq_self, comp_add, HomologicalComplex.comp_f, HomologicalComplex.add_f_apply, comp_id] exact (HigherFacesVanish.of_P q n).comp_Hσ_eq_zero (Nat.succ_le_iff.mp hqn) diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index a9402e71c3a01..3003a0457ca57 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -652,7 +652,7 @@ theorem eq_σ_comp_of_not_injective' {n : ℕ} {Δ' : SimplexCategory} (θ : mk cases' Nat.le.dest h' with c hc cases c · exfalso - simp only [Nat.zero_eq, add_zero, len_mk, Fin.coe_pred] at hc + simp only [add_zero, len_mk, Fin.coe_pred] at hc rw [hc] at h'' exact h'' rfl · rw [← hc] diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index cce923522de40..81b7d49ceba23 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -897,9 +897,9 @@ theorem Asymptotics.IsBigO.continuousMultilinearMap_apply_eq_zero {n : ℕ} {p : cases' n with n · exact norm_eq_zero.mp (by -- Porting note: the symmetric difference of the `simpa only` sets: - -- added `Nat.zero_eq, zero_add, pow_one` + -- added `zero_add, pow_one` -- removed `zero_pow, Ne.def, Nat.one_ne_zero, not_false_iff` - simpa only [Nat.zero_eq, fin0_apply_norm, norm_eq_zero, norm_zero, zero_add, pow_one, + simpa only [fin0_apply_norm, norm_eq_zero, norm_zero, zero_add, pow_one, mul_zero, norm_le_zero_iff] using ht 0 (δε (Metric.mem_ball_self δ_pos))) · refine Or.elim (Classical.em (y = 0)) (fun hy => by simpa only [hy] using p.map_zero) fun hy => ?_ diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index daa0f2c5172fc..7942e61582f38 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -183,7 +183,7 @@ lemma AnalyticAt.pow {f : E → A} {z : E} (hf : AnalyticAt 𝕜 f z) (n : ℕ) AnalyticAt 𝕜 (fun x ↦ f x ^ n) z := by induction n with | zero => - simp only [Nat.zero_eq, pow_zero] + simp only [pow_zero] apply analyticAt_const | succ m hm => simp only [pow_succ] diff --git a/Mathlib/Analysis/Analytic/Meromorphic.lean b/Mathlib/Analysis/Analytic/Meromorphic.lean index e2dd5221e27e3..a447d8a51c6d6 100644 --- a/Mathlib/Analysis/Analytic/Meromorphic.lean +++ b/Mathlib/Analysis/Analytic/Meromorphic.lean @@ -134,7 +134,7 @@ lemma div {f g : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (hg : Meromo lemma pow {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (n : ℕ) : MeromorphicAt (f ^ n) x := by induction' n with m hm - · simpa only [Nat.zero_eq, pow_zero] using MeromorphicAt.const 1 x + · simpa only [pow_zero] using MeromorphicAt.const 1 x · simpa only [pow_succ] using hm.mul hf lemma zpow {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (n : ℤ) : MeromorphicAt (f ^ n) x := by diff --git a/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean b/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean index 45672ad5d7b49..f0990ce3adde0 100644 --- a/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean +++ b/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean @@ -104,7 +104,7 @@ theorem SuperpolynomialDecay.mul_param (hf : SuperpolynomialDecay l k f) : theorem SuperpolynomialDecay.param_pow_mul (hf : SuperpolynomialDecay l k f) (n : ℕ) : SuperpolynomialDecay l k (k ^ n * f) := by induction n with - | zero => simpa only [Nat.zero_eq, one_mul, pow_zero] using hf + | zero => simpa only [one_mul, pow_zero] using hf | succ n hn => simpa only [pow_succ', mul_assoc] using hn.param_mul theorem SuperpolynomialDecay.mul_param_pow (hf : SuperpolynomialDecay l k f) (n : ℕ) : diff --git a/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean b/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean index d5bc8119aefce..dd973734df63e 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean @@ -130,7 +130,7 @@ theorem subbox_induction_on' {p : Box ι → Prop} (I : Box ι) have hJsub : ∀ m i, (J m).upper i - (J m).lower i = (I.upper i - I.lower i) / 2 ^ m := by intro m i induction' m with m ihm - · simp [J, Nat.zero_eq] + · simp [J] simp only [pow_succ, J_succ, upper_sub_lower_splitCenterBox, ihm, div_div] have h0 : J 0 = I := rfl clear_value J diff --git a/Mathlib/Analysis/CStarAlgebra/Basic.lean b/Mathlib/Analysis/CStarAlgebra/Basic.lean index f0cb7d301a058..34c2684caead7 100644 --- a/Mathlib/Analysis/CStarAlgebra/Basic.lean +++ b/Mathlib/Analysis/CStarAlgebra/Basic.lean @@ -233,7 +233,7 @@ end CStarRing theorem IsSelfAdjoint.nnnorm_pow_two_pow [NormedRing E] [StarRing E] [CStarRing E] {x : E} (hx : IsSelfAdjoint x) (n : ℕ) : ‖x ^ 2 ^ n‖₊ = ‖x‖₊ ^ 2 ^ n := by induction' n with k hk - · simp only [pow_zero, pow_one, Nat.zero_eq] + · simp only [pow_zero, pow_one] · rw [pow_succ', pow_mul', sq] nth_rw 1 [← selfAdjoint.mem_iff.mp hx] rw [← star_pow, CStarRing.nnnorm_star_mul_self, ← sq, hk, pow_mul'] diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index d45ba684d6e3d..bc047658d41d0 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -240,7 +240,7 @@ theorem ContinuousLinearEquiv.iteratedFDerivWithin_comp_left (g : F ≃L[𝕜] G (g : F →L[𝕜] G).compContinuousMultilinearMap (iteratedFDerivWithin 𝕜 i f s x) := by induction' i with i IH generalizing x · ext1 m - simp only [Nat.zero_eq, iteratedFDerivWithin_zero_apply, comp_apply, + simp only [iteratedFDerivWithin_zero_apply, comp_apply, ContinuousLinearMap.compContinuousMultilinearMap_coe, coe_coe] · ext1 m rw [iteratedFDerivWithin_succ_apply_left] @@ -382,7 +382,7 @@ theorem ContinuousLinearEquiv.iteratedFDerivWithin_comp_right (g : G ≃L[𝕜] (iteratedFDerivWithin 𝕜 i f s (g x)).compContinuousLinearMap fun _ => g := by induction' i with i IH generalizing x · ext1 - simp only [Nat.zero_eq, iteratedFDerivWithin_zero_apply, comp_apply, + simp only [iteratedFDerivWithin_zero_apply, comp_apply, ContinuousMultilinearMap.compContinuousLinearMap_apply] · ext1 m simp only [ContinuousMultilinearMap.compContinuousLinearMap_apply, diff --git a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean index 9d0be42f520d6..d5cce4ded4d53 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean @@ -50,7 +50,7 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear_aux {Du Eu original spaces, which explains why we assume in the lemma that all spaces live in the same universe. -/ induction' n with n IH generalizing Eu Fu Gu - · simp only [Nat.zero_eq, norm_iteratedFDerivWithin_zero, zero_add, Finset.range_one, + · 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] diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index f5ce7711b052d..1bde1bf2752a2 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -1493,7 +1493,7 @@ theorem iteratedFDerivWithin_of_isOpen (n : ℕ) (hs : IsOpen s) : | zero => intro x _ ext1 - simp only [Nat.zero_eq, iteratedFDerivWithin_zero_apply, iteratedFDeriv_zero_apply] + simp only [iteratedFDerivWithin_zero_apply, iteratedFDeriv_zero_apply] | succ n IH => intro x hx rw [iteratedFDeriv_succ_eq_comp_left, iteratedFDerivWithin_succ_eq_comp_left] diff --git a/Mathlib/Analysis/Calculus/Deriv/ZPow.lean b/Mathlib/Analysis/Calculus/Deriv/ZPow.lean index 29c9091921c56..c8b4e4b0dd3f7 100644 --- a/Mathlib/Analysis/Calculus/Deriv/ZPow.lean +++ b/Mathlib/Analysis/Calculus/Deriv/ZPow.lean @@ -96,7 +96,7 @@ theorem iter_deriv_zpow' (m : ℤ) (k : ℕ) : (deriv^[k] fun x : 𝕜 => x ^ m) = fun x => (∏ i ∈ Finset.range k, ((m : 𝕜) - i)) * x ^ (m - k) := by induction' k with k ihk - · simp only [Nat.zero_eq, one_mul, Int.ofNat_zero, id, sub_zero, Finset.prod_range_zero, + · simp only [one_mul, Int.ofNat_zero, id, sub_zero, Finset.prod_range_zero, Function.iterate_zero] · simp only [Function.iterate_succ_apply', ihk, deriv_const_mul_field', deriv_zpow', Finset.prod_range_succ, Int.ofNat_succ, ← sub_sub, Int.cast_sub, Int.cast_natCast, mul_assoc] diff --git a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean index a4e5c528314f8..2a99f4113e978 100644 --- a/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean +++ b/Mathlib/Analysis/Calculus/FormalMultilinearSeries.lean @@ -333,7 +333,7 @@ lemma constFormalMultilinearSeries_zero [NontriviallyNormedField 𝕜] [NormedAd simp only [FormalMultilinearSeries.zero_apply, ContinuousMultilinearMap.zero_apply, constFormalMultilinearSeries] induction n - · simp only [Nat.zero_eq, ContinuousMultilinearMap.curry0_apply] + · simp only [ContinuousMultilinearMap.curry0_apply] · simp only [constFormalMultilinearSeries.match_1.eq_2, ContinuousMultilinearMap.zero_apply] end Const diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean index f81f092709a46..9b8588e207bce 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Lemmas.lean @@ -117,7 +117,7 @@ lemma iteratedDeriv_neg (n : ℕ) (f : 𝕜 → F) (a : 𝕜) : lemma iteratedDeriv_comp_neg (n : ℕ) (f : 𝕜 → F) (a : 𝕜) : iteratedDeriv n (fun x ↦ f (-x)) a = (-1 : 𝕜) ^ n • iteratedDeriv n f (-a) := by induction' n with n ih generalizing a - · simp only [Nat.zero_eq, iteratedDeriv_zero, pow_zero, one_smul] + · simp only [iteratedDeriv_zero, pow_zero, one_smul] · have ih' : iteratedDeriv n (fun x ↦ f (-x)) = fun x ↦ (-1 : 𝕜) ^ n • iteratedDeriv n f (-x) := funext ih rw [iteratedDeriv_succ, iteratedDeriv_succ, ih', pow_succ', neg_mul, one_mul, diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index a4616a71f06cb..dd4ff0411f5f1 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -522,7 +522,7 @@ theorem fourierIntegral_iteratedFDeriv [FiniteDimensional ℝ V] induction n with | zero => ext w m - simp only [iteratedFDeriv_zero_apply, Nat.zero_eq, fourierPowSMulRight_apply, pow_zero, + simp only [iteratedFDeriv_zero_apply, fourierPowSMulRight_apply, pow_zero, Finset.univ_eq_empty, ContinuousLinearMap.neg_apply, ContinuousLinearMap.flip_apply, Finset.prod_empty, one_smul, fourierIntegral_continuousMultilinearMap_apply' ((h'f 0 bot_le))] | succ n ih => diff --git a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean index cc637d03ff57a..ad51b9d339160 100644 --- a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean +++ b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean @@ -233,7 +233,7 @@ theorem add_left (x y z : E) : inner_ 𝕜 (x + y) z = inner_ 𝕜 x z + inner_ theorem nat (n : ℕ) (x y : E) : inner_ 𝕜 ((n : 𝕜) • x) y = (n : 𝕜) * inner_ 𝕜 x y := by induction' n with n ih - · simp only [inner_, Nat.zero_eq, zero_sub, Nat.cast_zero, zero_mul, + · simp only [inner_, zero_sub, Nat.cast_zero, zero_mul, eq_self_iff_true, zero_smul, zero_add, mul_zero, sub_self, norm_neg, smul_zero] · simp only [Nat.cast_succ, add_smul, one_smul] rw [add_left, ih, add_mul, one_mul] diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index 560f636a6afa2..cef2a8e2bc38e 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -280,7 +280,7 @@ theorem hasFPowerSeriesOnBall_inverse_one_sub_smul [CompleteSpace A] (a : A) : le_radius_of_bound_nnreal _ (max 1 ‖(1 : A)‖₊) fun n => ?_ rw [← norm_toNNReal, norm_mkPiRing, norm_toNNReal] cases' n with n - · simp only [Nat.zero_eq, le_refl, mul_one, or_true_iff, le_max_iff, pow_zero] + · simp only [le_refl, mul_one, or_true_iff, le_max_iff, pow_zero] · refine le_trans (le_trans (mul_le_mul_right' (nnnorm_pow_le' a n.succ_pos) (r ^ n.succ)) ?_) (le_max_left _ _) diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index fbcfe54ff4ad9..c7963a31c6d62 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -412,7 +412,7 @@ theorem nnnorm_pow_le' (a : α) : ∀ {n : ℕ}, 0 < n → ‖a ^ n‖₊ ≤ /-- If `α` is a seminormed ring with `‖1‖₊ = 1`, then `‖a ^ n‖₊ ≤ ‖a‖₊ ^ n`. See also `nnnorm_pow_le'`. -/ theorem nnnorm_pow_le [NormOneClass α] (a : α) (n : ℕ) : ‖a ^ n‖₊ ≤ ‖a‖₊ ^ n := - Nat.recOn n (by simp only [Nat.zero_eq, pow_zero, nnnorm_one, le_rfl]) + Nat.recOn n (by simp only [pow_zero, nnnorm_one, le_rfl]) fun k _hk => nnnorm_pow_le' a k.succ_pos /-- If `α` is a seminormed ring, then `‖a ^ n‖ ≤ ‖a‖ ^ n` for `n > 0`. See also `norm_pow_le`. -/ @@ -422,7 +422,7 @@ theorem norm_pow_le' (a : α) {n : ℕ} (h : 0 < n) : ‖a ^ n‖ ≤ ‖a‖ ^ /-- If `α` is a seminormed ring with `‖1‖ = 1`, then `‖a ^ n‖ ≤ ‖a‖ ^ n`. See also `norm_pow_le'`. -/ theorem norm_pow_le [NormOneClass α] (a : α) (n : ℕ) : ‖a ^ n‖ ≤ ‖a‖ ^ n := - Nat.recOn n (by simp only [Nat.zero_eq, pow_zero, norm_one, le_rfl]) + Nat.recOn n (by simp only [pow_zero, norm_one, le_rfl]) fun n _hn => norm_pow_le' a n.succ_pos theorem eventually_norm_pow_le (a : α) : ∀ᶠ n : ℕ in atTop, ‖a ^ n‖ ≤ ‖a‖ ^ n := diff --git a/Mathlib/Analysis/Normed/Operator/Banach.lean b/Mathlib/Analysis/Normed/Operator/Banach.lean index 9a19ff37a6355..5d6eca063bbe1 100644 --- a/Mathlib/Analysis/Normed/Operator/Banach.lean +++ b/Mathlib/Analysis/Normed/Operator/Banach.lean @@ -178,7 +178,7 @@ theorem exists_preimage_norm_le (surj : Surjective f) : have hnle : ∀ n : ℕ, ‖h^[n] y‖ ≤ (1 / 2) ^ n * ‖y‖ := by intro n induction n with - | zero => simp only [one_div, Nat.zero_eq, one_mul, iterate_zero_apply, pow_zero, le_rfl] + | zero => simp only [one_div, one_mul, iterate_zero_apply, pow_zero, le_rfl] | succ n IH => rw [iterate_succ'] apply le_trans (hle _) _ diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index bd137a24eea7c..3c575bbb3d974 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -258,7 +258,7 @@ noncomputable def GammaAux : ℕ → ℂ → ℂ theorem GammaAux_recurrence1 (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : GammaAux n s = GammaAux n (s + 1) / s := by induction' n with n hn generalizing s - · simp only [Nat.zero_eq, CharP.cast_eq_zero, Left.neg_neg_iff] at h1 + · simp only [CharP.cast_eq_zero, Left.neg_neg_iff] at h1 dsimp only [GammaAux]; rw [GammaIntegral_add_one h1] rw [mul_comm, mul_div_cancel_right₀]; contrapose! h1; rw [h1] simp @@ -271,7 +271,7 @@ theorem GammaAux_recurrence1 (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : theorem GammaAux_recurrence2 (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : GammaAux n s = GammaAux (n + 1) s := by cases' n with n n - · simp only [Nat.zero_eq, CharP.cast_eq_zero, Left.neg_neg_iff] at h1 + · simp only [CharP.cast_eq_zero, Left.neg_neg_iff] at h1 dsimp only [GammaAux] rw [GammaIntegral_add_one h1, mul_div_cancel_left₀] rintro rfl diff --git a/Mathlib/Analysis/Subadditive.lean b/Mathlib/Analysis/Subadditive.lean index e3104ca58d187..00d92dc291067 100644 --- a/Mathlib/Analysis/Subadditive.lean +++ b/Mathlib/Analysis/Subadditive.lean @@ -46,7 +46,7 @@ theorem lim_le_div (hbdd : BddBelow (range fun n => u n / n)) {n : ℕ} (hn : n include h in theorem apply_mul_add_le (k n r) : u (k * n + r) ≤ k * u n + u r := by induction k with - | zero => simp only [Nat.zero_eq, Nat.cast_zero, zero_mul, zero_add]; rfl + | zero => simp only [Nat.cast_zero, zero_mul, zero_add]; rfl | succ k IH => calc u ((k + 1) * n + r) = u (n + (k * n + r)) := by congr 1; ring diff --git a/Mathlib/CategoryTheory/ComposableArrows.lean b/Mathlib/CategoryTheory/ComposableArrows.lean index 45883d4350c3a..d0547116852eb 100644 --- a/Mathlib/CategoryTheory/ComposableArrows.lean +++ b/Mathlib/CategoryTheory/ComposableArrows.lean @@ -185,7 +185,7 @@ def homMk {F G : ComposableArrows C n} (app : ∀ i, F.obj i ⟶ G.obj i) intro k induction' k with k hk · intro i j hj hj' - simp only [Nat.zero_eq, add_zero] at hj + simp only [add_zero] at hj obtain rfl := hj rw [F.map'_self i, G.map'_self i, id_comp, comp_id] · intro i j hj hj' diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index 13b38d99b0769..14f206f8513f1 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -463,7 +463,7 @@ instance sheafHomHasNSMul : SMul ℕ (P ⟶ Q) where { app := fun U => n • f.1.app U naturality := fun U V i => by induction n with - | zero => simp only [zero_smul, comp_zero, zero_comp, Nat.zero_eq] + | zero => simp only [zero_smul, comp_zero, zero_comp] | succ n ih => simp only [Nat.succ_eq_add_one, add_smul, ih, one_nsmul, comp_add, NatTrans.naturality, add_comp] } diff --git a/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean b/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean index 8c7637a070683..d9f2aef25fade 100644 --- a/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean +++ b/Mathlib/Computability/AkraBazzi/GrowsPolynomially.lean @@ -96,7 +96,7 @@ lemma eventually_zero_of_frequently_zero (hf : GrowsPolynomially f) (hf' : ∃ intro m induction m with | zero => - simp only [Nat.zero_eq, CharP.cast_eq_zero, neg_zero, zero_sub, zpow_zero, one_mul] at * + simp only [CharP.cast_eq_zero, neg_zero, zero_sub, zpow_zero, one_mul] at * specialize hx x₀ (le_of_max_le_left hx₀_ge) simp only [hx₀, mul_zero, Set.Icc_self, Set.mem_singleton_iff] at hx refine fun z _ hz => hx _ ?_ diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 70ae4513ecba5..99a30824f06d0 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -310,10 +310,10 @@ theorem ListBlank.nth_modifyNth {Γ} [Inhabited Γ] (f : Γ → Γ) (n i) (L : L (L.modifyNth f n).nth i = if i = n then f (L.nth i) else L.nth i := by induction' n with n IH generalizing i L · cases i <;> simp only [ListBlank.nth_zero, if_true, ListBlank.head_cons, ListBlank.modifyNth, - ListBlank.nth_succ, if_false, ListBlank.tail_cons, Nat.zero_eq] + ListBlank.nth_succ, if_false, ListBlank.tail_cons] · cases i · rw [if_neg (Nat.succ_ne_zero _).symm] - simp only [ListBlank.nth_zero, ListBlank.head_cons, ListBlank.modifyNth, Nat.zero_eq] + simp only [ListBlank.nth_zero, ListBlank.head_cons, ListBlank.modifyNth] · simp only [IH, ListBlank.modifyNth, ListBlank.nth_succ, ListBlank.tail_cons, Nat.succ.injEq] /-- A pointed map of `Inhabited` types is a map that sends one default value to the other. -/ @@ -553,7 +553,7 @@ theorem Tape.nth_zero {Γ} [Inhabited Γ] (T : Tape Γ) : T.nth 0 = T.1 := theorem Tape.right₀_nth {Γ} [Inhabited Γ] (T : Tape Γ) (n : ℕ) : T.right₀.nth n = T.nth n := by cases n <;> simp only [Tape.nth, Tape.right₀, Int.ofNat_zero, ListBlank.nth_zero, - ListBlank.nth_succ, ListBlank.head_cons, ListBlank.tail_cons, Nat.zero_eq] + ListBlank.nth_succ, ListBlank.head_cons, ListBlank.tail_cons] @[simp] theorem Tape.mk'_nth_nat {Γ} [Inhabited Γ] (L R : ListBlank Γ) (n : ℕ) : @@ -626,7 +626,7 @@ theorem Tape.write_move_right_n {Γ} [Inhabited Γ] (f : Γ → Γ) (L R : ListB ((Tape.move Dir.right)^[n] (Tape.mk' L R)).write (f (R.nth n)) = (Tape.move Dir.right)^[n] (Tape.mk' L (R.modifyNth f n)) := by induction' n with n IH generalizing L R - · simp only [ListBlank.nth_zero, ListBlank.modifyNth, iterate_zero_apply, Nat.zero_eq] + · simp only [ListBlank.nth_zero, ListBlank.modifyNth, iterate_zero_apply] rw [← Tape.write_mk', ListBlank.cons_head_tail] simp only [ListBlank.head_cons, ListBlank.nth_succ, ListBlank.modifyNth, Tape.move_right_mk', ListBlank.tail_cons, iterate_succ_apply, IH] diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index cd8cf284df159..ec2e6c8e48d03 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -1451,7 +1451,7 @@ theorem add_one_le_of_lt {n : ℕ} {a b : Fin (n + 1)} (h : a < b) : a + 1 ≤ b cases' a with a ha cases' b with b hb cases n - · simp only [Nat.zero_eq, Nat.zero_add, Nat.lt_one_iff] at ha hb + · simp only [Nat.zero_add, Nat.lt_one_iff] at ha hb simp [ha, hb] simp only [le_iff_val_le_val, val_add, lt_iff_val_lt_val, val_mk, val_one] at h ⊢ rwa [Nat.mod_eq_of_lt, Nat.succ_le_iff] @@ -1469,7 +1469,7 @@ theorem exists_eq_add_of_lt {n : ℕ} {a b : Fin (n + 1)} (h : a < b) : cases n · cases' a with a ha cases' b with b hb - simp only [Nat.zero_eq, Nat.zero_add, Nat.lt_one_iff] at ha hb + simp only [Nat.zero_add, Nat.lt_one_iff] at ha hb simp [ha, hb] at h obtain ⟨k, hk⟩ : ∃ k : ℕ, (b : ℕ) = a + k + 1 := Nat.exists_eq_add_of_lt h have hkb : k < b := by omega diff --git a/Mathlib/Data/Fin/VecNotation.lean b/Mathlib/Data/Fin/VecNotation.lean index 56651d1ba5421..fa80efe84fef3 100644 --- a/Mathlib/Data/Fin/VecNotation.lean +++ b/Mathlib/Data/Fin/VecNotation.lean @@ -317,7 +317,7 @@ theorem vecAlt1_vecAppend (v : Fin (n + 1) → α) : cases n with | zero => cases' i with i hi - simp only [Nat.zero_eq, Nat.zero_add, Nat.lt_one_iff] at hi; subst i; rfl + simp only [Nat.zero_add, Nat.lt_one_iff] at hi; subst i; rfl | succ n => split_ifs with h <;> congr · simp [Nat.mod_eq_of_lt, h] diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index d16894eab2b25..bcc952df08431 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -289,12 +289,12 @@ theorem prev_nthLe (l : List α) (h : Nodup l) (n : ℕ) (hn : n < l.length) : induction' l with y l hl generalizing n x · simp · rcases n with (_ | _ | n) - · simp [Nat.add_succ_sub_one, add_zero, List.prev_cons_cons_eq, Nat.zero_eq, List.length, + · simp [Nat.add_succ_sub_one, add_zero, List.prev_cons_cons_eq, List.length, List.nthLe, Nat.succ_add_sub_one, zero_add, getLast_eq_get, Nat.mod_eq_of_lt (Nat.succ_lt_succ l.length.lt_succ_self)] · simp only [mem_cons, nodup_cons] at h push_neg at h - simp only [List.prev_cons_cons_of_ne _ _ _ _ h.left.left.symm, Nat.zero_eq, List.length, + simp only [List.prev_cons_cons_of_ne _ _ _ _ h.left.left.symm, List.length, List.nthLe, add_comm, eq_self_iff_true, Nat.succ_add_sub_one, Nat.mod_self, zero_add, List.get] · rw [prev_ne_cons_cons] @@ -806,7 +806,7 @@ nonrec def Chain (r : α → α → Prop) (c : Cycle α) : Prop := · dsimp only cases' hab with n hn induction' n with d hd generalizing a b l m - · simp only [Nat.zero_eq, rotate_zero, cons.injEq] at hn + · simp only [rotate_zero, cons.injEq] at hn rw [hn.1, hn.2] · cases' l with c s · simp only [rotate_cons_succ, nil_append, rotate_singleton, cons.injEq] at hn diff --git a/Mathlib/Data/Nat/Hyperoperation.lean b/Mathlib/Data/Nat/Hyperoperation.lean index 2aa1fb4b12173..ea0cc84e74cc6 100644 --- a/Mathlib/Data/Nat/Hyperoperation.lean +++ b/Mathlib/Data/Nat/Hyperoperation.lean @@ -104,7 +104,7 @@ theorem hyperoperation_ge_four_zero (n k : ℕ) : hyperoperation (n + 4) 0 k = if Even k then 1 else 0 := by induction' k with kk kih · rw [hyperoperation_ge_three_eq_one] - simp only [Nat.zero_eq, even_zero, if_true] + simp only [even_zero, if_true] · rw [hyperoperation_recursion] rw [kih] simp_rw [Nat.even_add_one] diff --git a/Mathlib/Data/Nat/Totient.lean b/Mathlib/Data/Nat/Totient.lean index 3bfd3d40e7e5c..938f58268c44a 100644 --- a/Mathlib/Data/Nat/Totient.lean +++ b/Mathlib/Data/Nat/Totient.lean @@ -76,7 +76,7 @@ theorem Ico_filter_coprime_le {a : ℕ} (k n : ℕ) (a_pos : 0 < a) : induction' n / a with i ih · rw [← filter_coprime_Ico_eq_totient a k] simp only [add_zero, mul_one, mul_zero, le_of_lt (mod_lt n a_pos), - Nat.zero_eq, zero_add] + zero_add] gcongr exact Ico_subset_Ico rfl.le (add_le_add_left (le_of_lt (mod_lt n a_pos)) k) simp only [mul_succ] diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index 54d7f68a2f02a..c3669ef62772a 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -863,7 +863,7 @@ theorem castNum_testBit (m n) : testBit m n = Nat.testBit m n := by | pos m => rw [cast_pos] induction' n with n IH generalizing m <;> cases' m with m m - <;> dsimp only [PosNum.testBit, Nat.zero_eq] + <;> dsimp only [PosNum.testBit] · rfl · rw [PosNum.cast_bit1, ← two_mul, ← congr_fun Nat.bit_true, Nat.testBit_bit_zero] · rw [PosNum.cast_bit0, ← two_mul, ← congr_fun Nat.bit_false, Nat.testBit_bit_zero] diff --git a/Mathlib/Data/PNat/Factors.lean b/Mathlib/Data/PNat/Factors.lean index 80e5d15d4b64e..32cae0329ae6c 100644 --- a/Mathlib/Data/PNat/Factors.lean +++ b/Mathlib/Data/PNat/Factors.lean @@ -188,7 +188,7 @@ theorem prod_add (u v : PrimeMultiset) : (u + v).prod = u.prod * v.prod := by theorem prod_smul (d : ℕ) (u : PrimeMultiset) : (d • u).prod = u.prod ^ d := by induction d with - | zero => simp only [Nat.zero_eq, zero_nsmul, pow_zero, prod_zero] + | zero => simp only [zero_nsmul, pow_zero, prod_zero] | succ n ih => rw [succ_nsmul, prod_add, ih, pow_succ] end PrimeMultiset diff --git a/Mathlib/Data/Seq/Computation.lean b/Mathlib/Data/Seq/Computation.lean index 93868db28c7bc..9b4a265a6475b 100644 --- a/Mathlib/Data/Seq/Computation.lean +++ b/Mathlib/Data/Seq/Computation.lean @@ -695,7 +695,7 @@ theorem length_bind (s : Computation α) (f : α → Computation β) [_T1 : Term theorem of_results_bind {s : Computation α} {f : α → Computation β} {b k} : Results (bind s f) b k → ∃ a m n, Results s a m ∧ Results (f a) b n ∧ k = n + m := by induction' k with n IH generalizing s <;> apply recOn s (fun a => _) fun s' => _ <;> intro e h - · simp only [ret_bind, Nat.zero_eq] at h + · simp only [ret_bind] at h exact ⟨e, _, _, results_pure _, h, rfl⟩ · have := congr_arg head (eq_thinkN h) contradiction diff --git a/Mathlib/Data/Set/Enumerate.lean b/Mathlib/Data/Set/Enumerate.lean index 51a53f2a30a03..c6313777fa313 100644 --- a/Mathlib/Data/Set/Enumerate.lean +++ b/Mathlib/Data/Set/Enumerate.lean @@ -81,7 +81,7 @@ theorem enumerate_inj {n₁ n₂ : ℕ} {a : α} {s : Set α} (h_sel : ∀ s a, | zero => rfl | succ m => have h' : enumerate sel (s \ {a}) m = some a := by - simp_all only [enumerate, Nat.zero_eq, Nat.add_eq, zero_add]; exact h₂ + simp_all only [enumerate, Nat.add_eq, zero_add]; exact h₂ have : a ∈ s \ {a} := enumerate_mem sel h_sel h' simp_all [Set.mem_diff_singleton] | succ k ih => diff --git a/Mathlib/Data/Set/MemPartition.lean b/Mathlib/Data/Set/MemPartition.lean index eb38326f9234c..25b578031b0f6 100644 --- a/Mathlib/Data/Set/MemPartition.lean +++ b/Mathlib/Data/Set/MemPartition.lean @@ -54,7 +54,7 @@ lemma disjoint_memPartition (f : ℕ → Set α) (n : ℕ) {u v : Set α} induction n with | zero => intro u v hu hv huv - simp only [Nat.zero_eq, memPartition_zero, mem_insert_iff, mem_singleton_iff] at hu hv + simp only [memPartition_zero, mem_insert_iff, mem_singleton_iff] at hu hv rw [hu, hv] at huv exact absurd rfl huv | succ n ih => diff --git a/Mathlib/Data/Set/Pointwise/Finite.lean b/Mathlib/Data/Set/Pointwise/Finite.lean index 6e4932af9c42e..02e69997156ee 100644 --- a/Mathlib/Data/Set/Pointwise/Finite.lean +++ b/Mathlib/Data/Set/Pointwise/Finite.lean @@ -62,7 +62,7 @@ instance decidableMemMul [Fintype α] [DecidableEq α] [DecidablePred (· ∈ s) instance decidableMemPow [Fintype α] [DecidableEq α] [DecidablePred (· ∈ s)] (n : ℕ) : DecidablePred (· ∈ s ^ n) := by induction' n with n ih - · simp only [Nat.zero_eq, pow_zero, mem_one] + · simp only [pow_zero, mem_one] infer_instance · letI := ih rw [pow_succ] diff --git a/Mathlib/Data/Vector/Basic.lean b/Mathlib/Data/Vector/Basic.lean index f28e84a817f60..75c60fd7cda24 100644 --- a/Mathlib/Data/Vector/Basic.lean +++ b/Mathlib/Data/Vector/Basic.lean @@ -312,7 +312,7 @@ retrieved via `head`, is the starting value `b : β`. @[simp] theorem scanl_head : (scanl f b v).head = b := by cases n - · have : v = nil := by simp only [Nat.zero_eq, eq_iff_true_of_subsingleton] + · have : v = nil := by simp only [eq_iff_true_of_subsingleton] simp only [this, scanl_nil, head_cons] · rw [← cons_head_tail v] simp only [← get_zero, get_eq_get, toList_scanl, toList_cons, List.scanl, Fin.val_zero, diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 0a65bb9d6fac8..fed33c8dea188 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -100,7 +100,7 @@ instance charP (n : ℕ) : CharP (ZMod n) n where cast_eq_zero_iff' := by intro k cases' n with n - · simp [zero_dvd_iff, Int.natCast_eq_zero, Nat.zero_eq] + · simp [zero_dvd_iff, Int.natCast_eq_zero] · exact Fin.natCast_eq_zero @[simp] @@ -112,7 +112,7 @@ where `a ≠ 0` is `addOrderOf_coe'`. -/ @[simp] theorem addOrderOf_coe (a : ℕ) {n : ℕ} (n0 : n ≠ 0) : addOrderOf (a : ZMod n) = n / n.gcd a := by cases' a with a - · simp only [Nat.zero_eq, Nat.cast_zero, addOrderOf_zero, Nat.gcd_zero_right, + · simp only [Nat.cast_zero, addOrderOf_zero, Nat.gcd_zero_right, Nat.pos_of_ne_zero n0, Nat.div_self] rw [← Nat.smul_one_eq_cast, addOrderOf_nsmul' _ a.succ_ne_zero, ZMod.addOrderOf_one] diff --git a/Mathlib/FieldTheory/PerfectClosure.lean b/Mathlib/FieldTheory/PerfectClosure.lean index 8572b59904f89..924955c214e97 100644 --- a/Mathlib/FieldTheory/PerfectClosure.lean +++ b/Mathlib/FieldTheory/PerfectClosure.lean @@ -243,7 +243,7 @@ theorem R.sound (m n : ℕ) (x y : K) (H : (frobenius K p)^[m] x = y) : mk K p (n, x) = mk K p (m + n, y) := by subst H induction' m with m ih - · simp only [Nat.zero_eq, zero_add, iterate_zero_apply] + · simp only [zero_add, iterate_zero_apply] rw [ih, Nat.succ_add, iterate_succ'] apply Quot.sound apply R.intro diff --git a/Mathlib/GroupTheory/FiniteAbelian.lean b/Mathlib/GroupTheory/FiniteAbelian.lean index 6ec1d539f832d..5c79378d634fc 100644 --- a/Mathlib/GroupTheory/FiniteAbelian.lean +++ b/Mathlib/GroupTheory/FiniteAbelian.lean @@ -131,7 +131,7 @@ theorem equiv_directSum_zmod_of_finite [Finite G] : obtain ⟨n, ι, fι, p, hp, e, ⟨f⟩⟩ := equiv_free_prod_directSum_zmod G cases' n with n · have : Unique (Fin Nat.zero →₀ ℤ) := - { uniq := by simp only [Nat.zero_eq, eq_iff_true_of_subsingleton]; trivial } + { uniq := by simp only [eq_iff_true_of_subsingleton]; trivial } exact ⟨ι, fι, p, hp, e, ⟨f.trans AddEquiv.uniqueProd⟩⟩ · haveI := @Fintype.prodLeft _ _ _ (Fintype.ofEquiv G f.toEquiv) _ exact diff --git a/Mathlib/GroupTheory/Nilpotent.lean b/Mathlib/GroupTheory/Nilpotent.lean index 6ec49137ba22e..290faebdaff6e 100644 --- a/Mathlib/GroupTheory/Nilpotent.lean +++ b/Mathlib/GroupTheory/Nilpotent.lean @@ -446,7 +446,7 @@ theorem upperCentralSeries.map {H : Type*} [Group H] {f : G →* H} (h : Functio theorem lowerCentralSeries.map {H : Type*} [Group H] (f : G →* H) (n : ℕ) : Subgroup.map f (lowerCentralSeries G n) ≤ lowerCentralSeries H n := by induction' n with d hd - · simp [Nat.zero_eq] + · simp · rintro a ⟨x, hx : x ∈ lowerCentralSeries G d.succ, rfl⟩ refine closure_induction hx ?_ (by simp [f.map_one, Subgroup.one_mem _]) (fun y z hy hz => by simp [MonoidHom.map_mul, Subgroup.mul_mem _ hy hz]) (fun y hy => by @@ -534,7 +534,7 @@ private theorem comap_center_subst {H₁ H₂ : Subgroup G} [Normal H₁] [Norma theorem comap_upperCentralSeries_quotient_center (n : ℕ) : comap (mk' (center G)) (upperCentralSeries (G ⧸ center G) n) = upperCentralSeries G n.succ := by induction' n with n ih - · simp only [Nat.zero_eq, upperCentralSeries_zero, MonoidHom.comap_bot, ker_mk', + · simp only [upperCentralSeries_zero, MonoidHom.comap_bot, ker_mk', (upperCentralSeries_one G).symm] · let Hn := upperCentralSeries (G ⧸ center G) n calc diff --git a/Mathlib/GroupTheory/Perm/Closure.lean b/Mathlib/GroupTheory/Perm/Closure.lean index 9ee578518cde5..6fa1105deb178 100644 --- a/Mathlib/GroupTheory/Perm/Closure.lean +++ b/Mathlib/GroupTheory/Perm/Closure.lean @@ -57,7 +57,7 @@ theorem closure_cycle_adjacent_swap {σ : Perm α} (h1 : IsCycle σ) (h2 : σ.su intro n induction n with | zero => - simp only [Nat.zero_eq, pow_zero, coe_one, id_eq, swap_self, Set.mem_singleton_iff] + simp only [pow_zero, coe_one, id_eq, swap_self, Set.mem_singleton_iff] convert H.one_mem | succ n ih => by_cases h5 : x = (σ ^ n) x diff --git a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean index 924cd62e58286..8b3c477a43b28 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean @@ -301,7 +301,7 @@ theorem SameCycle.toList_isRotated {f : Perm α} {x y : α} (h : SameCycle f x y toList f x ~r toList f y := by by_cases hx : x ∈ f.support · obtain ⟨_ | k, _, hy⟩ := h.exists_pow_eq_of_mem_support hx - · simp only [coe_one, id, pow_zero, Nat.zero_eq] at hy + · simp only [coe_one, id, pow_zero] at hy -- Porting note: added `IsRotated.refl` simp [hy, IsRotated.refl] use k.succ diff --git a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean index a718ff65870f1..87dff651c4a05 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Factors.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Factors.lean @@ -281,7 +281,7 @@ theorem SameCycle.exists_pow_eq [DecidableEq α] [Fintype α] (f : Perm α) (h : · refine ⟨(f.cycleOf x).support.card, ?_, self_le_add_right _ _, ?_⟩ · refine zero_lt_one.trans (one_lt_card_support_of_ne_one ?_) simpa using hx - · simp only [Nat.zero_eq, pow_zero, coe_one, id_eq] at hk' + · simp only [pow_zero, coe_one, id_eq] at hk' subst hk' rw [← (isCycle_cycleOf _ <| mem_support.1 hx).orderOf, ← cycleOf_pow_apply_self, pow_orderOf_eq_one, one_apply] diff --git a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean index d5556a7dfff7f..4cc93bd2b2b32 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Zero.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Zero.lean @@ -191,7 +191,7 @@ lemma finrank_maxGenEigenspace (φ : Module.End K M) : generalize_proofs h' clear hx induction n with - | zero => simp only [Nat.zero_eq, pow_zero, one_apply] + | zero => simp only [pow_zero, one_apply] | succ n ih => simp only [pow_succ', LinearMap.mul_apply, ih, restrict_apply] end LinearMap diff --git a/Mathlib/LinearAlgebra/Matrix/Transvection.lean b/Mathlib/LinearAlgebra/Matrix/Transvection.lean index 1e06f78ec9172..24c27ffba8831 100644 --- a/Mathlib/LinearAlgebra/Matrix/Transvection.lean +++ b/Mathlib/LinearAlgebra/Matrix/Transvection.lean @@ -471,7 +471,7 @@ theorem mul_listTransvecRow_last_row (hM : M (inr unit) (inr unit) ≠ 0) (i : F simpa only [this, ite_eq_right_iff] using H r le_rfl intro k hk induction' k with n IH - · simp only [if_true, Matrix.mul_one, List.take_zero, zero_le', List.prod_nil, Nat.zero_eq] + · simp only [if_true, Matrix.mul_one, List.take_zero, zero_le', List.prod_nil] · have hnr : n < r := hk let n' : Fin r := ⟨n, hnr⟩ have A : diff --git a/Mathlib/MeasureTheory/Covering/Besicovitch.lean b/Mathlib/MeasureTheory/Covering/Besicovitch.lean index 46af5641e5aa1..29bfebfa3b723 100644 --- a/Mathlib/MeasureTheory/Covering/Besicovitch.lean +++ b/Mathlib/MeasureTheory/Covering/Besicovitch.lean @@ -743,7 +743,7 @@ theorem exists_disjoint_closedBall_covering_ae_of_finiteMeasure_aux (μ : Measur have Pu : ∀ n, P (u n) := by intro n induction' n with n IH - · simp only [P, u, Prod.forall, id, Function.iterate_zero, Nat.zero_eq] + · simp only [P, u, Prod.forall, id, Function.iterate_zero] simp only [Finset.not_mem_empty, IsEmpty.forall_iff, Finset.coe_empty, forall₂_true_iff, and_self_iff, pairwiseDisjoint_empty] · rw [u_succ] @@ -768,7 +768,7 @@ theorem exists_disjoint_closedBall_covering_ae_of_finiteMeasure_aux (μ : Measur intro n induction' n with n IH · simp only [u, le_refl, diff_empty, one_mul, iUnion_false, iUnion_empty, pow_zero, - Nat.zero_eq, Function.iterate_zero, id, Finset.not_mem_empty] + Function.iterate_zero, id, Finset.not_mem_empty] calc μ (s \ ⋃ (p : α × ℝ) (_ : p ∈ u n.succ), closedBall p.fst p.snd) ≤ N / (N + 1) * μ (s \ ⋃ (p : α × ℝ) (_ : p ∈ u n), closedBall p.fst p.snd) := by diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 76cad3cf36a87..8f21bd8c21ed8 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -767,7 +767,7 @@ def eapproxDiff (f : α → ℝ≥0∞) : ℕ → α →ₛ ℝ≥0 theorem sum_eapproxDiff (f : α → ℝ≥0∞) (n : ℕ) (a : α) : (∑ k ∈ Finset.range (n + 1), (eapproxDiff f k a : ℝ≥0∞)) = eapprox f n a := by induction' n with n IH - · simp only [Nat.zero_eq, Nat.zero_add, Finset.sum_singleton, Finset.range_one] + · simp only [Nat.zero_add, Finset.sum_singleton, Finset.range_one] rfl · erw [Finset.sum_range_succ, IH, eapproxDiff, coe_map, Function.comp_apply, coe_sub, Pi.sub_apply, ENNReal.coe_toNNReal, diff --git a/Mathlib/MeasureTheory/Group/Arithmetic.lean b/Mathlib/MeasureTheory/Group/Arithmetic.lean index 6fd334466abb6..d58f14fb0b7a4 100644 --- a/Mathlib/MeasureTheory/Group/Arithmetic.lean +++ b/Mathlib/MeasureTheory/Group/Arithmetic.lean @@ -166,7 +166,7 @@ instance Monoid.measurablePow (M : Type*) [Monoid M] [MeasurableSpace M] [Measur MeasurablePow M ℕ := ⟨measurable_from_prod_countable fun n => by induction' n with n ih - · simp only [Nat.zero_eq, pow_zero, ← Pi.one_def, measurable_one] + · simp only [pow_zero, ← Pi.one_def, measurable_one] · simp only [pow_succ] exact ih.mul measurable_id⟩ @@ -588,7 +588,7 @@ instance AddMonoid.measurableSMul_nat₂ (M : Type*) [AddMonoid M] [MeasurableSp suffices Measurable fun p : M × ℕ => p.2 • p.1 by apply this.comp measurable_swap refine measurable_from_prod_countable fun n => ?_ induction' n with n ih - · simp only [Nat.zero_eq, zero_smul, ← Pi.zero_def, measurable_zero] + · simp only [zero_smul, ← Pi.zero_def, measurable_zero] · simp only [succ_nsmul] exact ih.add measurable_id⟩ diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index f6a2650672d2e..4c2e837ace6f7 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -597,7 +597,7 @@ theorem measure_univ_of_isMulLeftInvariant [WeaklyLocallyCompactSpace G] [Noncom have M : ∀ n, μ (L n) = (n + 1 : ℕ) * μ K := by intro n induction' n with n IH - · simp only [L, one_mul, Nat.cast_one, iterate_zero, id, Nat.zero_eq, Nat.zero_add] + · simp only [L, one_mul, Nat.cast_one, iterate_zero, id, Nat.zero_add] · calc μ (L (n + 1)) = μ (L n) + μ (g (L n) • K) := by simp_rw [L, iterate_succ'] diff --git a/Mathlib/MeasureTheory/Integral/Pi.lean b/Mathlib/MeasureTheory/Integral/Pi.lean index b01b627974865..ad9202ef91631 100644 --- a/Mathlib/MeasureTheory/Integral/Pi.lean +++ b/Mathlib/MeasureTheory/Integral/Pi.lean @@ -28,7 +28,7 @@ theorem Integrable.fin_nat_prod {n : ℕ} {E : Fin n → Type*} {f : (i : Fin n) → E i → 𝕜} (hf : ∀ i, Integrable (f i)) : Integrable (fun (x : (i : Fin n) → E i) ↦ ∏ i, f i (x i)) := by induction n with - | zero => simp only [Nat.zero_eq, Finset.univ_eq_empty, Finset.prod_empty, volume_pi, + | zero => simp only [Finset.univ_eq_empty, Finset.prod_empty, volume_pi, integrable_const_iff, one_ne_zero, pi_empty_univ, ENNReal.one_lt_top, or_true] | succ n n_ih => have := ((measurePreserving_piFinSuccAbove (fun i => (volume : Measure (E i))) 0).symm) @@ -67,7 +67,7 @@ theorem integral_fin_nat_prod_eq_prod {n : ℕ} {E : Fin n → Type*} ∫ x : (i : Fin n) → E i, ∏ i, f i (x i) = ∏ i, ∫ x, f i x := by induction n with | zero => - simp only [Nat.zero_eq, volume_pi, Finset.univ_eq_empty, Finset.prod_empty, integral_const, + simp only [volume_pi, Finset.univ_eq_empty, Finset.prod_empty, integral_const, pi_empty_univ, ENNReal.one_toReal, smul_eq_mul, mul_one, pow_zero, one_smul] | succ n n_ih => calc diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 3469ca0689e07..9e2582523ac5f 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -373,7 +373,7 @@ lemma generateFrom_iUnion_memPartition (t : ℕ → Set α) : obtain ⟨n, hun⟩ := hu induction n generalizing u with | zero => - simp only [Nat.zero_eq, memPartition_zero, mem_insert_iff, mem_singleton_iff] at hun + simp only [memPartition_zero, mem_insert_iff, mem_singleton_iff] at hun rw [hun] exact MeasurableSet.univ | succ n ih => diff --git a/Mathlib/ModelTheory/Semantics.lean b/Mathlib/ModelTheory/Semantics.lean index 41d4021996c56..8ab1b65be5c42 100644 --- a/Mathlib/ModelTheory/Semantics.lean +++ b/Mathlib/ModelTheory/Semantics.lean @@ -152,7 +152,7 @@ theorem realize_constantsToVars [L[[α]].Structure M] [(lhomWithConstants L α). · simp · cases n · cases f - · simp only [realize, ih, Nat.zero_eq, constantsOn, mk₂_Functions] + · simp only [realize, ih, constantsOn, mk₂_Functions] -- Porting note: below lemma does not work with simp for some reason rw [withConstants_funMap_sum_inl] · simp only [realize, constantsToVars, Sum.elim_inl, funMap_eq_coe_constants] diff --git a/Mathlib/NumberTheory/Cyclotomic/Rat.lean b/Mathlib/NumberTheory/Cyclotomic/Rat.lean index 5c6aee417a142..230f3be84d42c 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Rat.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Rat.lean @@ -318,12 +318,12 @@ theorem zeta_sub_one_prime_of_two_pow [IsCyclotomicExtension {(2 : ℕ+) ^ (k + · convert Prime.neg Int.prime_two apply RingHom.injective_int (algebraMap ℤ ℚ) rw [← Algebra.norm_localization (Sₘ := K) ℤ (nonZeroDivisors ℤ)] - simp only [Nat.zero_eq, PNat.pow_coe, id.map_eq_id, RingHomCompTriple.comp_eq, RingHom.coe_coe, + simp only [PNat.pow_coe, id.map_eq_id, RingHomCompTriple.comp_eq, RingHom.coe_coe, Subalgebra.coe_val, algebraMap_int_eq, map_neg, map_ofNat] - simpa only [zero_add, pow_one, AddSubgroupClass.coe_sub, OneMemClass.coe_one, Nat.zero_eq, + simpa only [zero_add, pow_one, AddSubgroupClass.coe_sub, OneMemClass.coe_one, pow_zero] using hζ.norm_pow_sub_one_two (cyclotomic.irreducible_rat - (by simp only [Nat.zero_eq, zero_add, pow_one, Nat.ofNat_pos])) + (by simp only [zero_add, pow_one, Nat.ofNat_pos])) convert Int.prime_two apply RingHom.injective_int (algebraMap ℤ ℚ) rw [← Algebra.norm_localization (Sₘ := K) ℤ (nonZeroDivisors ℤ)] diff --git a/Mathlib/NumberTheory/DiophantineApproximation.lean b/Mathlib/NumberTheory/DiophantineApproximation.lean index 0e5d09c4d8fa2..e92bec64ba8ec 100644 --- a/Mathlib/NumberTheory/DiophantineApproximation.lean +++ b/Mathlib/NumberTheory/DiophantineApproximation.lean @@ -341,14 +341,14 @@ theorem convergent_succ (ξ : ℝ) (n : ℕ) : @[simp] theorem convergent_of_zero (n : ℕ) : convergent 0 n = 0 := by induction' n with n ih - · simp only [Nat.zero_eq, convergent_zero, floor_zero, cast_zero] + · simp only [convergent_zero, floor_zero, cast_zero] · simp only [ih, convergent_succ, floor_zero, cast_zero, fract_zero, add_zero, inv_zero] /-- If `ξ` is an integer, all its convergents equal `ξ`. -/ @[simp] theorem convergent_of_int {ξ : ℤ} (n : ℕ) : convergent ξ n = ξ := by cases n - · simp only [Nat.zero_eq, convergent_zero, floor_intCast] + · simp only [convergent_zero, floor_intCast] · simp only [convergent_succ, floor_intCast, fract_intCast, convergent_of_zero, add_zero, inv_zero] @@ -363,7 +363,7 @@ open GenContFract theorem convs_eq_convergent (ξ : ℝ) (n : ℕ) : (GenContFract.of ξ).convs n = ξ.convergent n := by induction' n with n ih generalizing ξ - · simp only [Nat.zero_eq, zeroth_conv_eq_h, of_h_eq_floor, convergent_zero, Rat.cast_intCast] + · simp only [zeroth_conv_eq_h, of_h_eq_floor, convergent_zero, Rat.cast_intCast] · rw [convs_succ, ih (fract ξ)⁻¹, convergent_succ, one_div] norm_cast diff --git a/Mathlib/NumberTheory/LSeries/Deriv.lean b/Mathlib/NumberTheory/LSeries/Deriv.lean index 3960fca45e188..d9a355faf2a9f 100644 --- a/Mathlib/NumberTheory/LSeries/Deriv.lean +++ b/Mathlib/NumberTheory/LSeries/Deriv.lean @@ -122,7 +122,7 @@ is the same as that of `f`. -/ lemma LSeries.absicssaOfAbsConv_logPowMul {f : ℕ → ℂ} {m : ℕ} : abscissaOfAbsConv (logMul^[m] f) = abscissaOfAbsConv f := by induction' m with n ih - · simp only [Nat.zero_eq, Function.iterate_zero, id_eq] + · simp only [Function.iterate_zero, id_eq] · simp only [ih, Function.iterate_succ', Function.comp_def, abscissaOfAbsConv_logMul] /-- If `re s` is greater than the abscissa of absolute convergence of `f`, then @@ -130,7 +130,7 @@ the `m`th derivative of this L-series is `(-1)^m` times the L-series of `log^m * lemma LSeries_iteratedDeriv {f : ℕ → ℂ} (m : ℕ) {s : ℂ} (h : abscissaOfAbsConv f < s.re) : iteratedDeriv m (LSeries f) s = (-1) ^ m * LSeries (logMul^[m] f) s := by induction' m with m ih generalizing s - · simp only [Nat.zero_eq, iteratedDeriv_zero, pow_zero, Function.iterate_zero, id_eq, one_mul] + · simp only [iteratedDeriv_zero, pow_zero, Function.iterate_zero, id_eq, one_mul] · have ih' : {s | abscissaOfAbsConv f < re s}.EqOn (iteratedDeriv m (LSeries f)) ((-1) ^ m * LSeries (logMul^[m] f)) := fun _ hs ↦ ih hs have := derivWithin_congr ih' (ih h) diff --git a/Mathlib/NumberTheory/LSeries/Linearity.lean b/Mathlib/NumberTheory/LSeries/Linearity.lean index 163fb8337bcce..edf33f6916742 100644 --- a/Mathlib/NumberTheory/LSeries/Linearity.lean +++ b/Mathlib/NumberTheory/LSeries/Linearity.lean @@ -48,7 +48,7 @@ lemma LSeries_add {f g : ℕ → ℂ} {s : ℂ} (hf : LSeriesSummable f s) (hg : lemma LSeries.term_neg (f : ℕ → ℂ) (s : ℂ) : term (-f) s = -term f s := by ext ⟨- | n⟩ - · simp only [Nat.zero_eq, term_zero, Pi.neg_apply, neg_zero] + · simp only [term_zero, Pi.neg_apply, neg_zero] · simp only [term_of_ne_zero (Nat.succ_ne_zero _), Pi.neg_apply, Nat.cast_succ, neg_div] lemma LSeries.term_neg_apply (f : ℕ → ℂ) (s : ℂ) (n : ℕ) : term (-f) s n = -term f s n := by @@ -105,7 +105,7 @@ lemma LSeries_sub {f g : ℕ → ℂ} {s : ℂ} (hf : LSeriesSummable f s) (hg : lemma LSeries.term_smul (f : ℕ → ℂ) (c s : ℂ) : term (c • f) s = c • term f s := by ext ⟨- | n⟩ - · simp only [Nat.zero_eq, term_zero, Pi.smul_apply, smul_eq_mul, mul_zero] + · simp only [term_zero, Pi.smul_apply, smul_eq_mul, mul_zero] · simp only [term_of_ne_zero (Nat.succ_ne_zero _), Pi.smul_apply, smul_eq_mul, Nat.cast_succ, mul_div_assoc] diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index 15ef0f0e6656e..84bfdd8eb4529 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -51,7 +51,7 @@ theorem dvd_geom_sum₂_self {x y : R} (h : ↑n ∣ x - y) : theorem sq_dvd_add_pow_sub_sub (p x : R) (n : ℕ) : p ^ 2 ∣ (x + p) ^ n - x ^ (n - 1) * p * n - x ^ n := by cases' n with n n - · simp only [pow_zero, Nat.cast_zero, sub_zero, sub_self, dvd_zero, Nat.zero_eq, mul_zero] + · simp only [pow_zero, Nat.cast_zero, sub_zero, sub_self, dvd_zero, mul_zero] · simp only [Nat.succ_sub_succ_eq_sub, tsub_zero, Nat.cast_succ, add_pow, Finset.sum_range_succ, Nat.choose_self, Nat.succ_sub _, tsub_self, pow_one, Nat.choose_succ_self_right, pow_zero, mul_one, Nat.cast_zero, zero_add, Nat.succ_eq_add_one, add_tsub_cancel_left] @@ -189,7 +189,7 @@ include hp hp1 theorem Int.pow_sub_pow {x y : ℤ} (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) (n : ℕ) : multiplicity (↑p) (x ^ n - y ^ n) = multiplicity (↑p) (x - y) + multiplicity p n := by cases' n with n - · simp only [multiplicity.zero, add_top, pow_zero, sub_self, Nat.zero_eq] + · simp only [multiplicity.zero, add_top, pow_zero, sub_self] have h : (multiplicity _ _).Dom := finite_nat_iff.mpr ⟨hp.ne_one, n.succ_pos⟩ simp only [Nat.succ_eq_add_one] at h rcases eq_coe_iff.mp (PartENat.natCast_get h).symm with ⟨⟨k, hk⟩, hpn⟩ @@ -239,7 +239,7 @@ end CommRing theorem pow_two_pow_sub_pow_two_pow [CommRing R] {x y : R} (n : ℕ) : x ^ 2 ^ n - y ^ 2 ^ n = (∏ i ∈ Finset.range n, (x ^ 2 ^ i + y ^ 2 ^ i)) * (x - y) := by induction' n with d hd - · simp only [pow_zero, pow_one, range_zero, prod_empty, one_mul, Nat.zero_eq] + · simp only [pow_zero, pow_one, range_zero, prod_empty, one_mul] · suffices x ^ 2 ^ d.succ - y ^ 2 ^ d.succ = (x ^ 2 ^ d + y ^ 2 ^ d) * (x ^ 2 ^ d - y ^ 2 ^ d) by rw [this, hd, Finset.prod_range_succ, ← mul_assoc, mul_comm (x ^ 2 ^ d + y ^ 2 ^ d)] rw [Nat.succ_eq_add_one] @@ -290,7 +290,7 @@ theorem Int.two_pow_sub_pow' {x y : ℤ} (n : ℕ) (hxy : 4 ∣ x - y) (hx : ¬2 have hxy_even : Even (x - y) := even_iff_two_dvd.mpr (dvd_trans (by decide) hxy) have hy_odd : Odd y := by simpa using hx_odd.sub_even hxy_even cases' n with n - · simp only [pow_zero, sub_self, multiplicity.zero, Int.ofNat_zero, Nat.zero_eq, add_top] + · simp only [pow_zero, sub_self, multiplicity.zero, Int.ofNat_zero, add_top] have h : (multiplicity 2 n.succ).Dom := multiplicity.finite_nat_iff.mpr ⟨by norm_num, n.succ_pos⟩ simp only [Nat.succ_eq_add_one] at h rcases multiplicity.eq_coe_iff.mp (PartENat.natCast_get h).symm with ⟨⟨k, hk⟩, hpn⟩ diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index 19128510fa244..4ffe309128a95 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -254,7 +254,7 @@ theorem y_mul_pos {a b : Solution₁ d} (hax : 0 < a.x) (hay : 0 < a.y) (hbx : 0 have positive `x`. -/ theorem x_pow_pos {a : Solution₁ d} (hax : 0 < a.x) (n : ℕ) : 0 < (a ^ n).x := by induction' n with n ih - · simp only [Nat.zero_eq, pow_zero, x_one, zero_lt_one] + · simp only [pow_zero, x_one, zero_lt_one] · rw [pow_succ] exact x_mul_pos ih hax diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index 94c2c0add091b..8d85a50b90b7e 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -192,7 +192,7 @@ theorem Ico_succ_left_eq_erase_Ico : Ico a.succ b = erase (Ico a b) a := by theorem mod_injOn_Ico (n a : ℕ) : Set.InjOn (· % a) (Finset.Ico n (n + a)) := by induction' n with n ih - · simp only [zero_add, Nat.zero_eq, Ico_zero_eq_range] + · simp only [zero_add, Ico_zero_eq_range] rintro k hk l hl (hkl : k % a = l % a) simp only [Finset.mem_range, Finset.mem_coe] at hk hl rwa [mod_eq_of_lt hk, mod_eq_of_lt hl] at hkl diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index 2c25896fdd070..f1b12c5c24bcf 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -359,7 +359,7 @@ def insertNth (p : RelSeries r) (i : Fin p.length) (a : α) rw [Fin.insertNth_apply_above] swap · exact hm.trans (lt_add_one _) - simp only [Fin.val_succ, Nat.zero_eq, Fin.pred_succ, eq_rec_constant, Fin.succ_mk] + simp only [Fin.val_succ, Fin.pred_succ, eq_rec_constant, Fin.succ_mk] congr exact Fin.ext <| Eq.symm <| Nat.succ_pred_eq_of_pos (lt_trans (Nat.zero_lt_succ _) hm) · convert connect_next diff --git a/Mathlib/Order/SuccPred/Basic.lean b/Mathlib/Order/SuccPred/Basic.lean index bf1c20483de89..8f35e8e5a9135 100644 --- a/Mathlib/Order/SuccPred/Basic.lean +++ b/Mathlib/Order/SuccPred/Basic.lean @@ -863,7 +863,7 @@ theorem pred_succ [NoMaxOrder α] (a : α) : pred (succ a) = a := theorem pred_succ_iterate_of_not_isMax (i : α) (n : ℕ) (hin : ¬IsMax (succ^[n - 1] i)) : pred^[n] (succ^[n] i) = i := by induction' n with n hn - · simp only [Nat.zero_eq, Function.iterate_zero, id] + · simp only [Function.iterate_zero, id] rw [Nat.succ_sub_succ_eq_sub, Nat.sub_zero] at hin have h_not_max : ¬IsMax (succ^[n - 1] i) := by cases' n with n diff --git a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean index d7540ab87b82e..3e7e6fbefc5d2 100644 --- a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean +++ b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean @@ -156,7 +156,7 @@ instance (priority := 100) LinearOrder.isPredArchimedean_of_isSuccArchimedean [S refine ⟨n, ?_⟩ rw [← hn_eq] induction' n with n - · simp only [Nat.zero_eq, Function.iterate_zero, id] + · simp only [Function.iterate_zero, id] · rw [pred_succ_iterate_of_not_isMax] rw [Nat.succ_sub_succ_eq_sub, tsub_zero] suffices succ^[n] i < succ^[n.succ] i from not_isMax_of_lt this @@ -240,7 +240,7 @@ theorem toZ_iterate_succ_of_not_isMax (n : ℕ) (hn : ¬IsMax (succ^[n] i0)) : theorem toZ_iterate_pred_of_not_isMin (n : ℕ) (hn : ¬IsMin (pred^[n] i0)) : toZ i0 (pred^[n] i0) = -n := by cases' n with n n - · simp only [Nat.zero_eq, Function.iterate_zero, id, toZ_of_eq, Nat.cast_zero, neg_zero]; rfl + · simp only [Function.iterate_zero, id, toZ_of_eq, Nat.cast_zero, neg_zero]; rfl have : pred^[n.succ] i0 < i0 := by refine lt_of_le_of_ne (pred_iterate_le _ _) fun h_pred_iterate_eq ↦ hn ?_ have h_pred_eq_pred : pred^[n.succ] i0 = pred^[0] i0 := by diff --git a/Mathlib/Probability/Martingale/Convergence.lean b/Mathlib/Probability/Martingale/Convergence.lean index 25d10c3496810..0757535ae6937 100644 --- a/Mathlib/Probability/Martingale/Convergence.lean +++ b/Mathlib/Probability/Martingale/Convergence.lean @@ -117,7 +117,7 @@ theorem not_frequently_of_upcrossings_lt_top (hab : a < b) (hω : upcrossings a push_neg intro k induction' k with k ih - · simp only [Nat.zero_eq, zero_le, exists_const] + · simp only [zero_le, exists_const] · obtain ⟨N, hN⟩ := ih obtain ⟨N₁, hN₁, hN₁'⟩ := h₁ N obtain ⟨N₂, hN₂, hN₂'⟩ := h₂ N₁ diff --git a/Mathlib/Probability/Martingale/Upcrossing.lean b/Mathlib/Probability/Martingale/Upcrossing.lean index 95f374d8b4aff..70a525bec0762 100644 --- a/Mathlib/Probability/Martingale/Upcrossing.lean +++ b/Mathlib/Probability/Martingale/Upcrossing.lean @@ -176,7 +176,7 @@ variable {a b : ℝ} {f : ι → Ω → ℝ} {N : ι} {n m : ℕ} {ω : Ω} theorem upperCrossingTime_le : upperCrossingTime a b f N n ω ≤ N := by cases n - · simp only [upperCrossingTime_zero, Pi.bot_apply, bot_le, Nat.zero_eq] + · simp only [upperCrossingTime_zero, Pi.bot_apply, bot_le] · simp only [upperCrossingTime_succ, hitting_le] @[simp] @@ -448,7 +448,7 @@ theorem crossing_eq_crossing_of_lowerCrossingTime_lt {M : ℕ} (hNM : N ≤ M) have h' : upperCrossingTime a b f N n ω < N := lt_of_le_of_lt upperCrossingTime_le_lowerCrossingTime h induction' n with k ih - · simp only [Nat.zero_eq, upperCrossingTime_zero, bot_eq_zero', eq_self_iff_true, + · simp only [upperCrossingTime_zero, bot_eq_zero', eq_self_iff_true, lowerCrossingTime_zero, true_and_iff, eq_comm] refine hitting_eq_hitting_of_exists hNM ?_ rw [lowerCrossingTime, hitting_lt_iff] at h @@ -627,9 +627,9 @@ theorem crossing_pos_eq (hab : a < b) : · refine ⟨rfl, ?_⟩ #adaptation_note /-- nightly-2024-03-16: simp was simp (config := { unfoldPartialApp := true }) only [lowerCrossingTime_zero, hitting, - Set.mem_Icc, Set.mem_Iic, Nat.zero_eq] -/ + Set.mem_Icc, Set.mem_Iic] -/ simp (config := { unfoldPartialApp := true }) only [lowerCrossingTime_zero, hitting_def, - Set.mem_Icc, Set.mem_Iic, Nat.zero_eq] + Set.mem_Icc, Set.mem_Iic] ext ω split_ifs with h₁ h₂ h₂ · simp_rw [hf'] diff --git a/Mathlib/RingTheory/AdicCompletion/Basic.lean b/Mathlib/RingTheory/AdicCompletion/Basic.lean index 97af754373c43..b1a5ab9bf4f6c 100644 --- a/Mathlib/RingTheory/AdicCompletion/Basic.lean +++ b/Mathlib/RingTheory/AdicCompletion/Basic.lean @@ -554,7 +554,7 @@ theorem le_jacobson_bot [IsAdicComplete I R] : I ≤ (⊥ : Ideal R).jacobson := convert Ideal.sub_mem _ this (Ideal.mul_mem_left _ (1 + -(x * y)) hL) using 1 ring cases n - · simp only [Ideal.one_eq_top, pow_zero, Nat.zero_eq, mem_top] + · simp only [Ideal.one_eq_top, pow_zero, mem_top] · rw [← neg_sub _ (1 : R), neg_mul, mul_geom_sum, neg_sub, sub_sub, add_comm, ← sub_sub, sub_self, zero_sub, @neg_mem_iff, mul_pow] exact Ideal.mul_mem_right _ (I ^ _) (Ideal.pow_mem_pow hx _) diff --git a/Mathlib/RingTheory/Binomial.lean b/Mathlib/RingTheory/Binomial.lean index 481d3ed558e54..31d44ff0e4ac7 100644 --- a/Mathlib/RingTheory/Binomial.lean +++ b/Mathlib/RingTheory/Binomial.lean @@ -169,7 +169,7 @@ theorem ascPochhammer_smeval_cast (R : Type*) [Semiring R] {S : Type*} [NonAssoc [Pow S ℕ] [Module R S] [IsScalarTower R S S] [NatPowAssoc S] (x : S) (n : ℕ) : (ascPochhammer R n).smeval x = (ascPochhammer ℕ n).smeval x := by induction n with - | zero => simp only [Nat.zero_eq, ascPochhammer_zero, smeval_one, one_smul] + | zero => simp only [ascPochhammer_zero, smeval_one, one_smul] | succ n hn => simp only [ascPochhammer_succ_right, mul_add, smeval_add, smeval_mul_X, ← Nat.cast_comm] simp only [← C_eq_natCast, smeval_C_mul, hn, Nat.cast_smul_eq_nsmul R n] diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index 3822701b81a5c..10ab27fe2e406 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -238,7 +238,7 @@ instance : Pow (NumDenSameDeg 𝒜 x) ℕ where ⟨n • c.deg, @GradedMonoid.GMonoid.gnpow _ (fun i => ↥(𝒜 i)) _ _ n _ c.num, @GradedMonoid.GMonoid.gnpow _ (fun i => ↥(𝒜 i)) _ _ n _ c.den, by induction' n with n ih - · simpa only [Nat.zero_eq, coe_gnpow, pow_zero] using Submonoid.one_mem _ + · simpa only [coe_gnpow, pow_zero] using Submonoid.one_mem _ · simpa only [pow_succ, coe_gnpow] using x.mul_mem ih c.den_mem⟩ @[simp] diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index 732438273dbdb..33be95d2952fa 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -53,7 +53,7 @@ theorem isPWO_iUnion_support_powers [LinearOrderedCancelAddCommMonoid Γ] [Ring (le_trans (le_of_lt hx) (orderTop_le_of_coeff_ne_zero hg)) refine Set.iUnion_subset fun n => ?_ induction' n with n ih <;> intro g hn - · simp only [Nat.zero_eq, pow_zero, support_one, Set.mem_singleton_iff] at hn + · simp only [pow_zero, support_one, Set.mem_singleton_iff] at hn rw [hn, SetLike.mem_coe] exact AddSubmonoid.zero_mem _ · obtain ⟨i, hi, j, hj, rfl⟩ := support_mul_subset_add_support hn diff --git a/Mathlib/RingTheory/Henselian.lean b/Mathlib/RingTheory/Henselian.lean index 85b5017c6df8f..18e8607dbc93a 100644 --- a/Mathlib/RingTheory/Henselian.lean +++ b/Mathlib/RingTheory/Henselian.lean @@ -201,7 +201,7 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] have hfcI : ∀ n, f.eval (c n) ∈ I ^ (n + 1) := by intro n induction' n with n ih - · simpa only [Nat.zero_eq, Nat.rec_zero, zero_add, pow_one] using h₁ + · simpa only [Nat.rec_zero, zero_add, pow_one] using h₁ rw [← taylor_eval_sub (c n), hc, sub_eq_add_neg, sub_eq_add_neg, add_neg_cancel_comm] rw [eval_eq_sum, sum_over_range' _ _ _ (lt_add_of_pos_right _ zero_lt_two), ← diff --git a/Mathlib/RingTheory/Polynomial/Pochhammer.lean b/Mathlib/RingTheory/Polynomial/Pochhammer.lean index 63e7f8d15a3bd..f4d2df4b5c382 100644 --- a/Mathlib/RingTheory/Polynomial/Pochhammer.lean +++ b/Mathlib/RingTheory/Polynomial/Pochhammer.lean @@ -178,7 +178,7 @@ variable {S : Type*} [StrictOrderedSemiring S] theorem ascPochhammer_pos (n : ℕ) (s : S) (h : 0 < s) : 0 < (ascPochhammer S n).eval s := by induction n with | zero => - simp only [Nat.zero_eq, ascPochhammer_zero, eval_one] + simp only [ascPochhammer_zero, eval_one] exact zero_lt_one | succ n ih => rw [ascPochhammer_succ_right, mul_add, eval_add, ← Nat.cast_comm, eval_natCast_mul, eval_mul_X, diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index b16ab068c4441..33cea1c3cd017 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -415,7 +415,7 @@ theorem isUnit_constantCoeff (φ : R⟦X⟧) (h : IsUnit φ) : IsUnit (constantC theorem eq_shift_mul_X_add_const (φ : R⟦X⟧) : φ = (mk fun p => coeff R (p + 1) φ) * X + C R (constantCoeff R φ) := by ext (_ | n) - · simp only [Nat.zero_eq, coeff_zero_eq_constantCoeff, map_add, map_mul, constantCoeff_X, + · simp only [coeff_zero_eq_constantCoeff, map_add, map_mul, constantCoeff_X, mul_zero, coeff_zero_C, zero_add] · simp only [coeff_succ_mul_X, coeff_mk, LinearMap.map_add, coeff_C, n.succ_ne_zero, sub_zero, if_false, add_zero] @@ -424,7 +424,7 @@ theorem eq_shift_mul_X_add_const (φ : R⟦X⟧) : theorem eq_X_mul_shift_add_const (φ : R⟦X⟧) : φ = (X * mk fun p => coeff R (p + 1) φ) + C R (constantCoeff R φ) := by ext (_ | n) - · simp only [Nat.zero_eq, coeff_zero_eq_constantCoeff, map_add, map_mul, constantCoeff_X, + · simp only [coeff_zero_eq_constantCoeff, map_add, map_mul, constantCoeff_X, zero_mul, coeff_zero_C, zero_add] · simp only [coeff_succ_X_mul, coeff_mk, LinearMap.map_add, coeff_C, n.succ_ne_zero, sub_zero, if_false, add_zero] diff --git a/Mathlib/RingTheory/WittVector/FrobeniusFractionField.lean b/Mathlib/RingTheory/WittVector/FrobeniusFractionField.lean index d45bf7b04ac15..46ee5f96de2b0 100644 --- a/Mathlib/RingTheory/WittVector/FrobeniusFractionField.lean +++ b/Mathlib/RingTheory/WittVector/FrobeniusFractionField.lean @@ -205,7 +205,7 @@ theorem frobenius_frobeniusRotation {a₁ a₂ : 𝕎 k} (ha₁ : a₁.coeff 0 ext n cases' n with n · simp only [WittVector.mul_coeff_zero, WittVector.coeff_frobenius_charP, frobeniusRotation, - frobeniusRotationCoeff, Nat.zero_eq] + frobeniusRotationCoeff] apply solution_spec' _ ha₁ · simp only [nthRemainder_spec, WittVector.coeff_frobenius_charP, frobeniusRotationCoeff, frobeniusRotation] diff --git a/Mathlib/RingTheory/WittVector/Identities.lean b/Mathlib/RingTheory/WittVector/Identities.lean index fd2d54f2430f1..4ca16f88555f7 100644 --- a/Mathlib/RingTheory/WittVector/Identities.lean +++ b/Mathlib/RingTheory/WittVector/Identities.lean @@ -52,7 +52,7 @@ variable (p R) theorem coeff_p_pow [CharP R p] (i : ℕ) : ((p : 𝕎 R) ^ i).coeff i = 1 := by induction' i with i h - · simp only [Nat.zero_eq, one_coeff_zero, Ne, pow_zero] + · simp only [one_coeff_zero, Ne, pow_zero] · rw [pow_succ, ← frobenius_verschiebung, coeff_frobenius_charP, verschiebung_coeff_succ, h, one_pow] diff --git a/Mathlib/RingTheory/WittVector/IsPoly.lean b/Mathlib/RingTheory/WittVector/IsPoly.lean index 4c912841041f4..bd28c2be06298 100644 --- a/Mathlib/RingTheory/WittVector/IsPoly.lean +++ b/Mathlib/RingTheory/WittVector/IsPoly.lean @@ -302,7 +302,7 @@ instance oneIsPoly [Fact p.Prime] : IsPoly p fun _ _ _ => 1 := ⟨⟨onePoly, by intros; funext n; cases n · -- Porting note: was `simp only [...]` but with slightly different `[...]`. - simp only [Nat.zero_eq, lt_self_iff_false, one_coeff_zero, onePoly, ite_true, map_one] + simp only [lt_self_iff_false, one_coeff_zero, onePoly, ite_true, map_one] · -- Porting note: was `simp only [...]` but with slightly different `[...]`. simp only [Nat.succ_pos', one_coeff_eq_of_pos, onePoly, Nat.succ_ne_zero, ite_false, map_zero] diff --git a/Mathlib/RingTheory/WittVector/StructurePolynomial.lean b/Mathlib/RingTheory/WittVector/StructurePolynomial.lean index 0558c0e06035e..ad4f137a05748 100644 --- a/Mathlib/RingTheory/WittVector/StructurePolynomial.lean +++ b/Mathlib/RingTheory/WittVector/StructurePolynomial.lean @@ -229,7 +229,7 @@ theorem C_p_pow_dvd_bind₁_rename_wittPolynomial_sub_sum (Φ : MvPolynomial idx ∑ i ∈ range n, C ((p : ℤ) ^ i) * wittStructureInt p Φ i ^ p ^ (n - i) := by cases' n with n · simp only [isUnit_one, Int.ofNat_zero, Int.ofNat_succ, zero_add, pow_zero, C_1, IsUnit.dvd, - Nat.cast_one, Nat.zero_eq] + Nat.cast_one] -- prepare a useful equation for rewriting have key := bind₁_rename_expand_wittPolynomial Φ n IH apply_fun map (Int.castRingHom (ZMod (p ^ (n + 1)))) at key diff --git a/Mathlib/RingTheory/WittVector/Verschiebung.lean b/Mathlib/RingTheory/WittVector/Verschiebung.lean index 17a60efd46d57..8d6566ba6c58a 100644 --- a/Mathlib/RingTheory/WittVector/Verschiebung.lean +++ b/Mathlib/RingTheory/WittVector/Verschiebung.lean @@ -112,8 +112,7 @@ noncomputable def verschiebung : 𝕎 R →+ 𝕎 R where map_add' := by dsimp ghost_calc _ _ - rintro ⟨⟩ <;> -- Uses the dumb induction principle, hence adding `Nat.zero_eq` to ghost_simps. - ghost_simp + rintro ⟨⟩ <;> ghost_simp /-- `WittVector.verschiebung` is a polynomial function. -/ @[is_poly] diff --git a/Mathlib/SetTheory/Surreal/Dyadic.lean b/Mathlib/SetTheory/Surreal/Dyadic.lean index c3b61c82a8071..f47d76361786e 100644 --- a/Mathlib/SetTheory/Surreal/Dyadic.lean +++ b/Mathlib/SetTheory/Surreal/Dyadic.lean @@ -169,7 +169,7 @@ theorem nsmul_pow_two_powHalf (n : ℕ) : 2 ^ n * powHalf n = 1 := by @[simp] theorem nsmul_pow_two_powHalf' (n k : ℕ) : 2 ^ n * powHalf (n + k) = powHalf k := by induction' k with k hk - · simp only [add_zero, Surreal.nsmul_pow_two_powHalf, Nat.zero_eq, eq_self_iff_true, + · simp only [add_zero, Surreal.nsmul_pow_two_powHalf, eq_self_iff_true, Surreal.powHalf_zero] · rw [← double_powHalf_succ_eq_powHalf (n + k), ← double_powHalf_succ_eq_powHalf k, ← mul_assoc, mul_comm (2 ^ n) 2, mul_assoc] at hk diff --git a/Mathlib/Tactic/Attr/Core.lean b/Mathlib/Tactic/Attr/Core.lean index e08139380e134..3da5505f2c84f 100644 --- a/Mathlib/Tactic/Attr/Core.lean +++ b/Mathlib/Tactic/Attr/Core.lean @@ -21,7 +21,4 @@ attribute [monad_norm] seq_eq_bind_map attribute [mfld_simps] id and_true true_and Function.comp_apply and_self eq_self not_false true_or or_true heq_eq_eq forall_const and_imp --- Porting note: until we change the default induction principle on `Nat`: -attribute [ghost_simps] Nat.zero_eq - attribute [nontriviality] eq_iff_true_of_subsingleton diff --git a/Mathlib/Topology/MetricSpace/Closeds.lean b/Mathlib/Topology/MetricSpace/Closeds.lean index 0c09e5375464a..67b9d4cbf3244 100644 --- a/Mathlib/Topology/MetricSpace/Closeds.lean +++ b/Mathlib/Topology/MetricSpace/Closeds.lean @@ -117,7 +117,7 @@ instance Closeds.completeSpace [CompleteSpace α] : CompleteSpace (Closeds α) : apply hs <;> simp exact ⟨⟨z', z'_mem⟩, le_of_lt hz'⟩ use fun k => Nat.recOn k ⟨x, hx⟩ fun l z => (this l z).choose - simp only [Nat.add_zero, Nat.zero_eq, Nat.rec_zero, Nat.rec_add_one, true_and] + simp only [Nat.add_zero, Nat.rec_zero, Nat.rec_add_one, true_and] exact fun k => (this k _).choose_spec -- it follows from the previous bound that `z` is a Cauchy sequence have : CauchySeq fun k => (z k : α) := cauchySeq_of_edist_le_geometric_two (B n) (B_ne_top n) hz From 925a304ce45efc493f60abf0194ac086d87cc0b8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Aug 2024 01:08:20 +0000 Subject: [PATCH 314/359] feat: Finset builder notation (#11582) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define elaborators (but no delaborators) to elaborate the following notations to a `Finset`: In `Data.Finset.Basic`, * `{x ∈ s | p x}` In `Data.Fintype.Basic`, * `{x | p x}` * `{x : α | p x}` * `{x ∉ s | p x}` * `{x ≠ a | p x}` In `Order.Interval.Finset.Basic`, * `{x ≤ a | p x}` * `{x ≥ a | p x}` * `{x < a | p x}` * `{x > a | p x}` The general heuristic for deciding whether to elaborate a given notation as a `Set` or `Finset` is: * Check whether the expected type is `Finset ?α`. * If it is, elaborate as a `Finset`. * If it isn't, check whether the expected type of `s` in the notation is `Finset ?α`. * If it is, elaborate as a `Finset`. * If it isn't or there is no `s` in the notation, elaborate as a `Set`. There is currently a gotcha that elaboration to `Set` is highly preferred because we can't postpone metavariable and override the set elaborator without breaking many existing uses of set builder notation. See [Zulip](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/.7Bx.20.E2.88.88.20s.20.7C.20p.20x.7D.20notation.20for.20finset). --- .../Enumerative/DoubleCounting.lean | 4 +- Mathlib/Data/Finset/Basic.lean | 61 ++++++++++++++++ Mathlib/Data/Fintype/Basic.lean | 69 +++++++++++++++++-- Mathlib/FieldTheory/ChevalleyWarning.lean | 8 +-- Mathlib/Init/Set.lean | 62 +++++++++++++++-- Mathlib/Order/Interval/Finset/Defs.lean | 45 ++++++++++++ scripts/nolints.json | 1 - scripts/style-exceptions.txt | 2 +- test/finset_builder.lean | 69 +++++++++++++++++++ 9 files changed, 300 insertions(+), 21 deletions(-) create mode 100644 test/finset_builder.lean diff --git a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean index 064efd1a8362a..abb6a33ac4f24 100644 --- a/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean +++ b/Mathlib/Combinatorics/Enumerative/DoubleCounting.lean @@ -53,10 +53,10 @@ theorem bipartiteBelow_swap : t.bipartiteBelow (swap r) a = t.bipartiteAbove r a theorem bipartiteAbove_swap : s.bipartiteAbove (swap r) b = s.bipartiteBelow r b := rfl @[simp, norm_cast] -theorem coe_bipartiteBelow : (s.bipartiteBelow r b : Set α) = { a ∈ s | r a b } := coe_filter _ _ +theorem coe_bipartiteBelow : s.bipartiteBelow r b = ({a ∈ s | r a b} : Set α) := coe_filter _ _ @[simp, norm_cast] -theorem coe_bipartiteAbove : (t.bipartiteAbove r a : Set β) = { b ∈ t | r a b } := coe_filter _ _ +theorem coe_bipartiteAbove : t.bipartiteAbove r a = ({b ∈ t | r a b} : Set β) := coe_filter _ _ variable {s t a a' b b'} diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index c010b93b78035..bf99110752d4a 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -2146,6 +2146,67 @@ as a `Finset α` (when a `DecidablePred (· ∈ t)` instance is available). -/ def filter (s : Finset α) : Finset α := ⟨_, s.2.filter p⟩ +end Finset.Filter + +namespace Mathlib.Meta +open Lean Elab Term Meta Batteries.ExtendedBinder + +/-- Return `true` if `expectedType?` is `some (Finset ?α)`, throws `throwUnsupportedSyntax` if it is +`some (Set ?α)`, and returns `false` otherwise. -/ +def knownToBeFinsetNotSet (expectedType? : Option Expr) : TermElabM Bool := + -- As we want to reason about the expected type, we would like to wait for it to be available. + -- However this means that if we fall back on `elabSetBuilder` we will have postponed. + -- This is undesirable as we want set builder notation to quickly elaborate to a `Set` when no + -- expected type is available. + -- tryPostponeIfNoneOrMVar expectedType? + match expectedType? with + | some expectedType => + match_expr expectedType with + -- If the expected type is known to be `Finset ?α`, return `true`. + | Finset _ => pure true + -- If the expected type is known to be `Set ?α`, give up. + | Set _ => throwUnsupportedSyntax + -- If the expected type is known to not be `Finset ?α` or `Set ?α`, return `false`. + | _ => pure false + -- If the expected type is not known, return `false`. + | none => pure false + +/-- Elaborate set builder notation for `Finset`. + +`{x ∈ s | p x}` is elaborated as `Finset.filter (fun x ↦ p x) s` if either the expected type is +`Finset ?α` or the expected type is not `Set ?α` and `s` has expected type `Finset ?α`. + +See also +* `Init.Set` for the `Set` builder notation elaborator that this elaborator partly overrides. +* `Data.Fintype.Basic` for the `Finset` builder notation elaborator handling syntax of the form + `{x | p x}`, `{x : α | p x}`, `{x ∉ s | p x}`, `{x ≠ a | p x}`. +* `Order.LocallyFinite.Basic` for the `Finset` builder notation elaborator handling syntax of the + form `{x ≤ a | p x}`, `{x ≥ a | p x}`, `{x < a | p x}`, `{x > a | p x}`. + +TODO: Write a delaborator +-/ +@[term_elab setBuilder] +def elabFinsetBuilderSep : TermElab + | `({ $x:ident ∈ $s:term | $p }), expectedType? => do + -- If the expected type is known to be `Set ?α`, give up. If it is not known to be `Set ?α` or + -- `Finset ?α`, check the expected type of `s`. + unless ← knownToBeFinsetNotSet expectedType? do + let ty ← try whnfR (← inferType (← elabTerm s none)) catch _ => throwUnsupportedSyntax + -- If the expected type of `s` is not known to be `Finset ?α`, give up. + match_expr ty with + | Finset _ => pure () + | _ => throwUnsupportedSyntax + -- Finally, we can elaborate the syntax as a finset. + -- TODO: Seems a bit wasteful to have computed the expected type but still use `expectedType?`. + elabTerm (← `(Finset.filter (fun $x:ident ↦ $p) $s)) expectedType? + | _, _ => throwUnsupportedSyntax + +end Mathlib.Meta + +namespace Finset +section Filter +variable (p q : α → Prop) [DecidablePred p] [DecidablePred q] {s t : Finset α} + @[simp] theorem filter_val (s : Finset α) : (filter p s).1 = s.1.filter p := rfl diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index eac49025de7cc..6b6d8e3c5377b 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -142,12 +142,71 @@ theorem codisjoint_left : Codisjoint s t ↔ ∀ ⦃a⦄, a ∉ s → a ∈ t := theorem codisjoint_right : Codisjoint s t ↔ ∀ ⦃a⦄, a ∉ t → a ∈ s := Codisjoint_comm.trans codisjoint_left -section BooleanAlgebra +instance booleanAlgebra [DecidableEq α] : BooleanAlgebra (Finset α) := + GeneralizedBooleanAlgebra.toBooleanAlgebra -variable [DecidableEq α] {a : α} +end Finset -instance booleanAlgebra : BooleanAlgebra (Finset α) := - GeneralizedBooleanAlgebra.toBooleanAlgebra +namespace Mathlib.Meta +open Lean Elab Term Meta Batteries.ExtendedBinder + +/-- Elaborate set builder notation for `Finset`. + +* `{x | p x}` is elaborated as `Finset.filter (fun x ↦ p x) Finset.univ` if the expected type is + `Finset ?α`. +* `{x : α | p x}` is elaborated as `Finset.filter (fun x : α ↦ p x) Finset.univ` if the expected + type is `Finset ?α`. +* `{x ∉ s | p x}` is elaborated as `Finset.filter (fun x ↦ p x) sᶜ` if either the expected type is + `Finset ?α` or the expected type is not `Set ?α` and `s` has expected type `Finset ?α`. +* `{x ≠ a | p x}` is elaborated as `Finset.filter (fun x ↦ p x) {a}ᶜ` if the expected type is + `Finset ?α`. + +See also +* `Init.Set` for the `Set` builder notation elaborator that this elaborator partly overrides. +* `Data.Finset.Basic` for the `Finset` builder notation elaborator partly overriding this one for + syntax of the form `{x ∈ s | p x}`. +* `Data.Fintype.Basic` for the `Finset` builder notation elaborator handling syntax of the form + `{x | p x}`, `{x : α | p x}`, `{x ∉ s | p x}`, `{x ≠ a | p x}`. +* `Order.LocallyFinite.Basic` for the `Finset` builder notation elaborator handling syntax of the + form `{x ≤ a | p x}`, `{x ≥ a | p x}`, `{x < a | p x}`, `{x > a | p x}`. + +TODO: Write a delaborator +-/ +@[term_elab setBuilder] +def elabFinsetBuilderSetOf : TermElab + | `({ $x:ident | $p }), expectedType? => do + -- If the expected type is not known to be `Finset ?α`, give up. + unless ← knownToBeFinsetNotSet expectedType? do throwUnsupportedSyntax + elabTerm (← `(Finset.filter (fun $x:ident ↦ $p) Finset.univ)) expectedType? + | `({ $x:ident : $t | $p }), expectedType? => do + -- If the expected type is not known to be `Finset ?α`, give up. + unless ← knownToBeFinsetNotSet expectedType? do throwUnsupportedSyntax + elabTerm (← `(Finset.filter (fun $x:ident : $t ↦ $p) Finset.univ)) expectedType? + | `({ $x:ident ∉ $s:term | $p }), expectedType? => do + -- If the expected type is known to be `Set ?α`, give up. If it is not known to be `Set ?α` or + -- `Finset ?α`, check the expected type of `s`. + unless ← knownToBeFinsetNotSet expectedType? do + let ty ← try whnfR (← inferType (← elabTerm s none)) catch _ => throwUnsupportedSyntax + -- If the expected type of `s` is not known to be `Finset ?α`, give up. + match_expr ty with + | Finset _ => pure () + | _ => throwUnsupportedSyntax + -- Finally, we can elaborate the syntax as a finset. + -- TODO: Seems a bit wasteful to have computed the expected type but still use `expectedType?`. + elabTerm (← `(Finset.filter (fun $x:ident ↦ $p) $sᶜ)) expectedType? + | `({ $x:ident ≠ $a | $p }), expectedType? => do + -- If the expected type is not known to be `Finset ?α`, give up. + unless ← knownToBeFinsetNotSet expectedType? do throwUnsupportedSyntax + elabTerm (← `(Finset.filter (fun $x:ident ↦ $p) (singleton $a)ᶜ)) expectedType? + | _, _ => throwUnsupportedSyntax + +end Mathlib.Meta + +namespace Finset +variable [Fintype α] {s t : Finset α} + +section BooleanAlgebra +variable [DecidableEq α] {a : α} theorem sdiff_eq_inter_compl (s t : Finset α) : s \ t = s ∩ tᶜ := sdiff_eq @@ -662,7 +721,7 @@ theorem toFinset_eq_univ [Fintype α] [Fintype s] : s.toFinset = Finset.univ ↔ @[simp] theorem toFinset_setOf [Fintype α] (p : α → Prop) [DecidablePred p] [Fintype { x | p x }] : - { x | p x }.toFinset = Finset.univ.filter p := by + Set.toFinset {x | p x} = Finset.univ.filter p := by ext simp diff --git a/Mathlib/FieldTheory/ChevalleyWarning.lean b/Mathlib/FieldTheory/ChevalleyWarning.lean index 7102e3a92e24b..134f4e650952c 100644 --- a/Mathlib/FieldTheory/ChevalleyWarning.lean +++ b/Mathlib/FieldTheory/ChevalleyWarning.lean @@ -105,11 +105,9 @@ theorem char_dvd_card_solutions_of_sum_lt {s : Finset ι} {f : ι → MvPolynomi (h : (∑ i ∈ s, (f i).totalDegree) < Fintype.card σ) : p ∣ Fintype.card { x : σ → K // ∀ i ∈ s, eval x (f i) = 0 } := by have hq : 0 < q - 1 := by rw [← Fintype.card_units, Fintype.card_pos_iff]; exact ⟨1⟩ - let S : Finset (σ → K) := { x ∈ univ | ∀ i ∈ s, eval x (f i) = 0 }.toFinset - have hS : ∀ x : σ → K, x ∈ S ↔ ∀ i : ι, i ∈ s → eval x (f i) = 0 := by - intro x - simp only [S, Set.toFinset_setOf, mem_univ, true_and, mem_filter] - /- The polynomial `F = ∏ i ∈ s, (1 - (f i)^(q - 1))` has the nice property + let S : Finset (σ → K) := {x | ∀ i ∈ s, eval x (f i) = 0} + have hS (x : σ → K) : x ∈ S ↔ ∀ i ∈ s, eval x (f i) = 0 := by simp [S] + /- The polynomial `F = ∏ i in s, (1 - (f i)^(q - 1))` has the nice property that it takes the value `1` on elements of `{x : σ → K // ∀ i ∈ s, (f i).eval x = 0}` while it is `0` outside that locus. Hence the sum of its values is equal to the cardinality of diff --git a/Mathlib/Init/Set.lean b/Mathlib/Init/Set.lean index 71458589bfde7..9f91aa04e1f49 100644 --- a/Mathlib/Init/Set.lean +++ b/Mathlib/Init/Set.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura -/ import Mathlib.Init import Batteries.Util.ExtendedBinder +import Lean.Elab.Term /-! # Note about `Mathlib/Init/` @@ -39,7 +40,7 @@ This file is a port of the core Lean 3 file `lib/lean/library/init/data/set.lean -/ -open Batteries.ExtendedBinder +open Lean Elab Term Meta Batteries.ExtendedBinder universe u variable {α : Type u} @@ -86,14 +87,57 @@ instance : HasSubset (Set α) := instance : EmptyCollection (Set α) := ⟨fun _ ↦ False⟩ -syntax "{" extBinder " | " term "}" : term +end Set + +namespace Mathlib.Meta -macro_rules - | `({ $x:ident | $p }) => `(setOf fun $x:ident ↦ $p) - | `({ $x:ident : $t | $p }) => `(setOf fun $x:ident : $t ↦ $p) - | `({ $x:ident $b:binderPred | $p }) => - `(setOf fun $x:ident ↦ satisfies_binder_pred% $x $b ∧ $p) +/-- Set builder syntax. This can be elaborated to either a `Set` or a `Finset` depending on context. +The elaborators for this syntax are located in: +* `Init.Set` for the `Set` builder notation elaborator for syntax of the form `{x | p x}`, + `{x : α | p x}`, `{binder x | p x}`. +* `Data.Finset.Basic` for the `Finset` builder notation elaborator for syntax of the form + `{x ∈ s | p x}`. +* `Data.Fintype.Basic` for the `Finset` builder notation elaborator for syntax of the form + `{x | p x}`, `{x : α | p x}`, `{x ∉ s | p x}`, `{x ≠ a | p x}`. +* `Order.LocallyFinite.Basic` for the `Finset` builder notation elaborator for syntax of the form + `{x ≤ a | p x}`, `{x ≥ a | p x}`, `{x < a | p x}`, `{x > a | p x}`. +-/ +syntax (name := setBuilder) "{" extBinder " | " term "}" : term + +/-- Elaborate set builder notation for `Set`. + +* `{x | p x}` is elaborated as `Set.setOf fun x ↦ p x` +* `{x : α | p x}` is elaborated as `Set.setOf fun x : α ↦ p x` +* `{binder x | p x}`, where `x` is bound by the `binder` binder, is elaborated as + `{x | binder x ∧ p x}`. The typical example is `{x ∈ s | p x}`, which is elaborated as + `{x | x ∈ s ∧ p x}`. The possible binders are + * `· ∈ s`, `· ∉ s` + * `· ⊆ s`, `· ⊂ s`, `· ⊇ s`, `· ⊃ s` + * `· ≤ a`, `· ≥ a`, `· < a`, `· > a`, `· ≠ a` + + More binders can be declared using the `binder_predicate` command, see `Init.BinderPredicates` for + more info. + +See also +* `Data.Finset.Basic` for the `Finset` builder notation elaborator partly overriding this one for + syntax of the form `{x ∈ s | p x}`. +* `Data.Fintype.Basic` for the `Finset` builder notation elaborator partly overriding this one for + syntax of the form `{x | p x}`, `{x : α | p x}`, `{x ∉ s | p x}`, `{x ≠ a | p x}`. +* `Order.LocallyFinite.Basic` for the `Finset` builder notation elaborator partly overriding this + one for syntax of the form `{x ≤ a | p x}`, `{x ≥ a | p x}`, `{x < a | p x}`, `{x > a | p x}`. +-/ +@[term_elab setBuilder] +def elabSetBuilder : TermElab + | `({ $x:ident | $p }), expectedType? => do + elabTerm (← `(setOf fun $x:ident ↦ $p)) expectedType? + | `({ $x:ident : $t | $p }), expectedType? => do + elabTerm (← `(setOf fun $x:ident : $t ↦ $p)) expectedType? + | `({ $x:ident $b:binderPred | $p }), expectedType? => do + elabTerm (← `(setOf fun $x:ident ↦ satisfies_binder_pred% $x $b ∧ $p)) expectedType? + | _, _ => throwUnsupportedSyntax + +/-- Unexpander for set builder notation. -/ @[app_unexpander setOf] def setOf.unexpander : Lean.PrettyPrinter.Unexpander | `($_ fun $x:ident ↦ $p) => `({ $x:ident | $p }) @@ -149,6 +193,10 @@ def setOfPatternMatchUnexpander : Lean.PrettyPrinter.Unexpander throw () | _ => throw () +end Mathlib.Meta + +namespace Set + /-- The universal set on a type `α` is the set containing all elements of `α`. This is conceptually the "same as" `α` (in set theory, it is actually the same), but type theory diff --git a/Mathlib/Order/Interval/Finset/Defs.lean b/Mathlib/Order/Interval/Finset/Defs.lean index 3a70b721e212a..4184e22a98c1e 100644 --- a/Mathlib/Order/Interval/Finset/Defs.lean +++ b/Mathlib/Order/Interval/Finset/Defs.lean @@ -452,6 +452,51 @@ end Lattice end Finset +namespace Mathlib.Meta +open Lean Elab Term Meta Batteries.ExtendedBinder + +/-- Elaborate set builder notation for `Finset`. + +* `{x ≤ a | p x}` is elaborated as `Finset.filter (fun x ↦ p x) (Finset.Iic a)` if the expected type + is `Finset ?α`. +* `{x ≥ a | p x}` is elaborated as `Finset.filter (fun x ↦ p x) (Finset.Ici a)` if the expected type + is `Finset ?α`. +* `{x < a | p x}` is elaborated as `Finset.filter (fun x ↦ p x) (Finset.Iio a)` if the expected type + is `Finset ?α`. +* `{x > a | p x}` is elaborated as `Finset.filter (fun x ↦ p x) (Finset.Ioi a)` if the expected type + is `Finset ?α`. + +See also +* `Init.Set` for the `Set` builder notation elaborator that this elaborator partly overrides. +* `Data.Finset.Basic` for the `Finset` builder notation elaborator partly overriding this one for + syntax of the form `{x ∈ s | p x}`. +* `Data.Fintype.Basic` for the `Finset` builder notation elaborator handling syntax of the form + `{x | p x}`, `{x : α | p x}`, `{x ∉ s | p x}`, `{x ≠ a | p x}`. + +TODO: Write a delaborator +-/ +@[term_elab setBuilder] +def elabFinsetBuilderIxx : TermElab + | `({ $x:ident ≤ $a | $p }), expectedType? => do + -- If the expected type is not known to be `Finset ?α`, give up. + unless ← knownToBeFinsetNotSet expectedType? do throwUnsupportedSyntax + elabTerm (← `(Finset.filter (fun $x:ident ↦ $p) (Finset.Iic $a))) expectedType? + | `({ $x:ident ≥ $a | $p }), expectedType? => do + -- If the expected type is not known to be `Finset ?α`, give up. + unless ← knownToBeFinsetNotSet expectedType? do throwUnsupportedSyntax + elabTerm (← `(Finset.filter (fun $x:ident ↦ $p) (Finset.Ici $a))) expectedType? + | `({ $x:ident < $a | $p }), expectedType? => do + -- If the expected type is not known to be `Finset ?α`, give up. + unless ← knownToBeFinsetNotSet expectedType? do throwUnsupportedSyntax + elabTerm (← `(Finset.filter (fun $x:ident ↦ $p) (Finset.Iio $a))) expectedType? + | `({ $x:ident > $a | $p }), expectedType? => do + -- If the expected type is not known to be `Finset ?α`, give up. + unless ← knownToBeFinsetNotSet expectedType? do throwUnsupportedSyntax + elabTerm (← `(Finset.filter (fun $x:ident ↦ $p) (Finset.Ioi $a))) expectedType? + | _, _ => throwUnsupportedSyntax + +end Mathlib.Meta + /-! ### Finiteness of `Set` intervals -/ diff --git a/scripts/nolints.json b/scripts/nolints.json index c8efbe16155ea..b2ed89639f44c 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -578,7 +578,6 @@ ["docBlame", "QuaternionAlgebra.Basis.k"], ["docBlame", "Sat.Clause.cons"], ["docBlame", "Sat.Clause.nil"], - ["docBlame", "Set.setOf.unexpander"], ["docBlame", "SimpleGraph.Subgraph.Adj"], ["docBlame", "SimpleGraph.Subgraph.verts"], ["docBlame", "SimpleGraph.Walk.notNilRec"], diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index cf67e34b72362..9edb3fdcd7751 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -13,7 +13,7 @@ Mathlib/Computability/TMToPartrec.lean : line 1 : ERR_NUM_LIN : 2200 file contai Mathlib/Computability/TuringMachine.lean : line 1 : ERR_NUM_LIN : 2700 file contains 2561 lines, try to split it up Mathlib/Data/DFinsupp/Basic.lean : line 1 : ERR_NUM_LIN : 2200 file contains 2078 lines, try to split it up Mathlib/Data/Fin/Basic.lean : line 1 : ERR_NUM_LIN : 1700 file contains 1524 lines, try to split it up -Mathlib/Data/Finset/Basic.lean : line 1 : ERR_NUM_LIN : 3100 file contains 2999 lines, try to split it up +Mathlib/Data/Finset/Basic.lean : line 1 : ERR_NUM_LIN : 3200 file contains 3057 lines, try to split it up Mathlib/Data/Finset/Lattice.lean : line 1 : ERR_NUM_LIN : 2100 file contains 1968 lines, try to split it up Mathlib/Data/Finset/Pointwise/Basic.lean : line 1 : ERR_NUM_LIN : 2100 file contains 1997 lines, try to split it up Mathlib/Data/Finsupp/Basic.lean : line 1 : ERR_NUM_LIN : 1900 file contains 1703 lines, try to split it up diff --git a/test/finset_builder.lean b/test/finset_builder.lean new file mode 100644 index 0000000000000..dadd88cf2edea --- /dev/null +++ b/test/finset_builder.lean @@ -0,0 +1,69 @@ +import Mathlib.Order.Interval.Finset.Basic + +/-! +# Examples of finset builder notation +-/ + +open Finset + +variable {α : Type*} (p : α → Prop) [DecidablePred p] + +/-! ## `Data.Finset.Basic` -/ + +example (s : Finset α) : {x ∈ s | p x} = s.filter p := rfl + +/-! ## `Data.Fintype.Basic` -/ + +section Fintype +variable [Fintype α] + +example : ({x | p x} : Finset α) = univ.filter p := rfl +example : ({x : α | p x} : Finset α) = univ.filter p := rfl + +-- If the type of `s` (or the entire expression) is `Finset ?α`, elaborate as `Finset`; +-- otherwise as `Set` +example (s : Finset α) : {x ∈ s | p x} = s.filter p := rfl +example (s : Finset α) : ({x ∈ s | p x} : Set α) = setOf fun x => x ∈ s ∧ p x := rfl +example (s : Finset α) : {x ∈ (s : Set α) | p x} = setOf fun x => x ∈ s ∧ p x := rfl + +-- elaborate as `Set` if no expected type present +example : {x | p x} = setOf p := rfl +example : {x : α | p x} = setOf p := rfl + +variable [DecidableEq α] + +example (s : Finset α) : {x ∉ s | p x} = sᶜ.filter p := rfl +example (a : α) : ({x ≠ a | p x} : Finset α) = ({a}ᶜ : Finset α).filter p := rfl + +-- elaborate as `Set` if the `s` or the expected type is not `Finset ?α` +example (s : Set α) : {x ∉ s | p x} = setOf fun x => x ∉ s ∧ p x := rfl +example (a : α) : {x ≠ a | p x} = setOf fun x => x ≠ a ∧ p x := rfl + +end Fintype + +/-! ## `Order.Interval.Finset.Basic` -/ + +section LocallyFiniteOrderBot +variable [Preorder α] [LocallyFiniteOrderBot α] + +example (a : α) : ({x ≤ a | p x} : Finset α) = (Iic a).filter p := rfl +example (a : α) : ({x < a | p x} : Finset α) = (Iio a).filter p := rfl + +-- elaborate as `Set` if the expected type is not `Finset ?α` +example (a : α) : {x ≤ a | p x} = setOf fun x => x ≤ a ∧ p x := rfl +example (a : α) : {x < a | p x} = setOf fun x => x < a ∧ p x := rfl + +end LocallyFiniteOrderBot + +section LocallyFiniteOrderTop +variable [Preorder α] [LocallyFiniteOrderTop α] + +example (a : α) : ({x ≥ a | p x} : Finset α) = (Ici a).filter p := rfl +example (a : α) : ({x > a | p x} : Finset α) = (Ioi a).filter p := rfl + +-- elaborate as `Set` if the expected type is not `Finset ?α` +example (a : α) : {x ≥ a | p x} = setOf fun x => x ≥ a ∧ p x := rfl +example (a : α) : {x > a | p x} = setOf fun x => x > a ∧ p x := rfl + +end LocallyFiniteOrderTop + From f2871ace112f5349bf43bfddc09c31e3ce52aee4 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 16 Aug 2024 01:08:21 +0000 Subject: [PATCH 315/359] chore: splitting up metric spaces files (#15790) --- Mathlib.lean | 6 + Mathlib/Analysis/BoxIntegral/Box/Basic.lean | 2 +- Mathlib/Analysis/Convex/Extrema.lean | 2 +- Mathlib/Analysis/Oscillation.lean | 2 +- Mathlib/InformationTheory/Hamming.lean | 1 - Mathlib/Topology/EMetricSpace/Basic.lean | 778 +----------------- Mathlib/Topology/EMetricSpace/Defs.lean | 694 ++++++++++++++++ Mathlib/Topology/EMetricSpace/Diam.lean | 149 ++++ Mathlib/Topology/EMetricSpace/Lipschitz.lean | 2 +- .../Topology/EMetricSpace/Paracompact.lean | 2 +- Mathlib/Topology/Instances/Discrete.lean | 6 - Mathlib/Topology/Instances/ENNReal.lean | 2 + Mathlib/Topology/Instances/Int.lean | 2 +- Mathlib/Topology/Instances/Nat.lean | 1 + Mathlib/Topology/Instances/RatLemmas.lean | 1 + Mathlib/Topology/Instances/Real.lean | 1 + Mathlib/Topology/MetricSpace/Basic.lean | 216 +---- Mathlib/Topology/MetricSpace/Bounded.lean | 1 + Mathlib/Topology/MetricSpace/Cauchy.lean | 1 + Mathlib/Topology/MetricSpace/Defs.lean | 253 ++++++ .../Topology/MetricSpace/Equicontinuity.lean | 2 +- .../Topology/MetricSpace/MetricSeparated.lean | 2 +- Mathlib/Topology/MetricSpace/ProperSpace.lean | 4 +- .../Topology/MetricSpace/Pseudo/Basic.lean | 231 ++++++ .../MetricSpace/Pseudo/Constructions.lean | 154 +--- Mathlib/Topology/MetricSpace/Pseudo/Defs.lean | 194 +---- .../Topology/MetricSpace/Pseudo/Lemmas.lean | 31 +- Mathlib/Topology/MetricSpace/Pseudo/Pi.lean | 160 ++++ Mathlib/Topology/MetricSpace/Pseudo/Real.lean | 45 + Mathlib/Topology/Metrizable/Uniformity.lean | 7 + 30 files changed, 1581 insertions(+), 1371 deletions(-) create mode 100644 Mathlib/Topology/EMetricSpace/Defs.lean create mode 100644 Mathlib/Topology/EMetricSpace/Diam.lean create mode 100644 Mathlib/Topology/MetricSpace/Defs.lean create mode 100644 Mathlib/Topology/MetricSpace/Pseudo/Basic.lean create mode 100644 Mathlib/Topology/MetricSpace/Pseudo/Pi.lean create mode 100644 Mathlib/Topology/MetricSpace/Pseudo/Real.lean diff --git a/Mathlib.lean b/Mathlib.lean index 710dcbc9c38a5..3d99e987052c1 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4436,6 +4436,8 @@ import Mathlib.Topology.DerivedSet import Mathlib.Topology.DiscreteQuotient import Mathlib.Topology.DiscreteSubset import Mathlib.Topology.EMetricSpace.Basic +import Mathlib.Topology.EMetricSpace.Defs +import Mathlib.Topology.EMetricSpace.Diam import Mathlib.Topology.EMetricSpace.Lipschitz import Mathlib.Topology.EMetricSpace.Paracompact import Mathlib.Topology.ExtendFrom @@ -4501,6 +4503,7 @@ import Mathlib.Topology.MetricSpace.Cauchy import Mathlib.Topology.MetricSpace.Closeds import Mathlib.Topology.MetricSpace.Completion import Mathlib.Topology.MetricSpace.Contracting +import Mathlib.Topology.MetricSpace.Defs import Mathlib.Topology.MetricSpace.Dilation import Mathlib.Topology.MetricSpace.DilationEquiv import Mathlib.Topology.MetricSpace.Equicontinuity @@ -4522,9 +4525,12 @@ import Mathlib.Topology.MetricSpace.PiNat import Mathlib.Topology.MetricSpace.Polish import Mathlib.Topology.MetricSpace.ProperSpace import Mathlib.Topology.MetricSpace.ProperSpace.Lemmas +import Mathlib.Topology.MetricSpace.Pseudo.Basic import Mathlib.Topology.MetricSpace.Pseudo.Constructions import Mathlib.Topology.MetricSpace.Pseudo.Defs import Mathlib.Topology.MetricSpace.Pseudo.Lemmas +import Mathlib.Topology.MetricSpace.Pseudo.Pi +import Mathlib.Topology.MetricSpace.Pseudo.Real import Mathlib.Topology.MetricSpace.Sequences import Mathlib.Topology.MetricSpace.ShrinkingLemma import Mathlib.Topology.MetricSpace.ThickenedIndicator diff --git a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean index 4561dd670d380..957bfe293bc0b 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean @@ -8,7 +8,7 @@ import Mathlib.Order.Interval.Set.Monotone import Mathlib.Topology.MetricSpace.Basic import Mathlib.Topology.MetricSpace.Bounded import Mathlib.Topology.Order.MonotoneConvergence -import Mathlib.Topology.MetricSpace.Pseudo.Lemmas +import Mathlib.Topology.MetricSpace.Pseudo.Real /-! # Rectangular boxes in `ℝⁿ` diff --git a/Mathlib/Analysis/Convex/Extrema.lean b/Mathlib/Analysis/Convex/Extrema.lean index 336609efc130a..53797cd0f3710 100644 --- a/Mathlib/Analysis/Convex/Extrema.lean +++ b/Mathlib/Analysis/Convex/Extrema.lean @@ -5,8 +5,8 @@ Authors: Frédéric Dupuis -/ import Mathlib.Analysis.Convex.Function import Mathlib.Topology.Algebra.Affine -import Mathlib.Topology.MetricSpace.Pseudo.Lemmas import Mathlib.Topology.Order.LocalExtr +import Mathlib.Topology.MetricSpace.Pseudo.Lemmas /-! # Minima and maxima of convex functions diff --git a/Mathlib/Analysis/Oscillation.lean b/Mathlib/Analysis/Oscillation.lean index 48e3533c9ef10..16bc8089b08b1 100644 --- a/Mathlib/Analysis/Oscillation.lean +++ b/Mathlib/Analysis/Oscillation.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 James Sundstrom. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: James Sundstrom -/ -import Mathlib.Topology.EMetricSpace.Basic +import Mathlib.Topology.EMetricSpace.Diam import Mathlib.Order.WellFoundedSet /-! diff --git a/Mathlib/InformationTheory/Hamming.lean b/Mathlib/InformationTheory/Hamming.lean index 0cd94971008a5..5f6eb9c5030e9 100644 --- a/Mathlib/InformationTheory/Hamming.lean +++ b/Mathlib/InformationTheory/Hamming.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Wrenna Robson -/ import Mathlib.Analysis.Normed.Group.Basic -import Mathlib.Topology.Instances.Discrete /-! # Hamming spaces diff --git a/Mathlib/Topology/EMetricSpace/Basic.lean b/Mathlib/Topology/EMetricSpace/Basic.lean index 7845b2fb5b6dd..79a2678c491d7 100644 --- a/Mathlib/Topology/EMetricSpace/Basic.lean +++ b/Mathlib/Topology/EMetricSpace/Basic.lean @@ -3,8 +3,8 @@ 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.ENNReal.Real import Mathlib.Order.Interval.Finset.Nat +import Mathlib.Topology.EMetricSpace.Defs import Mathlib.Topology.UniformSpace.Pi import Mathlib.Topology.UniformSpace.UniformConvergence import Mathlib.Topology.UniformSpace.UniformEmbedding @@ -12,121 +12,19 @@ import Mathlib.Topology.UniformSpace.UniformEmbedding /-! # Extended metric spaces -This file is devoted to the definition and study of `EMetricSpace`s, i.e., metric -spaces in which the distance is allowed to take the value ∞. This extended distance is -called `edist`, and takes values in `ℝ≥0∞`. - -Many definitions and theorems expected on emetric spaces are already introduced on uniform spaces -and topological spaces. For example: open and closed sets, compactness, completeness, continuity and -uniform continuity. - -The class `EMetricSpace` therefore extends `UniformSpace` (and `TopologicalSpace`). - -Since a lot of elementary properties don't require `eq_of_edist_eq_zero` we start setting up the -theory of `PseudoEMetricSpace`, where we don't require `edist x y = 0 → x = y` and we specialize -to `EMetricSpace` at the end. +Further results about extended metric spaces. -/ open Set Filter -open scoped Uniformity Topology NNReal ENNReal Pointwise universe u v w variable {α : Type u} {β : Type v} {X : Type*} -/-- Characterizing uniformities associated to a (generalized) distance function `D` -in terms of the elements of the uniformity. -/ -theorem uniformity_dist_of_mem_uniformity [LinearOrder β] {U : Filter (α × α)} (z : β) - (D : α → α → β) (H : ∀ s, s ∈ U ↔ ∃ ε > z, ∀ {a b : α}, D a b < ε → (a, b) ∈ s) : - U = ⨅ ε > z, 𝓟 { p : α × α | D p.1 p.2 < ε } := - HasBasis.eq_biInf ⟨fun s => by simp only [H, subset_def, Prod.forall, mem_setOf]⟩ - -/-- `EDist α` means that `α` is equipped with an extended distance. -/ -@[ext] -class EDist (α : Type*) where - edist : α → α → ℝ≥0∞ - -export EDist (edist) - -/-- Creating a uniform space from an extended distance. -/ -def uniformSpaceOfEDist (edist : α → α → ℝ≥0∞) (edist_self : ∀ x : α, edist x x = 0) - (edist_comm : ∀ x y : α, edist x y = edist y x) - (edist_triangle : ∀ x y z : α, edist x z ≤ edist x y + edist y z) : UniformSpace α := - .ofFun edist edist_self edist_comm edist_triangle fun ε ε0 => - ⟨ε / 2, ENNReal.half_pos ε0.ne', fun _ h₁ _ h₂ => - (ENNReal.add_lt_add h₁ h₂).trans_eq (ENNReal.add_halves _)⟩ - --- the uniform structure is embedded in the emetric space structure --- to avoid instance diamond issues. See Note [forgetful inheritance]. -/-- Extended (pseudo) metric spaces, with an extended distance `edist` possibly taking the -value ∞ - -Each pseudo_emetric space induces a canonical `UniformSpace` and hence a canonical -`TopologicalSpace`. -This is enforced in the type class definition, by extending the `UniformSpace` structure. When -instantiating a `PseudoEMetricSpace` structure, the uniformity fields are not necessary, they -will be filled in by default. There is a default value for the uniformity, that can be substituted -in cases of interest, for instance when instantiating a `PseudoEMetricSpace` structure -on a product. - -Continuity of `edist` is proved in `Topology.Instances.ENNReal` --/ -class PseudoEMetricSpace (α : Type u) extends EDist α : Type u where - edist_self : ∀ x : α, edist x x = 0 - edist_comm : ∀ x y : α, edist x y = edist y x - edist_triangle : ∀ x y z : α, edist x z ≤ edist x y + edist y z - toUniformSpace : UniformSpace α := uniformSpaceOfEDist edist edist_self edist_comm edist_triangle - uniformity_edist : 𝓤 α = ⨅ ε > 0, 𝓟 { p : α × α | edist p.1 p.2 < ε } := by rfl - -attribute [instance] PseudoEMetricSpace.toUniformSpace - -/- Pseudoemetric spaces are less common than metric spaces. Therefore, we work in a dedicated -namespace, while notions associated to metric spaces are mostly in the root namespace. -/ - -/-- Two pseudo emetric space structures with the same edistance function coincide. -/ -@[ext] -protected theorem PseudoEMetricSpace.ext {α : Type*} {m m' : PseudoEMetricSpace α} - (h : m.toEDist = m'.toEDist) : m = m' := by - cases' m with ed _ _ _ U hU - cases' m' with ed' _ _ _ U' hU' - congr 1 - exact UniformSpace.ext (((show ed = ed' from h) ▸ hU).trans hU'.symm) +open scoped Uniformity Topology NNReal ENNReal Pointwise variable [PseudoEMetricSpace α] -export PseudoEMetricSpace (edist_self edist_comm edist_triangle) - -attribute [simp] edist_self - -/-- Triangle inequality for the extended distance -/ -theorem edist_triangle_left (x y z : α) : edist x y ≤ edist z x + edist z y := by - rw [edist_comm z]; apply edist_triangle - -theorem edist_triangle_right (x y z : α) : edist x y ≤ edist x z + edist y z := by - rw [edist_comm y]; apply edist_triangle - -theorem edist_congr_right {x y z : α} (h : edist x y = 0) : edist x z = edist y z := by - apply le_antisymm - · rw [← zero_add (edist y z), ← h] - apply edist_triangle - · rw [edist_comm] at h - rw [← zero_add (edist x z), ← h] - apply edist_triangle - -theorem edist_congr_left {x y z : α} (h : edist x y = 0) : edist z x = edist z y := by - rw [edist_comm z x, edist_comm z y] - apply edist_congr_right h - --- new theorem -theorem edist_congr {w x y z : α} (hl : edist w x = 0) (hr : edist y z = 0) : - edist w y = edist x z := - (edist_congr_right hl).trans (edist_congr_left hr) - -theorem edist_triangle4 (x y z t : α) : edist x t ≤ edist x y + edist y z + edist z t := - calc - edist x t ≤ edist x z + edist z t := edist_triangle x z t - _ ≤ edist x y + edist y z + edist z t := add_le_add_right (edist_triangle x y z) _ - /-- The triangle (polygon) inequality for sequences of points; `Finset.Ico` version. -/ theorem edist_le_Ico_sum_edist (f : ℕ → α) {m n} (h : m ≤ n) : edist (f m) (f n) ≤ ∑ i ∈ Finset.Ico m n, edist (f i) (f (i + 1)) := by @@ -159,116 +57,8 @@ theorem edist_le_range_sum_of_edist_le {f : ℕ → α} (n : ℕ) {d : ℕ → edist (f 0) (f n) ≤ ∑ i ∈ Finset.range n, d i := Nat.Ico_zero_eq_range ▸ edist_le_Ico_sum_of_edist_le (zero_le n) fun _ => hd -/-- Reformulation of the uniform structure in terms of the extended distance -/ -theorem uniformity_pseudoedist : 𝓤 α = ⨅ ε > 0, 𝓟 { p : α × α | edist p.1 p.2 < ε } := - PseudoEMetricSpace.uniformity_edist - -theorem uniformSpace_edist : - ‹PseudoEMetricSpace α›.toUniformSpace = - uniformSpaceOfEDist edist edist_self edist_comm edist_triangle := - UniformSpace.ext uniformity_pseudoedist - -theorem uniformity_basis_edist : - (𝓤 α).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) fun ε => { p : α × α | edist p.1 p.2 < ε } := - (@uniformSpace_edist α _).symm ▸ UniformSpace.hasBasis_ofFun ⟨1, one_pos⟩ _ _ _ _ _ - -/-- Characterization of the elements of the uniformity in terms of the extended distance -/ -theorem mem_uniformity_edist {s : Set (α × α)} : - s ∈ 𝓤 α ↔ ∃ ε > 0, ∀ {a b : α}, edist a b < ε → (a, b) ∈ s := - uniformity_basis_edist.mem_uniformity_iff - -/-- Given `f : β → ℝ≥0∞`, if `f` sends `{i | p i}` to a set of positive numbers -accumulating to zero, then `f i`-neighborhoods of the diagonal form a basis of `𝓤 α`. - -For specific bases see `uniformity_basis_edist`, `uniformity_basis_edist'`, -`uniformity_basis_edist_nnreal`, and `uniformity_basis_edist_inv_nat`. -/ -protected theorem EMetric.mk_uniformity_basis {β : Type*} {p : β → Prop} {f : β → ℝ≥0∞} - (hf₀ : ∀ x, p x → 0 < f x) (hf : ∀ ε, 0 < ε → ∃ x, p x ∧ f x ≤ ε) : - (𝓤 α).HasBasis p fun x => { p : α × α | edist p.1 p.2 < f x } := by - refine ⟨fun s => uniformity_basis_edist.mem_iff.trans ?_⟩ - constructor - · rintro ⟨ε, ε₀, hε⟩ - rcases hf ε ε₀ with ⟨i, hi, H⟩ - exact ⟨i, hi, fun x hx => hε <| lt_of_lt_of_le hx.out H⟩ - · exact fun ⟨i, hi, H⟩ => ⟨f i, hf₀ i hi, H⟩ - -/-- Given `f : β → ℝ≥0∞`, if `f` sends `{i | p i}` to a set of positive numbers -accumulating to zero, then closed `f i`-neighborhoods of the diagonal form a basis of `𝓤 α`. - -For specific bases see `uniformity_basis_edist_le` and `uniformity_basis_edist_le'`. -/ -protected theorem EMetric.mk_uniformity_basis_le {β : Type*} {p : β → Prop} {f : β → ℝ≥0∞} - (hf₀ : ∀ x, p x → 0 < f x) (hf : ∀ ε, 0 < ε → ∃ x, p x ∧ f x ≤ ε) : - (𝓤 α).HasBasis p fun x => { p : α × α | edist p.1 p.2 ≤ f x } := by - refine ⟨fun s => uniformity_basis_edist.mem_iff.trans ?_⟩ - constructor - · rintro ⟨ε, ε₀, hε⟩ - rcases exists_between ε₀ with ⟨ε', hε'⟩ - rcases hf ε' hε'.1 with ⟨i, hi, H⟩ - exact ⟨i, hi, fun x hx => hε <| lt_of_le_of_lt (le_trans hx.out H) hε'.2⟩ - · exact fun ⟨i, hi, H⟩ => ⟨f i, hf₀ i hi, fun x hx => H (le_of_lt hx.out)⟩ - -theorem uniformity_basis_edist_le : - (𝓤 α).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) fun ε => { p : α × α | edist p.1 p.2 ≤ ε } := - EMetric.mk_uniformity_basis_le (fun _ => id) fun ε ε₀ => ⟨ε, ε₀, le_refl ε⟩ - -theorem uniformity_basis_edist' (ε' : ℝ≥0∞) (hε' : 0 < ε') : - (𝓤 α).HasBasis (fun ε : ℝ≥0∞ => ε ∈ Ioo 0 ε') fun ε => { p : α × α | edist p.1 p.2 < ε } := - EMetric.mk_uniformity_basis (fun _ => And.left) fun ε ε₀ => - let ⟨δ, hδ⟩ := exists_between hε' - ⟨min ε δ, ⟨lt_min ε₀ hδ.1, lt_of_le_of_lt (min_le_right _ _) hδ.2⟩, min_le_left _ _⟩ - -theorem uniformity_basis_edist_le' (ε' : ℝ≥0∞) (hε' : 0 < ε') : - (𝓤 α).HasBasis (fun ε : ℝ≥0∞ => ε ∈ Ioo 0 ε') fun ε => { p : α × α | edist p.1 p.2 ≤ ε } := - EMetric.mk_uniformity_basis_le (fun _ => And.left) fun ε ε₀ => - let ⟨δ, hδ⟩ := exists_between hε' - ⟨min ε δ, ⟨lt_min ε₀ hδ.1, lt_of_le_of_lt (min_le_right _ _) hδ.2⟩, min_le_left _ _⟩ - -theorem uniformity_basis_edist_nnreal : - (𝓤 α).HasBasis (fun ε : ℝ≥0 => 0 < ε) fun ε => { p : α × α | edist p.1 p.2 < ε } := - EMetric.mk_uniformity_basis (fun _ => ENNReal.coe_pos.2) fun _ε ε₀ => - let ⟨δ, hδ⟩ := ENNReal.lt_iff_exists_nnreal_btwn.1 ε₀ - ⟨δ, ENNReal.coe_pos.1 hδ.1, le_of_lt hδ.2⟩ - -theorem uniformity_basis_edist_nnreal_le : - (𝓤 α).HasBasis (fun ε : ℝ≥0 => 0 < ε) fun ε => { p : α × α | edist p.1 p.2 ≤ ε } := - EMetric.mk_uniformity_basis_le (fun _ => ENNReal.coe_pos.2) fun _ε ε₀ => - let ⟨δ, hδ⟩ := ENNReal.lt_iff_exists_nnreal_btwn.1 ε₀ - ⟨δ, ENNReal.coe_pos.1 hδ.1, le_of_lt hδ.2⟩ - -theorem uniformity_basis_edist_inv_nat : - (𝓤 α).HasBasis (fun _ => True) fun n : ℕ => { p : α × α | edist p.1 p.2 < (↑n)⁻¹ } := - EMetric.mk_uniformity_basis (fun n _ ↦ ENNReal.inv_pos.2 <| ENNReal.natCast_ne_top n) fun _ε ε₀ ↦ - let ⟨n, hn⟩ := ENNReal.exists_inv_nat_lt (ne_of_gt ε₀) - ⟨n, trivial, le_of_lt hn⟩ - -theorem uniformity_basis_edist_inv_two_pow : - (𝓤 α).HasBasis (fun _ => True) fun n : ℕ => { p : α × α | edist p.1 p.2 < 2⁻¹ ^ n } := - EMetric.mk_uniformity_basis (fun _ _ => ENNReal.pow_pos (ENNReal.inv_pos.2 ENNReal.two_ne_top) _) - fun _ε ε₀ => - let ⟨n, hn⟩ := ENNReal.exists_inv_two_pow_lt (ne_of_gt ε₀) - ⟨n, trivial, le_of_lt hn⟩ - -/-- Fixed size neighborhoods of the diagonal belong to the uniform structure -/ -theorem edist_mem_uniformity {ε : ℝ≥0∞} (ε0 : 0 < ε) : { p : α × α | edist p.1 p.2 < ε } ∈ 𝓤 α := - mem_uniformity_edist.2 ⟨ε, ε0, id⟩ - namespace EMetric -instance (priority := 900) instIsCountablyGeneratedUniformity : IsCountablyGenerated (𝓤 α) := - isCountablyGenerated_of_seq ⟨_, uniformity_basis_edist_inv_nat.eq_iInf⟩ - --- Porting note: changed explicit/implicit -/-- ε-δ characterization of uniform continuity on a set for pseudoemetric spaces -/ -theorem uniformContinuousOn_iff [PseudoEMetricSpace β] {f : α → β} {s : Set α} : - UniformContinuousOn f s ↔ - ∀ ε > 0, ∃ δ > 0, ∀ {a}, a ∈ s → ∀ {b}, b ∈ s → edist a b < δ → edist (f a) (f b) < ε := - uniformity_basis_edist.uniformContinuousOn_iff uniformity_basis_edist - -/-- ε-δ characterization of uniform continuity on pseudoemetric spaces -/ -theorem uniformContinuous_iff [PseudoEMetricSpace β] {f : α → β} : - UniformContinuous f ↔ ∀ ε > 0, ∃ δ > 0, ∀ {a b : α}, edist a b < δ → edist (f a) (f b) < ε := - uniformity_basis_edist.uniformContinuous_iff uniformity_basis_edist - theorem uniformInducing_iff [PseudoEMetricSpace β] {f : α → β} : UniformInducing f ↔ UniformContinuous f ∧ ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, edist (f a) (f b) < ε → edist a b < δ := @@ -348,83 +138,6 @@ end EMetric open EMetric -/-- Auxiliary function to replace the uniformity on a pseudoemetric space with -a uniformity which is equal to the original one, but maybe not defeq. -This is useful if one wants to construct a pseudoemetric space with a -specified uniformity. See Note [forgetful inheritance] explaining why having definitionally -the right uniformity is often important. --/ -def PseudoEMetricSpace.replaceUniformity {α} [U : UniformSpace α] (m : PseudoEMetricSpace α) - (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : PseudoEMetricSpace α where - edist := @edist _ m.toEDist - edist_self := edist_self - edist_comm := edist_comm - edist_triangle := edist_triangle - toUniformSpace := U - uniformity_edist := H.trans (@PseudoEMetricSpace.uniformity_edist α _) - -/-- The extended pseudometric induced by a function taking values in a pseudoemetric space. -/ -def PseudoEMetricSpace.induced {α β} (f : α → β) (m : PseudoEMetricSpace β) : - PseudoEMetricSpace α where - edist x y := edist (f x) (f y) - edist_self _ := edist_self _ - edist_comm _ _ := edist_comm _ _ - edist_triangle _ _ _ := edist_triangle _ _ _ - toUniformSpace := UniformSpace.comap f m.toUniformSpace - uniformity_edist := (uniformity_basis_edist.comap (Prod.map f f)).eq_biInf - -/-- Pseudoemetric space instance on subsets of pseudoemetric spaces -/ -instance {α : Type*} {p : α → Prop} [PseudoEMetricSpace α] : PseudoEMetricSpace (Subtype p) := - PseudoEMetricSpace.induced Subtype.val ‹_› - -/-- The extended pseudodistance on a subset of a pseudoemetric space is the restriction of -the original pseudodistance, by definition -/ -theorem Subtype.edist_eq {p : α → Prop} (x y : Subtype p) : edist x y = edist (x : α) y := rfl - -namespace MulOpposite - -/-- Pseudoemetric space instance on the multiplicative opposite of a pseudoemetric space. -/ -@[to_additive "Pseudoemetric space instance on the additive opposite of a pseudoemetric space."] -instance {α : Type*} [PseudoEMetricSpace α] : PseudoEMetricSpace αᵐᵒᵖ := - PseudoEMetricSpace.induced unop ‹_› - -@[to_additive] -theorem edist_unop (x y : αᵐᵒᵖ) : edist (unop x) (unop y) = edist x y := rfl - -@[to_additive] -theorem edist_op (x y : α) : edist (op x) (op y) = edist x y := rfl - -end MulOpposite - -section ULift - -instance : PseudoEMetricSpace (ULift α) := PseudoEMetricSpace.induced ULift.down ‹_› - -theorem ULift.edist_eq (x y : ULift α) : edist x y = edist x.down y.down := rfl - -@[simp] -theorem ULift.edist_up_up (x y : α) : edist (ULift.up x) (ULift.up y) = edist x y := rfl - -end ULift - -/-- The product of two pseudoemetric spaces, with the max distance, is an extended -pseudometric spaces. We make sure that the uniform structure thus constructed is the one -corresponding to the product of uniform spaces, to avoid diamond problems. -/ -instance Prod.pseudoEMetricSpaceMax [PseudoEMetricSpace β] : PseudoEMetricSpace (α × β) where - edist x y := edist x.1 y.1 ⊔ edist x.2 y.2 - edist_self x := by simp - edist_comm x y := by simp [edist_comm] - edist_triangle x y z := - max_le (le_trans (edist_triangle _ _ _) (add_le_add (le_max_left _ _) (le_max_left _ _))) - (le_trans (edist_triangle _ _ _) (add_le_add (le_max_right _ _) (le_max_right _ _))) - uniformity_edist := uniformity_prod.trans <| by - simp [PseudoEMetricSpace.uniformity_edist, ← iInf_inf_eq, setOf_and] - toUniformSpace := inferInstance - -theorem Prod.edist_eq [PseudoEMetricSpace β] (x y : α × β) : - edist x y = max (edist x.1 y.1) (edist x.2 y.2) := - rfl - section Pi open Finset @@ -478,173 +191,6 @@ namespace EMetric variable {x y z : α} {ε ε₁ ε₂ : ℝ≥0∞} {s t : Set α} -/-- `EMetric.ball x ε` is the set of all points `y` with `edist y x < ε` -/ -def ball (x : α) (ε : ℝ≥0∞) : Set α := - { y | edist y x < ε } - -@[simp] theorem mem_ball : y ∈ ball x ε ↔ edist y x < ε := Iff.rfl - -theorem mem_ball' : y ∈ ball x ε ↔ edist x y < ε := by rw [edist_comm, mem_ball] - -/-- `EMetric.closedBall x ε` is the set of all points `y` with `edist y x ≤ ε` -/ -def closedBall (x : α) (ε : ℝ≥0∞) := - { y | edist y x ≤ ε } - -@[simp] theorem mem_closedBall : y ∈ closedBall x ε ↔ edist y x ≤ ε := Iff.rfl - -theorem mem_closedBall' : y ∈ closedBall x ε ↔ edist x y ≤ ε := by rw [edist_comm, mem_closedBall] - -@[simp] -theorem closedBall_top (x : α) : closedBall x ∞ = univ := - eq_univ_of_forall fun _ => mem_setOf.2 le_top - -theorem ball_subset_closedBall : ball x ε ⊆ closedBall x ε := fun _ h => le_of_lt h.out - -theorem pos_of_mem_ball (hy : y ∈ ball x ε) : 0 < ε := - lt_of_le_of_lt (zero_le _) hy - -theorem mem_ball_self (h : 0 < ε) : x ∈ ball x ε := by - rwa [mem_ball, edist_self] - -theorem mem_closedBall_self : x ∈ closedBall x ε := by - rw [mem_closedBall, edist_self]; apply zero_le - -theorem mem_ball_comm : x ∈ ball y ε ↔ y ∈ ball x ε := by rw [mem_ball', mem_ball] - -theorem mem_closedBall_comm : x ∈ closedBall y ε ↔ y ∈ closedBall x ε := by - rw [mem_closedBall', mem_closedBall] - -@[gcongr] -theorem ball_subset_ball (h : ε₁ ≤ ε₂) : ball x ε₁ ⊆ ball x ε₂ := fun _y (yx : _ < ε₁) => - lt_of_lt_of_le yx h - -@[gcongr] -theorem closedBall_subset_closedBall (h : ε₁ ≤ ε₂) : closedBall x ε₁ ⊆ closedBall x ε₂ := - fun _y (yx : _ ≤ ε₁) => le_trans yx h - -theorem ball_disjoint (h : ε₁ + ε₂ ≤ edist x y) : Disjoint (ball x ε₁) (ball y ε₂) := - Set.disjoint_left.mpr fun z h₁ h₂ => - (edist_triangle_left x y z).not_lt <| (ENNReal.add_lt_add h₁ h₂).trans_le h - -theorem ball_subset (h : edist x y + ε₁ ≤ ε₂) (h' : edist x y ≠ ∞) : ball x ε₁ ⊆ ball y ε₂ := - fun z zx => - calc - edist z y ≤ edist z x + edist x y := edist_triangle _ _ _ - _ = edist x y + edist z x := add_comm _ _ - _ < edist x y + ε₁ := ENNReal.add_lt_add_left h' zx - _ ≤ ε₂ := h - -theorem exists_ball_subset_ball (h : y ∈ ball x ε) : ∃ ε' > 0, ball y ε' ⊆ ball x ε := by - have : 0 < ε - edist y x := by simpa using h - refine ⟨ε - edist y x, this, ball_subset ?_ (ne_top_of_lt h)⟩ - exact (add_tsub_cancel_of_le (mem_ball.mp h).le).le - -theorem ball_eq_empty_iff : ball x ε = ∅ ↔ ε = 0 := - eq_empty_iff_forall_not_mem.trans - ⟨fun h => le_bot_iff.1 (le_of_not_gt fun ε0 => h _ (mem_ball_self ε0)), fun ε0 _ h => - not_lt_of_le (le_of_eq ε0) (pos_of_mem_ball h)⟩ - -theorem ordConnected_setOf_closedBall_subset (x : α) (s : Set α) : - OrdConnected { r | closedBall x r ⊆ s } := - ⟨fun _ _ _ h₁ _ h₂ => (closedBall_subset_closedBall h₂.2).trans h₁⟩ - -theorem ordConnected_setOf_ball_subset (x : α) (s : Set α) : OrdConnected { r | ball x r ⊆ s } := - ⟨fun _ _ _ h₁ _ h₂ => (ball_subset_ball h₂.2).trans h₁⟩ - -/-- Relation “two points are at a finite edistance” is an equivalence relation. -/ -def edistLtTopSetoid : Setoid α where - r x y := edist x y < ⊤ - iseqv := - ⟨fun x => by rw [edist_self]; exact ENNReal.coe_lt_top, - fun h => by rwa [edist_comm], fun hxy hyz => - lt_of_le_of_lt (edist_triangle _ _ _) (ENNReal.add_lt_top.2 ⟨hxy, hyz⟩)⟩ - -@[simp] -theorem ball_zero : ball x 0 = ∅ := by rw [EMetric.ball_eq_empty_iff] - -theorem nhds_basis_eball : (𝓝 x).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) (ball x) := - nhds_basis_uniformity uniformity_basis_edist - -theorem nhdsWithin_basis_eball : (𝓝[s] x).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) fun ε => ball x ε ∩ s := - nhdsWithin_hasBasis nhds_basis_eball s - -theorem nhds_basis_closed_eball : (𝓝 x).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) (closedBall x) := - nhds_basis_uniformity uniformity_basis_edist_le - -theorem nhdsWithin_basis_closed_eball : - (𝓝[s] x).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) fun ε => closedBall x ε ∩ s := - nhdsWithin_hasBasis nhds_basis_closed_eball s - -theorem nhds_eq : 𝓝 x = ⨅ ε > 0, 𝓟 (ball x ε) := - nhds_basis_eball.eq_biInf - -theorem mem_nhds_iff : s ∈ 𝓝 x ↔ ∃ ε > 0, ball x ε ⊆ s := - nhds_basis_eball.mem_iff - -theorem mem_nhdsWithin_iff : s ∈ 𝓝[t] x ↔ ∃ ε > 0, ball x ε ∩ t ⊆ s := - nhdsWithin_basis_eball.mem_iff - -section - -variable [PseudoEMetricSpace β] {f : α → β} - -theorem tendsto_nhdsWithin_nhdsWithin {t : Set β} {a b} : - Tendsto f (𝓝[s] a) (𝓝[t] b) ↔ - ∀ ε > 0, ∃ δ > 0, ∀ ⦃x⦄, x ∈ s → edist x a < δ → f x ∈ t ∧ edist (f x) b < ε := - (nhdsWithin_basis_eball.tendsto_iff nhdsWithin_basis_eball).trans <| - forall₂_congr fun ε _ => exists_congr fun δ => and_congr_right fun _ => - forall_congr' fun x => by simp; tauto - -theorem tendsto_nhdsWithin_nhds {a b} : - Tendsto f (𝓝[s] a) (𝓝 b) ↔ - ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → edist x a < δ → edist (f x) b < ε := by - rw [← nhdsWithin_univ b, tendsto_nhdsWithin_nhdsWithin] - simp only [mem_univ, true_and_iff] - -theorem tendsto_nhds_nhds {a b} : - Tendsto f (𝓝 a) (𝓝 b) ↔ ∀ ε > 0, ∃ δ > 0, ∀ ⦃x⦄, edist x a < δ → edist (f x) b < ε := - nhds_basis_eball.tendsto_iff nhds_basis_eball - -end - -theorem isOpen_iff : IsOpen s ↔ ∀ x ∈ s, ∃ ε > 0, ball x ε ⊆ s := by - simp [isOpen_iff_nhds, mem_nhds_iff] - -theorem isOpen_ball : IsOpen (ball x ε) := - isOpen_iff.2 fun _ => exists_ball_subset_ball - -theorem isClosed_ball_top : IsClosed (ball x ⊤) := - isOpen_compl_iff.1 <| isOpen_iff.2 fun _y hy => - ⟨⊤, ENNReal.coe_lt_top, fun _z hzy hzx => - hy (edistLtTopSetoid.trans (edistLtTopSetoid.symm hzy) hzx)⟩ - -theorem ball_mem_nhds (x : α) {ε : ℝ≥0∞} (ε0 : 0 < ε) : ball x ε ∈ 𝓝 x := - isOpen_ball.mem_nhds (mem_ball_self ε0) - -theorem closedBall_mem_nhds (x : α) {ε : ℝ≥0∞} (ε0 : 0 < ε) : closedBall x ε ∈ 𝓝 x := - mem_of_superset (ball_mem_nhds x ε0) ball_subset_closedBall - -theorem ball_prod_same [PseudoEMetricSpace β] (x : α) (y : β) (r : ℝ≥0∞) : - ball x r ×ˢ ball y r = ball (x, y) r := - ext fun z => by simp [Prod.edist_eq] - -theorem closedBall_prod_same [PseudoEMetricSpace β] (x : α) (y : β) (r : ℝ≥0∞) : - closedBall x r ×ˢ closedBall y r = closedBall (x, y) r := - ext fun z => by simp [Prod.edist_eq] - -/-- ε-characterization of the closure in pseudoemetric spaces -/ -theorem mem_closure_iff : x ∈ closure s ↔ ∀ ε > 0, ∃ y ∈ s, edist x y < ε := - (mem_closure_iff_nhds_basis nhds_basis_eball).trans <| by simp only [mem_ball, edist_comm x] - -theorem tendsto_nhds {f : Filter β} {u : β → α} {a : α} : - Tendsto u f (𝓝 a) ↔ ∀ ε > 0, ∀ᶠ x in f, edist (u x) a < ε := - nhds_basis_eball.tendsto_right_iff - -theorem tendsto_atTop [Nonempty β] [SemilatticeSup β] {u : β → α} {a : α} : - Tendsto u atTop (𝓝 a) ↔ ∀ ε > 0, ∃ N, ∀ n ≥ N, edist (u n) a < ε := - (atTop_basis.tendsto_iff nhds_basis_eball).trans <| by - simp only [exists_prop, true_and_iff, mem_Ici, mem_ball] - theorem inseparable_iff : Inseparable x y ↔ edist x y = 0 := by simp [inseparable_iff_mem_closure, mem_closure_iff, edist_comm, forall_lt_iff_le'] @@ -682,61 +228,6 @@ theorem totallyBounded_iff' {s : Set α} : section Compact --- Porting note (#11215): TODO: generalize to a uniform space with metrizable uniformity -/-- For a set `s` in a pseudo emetric space, if for every `ε > 0` there exists a countable -set that is `ε`-dense in `s`, then there exists a countable subset `t ⊆ s` that is dense in `s`. -/ -theorem subset_countable_closure_of_almost_dense_set (s : Set α) - (hs : ∀ ε > 0, ∃ t : Set α, t.Countable ∧ s ⊆ ⋃ x ∈ t, closedBall x ε) : - ∃ t, t ⊆ s ∧ t.Countable ∧ s ⊆ closure t := by - rcases s.eq_empty_or_nonempty with (rfl | ⟨x₀, hx₀⟩) - · exact ⟨∅, empty_subset _, countable_empty, empty_subset _⟩ - choose! T hTc hsT using fun n : ℕ => hs n⁻¹ (by simp) - have : ∀ r x, ∃ y ∈ s, closedBall x r ∩ s ⊆ closedBall y (r * 2) := fun r x => by - rcases (closedBall x r ∩ s).eq_empty_or_nonempty with (he | ⟨y, hxy, hys⟩) - · refine ⟨x₀, hx₀, ?_⟩ - rw [he] - exact empty_subset _ - · refine ⟨y, hys, fun z hz => ?_⟩ - calc - edist z y ≤ edist z x + edist y x := edist_triangle_right _ _ _ - _ ≤ r + r := add_le_add hz.1 hxy - _ = r * 2 := (mul_two r).symm - choose f hfs hf using this - refine - ⟨⋃ n : ℕ, f n⁻¹ '' T n, iUnion_subset fun n => image_subset_iff.2 fun z _ => hfs _ _, - countable_iUnion fun n => (hTc n).image _, ?_⟩ - refine fun x hx => mem_closure_iff.2 fun ε ε0 => ?_ - rcases ENNReal.exists_inv_nat_lt (ENNReal.half_pos ε0.lt.ne').ne' with ⟨n, hn⟩ - rcases mem_iUnion₂.1 (hsT n hx) with ⟨y, hyn, hyx⟩ - refine ⟨f n⁻¹ y, mem_iUnion.2 ⟨n, mem_image_of_mem _ hyn⟩, ?_⟩ - calc - edist x (f n⁻¹ y) ≤ (n : ℝ≥0∞)⁻¹ * 2 := hf _ _ ⟨hyx, hx⟩ - _ < ε := ENNReal.mul_lt_of_lt_div hn - -open TopologicalSpace in -/-- If a set `s` is separable in a (pseudo extended) metric space, then it admits a countable dense -subset. This is not obvious, as the countable set whose closure covers `s` given by the definition -of separability does not need in general to be contained in `s`. -/ -theorem _root_.TopologicalSpace.IsSeparable.exists_countable_dense_subset - {s : Set α} (hs : IsSeparable s) : ∃ t, t ⊆ s ∧ t.Countable ∧ s ⊆ closure t := by - have : ∀ ε > 0, ∃ t : Set α, t.Countable ∧ s ⊆ ⋃ x ∈ t, closedBall x ε := fun ε ε0 => by - rcases hs with ⟨t, htc, hst⟩ - refine ⟨t, htc, hst.trans fun x hx => ?_⟩ - rcases mem_closure_iff.1 hx ε ε0 with ⟨y, hyt, hxy⟩ - exact mem_iUnion₂.2 ⟨y, hyt, mem_closedBall.2 hxy.le⟩ - exact subset_countable_closure_of_almost_dense_set _ this - -open TopologicalSpace in -/-- If a set `s` is separable, then the corresponding subtype is separable in a (pseudo extended) -metric space. This is not obvious, as the countable set whose closure covers `s` does not need in -general to be contained in `s`. -/ -theorem _root_.TopologicalSpace.IsSeparable.separableSpace {s : Set α} (hs : IsSeparable s) : - SeparableSpace s := by - rcases hs.exists_countable_dense_subset with ⟨t, hts, htc, hst⟩ - lift t to Set s using hts - refine ⟨⟨t, countable_of_injective_of_countable_image Subtype.coe_injective.injOn htc, ?_⟩⟩ - rwa [inducing_subtype_val.dense_iff, Subtype.forall] - -- Porting note (#11215): TODO: generalize to metrizable spaces /-- A compact set in a pseudo emetric space is separable, i.e., it is a subset of the closure of a countable set. -/ @@ -777,153 +268,10 @@ theorem secondCountable_of_almost_dense_set end SecondCountable -section Diam - -/-- The diameter of a set in a pseudoemetric space, named `EMetric.diam` -/ -noncomputable def diam (s : Set α) := - ⨆ (x ∈ s) (y ∈ s), edist x y - -theorem diam_eq_sSup (s : Set α) : diam s = sSup (image2 edist s s) := sSup_image2.symm - -theorem diam_le_iff {d : ℝ≥0∞} : diam s ≤ d ↔ ∀ x ∈ s, ∀ y ∈ s, edist x y ≤ d := by - simp only [diam, iSup_le_iff] - -theorem diam_image_le_iff {d : ℝ≥0∞} {f : β → α} {s : Set β} : - diam (f '' s) ≤ d ↔ ∀ x ∈ s, ∀ y ∈ s, edist (f x) (f y) ≤ d := by - simp only [diam_le_iff, forall_mem_image] - -theorem edist_le_of_diam_le {d} (hx : x ∈ s) (hy : y ∈ s) (hd : diam s ≤ d) : edist x y ≤ d := - diam_le_iff.1 hd x hx y hy - -/-- If two points belong to some set, their edistance is bounded by the diameter of the set -/ -theorem edist_le_diam_of_mem (hx : x ∈ s) (hy : y ∈ s) : edist x y ≤ diam s := - edist_le_of_diam_le hx hy le_rfl - -/-- If the distance between any two points in a set is bounded by some constant, this constant -bounds the diameter. -/ -theorem diam_le {d : ℝ≥0∞} (h : ∀ x ∈ s, ∀ y ∈ s, edist x y ≤ d) : diam s ≤ d := - diam_le_iff.2 h - -/-- The diameter of a subsingleton vanishes. -/ -theorem diam_subsingleton (hs : s.Subsingleton) : diam s = 0 := - nonpos_iff_eq_zero.1 <| diam_le fun _x hx y hy => (hs hx hy).symm ▸ edist_self y ▸ le_rfl - -/-- The diameter of the empty set vanishes -/ -@[simp] -theorem diam_empty : diam (∅ : Set α) = 0 := - diam_subsingleton subsingleton_empty - -/-- The diameter of a singleton vanishes -/ -@[simp] -theorem diam_singleton : diam ({x} : Set α) = 0 := - diam_subsingleton subsingleton_singleton - -@[to_additive (attr := simp)] -theorem diam_one [One α] : diam (1 : Set α) = 0 := - diam_singleton - -theorem diam_iUnion_mem_option {ι : Type*} (o : Option ι) (s : ι → Set α) : - diam (⋃ i ∈ o, s i) = ⨆ i ∈ o, diam (s i) := by cases o <;> simp - -theorem diam_insert : diam (insert x s) = max (⨆ y ∈ s, edist x y) (diam s) := - eq_of_forall_ge_iff fun d => by - simp only [diam_le_iff, forall_mem_insert, edist_self, edist_comm x, max_le_iff, iSup_le_iff, - zero_le, true_and_iff, forall_and, and_self_iff, ← and_assoc] - -theorem diam_pair : diam ({x, y} : Set α) = edist x y := by - simp only [iSup_singleton, diam_insert, diam_singleton, ENNReal.max_zero_right] - -theorem diam_triple : diam ({x, y, z} : Set α) = max (max (edist x y) (edist x z)) (edist y z) := by - simp only [diam_insert, iSup_insert, iSup_singleton, diam_singleton, ENNReal.max_zero_right, - ENNReal.sup_eq_max] - -/-- The diameter is monotonous with respect to inclusion -/ -@[gcongr] -theorem diam_mono {s t : Set α} (h : s ⊆ t) : diam s ≤ diam t := - diam_le fun _x hx _y hy => edist_le_diam_of_mem (h hx) (h hy) - -/-- The diameter of a union is controlled by the diameter of the sets, and the edistance -between two points in the sets. -/ -theorem diam_union {t : Set α} (xs : x ∈ s) (yt : y ∈ t) : - diam (s ∪ t) ≤ diam s + edist x y + diam t := by - have A : ∀ a ∈ s, ∀ b ∈ t, edist a b ≤ diam s + edist x y + diam t := fun a ha b hb => - calc - edist a b ≤ edist a x + edist x y + edist y b := edist_triangle4 _ _ _ _ - _ ≤ diam s + edist x y + diam t := - add_le_add (add_le_add (edist_le_diam_of_mem ha xs) le_rfl) (edist_le_diam_of_mem yt hb) - refine diam_le fun a ha b hb => ?_ - cases' (mem_union _ _ _).1 ha with h'a h'a <;> cases' (mem_union _ _ _).1 hb with h'b h'b - · calc - edist a b ≤ diam s := edist_le_diam_of_mem h'a h'b - _ ≤ diam s + (edist x y + diam t) := le_self_add - _ = diam s + edist x y + diam t := (add_assoc _ _ _).symm - · exact A a h'a b h'b - · have Z := A b h'b a h'a - rwa [edist_comm] at Z - · calc - edist a b ≤ diam t := edist_le_diam_of_mem h'a h'b - _ ≤ diam s + edist x y + diam t := le_add_self - -theorem diam_union' {t : Set α} (h : (s ∩ t).Nonempty) : diam (s ∪ t) ≤ diam s + diam t := by - let ⟨x, ⟨xs, xt⟩⟩ := h - simpa using diam_union xs xt - -theorem diam_closedBall {r : ℝ≥0∞} : diam (closedBall x r) ≤ 2 * r := - diam_le fun a ha b hb => - calc - edist a b ≤ edist a x + edist b x := edist_triangle_right _ _ _ - _ ≤ r + r := add_le_add ha hb - _ = 2 * r := (two_mul r).symm - -theorem diam_ball {r : ℝ≥0∞} : diam (ball x r) ≤ 2 * r := - le_trans (diam_mono ball_subset_closedBall) diam_closedBall - -theorem diam_pi_le_of_le {π : β → Type*} [Fintype β] [∀ b, PseudoEMetricSpace (π b)] - {s : ∀ b : β, Set (π b)} {c : ℝ≥0∞} (h : ∀ b, diam (s b) ≤ c) : diam (Set.pi univ s) ≤ c := by - refine diam_le fun x hx y hy => edist_pi_le_iff.mpr ?_ - rw [mem_univ_pi] at hx hy - exact fun b => diam_le_iff.1 (h b) (x b) (hx b) (y b) (hy b) - -end Diam - end EMetric ---namespace -/-- We now define `EMetricSpace`, extending `PseudoEMetricSpace`. -/ -class EMetricSpace (α : Type u) extends PseudoEMetricSpace α : Type u where - eq_of_edist_eq_zero : ∀ {x y : α}, edist x y = 0 → x = y - -@[ext] -protected theorem EMetricSpace.ext - {α : Type*} {m m' : EMetricSpace α} (h : m.toEDist = m'.toEDist) : m = m' := by - cases m - cases m' - congr - ext1 - assumption - variable {γ : Type w} [EMetricSpace γ] -export EMetricSpace (eq_of_edist_eq_zero) - -/-- Characterize the equality of points by the vanishing of their extended distance -/ -@[simp] -theorem edist_eq_zero {x y : γ} : edist x y = 0 ↔ x = y := - ⟨eq_of_edist_eq_zero, fun h => h ▸ edist_self _⟩ - -@[simp] -theorem zero_eq_edist {x y : γ} : 0 = edist x y ↔ x = y := eq_comm.trans edist_eq_zero - -theorem edist_le_zero {x y : γ} : edist x y ≤ 0 ↔ x = y := - nonpos_iff_eq_zero.trans edist_eq_zero - -@[simp] -theorem edist_pos {x y : γ} : 0 < edist x y ↔ x ≠ y := by simp [← not_le] - -/-- Two points coincide if their distance is `< ε` for all positive ε -/ -theorem eq_of_forall_edist_le {x y : γ} (h : ∀ ε > 0, edist x y ≤ ε) : x = y := - eq_of_edist_eq_zero (eq_of_le_of_forall_le_of_dense bot_le h) - -- see Note [lower instance priority] /-- An emetric space is separated -/ instance (priority := 100) EMetricSpace.instT0Space : T0Space γ where @@ -945,50 +293,12 @@ abbrev EMetricSpace.ofT0PseudoEMetricSpace (α : Type*) [PseudoEMetricSpace α] { ‹PseudoEMetricSpace α› with eq_of_edist_eq_zero := fun h => (EMetric.inseparable_iff.2 h).eq } -/-- Auxiliary function to replace the uniformity on an emetric space with -a uniformity which is equal to the original one, but maybe not defeq. -This is useful if one wants to construct an emetric space with a -specified uniformity. See Note [forgetful inheritance] explaining why having definitionally -the right uniformity is often important. --/ -def EMetricSpace.replaceUniformity {γ} [U : UniformSpace γ] (m : EMetricSpace γ) - (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : EMetricSpace γ where - edist := @edist _ m.toEDist - edist_self := edist_self - eq_of_edist_eq_zero := @eq_of_edist_eq_zero _ _ - edist_comm := edist_comm - edist_triangle := edist_triangle - toUniformSpace := U - uniformity_edist := H.trans (@PseudoEMetricSpace.uniformity_edist γ _) - -/-- The extended metric induced by an injective function taking values in an emetric space. -/ -def EMetricSpace.induced {γ β} (f : γ → β) (hf : Function.Injective f) (m : EMetricSpace β) : - EMetricSpace γ := - { PseudoEMetricSpace.induced f m.toPseudoEMetricSpace with - eq_of_edist_eq_zero := fun h => hf (edist_eq_zero.1 h) } - -/-- EMetric space instance on subsets of emetric spaces -/ -instance {α : Type*} {p : α → Prop} [EMetricSpace α] : EMetricSpace (Subtype p) := - EMetricSpace.induced Subtype.val Subtype.coe_injective ‹_› - -/-- EMetric space instance on the multiplicative opposite of an emetric space. -/ -@[to_additive "EMetric space instance on the additive opposite of an emetric space."] -instance {α : Type*} [EMetricSpace α] : EMetricSpace αᵐᵒᵖ := - EMetricSpace.induced MulOpposite.unop MulOpposite.unop_injective ‹_› - -instance {α : Type*} [EMetricSpace α] : EMetricSpace (ULift α) := - EMetricSpace.induced ULift.down ULift.down_injective ‹_› - /-- The product of two emetric spaces, with the max distance, is an extended metric spaces. We make sure that the uniform structure thus constructed is the one corresponding to the product of uniform spaces, to avoid diamond problems. -/ instance Prod.emetricSpaceMax [EMetricSpace β] : EMetricSpace (γ × β) := .ofT0PseudoEMetricSpace _ -/-- Reformulation of the uniform structure in terms of the extended distance -/ -theorem uniformity_edist : 𝓤 γ = ⨅ ε > 0, 𝓟 { p : γ × γ | edist p.1 p.2 < ε } := - PseudoEMetricSpace.uniformity_edist - section Pi open Finset @@ -1013,21 +323,6 @@ theorem countable_closure_of_compact {s : Set γ} (hs : IsCompact s) : rcases subset_countable_closure_of_compact hs with ⟨t, hts, htc, hsub⟩ exact ⟨t, hts, htc, hsub.antisymm (closure_minimal hts hs.isClosed)⟩ -section Diam - -variable {s : Set γ} - -theorem diam_eq_zero_iff : diam s = 0 ↔ s.Subsingleton := - ⟨fun h _x hx _y hy => edist_le_zero.1 <| h ▸ edist_le_diam_of_mem hx hy, diam_subsingleton⟩ - -theorem diam_pos_iff : 0 < diam s ↔ s.Nontrivial := by - simp only [pos_iff_ne_zero, Ne, diam_eq_zero_iff, Set.not_subsingleton_iff] - -theorem diam_pos_iff' : 0 < diam s ↔ ∃ x ∈ s, ∃ y ∈ s, x ≠ y := by - simp only [diam_pos_iff, Set.Nontrivial, exists_prop] - -end Diam - end EMetric /-! @@ -1051,70 +346,3 @@ instance [PseudoEMetricSpace X] : EMetricSpace (SeparationQuotient X) := toUniformSpace := inferInstance, uniformity_edist := comap_injective (surjective_mk.prodMap surjective_mk) <| by simp [comap_mk_uniformity, PseudoEMetricSpace.uniformity_edist] } _ - -/-! -### `Additive`, `Multiplicative` - -The distance on those type synonyms is inherited without change. --/ - - -open Additive Multiplicative - -section - -variable [EDist X] - -instance : EDist (Additive X) := ‹EDist X› -instance : EDist (Multiplicative X) := ‹EDist X› - -@[simp] -theorem edist_ofMul (a b : X) : edist (ofMul a) (ofMul b) = edist a b := - rfl - -@[simp] -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 := - rfl - -@[simp] -theorem edist_toAdd (a b : Multiplicative X) : edist (toAdd a) (toAdd b) = edist a b := - rfl - -end - -instance [PseudoEMetricSpace X] : PseudoEMetricSpace (Additive X) := ‹PseudoEMetricSpace X› -instance [PseudoEMetricSpace X] : PseudoEMetricSpace (Multiplicative X) := ‹PseudoEMetricSpace X› -instance [EMetricSpace X] : EMetricSpace (Additive X) := ‹EMetricSpace X› -instance [EMetricSpace X] : EMetricSpace (Multiplicative X) := ‹EMetricSpace X› - -/-! -### Order dual - -The distance on this type synonym is inherited without change. --/ - - -open OrderDual - -section - -variable [EDist X] - -instance : EDist Xᵒᵈ := ‹EDist X› - -@[simp] -theorem edist_toDual (a b : X) : edist (toDual a) (toDual b) = edist a b := - rfl - -@[simp] -theorem edist_ofDual (a b : Xᵒᵈ) : edist (ofDual a) (ofDual b) = edist a b := - rfl - -end - -instance [PseudoEMetricSpace X] : PseudoEMetricSpace Xᵒᵈ := ‹PseudoEMetricSpace X› -instance [EMetricSpace X] : EMetricSpace Xᵒᵈ := ‹EMetricSpace X› diff --git a/Mathlib/Topology/EMetricSpace/Defs.lean b/Mathlib/Topology/EMetricSpace/Defs.lean new file mode 100644 index 0000000000000..aed2614cb3ad5 --- /dev/null +++ b/Mathlib/Topology/EMetricSpace/Defs.lean @@ -0,0 +1,694 @@ +/- +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.ENNReal.Inv +import Mathlib.Topology.UniformSpace.Basic + +/-! +# Extended metric spaces + +This file is devoted to the definition and study of `EMetricSpace`s, i.e., metric +spaces in which the distance is allowed to take the value ∞. This extended distance is +called `edist`, and takes values in `ℝ≥0∞`. + +Many definitions and theorems expected on emetric spaces are already introduced on uniform spaces +and topological spaces. For example: open and closed sets, compactness, completeness, continuity and +uniform continuity. + +The class `EMetricSpace` therefore extends `UniformSpace` (and `TopologicalSpace`). + +Since a lot of elementary properties don't require `eq_of_edist_eq_zero` we start setting up the +theory of `PseudoEMetricSpace`, where we don't require `edist x y = 0 → x = y` and we specialize +to `EMetricSpace` at the end. +-/ + +assert_not_exists Nat.instLocallyFiniteOrder +assert_not_exists UniformEmbedding +assert_not_exists TendstoUniformlyOnFilter + +open Set Filter + +universe u v w + +variable {α : Type u} {β : Type v} {X : Type*} + +/-- Characterizing uniformities associated to a (generalized) distance function `D` +in terms of the elements of the uniformity. -/ +theorem uniformity_dist_of_mem_uniformity [LinearOrder β] {U : Filter (α × α)} (z : β) + (D : α → α → β) (H : ∀ s, s ∈ U ↔ ∃ ε > z, ∀ {a b : α}, D a b < ε → (a, b) ∈ s) : + U = ⨅ ε > z, 𝓟 { p : α × α | D p.1 p.2 < ε } := + HasBasis.eq_biInf ⟨fun s => by simp only [H, subset_def, Prod.forall, mem_setOf]⟩ + +open scoped Uniformity Topology Filter NNReal ENNReal Pointwise + +/-- `EDist α` means that `α` is equipped with an extended distance. -/ +@[ext] +class EDist (α : Type*) where + edist : α → α → ℝ≥0∞ + +export EDist (edist) + +/-- Creating a uniform space from an extended distance. -/ +def uniformSpaceOfEDist (edist : α → α → ℝ≥0∞) (edist_self : ∀ x : α, edist x x = 0) + (edist_comm : ∀ x y : α, edist x y = edist y x) + (edist_triangle : ∀ x y z : α, edist x z ≤ edist x y + edist y z) : UniformSpace α := + .ofFun edist edist_self edist_comm edist_triangle fun ε ε0 => + ⟨ε / 2, ENNReal.half_pos ε0.ne', fun _ h₁ _ h₂ => + (ENNReal.add_lt_add h₁ h₂).trans_eq (ENNReal.add_halves _)⟩ + +-- the uniform structure is embedded in the emetric space structure +-- to avoid instance diamond issues. See Note [forgetful inheritance]. +/-- Extended (pseudo) metric spaces, with an extended distance `edist` possibly taking the +value ∞ + +Each pseudo_emetric space induces a canonical `UniformSpace` and hence a canonical +`TopologicalSpace`. +This is enforced in the type class definition, by extending the `UniformSpace` structure. When +instantiating a `PseudoEMetricSpace` structure, the uniformity fields are not necessary, they +will be filled in by default. There is a default value for the uniformity, that can be substituted +in cases of interest, for instance when instantiating a `PseudoEMetricSpace` structure +on a product. + +Continuity of `edist` is proved in `Topology.Instances.ENNReal` +-/ +class PseudoEMetricSpace (α : Type u) extends EDist α : Type u where + edist_self : ∀ x : α, edist x x = 0 + edist_comm : ∀ x y : α, edist x y = edist y x + edist_triangle : ∀ x y z : α, edist x z ≤ edist x y + edist y z + toUniformSpace : UniformSpace α := uniformSpaceOfEDist edist edist_self edist_comm edist_triangle + uniformity_edist : 𝓤 α = ⨅ ε > 0, 𝓟 { p : α × α | edist p.1 p.2 < ε } := by rfl + +attribute [instance] PseudoEMetricSpace.toUniformSpace + +/- Pseudoemetric spaces are less common than metric spaces. Therefore, we work in a dedicated +namespace, while notions associated to metric spaces are mostly in the root namespace. -/ + +/-- Two pseudo emetric space structures with the same edistance function coincide. -/ +@[ext] +protected theorem PseudoEMetricSpace.ext {α : Type*} {m m' : PseudoEMetricSpace α} + (h : m.toEDist = m'.toEDist) : m = m' := by + cases' m with ed _ _ _ U hU + cases' m' with ed' _ _ _ U' hU' + congr 1 + exact UniformSpace.ext (((show ed = ed' from h) ▸ hU).trans hU'.symm) + +variable [PseudoEMetricSpace α] + +export PseudoEMetricSpace (edist_self edist_comm edist_triangle) + +attribute [simp] edist_self + +/-- Triangle inequality for the extended distance -/ +theorem edist_triangle_left (x y z : α) : edist x y ≤ edist z x + edist z y := by + rw [edist_comm z]; apply edist_triangle + +theorem edist_triangle_right (x y z : α) : edist x y ≤ edist x z + edist y z := by + rw [edist_comm y]; apply edist_triangle + +theorem edist_congr_right {x y z : α} (h : edist x y = 0) : edist x z = edist y z := by + apply le_antisymm + · rw [← zero_add (edist y z), ← h] + apply edist_triangle + · rw [edist_comm] at h + rw [← zero_add (edist x z), ← h] + apply edist_triangle + +theorem edist_congr_left {x y z : α} (h : edist x y = 0) : edist z x = edist z y := by + rw [edist_comm z x, edist_comm z y] + apply edist_congr_right h + +-- new theorem +theorem edist_congr {w x y z : α} (hl : edist w x = 0) (hr : edist y z = 0) : + edist w y = edist x z := + (edist_congr_right hl).trans (edist_congr_left hr) + +theorem edist_triangle4 (x y z t : α) : edist x t ≤ edist x y + edist y z + edist z t := + calc + edist x t ≤ edist x z + edist z t := edist_triangle x z t + _ ≤ edist x y + edist y z + edist z t := add_le_add_right (edist_triangle x y z) _ + +/-- Reformulation of the uniform structure in terms of the extended distance -/ +theorem uniformity_pseudoedist : 𝓤 α = ⨅ ε > 0, 𝓟 { p : α × α | edist p.1 p.2 < ε } := + PseudoEMetricSpace.uniformity_edist + +theorem uniformSpace_edist : + ‹PseudoEMetricSpace α›.toUniformSpace = + uniformSpaceOfEDist edist edist_self edist_comm edist_triangle := + UniformSpace.ext uniformity_pseudoedist + +theorem uniformity_basis_edist : + (𝓤 α).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) fun ε => { p : α × α | edist p.1 p.2 < ε } := + (@uniformSpace_edist α _).symm ▸ UniformSpace.hasBasis_ofFun ⟨1, one_pos⟩ _ _ _ _ _ + +/-- Characterization of the elements of the uniformity in terms of the extended distance -/ +theorem mem_uniformity_edist {s : Set (α × α)} : + s ∈ 𝓤 α ↔ ∃ ε > 0, ∀ {a b : α}, edist a b < ε → (a, b) ∈ s := + uniformity_basis_edist.mem_uniformity_iff + +/-- Given `f : β → ℝ≥0∞`, if `f` sends `{i | p i}` to a set of positive numbers +accumulating to zero, then `f i`-neighborhoods of the diagonal form a basis of `𝓤 α`. + +For specific bases see `uniformity_basis_edist`, `uniformity_basis_edist'`, +`uniformity_basis_edist_nnreal`, and `uniformity_basis_edist_inv_nat`. -/ +protected theorem EMetric.mk_uniformity_basis {β : Type*} {p : β → Prop} {f : β → ℝ≥0∞} + (hf₀ : ∀ x, p x → 0 < f x) (hf : ∀ ε, 0 < ε → ∃ x, p x ∧ f x ≤ ε) : + (𝓤 α).HasBasis p fun x => { p : α × α | edist p.1 p.2 < f x } := by + refine ⟨fun s => uniformity_basis_edist.mem_iff.trans ?_⟩ + constructor + · rintro ⟨ε, ε₀, hε⟩ + rcases hf ε ε₀ with ⟨i, hi, H⟩ + exact ⟨i, hi, fun x hx => hε <| lt_of_lt_of_le hx.out H⟩ + · exact fun ⟨i, hi, H⟩ => ⟨f i, hf₀ i hi, H⟩ + +/-- Given `f : β → ℝ≥0∞`, if `f` sends `{i | p i}` to a set of positive numbers +accumulating to zero, then closed `f i`-neighborhoods of the diagonal form a basis of `𝓤 α`. + +For specific bases see `uniformity_basis_edist_le` and `uniformity_basis_edist_le'`. -/ +protected theorem EMetric.mk_uniformity_basis_le {β : Type*} {p : β → Prop} {f : β → ℝ≥0∞} + (hf₀ : ∀ x, p x → 0 < f x) (hf : ∀ ε, 0 < ε → ∃ x, p x ∧ f x ≤ ε) : + (𝓤 α).HasBasis p fun x => { p : α × α | edist p.1 p.2 ≤ f x } := by + refine ⟨fun s => uniformity_basis_edist.mem_iff.trans ?_⟩ + constructor + · rintro ⟨ε, ε₀, hε⟩ + rcases exists_between ε₀ with ⟨ε', hε'⟩ + rcases hf ε' hε'.1 with ⟨i, hi, H⟩ + exact ⟨i, hi, fun x hx => hε <| lt_of_le_of_lt (le_trans hx.out H) hε'.2⟩ + · exact fun ⟨i, hi, H⟩ => ⟨f i, hf₀ i hi, fun x hx => H (le_of_lt hx.out)⟩ + +theorem uniformity_basis_edist_le : + (𝓤 α).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) fun ε => { p : α × α | edist p.1 p.2 ≤ ε } := + EMetric.mk_uniformity_basis_le (fun _ => id) fun ε ε₀ => ⟨ε, ε₀, le_refl ε⟩ + +theorem uniformity_basis_edist' (ε' : ℝ≥0∞) (hε' : 0 < ε') : + (𝓤 α).HasBasis (fun ε : ℝ≥0∞ => ε ∈ Ioo 0 ε') fun ε => { p : α × α | edist p.1 p.2 < ε } := + EMetric.mk_uniformity_basis (fun _ => And.left) fun ε ε₀ => + let ⟨δ, hδ⟩ := exists_between hε' + ⟨min ε δ, ⟨lt_min ε₀ hδ.1, lt_of_le_of_lt (min_le_right _ _) hδ.2⟩, min_le_left _ _⟩ + +theorem uniformity_basis_edist_le' (ε' : ℝ≥0∞) (hε' : 0 < ε') : + (𝓤 α).HasBasis (fun ε : ℝ≥0∞ => ε ∈ Ioo 0 ε') fun ε => { p : α × α | edist p.1 p.2 ≤ ε } := + EMetric.mk_uniformity_basis_le (fun _ => And.left) fun ε ε₀ => + let ⟨δ, hδ⟩ := exists_between hε' + ⟨min ε δ, ⟨lt_min ε₀ hδ.1, lt_of_le_of_lt (min_le_right _ _) hδ.2⟩, min_le_left _ _⟩ + +theorem uniformity_basis_edist_nnreal : + (𝓤 α).HasBasis (fun ε : ℝ≥0 => 0 < ε) fun ε => { p : α × α | edist p.1 p.2 < ε } := + EMetric.mk_uniformity_basis (fun _ => ENNReal.coe_pos.2) fun _ε ε₀ => + let ⟨δ, hδ⟩ := ENNReal.lt_iff_exists_nnreal_btwn.1 ε₀ + ⟨δ, ENNReal.coe_pos.1 hδ.1, le_of_lt hδ.2⟩ + +theorem uniformity_basis_edist_nnreal_le : + (𝓤 α).HasBasis (fun ε : ℝ≥0 => 0 < ε) fun ε => { p : α × α | edist p.1 p.2 ≤ ε } := + EMetric.mk_uniformity_basis_le (fun _ => ENNReal.coe_pos.2) fun _ε ε₀ => + let ⟨δ, hδ⟩ := ENNReal.lt_iff_exists_nnreal_btwn.1 ε₀ + ⟨δ, ENNReal.coe_pos.1 hδ.1, le_of_lt hδ.2⟩ + +theorem uniformity_basis_edist_inv_nat : + (𝓤 α).HasBasis (fun _ => True) fun n : ℕ => { p : α × α | edist p.1 p.2 < (↑n)⁻¹ } := + EMetric.mk_uniformity_basis (fun n _ ↦ ENNReal.inv_pos.2 <| ENNReal.natCast_ne_top n) fun _ε ε₀ ↦ + let ⟨n, hn⟩ := ENNReal.exists_inv_nat_lt (ne_of_gt ε₀) + ⟨n, trivial, le_of_lt hn⟩ + +theorem uniformity_basis_edist_inv_two_pow : + (𝓤 α).HasBasis (fun _ => True) fun n : ℕ => { p : α × α | edist p.1 p.2 < 2⁻¹ ^ n } := + EMetric.mk_uniformity_basis (fun _ _ => ENNReal.pow_pos (ENNReal.inv_pos.2 ENNReal.two_ne_top) _) + fun _ε ε₀ => + let ⟨n, hn⟩ := ENNReal.exists_inv_two_pow_lt (ne_of_gt ε₀) + ⟨n, trivial, le_of_lt hn⟩ + +/-- Fixed size neighborhoods of the diagonal belong to the uniform structure -/ +theorem edist_mem_uniformity {ε : ℝ≥0∞} (ε0 : 0 < ε) : { p : α × α | edist p.1 p.2 < ε } ∈ 𝓤 α := + mem_uniformity_edist.2 ⟨ε, ε0, id⟩ + +namespace EMetric + +instance (priority := 900) instIsCountablyGeneratedUniformity : IsCountablyGenerated (𝓤 α) := + isCountablyGenerated_of_seq ⟨_, uniformity_basis_edist_inv_nat.eq_iInf⟩ + +-- Porting note: changed explicit/implicit +/-- ε-δ characterization of uniform continuity on a set for pseudoemetric spaces -/ +theorem uniformContinuousOn_iff [PseudoEMetricSpace β] {f : α → β} {s : Set α} : + UniformContinuousOn f s ↔ + ∀ ε > 0, ∃ δ > 0, ∀ {a}, a ∈ s → ∀ {b}, b ∈ s → edist a b < δ → edist (f a) (f b) < ε := + uniformity_basis_edist.uniformContinuousOn_iff uniformity_basis_edist + +/-- ε-δ characterization of uniform continuity on pseudoemetric spaces -/ +theorem uniformContinuous_iff [PseudoEMetricSpace β] {f : α → β} : + UniformContinuous f ↔ ∀ ε > 0, ∃ δ > 0, ∀ {a b : α}, edist a b < δ → edist (f a) (f b) < ε := + uniformity_basis_edist.uniformContinuous_iff uniformity_basis_edist + +end EMetric + +open EMetric + +/-- Auxiliary function to replace the uniformity on a pseudoemetric space with +a uniformity which is equal to the original one, but maybe not defeq. +This is useful if one wants to construct a pseudoemetric space with a +specified uniformity. See Note [forgetful inheritance] explaining why having definitionally +the right uniformity is often important. +-/ +def PseudoEMetricSpace.replaceUniformity {α} [U : UniformSpace α] (m : PseudoEMetricSpace α) + (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : PseudoEMetricSpace α where + edist := @edist _ m.toEDist + edist_self := edist_self + edist_comm := edist_comm + edist_triangle := edist_triangle + toUniformSpace := U + uniformity_edist := H.trans (@PseudoEMetricSpace.uniformity_edist α _) + +/-- The extended pseudometric induced by a function taking values in a pseudoemetric space. -/ +def PseudoEMetricSpace.induced {α β} (f : α → β) (m : PseudoEMetricSpace β) : + PseudoEMetricSpace α where + edist x y := edist (f x) (f y) + edist_self _ := edist_self _ + edist_comm _ _ := edist_comm _ _ + edist_triangle _ _ _ := edist_triangle _ _ _ + toUniformSpace := UniformSpace.comap f m.toUniformSpace + uniformity_edist := (uniformity_basis_edist.comap (Prod.map f f)).eq_biInf + +/-- Pseudoemetric space instance on subsets of pseudoemetric spaces -/ +instance {α : Type*} {p : α → Prop} [PseudoEMetricSpace α] : PseudoEMetricSpace (Subtype p) := + PseudoEMetricSpace.induced Subtype.val ‹_› + +/-- The extended pseudodistance on a subset of a pseudoemetric space is the restriction of +the original pseudodistance, by definition -/ +theorem Subtype.edist_eq {p : α → Prop} (x y : Subtype p) : edist x y = edist (x : α) y := rfl + +namespace MulOpposite + +/-- Pseudoemetric space instance on the multiplicative opposite of a pseudoemetric space. -/ +@[to_additive "Pseudoemetric space instance on the additive opposite of a pseudoemetric space."] +instance {α : Type*} [PseudoEMetricSpace α] : PseudoEMetricSpace αᵐᵒᵖ := + PseudoEMetricSpace.induced unop ‹_› + +@[to_additive] +theorem edist_unop (x y : αᵐᵒᵖ) : edist (unop x) (unop y) = edist x y := rfl + +@[to_additive] +theorem edist_op (x y : α) : edist (op x) (op y) = edist x y := rfl + +end MulOpposite + +section ULift + +instance : PseudoEMetricSpace (ULift α) := PseudoEMetricSpace.induced ULift.down ‹_› + +theorem ULift.edist_eq (x y : ULift α) : edist x y = edist x.down y.down := rfl + +@[simp] +theorem ULift.edist_up_up (x y : α) : edist (ULift.up x) (ULift.up y) = edist x y := rfl + +end ULift + +/-- The product of two pseudoemetric spaces, with the max distance, is an extended +pseudometric spaces. We make sure that the uniform structure thus constructed is the one +corresponding to the product of uniform spaces, to avoid diamond problems. -/ +instance Prod.pseudoEMetricSpaceMax [PseudoEMetricSpace β] : PseudoEMetricSpace (α × β) where + edist x y := edist x.1 y.1 ⊔ edist x.2 y.2 + edist_self x := by simp + edist_comm x y := by simp [edist_comm] + edist_triangle x y z := + max_le (le_trans (edist_triangle _ _ _) (add_le_add (le_max_left _ _) (le_max_left _ _))) + (le_trans (edist_triangle _ _ _) (add_le_add (le_max_right _ _) (le_max_right _ _))) + uniformity_edist := uniformity_prod.trans <| by + simp [PseudoEMetricSpace.uniformity_edist, ← iInf_inf_eq, setOf_and] + toUniformSpace := inferInstance + +theorem Prod.edist_eq [PseudoEMetricSpace β] (x y : α × β) : + edist x y = max (edist x.1 y.1) (edist x.2 y.2) := + rfl + +namespace EMetric + +variable {x y z : α} {ε ε₁ ε₂ : ℝ≥0∞} {s t : Set α} + +/-- `EMetric.ball x ε` is the set of all points `y` with `edist y x < ε` -/ +def ball (x : α) (ε : ℝ≥0∞) : Set α := + { y | edist y x < ε } + +@[simp] theorem mem_ball : y ∈ ball x ε ↔ edist y x < ε := Iff.rfl + +theorem mem_ball' : y ∈ ball x ε ↔ edist x y < ε := by rw [edist_comm, mem_ball] + +/-- `EMetric.closedBall x ε` is the set of all points `y` with `edist y x ≤ ε` -/ +def closedBall (x : α) (ε : ℝ≥0∞) := + { y | edist y x ≤ ε } + +@[simp] theorem mem_closedBall : y ∈ closedBall x ε ↔ edist y x ≤ ε := Iff.rfl + +theorem mem_closedBall' : y ∈ closedBall x ε ↔ edist x y ≤ ε := by rw [edist_comm, mem_closedBall] + +@[simp] +theorem closedBall_top (x : α) : closedBall x ∞ = univ := + eq_univ_of_forall fun _ => mem_setOf.2 le_top + +theorem ball_subset_closedBall : ball x ε ⊆ closedBall x ε := fun _ h => le_of_lt h.out + +theorem pos_of_mem_ball (hy : y ∈ ball x ε) : 0 < ε := + lt_of_le_of_lt (zero_le _) hy + +theorem mem_ball_self (h : 0 < ε) : x ∈ ball x ε := by + rwa [mem_ball, edist_self] + +theorem mem_closedBall_self : x ∈ closedBall x ε := by + rw [mem_closedBall, edist_self]; apply zero_le + +theorem mem_ball_comm : x ∈ ball y ε ↔ y ∈ ball x ε := by rw [mem_ball', mem_ball] + +theorem mem_closedBall_comm : x ∈ closedBall y ε ↔ y ∈ closedBall x ε := by + rw [mem_closedBall', mem_closedBall] + +@[gcongr] +theorem ball_subset_ball (h : ε₁ ≤ ε₂) : ball x ε₁ ⊆ ball x ε₂ := fun _y (yx : _ < ε₁) => + lt_of_lt_of_le yx h + +@[gcongr] +theorem closedBall_subset_closedBall (h : ε₁ ≤ ε₂) : closedBall x ε₁ ⊆ closedBall x ε₂ := + fun _y (yx : _ ≤ ε₁) => le_trans yx h + +theorem ball_disjoint (h : ε₁ + ε₂ ≤ edist x y) : Disjoint (ball x ε₁) (ball y ε₂) := + Set.disjoint_left.mpr fun z h₁ h₂ => + (edist_triangle_left x y z).not_lt <| (ENNReal.add_lt_add h₁ h₂).trans_le h + +theorem ball_subset (h : edist x y + ε₁ ≤ ε₂) (h' : edist x y ≠ ∞) : ball x ε₁ ⊆ ball y ε₂ := + fun z zx => + calc + edist z y ≤ edist z x + edist x y := edist_triangle _ _ _ + _ = edist x y + edist z x := add_comm _ _ + _ < edist x y + ε₁ := ENNReal.add_lt_add_left h' zx + _ ≤ ε₂ := h + +theorem exists_ball_subset_ball (h : y ∈ ball x ε) : ∃ ε' > 0, ball y ε' ⊆ ball x ε := by + have : 0 < ε - edist y x := by simpa using h + refine ⟨ε - edist y x, this, ball_subset ?_ (ne_top_of_lt h)⟩ + exact (add_tsub_cancel_of_le (mem_ball.mp h).le).le + +theorem ball_eq_empty_iff : ball x ε = ∅ ↔ ε = 0 := + eq_empty_iff_forall_not_mem.trans + ⟨fun h => le_bot_iff.1 (le_of_not_gt fun ε0 => h _ (mem_ball_self ε0)), fun ε0 _ h => + not_lt_of_le (le_of_eq ε0) (pos_of_mem_ball h)⟩ + +theorem ordConnected_setOf_closedBall_subset (x : α) (s : Set α) : + OrdConnected { r | closedBall x r ⊆ s } := + ⟨fun _ _ _ h₁ _ h₂ => (closedBall_subset_closedBall h₂.2).trans h₁⟩ + +theorem ordConnected_setOf_ball_subset (x : α) (s : Set α) : OrdConnected { r | ball x r ⊆ s } := + ⟨fun _ _ _ h₁ _ h₂ => (ball_subset_ball h₂.2).trans h₁⟩ + +/-- Relation “two points are at a finite edistance” is an equivalence relation. -/ +def edistLtTopSetoid : Setoid α where + r x y := edist x y < ⊤ + iseqv := + ⟨fun x => by rw [edist_self]; exact ENNReal.coe_lt_top, + fun h => by rwa [edist_comm], fun hxy hyz => + lt_of_le_of_lt (edist_triangle _ _ _) (ENNReal.add_lt_top.2 ⟨hxy, hyz⟩)⟩ + +@[simp] +theorem ball_zero : ball x 0 = ∅ := by rw [EMetric.ball_eq_empty_iff] + +theorem nhds_basis_eball : (𝓝 x).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) (ball x) := + nhds_basis_uniformity uniformity_basis_edist + +theorem nhdsWithin_basis_eball : (𝓝[s] x).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) fun ε => ball x ε ∩ s := + nhdsWithin_hasBasis nhds_basis_eball s + +theorem nhds_basis_closed_eball : (𝓝 x).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) (closedBall x) := + nhds_basis_uniformity uniformity_basis_edist_le + +theorem nhdsWithin_basis_closed_eball : + (𝓝[s] x).HasBasis (fun ε : ℝ≥0∞ => 0 < ε) fun ε => closedBall x ε ∩ s := + nhdsWithin_hasBasis nhds_basis_closed_eball s + +theorem nhds_eq : 𝓝 x = ⨅ ε > 0, 𝓟 (ball x ε) := + nhds_basis_eball.eq_biInf + +theorem mem_nhds_iff : s ∈ 𝓝 x ↔ ∃ ε > 0, ball x ε ⊆ s := + nhds_basis_eball.mem_iff + +theorem mem_nhdsWithin_iff : s ∈ 𝓝[t] x ↔ ∃ ε > 0, ball x ε ∩ t ⊆ s := + nhdsWithin_basis_eball.mem_iff + +section + +variable [PseudoEMetricSpace β] {f : α → β} + +theorem tendsto_nhdsWithin_nhdsWithin {t : Set β} {a b} : + Tendsto f (𝓝[s] a) (𝓝[t] b) ↔ + ∀ ε > 0, ∃ δ > 0, ∀ ⦃x⦄, x ∈ s → edist x a < δ → f x ∈ t ∧ edist (f x) b < ε := + (nhdsWithin_basis_eball.tendsto_iff nhdsWithin_basis_eball).trans <| + forall₂_congr fun ε _ => exists_congr fun δ => and_congr_right fun _ => + forall_congr' fun x => by simp; tauto + +theorem tendsto_nhdsWithin_nhds {a b} : + Tendsto f (𝓝[s] a) (𝓝 b) ↔ + ∀ ε > 0, ∃ δ > 0, ∀ {x : α}, x ∈ s → edist x a < δ → edist (f x) b < ε := by + rw [← nhdsWithin_univ b, tendsto_nhdsWithin_nhdsWithin] + simp only [mem_univ, true_and_iff] + +theorem tendsto_nhds_nhds {a b} : + Tendsto f (𝓝 a) (𝓝 b) ↔ ∀ ε > 0, ∃ δ > 0, ∀ ⦃x⦄, edist x a < δ → edist (f x) b < ε := + nhds_basis_eball.tendsto_iff nhds_basis_eball + +end + +theorem isOpen_iff : IsOpen s ↔ ∀ x ∈ s, ∃ ε > 0, ball x ε ⊆ s := by + simp [isOpen_iff_nhds, mem_nhds_iff] + +theorem isOpen_ball : IsOpen (ball x ε) := + isOpen_iff.2 fun _ => exists_ball_subset_ball + +theorem isClosed_ball_top : IsClosed (ball x ⊤) := + isOpen_compl_iff.1 <| isOpen_iff.2 fun _y hy => + ⟨⊤, ENNReal.coe_lt_top, fun _z hzy hzx => + hy (edistLtTopSetoid.trans (edistLtTopSetoid.symm hzy) hzx)⟩ + +theorem ball_mem_nhds (x : α) {ε : ℝ≥0∞} (ε0 : 0 < ε) : ball x ε ∈ 𝓝 x := + isOpen_ball.mem_nhds (mem_ball_self ε0) + +theorem closedBall_mem_nhds (x : α) {ε : ℝ≥0∞} (ε0 : 0 < ε) : closedBall x ε ∈ 𝓝 x := + mem_of_superset (ball_mem_nhds x ε0) ball_subset_closedBall + +theorem ball_prod_same [PseudoEMetricSpace β] (x : α) (y : β) (r : ℝ≥0∞) : + ball x r ×ˢ ball y r = ball (x, y) r := + ext fun z => by simp [Prod.edist_eq] + +theorem closedBall_prod_same [PseudoEMetricSpace β] (x : α) (y : β) (r : ℝ≥0∞) : + closedBall x r ×ˢ closedBall y r = closedBall (x, y) r := + ext fun z => by simp [Prod.edist_eq] + +/-- ε-characterization of the closure in pseudoemetric spaces -/ +theorem mem_closure_iff : x ∈ closure s ↔ ∀ ε > 0, ∃ y ∈ s, edist x y < ε := + (mem_closure_iff_nhds_basis nhds_basis_eball).trans <| by simp only [mem_ball, edist_comm x] + +theorem tendsto_nhds {f : Filter β} {u : β → α} {a : α} : + Tendsto u f (𝓝 a) ↔ ∀ ε > 0, ∀ᶠ x in f, edist (u x) a < ε := + nhds_basis_eball.tendsto_right_iff + +theorem tendsto_atTop [Nonempty β] [SemilatticeSup β] {u : β → α} {a : α} : + Tendsto u atTop (𝓝 a) ↔ ∀ ε > 0, ∃ N, ∀ n ≥ N, edist (u n) a < ε := + (atTop_basis.tendsto_iff nhds_basis_eball).trans <| by + simp only [exists_prop, true_and_iff, mem_Ici, mem_ball] + +section Compact + +-- Porting note (#11215): TODO: generalize to a uniform space with metrizable uniformity +/-- For a set `s` in a pseudo emetric space, if for every `ε > 0` there exists a countable +set that is `ε`-dense in `s`, then there exists a countable subset `t ⊆ s` that is dense in `s`. -/ +theorem subset_countable_closure_of_almost_dense_set (s : Set α) + (hs : ∀ ε > 0, ∃ t : Set α, t.Countable ∧ s ⊆ ⋃ x ∈ t, closedBall x ε) : + ∃ t, t ⊆ s ∧ t.Countable ∧ s ⊆ closure t := by + rcases s.eq_empty_or_nonempty with (rfl | ⟨x₀, hx₀⟩) + · exact ⟨∅, empty_subset _, countable_empty, empty_subset _⟩ + choose! T hTc hsT using fun n : ℕ => hs n⁻¹ (by simp) + have : ∀ r x, ∃ y ∈ s, closedBall x r ∩ s ⊆ closedBall y (r * 2) := fun r x => by + rcases (closedBall x r ∩ s).eq_empty_or_nonempty with (he | ⟨y, hxy, hys⟩) + · refine ⟨x₀, hx₀, ?_⟩ + rw [he] + exact empty_subset _ + · refine ⟨y, hys, fun z hz => ?_⟩ + calc + edist z y ≤ edist z x + edist y x := edist_triangle_right _ _ _ + _ ≤ r + r := add_le_add hz.1 hxy + _ = r * 2 := (mul_two r).symm + choose f hfs hf using this + refine + ⟨⋃ n : ℕ, f n⁻¹ '' T n, iUnion_subset fun n => image_subset_iff.2 fun z _ => hfs _ _, + countable_iUnion fun n => (hTc n).image _, ?_⟩ + refine fun x hx => mem_closure_iff.2 fun ε ε0 => ?_ + rcases ENNReal.exists_inv_nat_lt (ENNReal.half_pos ε0.lt.ne').ne' with ⟨n, hn⟩ + rcases mem_iUnion₂.1 (hsT n hx) with ⟨y, hyn, hyx⟩ + refine ⟨f n⁻¹ y, mem_iUnion.2 ⟨n, mem_image_of_mem _ hyn⟩, ?_⟩ + calc + edist x (f n⁻¹ y) ≤ (n : ℝ≥0∞)⁻¹ * 2 := hf _ _ ⟨hyx, hx⟩ + _ < ε := ENNReal.mul_lt_of_lt_div hn + +open TopologicalSpace in +/-- If a set `s` is separable in a (pseudo extended) metric space, then it admits a countable dense +subset. This is not obvious, as the countable set whose closure covers `s` given by the definition +of separability does not need in general to be contained in `s`. -/ +theorem _root_.TopologicalSpace.IsSeparable.exists_countable_dense_subset + {s : Set α} (hs : IsSeparable s) : ∃ t, t ⊆ s ∧ t.Countable ∧ s ⊆ closure t := by + have : ∀ ε > 0, ∃ t : Set α, t.Countable ∧ s ⊆ ⋃ x ∈ t, closedBall x ε := fun ε ε0 => by + rcases hs with ⟨t, htc, hst⟩ + refine ⟨t, htc, hst.trans fun x hx => ?_⟩ + rcases mem_closure_iff.1 hx ε ε0 with ⟨y, hyt, hxy⟩ + exact mem_iUnion₂.2 ⟨y, hyt, mem_closedBall.2 hxy.le⟩ + exact subset_countable_closure_of_almost_dense_set _ this + +open TopologicalSpace in +/-- If a set `s` is separable, then the corresponding subtype is separable in a (pseudo extended) +metric space. This is not obvious, as the countable set whose closure covers `s` does not need in +general to be contained in `s`. -/ +theorem _root_.TopologicalSpace.IsSeparable.separableSpace {s : Set α} (hs : IsSeparable s) : + SeparableSpace s := by + rcases hs.exists_countable_dense_subset with ⟨t, hts, htc, hst⟩ + lift t to Set s using hts + refine ⟨⟨t, countable_of_injective_of_countable_image Subtype.coe_injective.injOn htc, ?_⟩⟩ + rwa [inducing_subtype_val.dense_iff, Subtype.forall] + +end Compact + +end EMetric + +--namespace +/-- We now define `EMetricSpace`, extending `PseudoEMetricSpace`. -/ +class EMetricSpace (α : Type u) extends PseudoEMetricSpace α : Type u where + eq_of_edist_eq_zero : ∀ {x y : α}, edist x y = 0 → x = y + +@[ext] +protected theorem EMetricSpace.ext + {α : Type*} {m m' : EMetricSpace α} (h : m.toEDist = m'.toEDist) : m = m' := by + cases m + cases m' + congr + ext1 + assumption + +variable {γ : Type w} [EMetricSpace γ] + +export EMetricSpace (eq_of_edist_eq_zero) + +/-- Characterize the equality of points by the vanishing of their extended distance -/ +@[simp] +theorem edist_eq_zero {x y : γ} : edist x y = 0 ↔ x = y := + ⟨eq_of_edist_eq_zero, fun h => h ▸ edist_self _⟩ + +@[simp] +theorem zero_eq_edist {x y : γ} : 0 = edist x y ↔ x = y := eq_comm.trans edist_eq_zero + +theorem edist_le_zero {x y : γ} : edist x y ≤ 0 ↔ x = y := + nonpos_iff_eq_zero.trans edist_eq_zero + +@[simp] +theorem edist_pos {x y : γ} : 0 < edist x y ↔ x ≠ y := by simp [← not_le] + +/-- Two points coincide if their distance is `< ε` for all positive ε -/ +theorem eq_of_forall_edist_le {x y : γ} (h : ∀ ε > 0, edist x y ≤ ε) : x = y := + eq_of_edist_eq_zero (eq_of_le_of_forall_le_of_dense bot_le h) + +/-- Auxiliary function to replace the uniformity on an emetric space with +a uniformity which is equal to the original one, but maybe not defeq. +This is useful if one wants to construct an emetric space with a +specified uniformity. See Note [forgetful inheritance] explaining why having definitionally +the right uniformity is often important. +-/ +def EMetricSpace.replaceUniformity {γ} [U : UniformSpace γ] (m : EMetricSpace γ) + (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : EMetricSpace γ where + edist := @edist _ m.toEDist + edist_self := edist_self + eq_of_edist_eq_zero := @eq_of_edist_eq_zero _ _ + edist_comm := edist_comm + edist_triangle := edist_triangle + toUniformSpace := U + uniformity_edist := H.trans (@PseudoEMetricSpace.uniformity_edist γ _) + +/-- The extended metric induced by an injective function taking values in an emetric space. -/ +def EMetricSpace.induced {γ β} (f : γ → β) (hf : Function.Injective f) (m : EMetricSpace β) : + EMetricSpace γ := + { PseudoEMetricSpace.induced f m.toPseudoEMetricSpace with + eq_of_edist_eq_zero := fun h => hf (edist_eq_zero.1 h) } + +/-- EMetric space instance on subsets of emetric spaces -/ +instance {α : Type*} {p : α → Prop} [EMetricSpace α] : EMetricSpace (Subtype p) := + EMetricSpace.induced Subtype.val Subtype.coe_injective ‹_› + +/-- EMetric space instance on the multiplicative opposite of an emetric space. -/ +@[to_additive "EMetric space instance on the additive opposite of an emetric space."] +instance {α : Type*} [EMetricSpace α] : EMetricSpace αᵐᵒᵖ := + EMetricSpace.induced MulOpposite.unop MulOpposite.unop_injective ‹_› + +instance {α : Type*} [EMetricSpace α] : EMetricSpace (ULift α) := + EMetricSpace.induced ULift.down ULift.down_injective ‹_› + +/-- Reformulation of the uniform structure in terms of the extended distance -/ +theorem uniformity_edist : 𝓤 γ = ⨅ ε > 0, 𝓟 { p : γ × γ | edist p.1 p.2 < ε } := + PseudoEMetricSpace.uniformity_edist + +/-! +### `Additive`, `Multiplicative` + +The distance on those type synonyms is inherited without change. +-/ + + +open Additive Multiplicative + +section + +variable [EDist X] + +instance : EDist (Additive X) := ‹EDist X› +instance : EDist (Multiplicative X) := ‹EDist X› + +@[simp] +theorem edist_ofMul (a b : X) : edist (ofMul a) (ofMul b) = edist a b := + rfl + +@[simp] +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 := + rfl + +@[simp] +theorem edist_toAdd (a b : Multiplicative X) : edist (toAdd a) (toAdd b) = edist a b := + rfl + +end + +instance [PseudoEMetricSpace X] : PseudoEMetricSpace (Additive X) := ‹PseudoEMetricSpace X› +instance [PseudoEMetricSpace X] : PseudoEMetricSpace (Multiplicative X) := ‹PseudoEMetricSpace X› +instance [EMetricSpace X] : EMetricSpace (Additive X) := ‹EMetricSpace X› +instance [EMetricSpace X] : EMetricSpace (Multiplicative X) := ‹EMetricSpace X› + +/-! +### Order dual + +The distance on this type synonym is inherited without change. +-/ + + +open OrderDual + +section + +variable [EDist X] + +instance : EDist Xᵒᵈ := ‹EDist X› + +@[simp] +theorem edist_toDual (a b : X) : edist (toDual a) (toDual b) = edist a b := + rfl + +@[simp] +theorem edist_ofDual (a b : Xᵒᵈ) : edist (ofDual a) (ofDual b) = edist a b := + rfl + +end + +instance [PseudoEMetricSpace X] : PseudoEMetricSpace Xᵒᵈ := ‹PseudoEMetricSpace X› +instance [EMetricSpace X] : EMetricSpace Xᵒᵈ := ‹EMetricSpace X› diff --git a/Mathlib/Topology/EMetricSpace/Diam.lean b/Mathlib/Topology/EMetricSpace/Diam.lean new file mode 100644 index 0000000000000..d1ff6d783af68 --- /dev/null +++ b/Mathlib/Topology/EMetricSpace/Diam.lean @@ -0,0 +1,149 @@ +/- +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.Topology.EMetricSpace.Basic +import Mathlib.Data.ENNReal.Real + +/-! +# Diameters of sets in extended metric spaces + +-/ + + +open Set Filter Classical + +open scoped Uniformity Topology Filter NNReal ENNReal Pointwise + +universe u v w + +variable {α β : Type*} {s : Set α} {x y z : α} + +namespace EMetric + +section +variable [PseudoEMetricSpace α] + +/-- The diameter of a set in a pseudoemetric space, named `EMetric.diam` -/ +noncomputable def diam (s : Set α) := + ⨆ (x ∈ s) (y ∈ s), edist x y + +theorem diam_eq_sSup (s : Set α) : diam s = sSup (image2 edist s s) := sSup_image2.symm + +theorem diam_le_iff {d : ℝ≥0∞} : diam s ≤ d ↔ ∀ x ∈ s, ∀ y ∈ s, edist x y ≤ d := by + simp only [diam, iSup_le_iff] + +theorem diam_image_le_iff {d : ℝ≥0∞} {f : β → α} {s : Set β} : + diam (f '' s) ≤ d ↔ ∀ x ∈ s, ∀ y ∈ s, edist (f x) (f y) ≤ d := by + simp only [diam_le_iff, forall_mem_image] + +theorem edist_le_of_diam_le {d} (hx : x ∈ s) (hy : y ∈ s) (hd : diam s ≤ d) : edist x y ≤ d := + diam_le_iff.1 hd x hx y hy + +/-- If two points belong to some set, their edistance is bounded by the diameter of the set -/ +theorem edist_le_diam_of_mem (hx : x ∈ s) (hy : y ∈ s) : edist x y ≤ diam s := + edist_le_of_diam_le hx hy le_rfl + +/-- If the distance between any two points in a set is bounded by some constant, this constant +bounds the diameter. -/ +theorem diam_le {d : ℝ≥0∞} (h : ∀ x ∈ s, ∀ y ∈ s, edist x y ≤ d) : diam s ≤ d := + diam_le_iff.2 h + +/-- The diameter of a subsingleton vanishes. -/ +theorem diam_subsingleton (hs : s.Subsingleton) : diam s = 0 := + nonpos_iff_eq_zero.1 <| diam_le fun _x hx y hy => (hs hx hy).symm ▸ edist_self y ▸ le_rfl + +/-- The diameter of the empty set vanishes -/ +@[simp] +theorem diam_empty : diam (∅ : Set α) = 0 := + diam_subsingleton subsingleton_empty + +/-- The diameter of a singleton vanishes -/ +@[simp] +theorem diam_singleton : diam ({x} : Set α) = 0 := + diam_subsingleton subsingleton_singleton + +@[to_additive (attr := simp)] +theorem diam_one [One α] : diam (1 : Set α) = 0 := + diam_singleton + +theorem diam_iUnion_mem_option {ι : Type*} (o : Option ι) (s : ι → Set α) : + diam (⋃ i ∈ o, s i) = ⨆ i ∈ o, diam (s i) := by cases o <;> simp + +theorem diam_insert : diam (insert x s) = max (⨆ y ∈ s, edist x y) (diam s) := + eq_of_forall_ge_iff fun d => by + simp only [diam_le_iff, forall_mem_insert, edist_self, edist_comm x, max_le_iff, iSup_le_iff, + zero_le, true_and_iff, forall_and, and_self_iff, ← and_assoc] + +theorem diam_pair : diam ({x, y} : Set α) = edist x y := by + simp only [iSup_singleton, diam_insert, diam_singleton, ENNReal.max_zero_right] + +theorem diam_triple : diam ({x, y, z} : Set α) = max (max (edist x y) (edist x z)) (edist y z) := by + simp only [diam_insert, iSup_insert, iSup_singleton, diam_singleton, ENNReal.max_zero_right, + ENNReal.sup_eq_max] + +/-- The diameter is monotonous with respect to inclusion -/ +@[gcongr] +theorem diam_mono {s t : Set α} (h : s ⊆ t) : diam s ≤ diam t := + diam_le fun _x hx _y hy => edist_le_diam_of_mem (h hx) (h hy) + +/-- The diameter of a union is controlled by the diameter of the sets, and the edistance +between two points in the sets. -/ +theorem diam_union {t : Set α} (xs : x ∈ s) (yt : y ∈ t) : + diam (s ∪ t) ≤ diam s + edist x y + diam t := by + have A : ∀ a ∈ s, ∀ b ∈ t, edist a b ≤ diam s + edist x y + diam t := fun a ha b hb => + calc + edist a b ≤ edist a x + edist x y + edist y b := edist_triangle4 _ _ _ _ + _ ≤ diam s + edist x y + diam t := + add_le_add (add_le_add (edist_le_diam_of_mem ha xs) le_rfl) (edist_le_diam_of_mem yt hb) + refine diam_le fun a ha b hb => ?_ + cases' (mem_union _ _ _).1 ha with h'a h'a <;> cases' (mem_union _ _ _).1 hb with h'b h'b + · calc + edist a b ≤ diam s := edist_le_diam_of_mem h'a h'b + _ ≤ diam s + (edist x y + diam t) := le_self_add + _ = diam s + edist x y + diam t := (add_assoc _ _ _).symm + · exact A a h'a b h'b + · have Z := A b h'b a h'a + rwa [edist_comm] at Z + · calc + edist a b ≤ diam t := edist_le_diam_of_mem h'a h'b + _ ≤ diam s + edist x y + diam t := le_add_self + +theorem diam_union' {t : Set α} (h : (s ∩ t).Nonempty) : diam (s ∪ t) ≤ diam s + diam t := by + let ⟨x, ⟨xs, xt⟩⟩ := h + simpa using diam_union xs xt + +theorem diam_closedBall {r : ℝ≥0∞} : diam (closedBall x r) ≤ 2 * r := + diam_le fun a ha b hb => + calc + edist a b ≤ edist a x + edist b x := edist_triangle_right _ _ _ + _ ≤ r + r := add_le_add ha hb + _ = 2 * r := (two_mul r).symm + +theorem diam_ball {r : ℝ≥0∞} : diam (ball x r) ≤ 2 * r := + le_trans (diam_mono ball_subset_closedBall) diam_closedBall + +theorem diam_pi_le_of_le {π : β → Type*} [Fintype β] [∀ b, PseudoEMetricSpace (π b)] + {s : ∀ b : β, Set (π b)} {c : ℝ≥0∞} (h : ∀ b, diam (s b) ≤ c) : diam (Set.pi univ s) ≤ c := by + refine diam_le fun x hx y hy => edist_pi_le_iff.mpr ?_ + rw [mem_univ_pi] at hx hy + exact fun b => diam_le_iff.1 (h b) (x b) (hx b) (y b) (hy b) + +end + +section +variable [EMetricSpace β] {s : Set β} + +theorem diam_eq_zero_iff : diam s = 0 ↔ s.Subsingleton := + ⟨fun h _x hx _y hy => edist_le_zero.1 <| h ▸ edist_le_diam_of_mem hx hy, diam_subsingleton⟩ + +theorem diam_pos_iff : 0 < diam s ↔ s.Nontrivial := by + simp only [pos_iff_ne_zero, Ne, diam_eq_zero_iff, Set.not_subsingleton_iff] + +theorem diam_pos_iff' : 0 < diam s ↔ ∃ x ∈ s, ∃ y ∈ s, x ≠ y := by + simp only [diam_pos_iff, Set.Nontrivial, exists_prop] + +end + +end EMetric diff --git a/Mathlib/Topology/EMetricSpace/Lipschitz.lean b/Mathlib/Topology/EMetricSpace/Lipschitz.lean index b36604dea713c..4240b7841e469 100644 --- a/Mathlib/Topology/EMetricSpace/Lipschitz.lean +++ b/Mathlib/Topology/EMetricSpace/Lipschitz.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rohan Mitta, Kevin Buzzard, Alistair Tucker, Johannes Hölzl, Yury Kudryashov -/ import Mathlib.Logic.Function.Iterate -import Mathlib.Topology.EMetricSpace.Basic +import Mathlib.Topology.EMetricSpace.Diam import Mathlib.Tactic.GCongr /-! diff --git a/Mathlib/Topology/EMetricSpace/Paracompact.lean b/Mathlib/Topology/EMetricSpace/Paracompact.lean index db4fdd6346d82..dc2427907aac4 100644 --- a/Mathlib/Topology/EMetricSpace/Paracompact.lean +++ b/Mathlib/Topology/EMetricSpace/Paracompact.lean @@ -5,8 +5,8 @@ Authors: Yury Kudryashov -/ import Mathlib.SetTheory.Ordinal.Basic import Mathlib.Tactic.GCongr -import Mathlib.Topology.EMetricSpace.Basic import Mathlib.Topology.Compactness.Paracompact +import Mathlib.Topology.EMetricSpace.Basic /-! # (Extended) metric spaces are paracompact diff --git a/Mathlib/Topology/Instances/Discrete.lean b/Mathlib/Topology/Instances/Discrete.lean index b937256b9520c..f8dd829077e96 100644 --- a/Mathlib/Topology/Instances/Discrete.lean +++ b/Mathlib/Topology/Instances/Discrete.lean @@ -5,7 +5,6 @@ Authors: Rémy Degenne -/ import Mathlib.Order.SuccPred.Basic import Mathlib.Topology.Order.Basic -import Mathlib.Topology.Metrizable.Uniformity /-! # Instances related to the discrete topology @@ -110,8 +109,3 @@ theorem discreteTopology_iff_orderTopology_of_pred_succ [LinearOrder α] [PredOr instance (priority := 100) DiscreteTopology.orderTopology_of_pred_succ [h : DiscreteTopology α] [LinearOrder α] [PredOrder α] [SuccOrder α] : OrderTopology α := discreteTopology_iff_orderTopology_of_pred_succ.mp h - -instance (priority := 100) DiscreteTopology.metrizableSpace [DiscreteTopology α] : - MetrizableSpace α := by - obtain rfl := DiscreteTopology.eq_bot (α := α) - exact @UniformSpace.metrizableSpace α ⊥ (isCountablyGenerated_principal _) _ diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 7d4f1adcd31f9..bfe23ec8d795c 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -9,6 +9,8 @@ import Mathlib.Topology.Instances.NNReal import Mathlib.Topology.EMetricSpace.Lipschitz import Mathlib.Topology.Metrizable.Basic import Mathlib.Topology.Order.T5 +import Mathlib.Topology.MetricSpace.Pseudo.Real +import Mathlib.Topology.Metrizable.Uniformity /-! # Topology on extended non-negative reals diff --git a/Mathlib/Topology/Instances/Int.lean b/Mathlib/Topology/Instances/Int.lean index 1c532fee6a853..eae87217dd5dc 100644 --- a/Mathlib/Topology/Instances/Int.lean +++ b/Mathlib/Topology/Instances/Int.lean @@ -8,8 +8,8 @@ import Mathlib.Data.Int.SuccPred import Mathlib.Data.Int.ConditionallyCompleteOrder import Mathlib.Topology.Instances.Discrete import Mathlib.Topology.MetricSpace.Bounded -import Mathlib.Topology.MetricSpace.Pseudo.Lemmas import Mathlib.Order.Filter.Archimedean +import Mathlib.Topology.MetricSpace.Basic /-! # Topology on the integers diff --git a/Mathlib/Topology/Instances/Nat.lean b/Mathlib/Topology/Instances/Nat.lean index def7fa44e4af2..2e4857faf7c51 100644 --- a/Mathlib/Topology/Instances/Nat.lean +++ b/Mathlib/Topology/Instances/Nat.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Topology.Instances.Int +import Mathlib.Data.Nat.Lattice /-! # Topology on the natural numbers diff --git a/Mathlib/Topology/Instances/RatLemmas.lean b/Mathlib/Topology/Instances/RatLemmas.lean index b85ff2b228d17..b94351980b416 100644 --- a/Mathlib/Topology/Instances/RatLemmas.lean +++ b/Mathlib/Topology/Instances/RatLemmas.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov import Mathlib.Topology.Instances.Irrational import Mathlib.Topology.Instances.Rat import Mathlib.Topology.Compactification.OnePoint +import Mathlib.Topology.Metrizable.Uniformity /-! # Additional lemmas about the topology on rational numbers diff --git a/Mathlib/Topology/Instances/Real.lean b/Mathlib/Topology/Instances/Real.lean index 139a67e9753bb..5b895c2cdeb63 100644 --- a/Mathlib/Topology/Instances/Real.lean +++ b/Mathlib/Topology/Instances/Real.lean @@ -12,6 +12,7 @@ import Mathlib.Topology.Algebra.UniformMulAction import Mathlib.Topology.Algebra.Star import Mathlib.Topology.Instances.Int import Mathlib.Topology.Order.Bornology +import Mathlib.Topology.Metrizable.Basic /-! # Topological properties of ℝ diff --git a/Mathlib/Topology/MetricSpace/Basic.lean b/Mathlib/Topology/MetricSpace/Basic.lean index 692a477a5e8a4..f3b2468b6e50b 100644 --- a/Mathlib/Topology/MetricSpace/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Basic.lean @@ -3,26 +3,14 @@ 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.Topology.MetricSpace.Pseudo.Basic import Mathlib.Topology.MetricSpace.Pseudo.Lemmas +import Mathlib.Topology.MetricSpace.Pseudo.Pi +import Mathlib.Topology.MetricSpace.Defs /-! -# Metric spaces +# Basic properties of metric spaces, and instances. -This file defines metric spaces and shows some of their basic properties. - -Many definitions and theorems expected on metric spaces are already introduced on uniform spaces and -topological spaces. This includes open and closed sets, compactness, completeness, continuity -and uniform continuity. - -TODO (anyone): Add "Main results" section. - -## Implementation notes -A lot of elementary properties don't require `eq_of_dist_eq_zero`, hence are stated and proven -for `PseudoMetricSpace`s in `PseudoMetric.lean`. - -## Tags - -metric, pseudo_metric, dist -/ open Set Filter Bornology @@ -30,87 +18,14 @@ open scoped NNReal Uniformity universe u v w -variable {α : Type u} {β : Type v} {X ι : Type*} +variable {α : Type u} {β : Type v} {X : Type*} variable [PseudoMetricSpace α] - -/-- We now define `MetricSpace`, extending `PseudoMetricSpace`. -/ -class MetricSpace (α : Type u) extends PseudoMetricSpace α : Type u where - eq_of_dist_eq_zero : ∀ {x y : α}, dist x y = 0 → x = y - -/-- Two metric space structures with the same distance coincide. -/ -@[ext] -theorem MetricSpace.ext {α : Type*} {m m' : MetricSpace α} (h : m.toDist = m'.toDist) : - m = m' := by - cases m; cases m'; congr; ext1; assumption - -/-- Construct a metric space structure whose underlying topological space structure -(definitionally) agrees which a pre-existing topology which is compatible with a given distance -function. -/ -def MetricSpace.ofDistTopology {α : Type u} [TopologicalSpace α] (dist : α → α → ℝ) - (dist_self : ∀ x : α, dist x x = 0) (dist_comm : ∀ x y : α, dist x y = dist y x) - (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) - (H : ∀ s : Set α, IsOpen s ↔ ∀ x ∈ s, ∃ ε > 0, ∀ y, dist x y < ε → y ∈ s) - (eq_of_dist_eq_zero : ∀ x y : α, dist x y = 0 → x = y) : MetricSpace α := - { PseudoMetricSpace.ofDistTopology dist dist_self dist_comm dist_triangle H with - eq_of_dist_eq_zero := eq_of_dist_eq_zero _ _ } - variable {γ : Type w} [MetricSpace γ] -theorem eq_of_dist_eq_zero {x y : γ} : dist x y = 0 → x = y := - MetricSpace.eq_of_dist_eq_zero - -@[simp] -theorem dist_eq_zero {x y : γ} : dist x y = 0 ↔ x = y := - Iff.intro eq_of_dist_eq_zero fun this => this ▸ dist_self _ - -@[simp] -theorem zero_eq_dist {x y : γ} : 0 = dist x y ↔ x = y := by rw [eq_comm, dist_eq_zero] - -theorem dist_ne_zero {x y : γ} : dist x y ≠ 0 ↔ x ≠ y := by - simpa only [not_iff_not] using dist_eq_zero - -@[simp] -theorem dist_le_zero {x y : γ} : dist x y ≤ 0 ↔ x = y := by - simpa [le_antisymm_iff, dist_nonneg] using @dist_eq_zero _ _ x y - -@[simp] -theorem dist_pos {x y : γ} : 0 < dist x y ↔ x ≠ y := by - simpa only [not_le] using not_congr dist_le_zero - -theorem eq_of_forall_dist_le {x y : γ} (h : ∀ ε > 0, dist x y ≤ ε) : x = y := - eq_of_dist_eq_zero (eq_of_le_of_forall_le_of_dense dist_nonneg h) - -/-- Deduce the equality of points from the vanishing of the nonnegative distance-/ -theorem eq_of_nndist_eq_zero {x y : γ} : nndist x y = 0 → x = y := by - simp only [NNReal.eq_iff, ← dist_nndist, imp_self, NNReal.coe_zero, dist_eq_zero] - -/-- Characterize the equality of points as the vanishing of the nonnegative distance-/ -@[simp] -theorem nndist_eq_zero {x y : γ} : nndist x y = 0 ↔ x = y := by - simp only [NNReal.eq_iff, ← dist_nndist, imp_self, NNReal.coe_zero, dist_eq_zero] - -@[simp] -theorem zero_eq_nndist {x y : γ} : 0 = nndist x y ↔ x = y := by - simp only [NNReal.eq_iff, ← dist_nndist, imp_self, NNReal.coe_zero, zero_eq_dist] - namespace Metric variable {x : γ} {s : Set γ} -@[simp] theorem closedBall_zero : closedBall x 0 = {x} := Set.ext fun _ => dist_le_zero - -@[simp] theorem sphere_zero : sphere x 0 = {x} := Set.ext fun _ => dist_eq_zero - -theorem subsingleton_closedBall (x : γ) {r : ℝ} (hr : r ≤ 0) : (closedBall x r).Subsingleton := by - rcases hr.lt_or_eq with (hr | rfl) - · rw [closedBall_eq_empty.2 hr] - exact subsingleton_empty - · rw [closedBall_zero] - exact subsingleton_singleton - -theorem subsingleton_sphere (x : γ) {r : ℝ} (hr : r ≤ 0) : (sphere x r).Subsingleton := - (subsingleton_closedBall x hr).anti sphere_subset_closedBall - -- see Note [lower instance priority] instance (priority := 100) _root_.MetricSpace.instT0Space : T0Space γ where t0 _ _ h := eq_of_dist_eq_zero <| Metric.inseparable_iff.1 h @@ -152,32 +67,6 @@ theorem uniformEmbedding_bot_of_pairwise_le_dist {β : Type*} {ε : ℝ} (hε : end Metric -/-- Build a new metric space from an old one where the bundled uniform structure is provably -(but typically non-definitionaly) equal to some given uniform structure. -See Note [forgetful inheritance]. --/ -def MetricSpace.replaceUniformity {γ} [U : UniformSpace γ] (m : MetricSpace γ) - (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : MetricSpace γ where - toPseudoMetricSpace := PseudoMetricSpace.replaceUniformity m.toPseudoMetricSpace H - eq_of_dist_eq_zero := @eq_of_dist_eq_zero _ _ - -theorem MetricSpace.replaceUniformity_eq {γ} [U : UniformSpace γ] (m : MetricSpace γ) - (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : m.replaceUniformity H = m := by - ext; rfl - -/-- Build a new metric space from an old one where the bundled topological structure is provably -(but typically non-definitionaly) equal to some given topological structure. -See Note [forgetful inheritance]. --/ -abbrev MetricSpace.replaceTopology {γ} [U : TopologicalSpace γ] (m : MetricSpace γ) - (H : U = m.toPseudoMetricSpace.toUniformSpace.toTopologicalSpace) : MetricSpace γ := - @MetricSpace.replaceUniformity γ (m.toUniformSpace.replaceTopology H) m rfl - -theorem MetricSpace.replaceTopology_eq {γ} [U : TopologicalSpace γ] (m : MetricSpace γ) - (H : U = m.toPseudoMetricSpace.toUniformSpace.toTopologicalSpace) : - m.replaceTopology H = m := by - ext; rfl - /-- One gets a metric space from an emetric space if the edistance is everywhere finite, by pushing the edistance to reals. We set it up so that the edist and the uniformity are defeq in the metric space and the emetric space. In this definition, the distance @@ -196,20 +85,6 @@ def EMetricSpace.toMetricSpace {α : Type u} [EMetricSpace α] (h : ∀ x y : α MetricSpace α := EMetricSpace.toMetricSpaceOfDist (fun x y => ENNReal.toReal (edist x y)) h fun _ _ => rfl -/-- Build a new metric space from an old one where the bundled bornology structure is provably -(but typically non-definitionaly) equal to some given bornology structure. -See Note [forgetful inheritance]. --/ -def MetricSpace.replaceBornology {α} [B : Bornology α] (m : MetricSpace α) - (H : ∀ s, @IsBounded _ B s ↔ @IsBounded _ PseudoMetricSpace.toBornology s) : MetricSpace α := - { PseudoMetricSpace.replaceBornology _ H, m with toBornology := B } - -theorem MetricSpace.replaceBornology_eq {α} [m : MetricSpace α] [B : Bornology α] - (H : ∀ s, @IsBounded _ B s ↔ @IsBounded _ PseudoMetricSpace.toBornology s) : - MetricSpace.replaceBornology _ H = m := by - ext - rfl - /-- Metric space structure pulled back by an injective function. Injectivity is necessary to ensure that `dist x y = 0` only if `x = y`. -/ abbrev MetricSpace.induced {γ β} (f : γ → β) (hf : Function.Injective f) (m : MetricSpace β) : @@ -237,27 +112,6 @@ instance Subtype.metricSpace {α : Type*} {p : α → Prop} [MetricSpace α] : instance {α : Type*} [MetricSpace α] : MetricSpace αᵐᵒᵖ := MetricSpace.induced MulOpposite.unop MulOpposite.unop_injective ‹_› -instance : MetricSpace Empty where - dist _ _ := 0 - dist_self _ := rfl - dist_comm _ _ := rfl - edist _ _ := 0 - eq_of_dist_eq_zero _ := Subsingleton.elim _ _ - dist_triangle _ _ _ := show (0 : ℝ) ≤ 0 + 0 by rw [add_zero] - toUniformSpace := inferInstance - uniformity_dist := Subsingleton.elim _ _ - -instance : MetricSpace PUnit.{u + 1} where - dist _ _ := 0 - dist_self _ := rfl - dist_comm _ _ := rfl - edist _ _ := 0 - eq_of_dist_eq_zero _ := Subsingleton.elim _ _ - dist_triangle _ _ _ := show (0 : ℝ) ≤ 0 + 0 by rw [add_zero] - toUniformSpace := inferInstance - uniformity_dist := by - simp (config := { contextual := true }) [principal_univ, eq_top_of_neBot (𝓤 PUnit)] - section Real /-- Instantiate the reals as a metric space. -/ @@ -342,44 +196,8 @@ end EqRel The distance on those type synonyms is inherited without change. -/ - open Additive Multiplicative -section - -variable [Dist X] - -instance : Dist (Additive X) := ‹Dist X› -instance : Dist (Multiplicative X) := ‹Dist X› - -@[simp] theorem dist_ofMul (a b : X) : dist (ofMul a) (ofMul b) = dist a b := rfl - -@[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_toAdd (a b : Multiplicative X) : dist (toAdd a) (toAdd b) = dist a b := rfl - -end - -section - -variable [PseudoMetricSpace X] - -instance : PseudoMetricSpace (Additive X) := ‹PseudoMetricSpace X› -instance : PseudoMetricSpace (Multiplicative X) := ‹PseudoMetricSpace X› - -@[simp] theorem nndist_ofMul (a b : X) : nndist (ofMul a) (ofMul b) = nndist a b := rfl - -@[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_toAdd (a b : Multiplicative X) : nndist (toAdd a) (toAdd b) = nndist a b := rfl - -end - instance [MetricSpace X] : MetricSpace (Additive X) := ‹MetricSpace X› instance [MetricSpace X] : MetricSpace (Multiplicative X) := ‹MetricSpace X› @@ -394,28 +212,4 @@ The distance on this type synonym is inherited without change. open OrderDual -section - -variable [Dist X] - -instance : Dist Xᵒᵈ := ‹Dist X› - -@[simp] theorem dist_toDual (a b : X) : dist (toDual a) (toDual b) = dist a b := rfl - -@[simp] theorem dist_ofDual (a b : Xᵒᵈ) : dist (ofDual a) (ofDual b) = dist a b := rfl - -end - -section - -variable [PseudoMetricSpace X] - -instance : PseudoMetricSpace Xᵒᵈ := ‹PseudoMetricSpace X› - -@[simp] theorem nndist_toDual (a b : X) : nndist (toDual a) (toDual b) = nndist a b := rfl - -@[simp] theorem nndist_ofDual (a b : Xᵒᵈ) : nndist (ofDual a) (ofDual b) = nndist a b := rfl - -end - instance [MetricSpace X] : MetricSpace Xᵒᵈ := ‹MetricSpace X› diff --git a/Mathlib/Topology/MetricSpace/Bounded.lean b/Mathlib/Topology/MetricSpace/Bounded.lean index 19380e29c112e..0bbc10d2bcc9f 100644 --- a/Mathlib/Topology/MetricSpace/Bounded.lean +++ b/Mathlib/Topology/MetricSpace/Bounded.lean @@ -6,6 +6,7 @@ Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébas import Mathlib.Topology.Algebra.Order.Compact import Mathlib.Topology.MetricSpace.ProperSpace import Mathlib.Topology.MetricSpace.Cauchy +import Mathlib.Topology.EMetricSpace.Diam /-! ## Boundedness in (pseudo)-metric spaces diff --git a/Mathlib/Topology/MetricSpace/Cauchy.lean b/Mathlib/Topology/MetricSpace/Cauchy.lean index f6af5f9a89078..7a51a0300765e 100644 --- a/Mathlib/Topology/MetricSpace/Cauchy.lean +++ b/Mathlib/Topology/MetricSpace/Cauchy.lean @@ -4,6 +4,7 @@ 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.Topology.MetricSpace.Pseudo.Lemmas +import Mathlib.Topology.EMetricSpace.Basic /-! ## Cauchy sequences in (pseudo-)metric spaces diff --git a/Mathlib/Topology/MetricSpace/Defs.lean b/Mathlib/Topology/MetricSpace/Defs.lean new file mode 100644 index 0000000000000..fa1af58dc2f89 --- /dev/null +++ b/Mathlib/Topology/MetricSpace/Defs.lean @@ -0,0 +1,253 @@ +/- +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.Topology.MetricSpace.Pseudo.Defs + +/-! +# Metric spaces + +This file defines metric spaces and shows some of their basic properties. + +Many definitions and theorems expected on metric spaces are already introduced on uniform spaces and +topological spaces. This includes open and closed sets, compactness, completeness, continuity +and uniform continuity. + +TODO (anyone): Add "Main results" section. + +## Implementation notes +A lot of elementary properties don't require `eq_of_dist_eq_zero`, hence are stated and proven +for `PseudoMetricSpace`s in `PseudoMetric.lean`. + +## Tags + +metric, pseudo_metric, dist +-/ + +open Set Filter Bornology +open scoped NNReal Uniformity + +universe u v w + +variable {α : Type u} {β : Type v} {X ι : Type*} +variable [PseudoMetricSpace α] + +/-- We now define `MetricSpace`, extending `PseudoMetricSpace`. -/ +class MetricSpace (α : Type u) extends PseudoMetricSpace α : Type u where + eq_of_dist_eq_zero : ∀ {x y : α}, dist x y = 0 → x = y + +/-- Two metric space structures with the same distance coincide. -/ +@[ext] +theorem MetricSpace.ext {α : Type*} {m m' : MetricSpace α} (h : m.toDist = m'.toDist) : + m = m' := by + cases m; cases m'; congr; ext1; assumption + +/-- Construct a metric space structure whose underlying topological space structure +(definitionally) agrees which a pre-existing topology which is compatible with a given distance +function. -/ +def MetricSpace.ofDistTopology {α : Type u} [TopologicalSpace α] (dist : α → α → ℝ) + (dist_self : ∀ x : α, dist x x = 0) (dist_comm : ∀ x y : α, dist x y = dist y x) + (dist_triangle : ∀ x y z : α, dist x z ≤ dist x y + dist y z) + (H : ∀ s : Set α, IsOpen s ↔ ∀ x ∈ s, ∃ ε > 0, ∀ y, dist x y < ε → y ∈ s) + (eq_of_dist_eq_zero : ∀ x y : α, dist x y = 0 → x = y) : MetricSpace α := + { PseudoMetricSpace.ofDistTopology dist dist_self dist_comm dist_triangle H with + eq_of_dist_eq_zero := eq_of_dist_eq_zero _ _ } + +variable {γ : Type w} [MetricSpace γ] + +theorem eq_of_dist_eq_zero {x y : γ} : dist x y = 0 → x = y := + MetricSpace.eq_of_dist_eq_zero + +@[simp] +theorem dist_eq_zero {x y : γ} : dist x y = 0 ↔ x = y := + Iff.intro eq_of_dist_eq_zero fun this => this ▸ dist_self _ + +@[simp] +theorem zero_eq_dist {x y : γ} : 0 = dist x y ↔ x = y := by rw [eq_comm, dist_eq_zero] + +theorem dist_ne_zero {x y : γ} : dist x y ≠ 0 ↔ x ≠ y := by + simpa only [not_iff_not] using dist_eq_zero + +@[simp] +theorem dist_le_zero {x y : γ} : dist x y ≤ 0 ↔ x = y := by + simpa [le_antisymm_iff, dist_nonneg] using @dist_eq_zero _ _ x y + +@[simp] +theorem dist_pos {x y : γ} : 0 < dist x y ↔ x ≠ y := by + simpa only [not_le] using not_congr dist_le_zero + +theorem eq_of_forall_dist_le {x y : γ} (h : ∀ ε > 0, dist x y ≤ ε) : x = y := + eq_of_dist_eq_zero (eq_of_le_of_forall_le_of_dense dist_nonneg h) + +/-- Deduce the equality of points from the vanishing of the nonnegative distance-/ +theorem eq_of_nndist_eq_zero {x y : γ} : nndist x y = 0 → x = y := by + simp only [NNReal.eq_iff, ← dist_nndist, imp_self, NNReal.coe_zero, dist_eq_zero] + +/-- Characterize the equality of points as the vanishing of the nonnegative distance-/ +@[simp] +theorem nndist_eq_zero {x y : γ} : nndist x y = 0 ↔ x = y := by + simp only [NNReal.eq_iff, ← dist_nndist, imp_self, NNReal.coe_zero, dist_eq_zero] + +@[simp] +theorem zero_eq_nndist {x y : γ} : 0 = nndist x y ↔ x = y := by + simp only [NNReal.eq_iff, ← dist_nndist, imp_self, NNReal.coe_zero, zero_eq_dist] + +namespace Metric + +variable {x : γ} {s : Set γ} + +@[simp] theorem closedBall_zero : closedBall x 0 = {x} := Set.ext fun _ => dist_le_zero + +@[simp] theorem sphere_zero : sphere x 0 = {x} := Set.ext fun _ => dist_eq_zero + +theorem subsingleton_closedBall (x : γ) {r : ℝ} (hr : r ≤ 0) : (closedBall x r).Subsingleton := by + rcases hr.lt_or_eq with (hr | rfl) + · rw [closedBall_eq_empty.2 hr] + exact subsingleton_empty + · rw [closedBall_zero] + exact subsingleton_singleton + +theorem subsingleton_sphere (x : γ) {r : ℝ} (hr : r ≤ 0) : (sphere x r).Subsingleton := + (subsingleton_closedBall x hr).anti sphere_subset_closedBall + +end Metric + +/-- Build a new metric space from an old one where the bundled uniform structure is provably +(but typically non-definitionaly) equal to some given uniform structure. +See Note [forgetful inheritance]. +-/ +def MetricSpace.replaceUniformity {γ} [U : UniformSpace γ] (m : MetricSpace γ) + (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : MetricSpace γ where + toPseudoMetricSpace := PseudoMetricSpace.replaceUniformity m.toPseudoMetricSpace H + eq_of_dist_eq_zero := @eq_of_dist_eq_zero _ _ + +theorem MetricSpace.replaceUniformity_eq {γ} [U : UniformSpace γ] (m : MetricSpace γ) + (H : 𝓤[U] = 𝓤[PseudoEMetricSpace.toUniformSpace]) : m.replaceUniformity H = m := by + ext; rfl + +/-- Build a new metric space from an old one where the bundled topological structure is provably +(but typically non-definitionaly) equal to some given topological structure. +See Note [forgetful inheritance]. +-/ +abbrev MetricSpace.replaceTopology {γ} [U : TopologicalSpace γ] (m : MetricSpace γ) + (H : U = m.toPseudoMetricSpace.toUniformSpace.toTopologicalSpace) : MetricSpace γ := + @MetricSpace.replaceUniformity γ (m.toUniformSpace.replaceTopology H) m rfl + +theorem MetricSpace.replaceTopology_eq {γ} [U : TopologicalSpace γ] (m : MetricSpace γ) + (H : U = m.toPseudoMetricSpace.toUniformSpace.toTopologicalSpace) : + m.replaceTopology H = m := by + ext; rfl + +/-- Build a new metric space from an old one where the bundled bornology structure is provably +(but typically non-definitionaly) equal to some given bornology structure. +See Note [forgetful inheritance]. +-/ +def MetricSpace.replaceBornology {α} [B : Bornology α] (m : MetricSpace α) + (H : ∀ s, @IsBounded _ B s ↔ @IsBounded _ PseudoMetricSpace.toBornology s) : MetricSpace α := + { PseudoMetricSpace.replaceBornology _ H, m with toBornology := B } + +theorem MetricSpace.replaceBornology_eq {α} [m : MetricSpace α] [B : Bornology α] + (H : ∀ s, @IsBounded _ B s ↔ @IsBounded _ PseudoMetricSpace.toBornology s) : + MetricSpace.replaceBornology _ H = m := by + ext + rfl + +instance : MetricSpace Empty where + dist _ _ := 0 + dist_self _ := rfl + dist_comm _ _ := rfl + edist _ _ := 0 + eq_of_dist_eq_zero _ := Subsingleton.elim _ _ + dist_triangle _ _ _ := show (0 : ℝ) ≤ 0 + 0 by rw [add_zero] + toUniformSpace := inferInstance + uniformity_dist := Subsingleton.elim _ _ + +instance : MetricSpace PUnit.{u + 1} where + dist _ _ := 0 + dist_self _ := rfl + dist_comm _ _ := rfl + edist _ _ := 0 + eq_of_dist_eq_zero _ := Subsingleton.elim _ _ + dist_triangle _ _ _ := show (0 : ℝ) ≤ 0 + 0 by rw [add_zero] + toUniformSpace := inferInstance + uniformity_dist := by + simp (config := { contextual := true }) [principal_univ, eq_top_of_neBot (𝓤 PUnit)] + +/-! +### `Additive`, `Multiplicative` + +The distance on those type synonyms is inherited without change. +-/ + + +open Additive Multiplicative + +section + +variable [Dist X] + +instance : Dist (Additive X) := ‹Dist X› +instance : Dist (Multiplicative X) := ‹Dist X› + +@[simp] theorem dist_ofMul (a b : X) : dist (ofMul a) (ofMul b) = dist a b := rfl + +@[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_toAdd (a b : Multiplicative X) : dist (toAdd a) (toAdd b) = dist a b := rfl + +end + +section + +variable [PseudoMetricSpace X] + +@[simp] theorem nndist_ofMul (a b : X) : nndist (ofMul a) (ofMul b) = nndist a b := rfl + +@[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_toAdd (a b : Multiplicative X) : nndist (toAdd a) (toAdd b) = nndist a b := rfl + +end + +instance [MetricSpace X] : MetricSpace (Additive X) := ‹MetricSpace X› +instance [MetricSpace X] : MetricSpace (Multiplicative X) := ‹MetricSpace X› + +/-! +### Order dual + +The distance on this type synonym is inherited without change. +-/ + +open OrderDual + +section + +variable [Dist X] + +instance : Dist Xᵒᵈ := ‹Dist X› + +@[simp] theorem dist_toDual (a b : X) : dist (toDual a) (toDual b) = dist a b := rfl + +@[simp] theorem dist_ofDual (a b : Xᵒᵈ) : dist (ofDual a) (ofDual b) = dist a b := rfl + +end + +section + +variable [PseudoMetricSpace X] + +instance : PseudoMetricSpace Xᵒᵈ := ‹PseudoMetricSpace X› + +@[simp] theorem nndist_toDual (a b : X) : nndist (toDual a) (toDual b) = nndist a b := rfl + +@[simp] theorem nndist_ofDual (a b : Xᵒᵈ) : nndist (ofDual a) (ofDual b) = nndist a b := rfl + +end + +instance [MetricSpace X] : MetricSpace Xᵒᵈ := ‹MetricSpace X› diff --git a/Mathlib/Topology/MetricSpace/Equicontinuity.lean b/Mathlib/Topology/MetricSpace/Equicontinuity.lean index 6f24d5e53aa5a..a6b57e7e49085 100644 --- a/Mathlib/Topology/MetricSpace/Equicontinuity.lean +++ b/Mathlib/Topology/MetricSpace/Equicontinuity.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ -import Mathlib.Topology.MetricSpace.Pseudo.Lemmas import Mathlib.Topology.UniformSpace.Equicontinuity +import Mathlib.Topology.MetricSpace.Pseudo.Lemmas /-! # Equicontinuity in metric spaces diff --git a/Mathlib/Topology/MetricSpace/MetricSeparated.lean b/Mathlib/Topology/MetricSpace/MetricSeparated.lean index 61ef5d09400f4..2529db5502ed7 100644 --- a/Mathlib/Topology/MetricSpace/MetricSeparated.lean +++ b/Mathlib/Topology/MetricSpace/MetricSeparated.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Topology.EMetricSpace.Basic +import Mathlib.Topology.EMetricSpace.Defs /-! # Metric separated pairs of sets diff --git a/Mathlib/Topology/MetricSpace/ProperSpace.lean b/Mathlib/Topology/MetricSpace/ProperSpace.lean index b44aeb37c523e..e8df38ca5e7a0 100644 --- a/Mathlib/Topology/MetricSpace/ProperSpace.lean +++ b/Mathlib/Topology/MetricSpace/ProperSpace.lean @@ -4,9 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Topology.MetricSpace.Pseudo.Lemmas import Mathlib.Topology.Order.IsLUB import Mathlib.Topology.Support +import Mathlib.Topology.MetricSpace.Pseudo.Basic +import Mathlib.Topology.MetricSpace.Pseudo.Lemmas +import Mathlib.Topology.MetricSpace.Pseudo.Pi /-! ## Proper spaces diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean b/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean new file mode 100644 index 0000000000000..5793a4509ed7c --- /dev/null +++ b/Mathlib/Topology/MetricSpace/Pseudo/Basic.lean @@ -0,0 +1,231 @@ +/- +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.ENNReal.Real +import Mathlib.Tactic.Bound.Attribute +import Mathlib.Topology.EMetricSpace.Basic +import Mathlib.Topology.MetricSpace.Pseudo.Defs + +/-! +## Pseudo-metric spaces + +Further results about pseudo-metric spaces. + +-/ + +open Set Filter TopologicalSpace Bornology +open scoped ENNReal NNReal Uniformity Topology + +universe u v w + +variable {α : Type u} {β : Type v} {X ι : Type*} + +variable [PseudoMetricSpace α] + +/-- The triangle (polygon) inequality for sequences of points; `Finset.Ico` version. -/ +theorem dist_le_Ico_sum_dist (f : ℕ → α) {m n} (h : m ≤ n) : + dist (f m) (f n) ≤ ∑ i ∈ Finset.Ico m n, dist (f i) (f (i + 1)) := by + induction n, h using Nat.le_induction with + | base => rw [Finset.Ico_self, Finset.sum_empty, dist_self] + | succ n hle ihn => + calc + dist (f m) (f (n + 1)) ≤ dist (f m) (f n) + dist (f n) (f (n + 1)) := dist_triangle _ _ _ + _ ≤ (∑ i ∈ Finset.Ico m n, _) + _ := add_le_add ihn le_rfl + _ = ∑ i ∈ Finset.Ico m (n + 1), _ := by + { rw [Nat.Ico_succ_right_eq_insert_Ico hle, Finset.sum_insert, add_comm]; simp } + +/-- The triangle (polygon) inequality for sequences of points; `Finset.range` version. -/ +theorem dist_le_range_sum_dist (f : ℕ → α) (n : ℕ) : + dist (f 0) (f n) ≤ ∑ i ∈ Finset.range n, dist (f i) (f (i + 1)) := + Nat.Ico_zero_eq_range ▸ dist_le_Ico_sum_dist f (Nat.zero_le n) + +/-- A version of `dist_le_Ico_sum_dist` with each intermediate distance replaced +with an upper estimate. -/ +theorem dist_le_Ico_sum_of_dist_le {f : ℕ → α} {m n} (hmn : m ≤ n) {d : ℕ → ℝ} + (hd : ∀ {k}, m ≤ k → k < n → dist (f k) (f (k + 1)) ≤ d k) : + dist (f m) (f n) ≤ ∑ i ∈ Finset.Ico m n, d i := + le_trans (dist_le_Ico_sum_dist f hmn) <| + Finset.sum_le_sum fun _k hk => hd (Finset.mem_Ico.1 hk).1 (Finset.mem_Ico.1 hk).2 + +/-- A version of `dist_le_range_sum_dist` with each intermediate distance replaced +with an upper estimate. -/ +theorem dist_le_range_sum_of_dist_le {f : ℕ → α} (n : ℕ) {d : ℕ → ℝ} + (hd : ∀ {k}, k < n → dist (f k) (f (k + 1)) ≤ d k) : + dist (f 0) (f n) ≤ ∑ i ∈ Finset.range n, d i := + Nat.Ico_zero_eq_range ▸ dist_le_Ico_sum_of_dist_le (zero_le n) fun _ => hd + +namespace Metric + +-- instantiate pseudometric space as a topology +variable {x y z : α} {δ ε ε₁ ε₂ : ℝ} {s : Set α} + +nonrec theorem uniformInducing_iff [PseudoMetricSpace β] {f : α → β} : + UniformInducing f ↔ UniformContinuous f ∧ + ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := + uniformInducing_iff'.trans <| Iff.rfl.and <| + ((uniformity_basis_dist.comap _).le_basis_iff uniformity_basis_dist).trans <| by + simp only [subset_def, Prod.forall, gt_iff_lt, preimage_setOf_eq, Prod.map_apply, mem_setOf] + +nonrec theorem uniformEmbedding_iff [PseudoMetricSpace β] {f : α → β} : + UniformEmbedding f ↔ Function.Injective f ∧ UniformContinuous f ∧ + ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := by + rw [uniformEmbedding_iff, and_comm, uniformInducing_iff] + +/-- If a map between pseudometric spaces is a uniform embedding then the distance between `f x` +and `f y` is controlled in terms of the distance between `x` and `y`. -/ +theorem controlled_of_uniformEmbedding [PseudoMetricSpace β] {f : α → β} (h : UniformEmbedding f) : + (∀ ε > 0, ∃ δ > 0, ∀ {a b : α}, dist a b < δ → dist (f a) (f b) < ε) ∧ + ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := + ⟨uniformContinuous_iff.1 h.uniformContinuous, (uniformEmbedding_iff.1 h).2.2⟩ + +theorem totallyBounded_iff {s : Set α} : + TotallyBounded s ↔ ∀ ε > 0, ∃ t : Set α, t.Finite ∧ s ⊆ ⋃ y ∈ t, ball y ε := + uniformity_basis_dist.totallyBounded_iff + +/-- A pseudometric space is totally bounded if one can reconstruct up to any ε>0 any element of the +space from finitely many data. -/ +theorem totallyBounded_of_finite_discretization {s : Set α} + (H : ∀ ε > (0 : ℝ), + ∃ (β : Type u) (_ : Fintype β) (F : s → β), ∀ x y, F x = F y → dist (x : α) y < ε) : + TotallyBounded s := by + rcases s.eq_empty_or_nonempty with hs | hs + · rw [hs] + exact totallyBounded_empty + rcases hs with ⟨x0, hx0⟩ + haveI : Inhabited s := ⟨⟨x0, hx0⟩⟩ + refine totallyBounded_iff.2 fun ε ε0 => ?_ + rcases H ε ε0 with ⟨β, fβ, F, hF⟩ + let Finv := Function.invFun F + refine ⟨range (Subtype.val ∘ Finv), finite_range _, fun x xs => ?_⟩ + let x' := Finv (F ⟨x, xs⟩) + have : F x' = F ⟨x, xs⟩ := Function.invFun_eq ⟨⟨x, xs⟩, rfl⟩ + simp only [Set.mem_iUnion, Set.mem_range] + exact ⟨_, ⟨F ⟨x, xs⟩, rfl⟩, hF _ _ this.symm⟩ + +theorem finite_approx_of_totallyBounded {s : Set α} (hs : TotallyBounded s) : + ∀ ε > 0, ∃ t, t ⊆ s ∧ Set.Finite t ∧ s ⊆ ⋃ y ∈ t, ball y ε := by + intro ε ε_pos + rw [totallyBounded_iff_subset] at hs + exact hs _ (dist_mem_uniformity ε_pos) + +/-- Expressing uniform convergence using `dist` -/ +theorem tendstoUniformlyOnFilter_iff {F : ι → β → α} {f : β → α} {p : Filter ι} {p' : Filter β} : + TendstoUniformlyOnFilter F f p p' ↔ + ∀ ε > 0, ∀ᶠ n : ι × β in p ×ˢ p', dist (f n.snd) (F n.fst n.snd) < ε := by + refine ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu => ?_⟩ + rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ + exact (H ε εpos).mono fun n hn => hε hn + +/-- Expressing locally uniform convergence on a set using `dist`. -/ +theorem tendstoLocallyUniformlyOn_iff [TopologicalSpace β] {F : ι → β → α} {f : β → α} + {p : Filter ι} {s : Set β} : + TendstoLocallyUniformlyOn F f p s ↔ + ∀ ε > 0, ∀ x ∈ s, ∃ t ∈ 𝓝[s] x, ∀ᶠ n in p, ∀ y ∈ t, dist (f y) (F n y) < ε := by + refine ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu x hx => ?_⟩ + rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ + rcases H ε εpos x hx with ⟨t, ht, Ht⟩ + exact ⟨t, ht, Ht.mono fun n hs x hx => hε (hs x hx)⟩ + +/-- Expressing uniform convergence on a set using `dist`. -/ +theorem tendstoUniformlyOn_iff {F : ι → β → α} {f : β → α} {p : Filter ι} {s : Set β} : + TendstoUniformlyOn F f p s ↔ ∀ ε > 0, ∀ᶠ n in p, ∀ x ∈ s, dist (f x) (F n x) < ε := by + refine ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu => ?_⟩ + rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ + exact (H ε εpos).mono fun n hs x hx => hε (hs x hx) + +/-- Expressing locally uniform convergence using `dist`. -/ +theorem tendstoLocallyUniformly_iff [TopologicalSpace β] {F : ι → β → α} {f : β → α} + {p : Filter ι} : + TendstoLocallyUniformly F f p ↔ + ∀ ε > 0, ∀ x : β, ∃ t ∈ 𝓝 x, ∀ᶠ n in p, ∀ y ∈ t, dist (f y) (F n y) < ε := by + simp only [← tendstoLocallyUniformlyOn_univ, tendstoLocallyUniformlyOn_iff, nhdsWithin_univ, + mem_univ, forall_const, exists_prop] + +/-- Expressing uniform convergence using `dist`. -/ +theorem tendstoUniformly_iff {F : ι → β → α} {f : β → α} {p : Filter ι} : + TendstoUniformly F f p ↔ ∀ ε > 0, ∀ᶠ n in p, ∀ x, dist (f x) (F n x) < ε := by + rw [← tendstoUniformlyOn_univ, tendstoUniformlyOn_iff] + simp + +protected theorem cauchy_iff {f : Filter α} : + Cauchy f ↔ NeBot f ∧ ∀ ε > 0, ∃ t ∈ f, ∀ x ∈ t, ∀ y ∈ t, dist x y < ε := + uniformity_basis_dist.cauchy_iff + +/-- Given a point `x` in a discrete subset `s` of a pseudometric space, there is an open ball +centered at `x` and intersecting `s` only at `x`. -/ +theorem exists_ball_inter_eq_singleton_of_mem_discrete [DiscreteTopology s] {x : α} (hx : x ∈ s) : + ∃ ε > 0, Metric.ball x ε ∩ s = {x} := + nhds_basis_ball.exists_inter_eq_singleton_of_mem_discrete hx + +/-- Given a point `x` in a discrete subset `s` of a pseudometric space, there is a closed ball +of positive radius centered at `x` and intersecting `s` only at `x`. -/ +theorem exists_closedBall_inter_eq_singleton_of_discrete [DiscreteTopology s] {x : α} (hx : x ∈ s) : + ∃ ε > 0, Metric.closedBall x ε ∩ s = {x} := + nhds_basis_closedBall.exists_inter_eq_singleton_of_mem_discrete hx + +end Metric + +open Metric + +theorem Metric.inseparable_iff {x y : α} : Inseparable x y ↔ dist x y = 0 := by + rw [EMetric.inseparable_iff, edist_nndist, dist_nndist, ENNReal.coe_eq_zero, NNReal.coe_eq_zero] + +section Real + +theorem cauchySeq_iff_tendsto_dist_atTop_0 [Nonempty β] [SemilatticeSup β] {u : β → α} : + CauchySeq u ↔ Tendsto (fun n : β × β => dist (u n.1) (u n.2)) atTop (𝓝 0) := by + rw [cauchySeq_iff_tendsto, Metric.uniformity_eq_comap_nhds_zero, tendsto_comap_iff, + Function.comp_def] + simp_rw [Prod.map_fst, Prod.map_snd] + +end Real + +namespace Metric + +variable {x y z : α} {ε ε₁ ε₂ : ℝ} {s : Set α} + +-- Porting note: `TopologicalSpace.IsSeparable.separableSpace` moved to `EMetricSpace` + +/-- The preimage of a separable set by an inducing map is separable. -/ +protected theorem _root_.Inducing.isSeparable_preimage {f : β → α} [TopologicalSpace β] + (hf : Inducing f) {s : Set α} (hs : IsSeparable s) : IsSeparable (f ⁻¹' s) := by + have : SeparableSpace s := hs.separableSpace + have : SecondCountableTopology s := UniformSpace.secondCountable_of_separable _ + have : Inducing ((mapsTo_preimage f s).restrict _ _ _) := + (hf.comp inducing_subtype_val).codRestrict _ + have := this.secondCountableTopology + exact .of_subtype _ + +protected theorem _root_.Embedding.isSeparable_preimage {f : β → α} [TopologicalSpace β] + (hf : Embedding f) {s : Set α} (hs : IsSeparable s) : IsSeparable (f ⁻¹' s) := + hf.toInducing.isSeparable_preimage hs + +end Metric + +/-- A compact set is separable. -/ +theorem IsCompact.isSeparable {s : Set α} (hs : IsCompact s) : IsSeparable s := + haveI : CompactSpace s := isCompact_iff_compactSpace.mp hs + .of_subtype s + +namespace Metric + +section SecondCountable + +open TopologicalSpace + +/-- A pseudometric space is second countable if, for every `ε > 0`, there is a countable set which +is `ε`-dense. -/ +theorem secondCountable_of_almost_dense_set + (H : ∀ ε > (0 : ℝ), ∃ s : Set α, s.Countable ∧ ∀ x, ∃ y ∈ s, dist x y ≤ ε) : + SecondCountableTopology α := by + refine EMetric.secondCountable_of_almost_dense_set fun ε ε0 => ?_ + rcases ENNReal.lt_iff_exists_nnreal_btwn.1 ε0 with ⟨ε', ε'0, ε'ε⟩ + choose s hsc y hys hyx using H ε' (mod_cast ε'0) + refine ⟨s, hsc, iUnion₂_eq_univ_iff.2 fun x => ⟨y x, hys _, le_trans ?_ ε'ε.le⟩⟩ + exact mod_cast hyx x + +end SecondCountable + +end Metric diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean b/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean index 06c84c99cc8fb..3ad6637ff4565 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Constructions.lean @@ -5,11 +5,12 @@ Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébas -/ import Mathlib.Topology.Bornology.Constructions import Mathlib.Topology.MetricSpace.Pseudo.Defs +import Mathlib.Topology.UniformSpace.UniformEmbedding /-! -# Product of pseudo-metric spaces and other constructions +# Products of pseudometric spaces and other constructions -This file constructs the infinity distance on finite products of normed groups and provides +This file constructs the supremum distance on binary products of pseudometric spaces and provides instances for type synonyms. -/ @@ -232,152 +233,3 @@ protected lemma Filter.Tendsto.nndist {f g : β → α} {x : Filter β} {a b : (hf : Tendsto f x (𝓝 a)) (hg : Tendsto g x (𝓝 b)) : Tendsto (fun x => nndist (f x) (g x)) x (𝓝 (nndist a b)) := (continuous_nndist.tendsto (a, b)).comp (hf.prod_mk_nhds hg) - -section Pi - -open Finset - -variable {π : β → Type*} [Fintype β] [∀ b, PseudoMetricSpace (π b)] - -/-- A finite product of pseudometric spaces is a pseudometric space, with the sup distance. -/ -instance pseudoMetricSpacePi : PseudoMetricSpace (∀ b, π b) := by - /- we construct the instance from the pseudoemetric space instance to avoid checking again that - the uniformity is the same as the product uniformity, but we register nevertheless a nice - formula for the distance -/ - let i := PseudoEMetricSpace.toPseudoMetricSpaceOfDist - (fun f g : ∀ b, π b => ((sup univ fun b => nndist (f b) (g b) : ℝ≥0) : ℝ)) - (fun f g => ((Finset.sup_lt_iff bot_lt_top).2 fun b _ => edist_lt_top _ _).ne) - (fun f g => by - simp only [edist_pi_def, edist_nndist, ← ENNReal.coe_finset_sup, ENNReal.coe_toReal]) - refine i.replaceBornology fun s => ?_ - simp only [← isBounded_def, isBounded_iff_eventually, ← forall_isBounded_image_eval_iff, - forall_mem_image, ← Filter.eventually_all, Function.eval_apply, @dist_nndist (π _)] - refine eventually_congr ((eventually_ge_atTop 0).mono fun C hC ↦ ?_) - lift C to ℝ≥0 using hC - refine ⟨fun H x hx y hy ↦ NNReal.coe_le_coe.2 <| Finset.sup_le fun b _ ↦ H b hx hy, - fun H b x hx y hy ↦ NNReal.coe_le_coe.2 ?_⟩ - simpa only using Finset.sup_le_iff.1 (NNReal.coe_le_coe.1 <| H hx hy) b (Finset.mem_univ b) - -lemma nndist_pi_def (f g : ∀ b, π b) : nndist f g = sup univ fun b => nndist (f b) (g b) := rfl - -lemma dist_pi_def (f g : ∀ b, π b) : dist f g = (sup univ fun b => nndist (f b) (g b) : ℝ≥0) := rfl - -lemma nndist_pi_le_iff {f g : ∀ b, π b} {r : ℝ≥0} : - nndist f g ≤ r ↔ ∀ b, nndist (f b) (g b) ≤ r := by simp [nndist_pi_def] - -lemma nndist_pi_lt_iff {f g : ∀ b, π b} {r : ℝ≥0} (hr : 0 < r) : - nndist f g < r ↔ ∀ b, nndist (f b) (g b) < r := by - rw [← bot_eq_zero'] at hr - simp [nndist_pi_def, Finset.sup_lt_iff hr] - -lemma nndist_pi_eq_iff {f g : ∀ b, π b} {r : ℝ≥0} (hr : 0 < r) : - nndist f g = r ↔ (∃ i, nndist (f i) (g i) = r) ∧ ∀ b, nndist (f b) (g b) ≤ r := by - rw [eq_iff_le_not_lt, nndist_pi_lt_iff hr, nndist_pi_le_iff, not_forall, and_comm] - simp_rw [not_lt, and_congr_left_iff, le_antisymm_iff] - intro h - refine exists_congr fun b => ?_ - apply (and_iff_right <| h _).symm - -lemma dist_pi_lt_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 < r) : - dist f g < r ↔ ∀ b, dist (f b) (g b) < r := by - lift r to ℝ≥0 using hr.le - exact nndist_pi_lt_iff hr - -lemma dist_pi_le_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 ≤ r) : - dist f g ≤ r ↔ ∀ b, dist (f b) (g b) ≤ r := by - lift r to ℝ≥0 using hr - exact nndist_pi_le_iff - -lemma dist_pi_eq_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 < r) : - dist f g = r ↔ (∃ i, dist (f i) (g i) = r) ∧ ∀ b, dist (f b) (g b) ≤ r := by - lift r to ℝ≥0 using hr.le - simp_rw [← coe_nndist, NNReal.coe_inj, nndist_pi_eq_iff hr, NNReal.coe_le_coe] - -lemma dist_pi_le_iff' [Nonempty β] {f g : ∀ b, π b} {r : ℝ} : - dist f g ≤ r ↔ ∀ b, dist (f b) (g b) ≤ r := by - by_cases hr : 0 ≤ r - · exact dist_pi_le_iff hr - · exact iff_of_false (fun h => hr <| dist_nonneg.trans h) fun h => - hr <| dist_nonneg.trans <| h <| Classical.arbitrary _ - -lemma dist_pi_const_le (a b : α) : (dist (fun _ : β => a) fun _ => b) ≤ dist a b := - (dist_pi_le_iff dist_nonneg).2 fun _ => le_rfl - -lemma nndist_pi_const_le (a b : α) : (nndist (fun _ : β => a) fun _ => b) ≤ nndist a b := - nndist_pi_le_iff.2 fun _ => le_rfl - -@[simp] -lemma dist_pi_const [Nonempty β] (a b : α) : (dist (fun _ : β => a) fun _ => b) = dist a b := by - simpa only [dist_edist] using congr_arg ENNReal.toReal (edist_pi_const a b) - -@[simp] -lemma nndist_pi_const [Nonempty β] (a b : α) : (nndist (fun _ : β => a) fun _ => b) = nndist a b := - NNReal.eq <| dist_pi_const a b - -lemma nndist_le_pi_nndist (f g : ∀ b, π b) (b : β) : nndist (f b) (g b) ≤ nndist f g := by - rw [← ENNReal.coe_le_coe, ← edist_nndist, ← edist_nndist] - exact edist_le_pi_edist f g b - -lemma dist_le_pi_dist (f g : ∀ b, π b) (b : β) : dist (f b) (g b) ≤ dist f g := by - simp only [dist_nndist, NNReal.coe_le_coe, nndist_le_pi_nndist f g b] - -/-- An open ball in a product space is a product of open balls. See also `ball_pi'` -for a version assuming `Nonempty β` instead of `0 < r`. -/ -lemma ball_pi (x : ∀ b, π b) {r : ℝ} (hr : 0 < r) : - ball x r = Set.pi univ fun b => ball (x b) r := by - ext p - simp [dist_pi_lt_iff hr] - -/-- An open ball in a product space is a product of open balls. See also `ball_pi` -for a version assuming `0 < r` instead of `Nonempty β`. -/ -lemma ball_pi' [Nonempty β] (x : ∀ b, π b) (r : ℝ) : - ball x r = Set.pi univ fun b => ball (x b) r := - (lt_or_le 0 r).elim (ball_pi x) fun hr => by simp [ball_eq_empty.2 hr] - -/-- A closed ball in a product space is a product of closed balls. See also `closedBall_pi'` -for a version assuming `Nonempty β` instead of `0 ≤ r`. -/ -lemma closedBall_pi (x : ∀ b, π b) {r : ℝ} (hr : 0 ≤ r) : - closedBall x r = Set.pi univ fun b => closedBall (x b) r := by - ext p - simp [dist_pi_le_iff hr] - -/-- A closed ball in a product space is a product of closed balls. See also `closedBall_pi` -for a version assuming `0 ≤ r` instead of `Nonempty β`. -/ -lemma closedBall_pi' [Nonempty β] (x : ∀ b, π b) (r : ℝ) : - closedBall x r = Set.pi univ fun b => closedBall (x b) r := - (le_or_lt 0 r).elim (closedBall_pi x) fun hr => by simp [closedBall_eq_empty.2 hr] - -/-- A sphere in a product space is a union of spheres on each component restricted to the closed -ball. -/ -lemma sphere_pi (x : ∀ b, π b) {r : ℝ} (h : 0 < r ∨ Nonempty β) : - sphere x r = (⋃ i : β, Function.eval i ⁻¹' sphere (x i) r) ∩ closedBall x r := by - obtain hr | rfl | hr := lt_trichotomy r 0 - · simp [hr] - · rw [closedBall_eq_sphere_of_nonpos le_rfl, eq_comm, Set.inter_eq_right] - letI := h.resolve_left (lt_irrefl _) - inhabit β - refine subset_iUnion_of_subset default ?_ - intro x hx - replace hx := hx.le - rw [dist_pi_le_iff le_rfl] at hx - exact le_antisymm (hx default) dist_nonneg - · ext - simp [dist_pi_eq_iff hr, dist_pi_le_iff hr.le] - -@[simp] -lemma Fin.nndist_insertNth_insertNth {n : ℕ} {α : Fin (n + 1) → Type*} - [∀ i, PseudoMetricSpace (α i)] (i : Fin (n + 1)) (x y : α i) (f g : ∀ j, α (i.succAbove j)) : - nndist (i.insertNth x f) (i.insertNth y g) = max (nndist x y) (nndist f g) := - eq_of_forall_ge_iff fun c => by simp [nndist_pi_le_iff, i.forall_iff_succAbove] - -@[simp] -lemma Fin.dist_insertNth_insertNth {n : ℕ} {α : Fin (n + 1) → Type*} - [∀ i, PseudoMetricSpace (α i)] (i : Fin (n + 1)) (x y : α i) (f g : ∀ j, α (i.succAbove j)) : - dist (i.insertNth x f) (i.insertNth y g) = max (dist x y) (dist f g) := by - simp only [dist_nndist, Fin.nndist_insertNth_insertNth, NNReal.coe_max] - -end Pi - -instance : PseudoMetricSpace (Additive α) := ‹_› -instance : PseudoMetricSpace (Multiplicative α) := ‹_› -instance : PseudoMetricSpace αᵒᵈ := ‹_› diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean index 282e00a70e43b..04ae492896a9c 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Defs.lean @@ -3,8 +3,9 @@ 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.Topology.EMetricSpace.Basic +import Mathlib.Data.ENNReal.Real import Mathlib.Tactic.Bound.Attribute +import Mathlib.Topology.EMetricSpace.Defs /-! ## Pseudo-metric spaces @@ -195,38 +196,6 @@ theorem dist_triangle4_right (x₁ y₁ x₂ y₂ : α) : rw [add_right_comm, dist_comm y₁] apply dist_triangle4 -/-- The triangle (polygon) inequality for sequences of points; `Finset.Ico` version. -/ -theorem dist_le_Ico_sum_dist (f : ℕ → α) {m n} (h : m ≤ n) : - dist (f m) (f n) ≤ ∑ i ∈ Finset.Ico m n, dist (f i) (f (i + 1)) := by - induction n, h using Nat.le_induction with - | base => rw [Finset.Ico_self, Finset.sum_empty, dist_self] - | succ n hle ihn => - calc - dist (f m) (f (n + 1)) ≤ dist (f m) (f n) + dist (f n) (f (n + 1)) := dist_triangle _ _ _ - _ ≤ (∑ i ∈ Finset.Ico m n, _) + _ := add_le_add ihn le_rfl - _ = ∑ i ∈ Finset.Ico m (n + 1), _ := by - { rw [Nat.Ico_succ_right_eq_insert_Ico hle, Finset.sum_insert, add_comm]; simp } - -/-- The triangle (polygon) inequality for sequences of points; `Finset.range` version. -/ -theorem dist_le_range_sum_dist (f : ℕ → α) (n : ℕ) : - dist (f 0) (f n) ≤ ∑ i ∈ Finset.range n, dist (f i) (f (i + 1)) := - Nat.Ico_zero_eq_range ▸ dist_le_Ico_sum_dist f (Nat.zero_le n) - -/-- A version of `dist_le_Ico_sum_dist` with each intermediate distance replaced -with an upper estimate. -/ -theorem dist_le_Ico_sum_of_dist_le {f : ℕ → α} {m n} (hmn : m ≤ n) {d : ℕ → ℝ} - (hd : ∀ {k}, m ≤ k → k < n → dist (f k) (f (k + 1)) ≤ d k) : - dist (f m) (f n) ≤ ∑ i ∈ Finset.Ico m n, d i := - le_trans (dist_le_Ico_sum_dist f hmn) <| - Finset.sum_le_sum fun _k hk => hd (Finset.mem_Ico.1 hk).1 (Finset.mem_Ico.1 hk).2 - -/-- A version of `dist_le_range_sum_dist` with each intermediate distance replaced -with an upper estimate. -/ -theorem dist_le_range_sum_of_dist_le {f : ℕ → α} (n : ℕ) {d : ℕ → ℝ} - (hd : ∀ {k}, k < n → dist (f k) (f (k + 1)) ≤ d k) : - dist (f 0) (f n) ≤ ∑ i ∈ Finset.range n, d i := - Nat.Ico_zero_eq_range ▸ dist_le_Ico_sum_of_dist_le (zero_le n) fun _ => hd - theorem swap_dist : Function.swap (@dist α _) = dist := by funext x y; exact dist_comm _ _ theorem abs_dist_sub_le (x y z : α) : |dist x z - dist y z| ≤ dist x y := @@ -713,98 +682,6 @@ theorem uniformContinuousOn_iff_le [PseudoMetricSpace β] {f : α → β} {s : S ∀ ε > 0, ∃ δ > 0, ∀ x ∈ s, ∀ y ∈ s, dist x y ≤ δ → dist (f x) (f y) ≤ ε := Metric.uniformity_basis_dist_le.uniformContinuousOn_iff Metric.uniformity_basis_dist_le -nonrec theorem uniformInducing_iff [PseudoMetricSpace β] {f : α → β} : - UniformInducing f ↔ UniformContinuous f ∧ - ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := - uniformInducing_iff'.trans <| Iff.rfl.and <| - ((uniformity_basis_dist.comap _).le_basis_iff uniformity_basis_dist).trans <| by - simp only [subset_def, Prod.forall, gt_iff_lt, preimage_setOf_eq, Prod.map_apply, mem_setOf] - -nonrec theorem uniformEmbedding_iff [PseudoMetricSpace β] {f : α → β} : - UniformEmbedding f ↔ Function.Injective f ∧ UniformContinuous f ∧ - ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := by - rw [uniformEmbedding_iff, and_comm, uniformInducing_iff] - -/-- If a map between pseudometric spaces is a uniform embedding then the distance between `f x` -and `f y` is controlled in terms of the distance between `x` and `y`. -/ -theorem controlled_of_uniformEmbedding [PseudoMetricSpace β] {f : α → β} (h : UniformEmbedding f) : - (∀ ε > 0, ∃ δ > 0, ∀ {a b : α}, dist a b < δ → dist (f a) (f b) < ε) ∧ - ∀ δ > 0, ∃ ε > 0, ∀ {a b : α}, dist (f a) (f b) < ε → dist a b < δ := - ⟨uniformContinuous_iff.1 h.uniformContinuous, (uniformEmbedding_iff.1 h).2.2⟩ - -theorem totallyBounded_iff {s : Set α} : - TotallyBounded s ↔ ∀ ε > 0, ∃ t : Set α, t.Finite ∧ s ⊆ ⋃ y ∈ t, ball y ε := - uniformity_basis_dist.totallyBounded_iff - -/-- A pseudometric space is totally bounded if one can reconstruct up to any ε>0 any element of the -space from finitely many data. -/ -theorem totallyBounded_of_finite_discretization {s : Set α} - (H : ∀ ε > (0 : ℝ), - ∃ (β : Type u) (_ : Fintype β) (F : s → β), ∀ x y, F x = F y → dist (x : α) y < ε) : - TotallyBounded s := by - rcases s.eq_empty_or_nonempty with hs | hs - · rw [hs] - exact totallyBounded_empty - rcases hs with ⟨x0, hx0⟩ - haveI : Inhabited s := ⟨⟨x0, hx0⟩⟩ - refine totallyBounded_iff.2 fun ε ε0 => ?_ - rcases H ε ε0 with ⟨β, fβ, F, hF⟩ - let Finv := Function.invFun F - refine ⟨range (Subtype.val ∘ Finv), finite_range _, fun x xs => ?_⟩ - let x' := Finv (F ⟨x, xs⟩) - have : F x' = F ⟨x, xs⟩ := Function.invFun_eq ⟨⟨x, xs⟩, rfl⟩ - simp only [Set.mem_iUnion, Set.mem_range] - exact ⟨_, ⟨F ⟨x, xs⟩, rfl⟩, hF _ _ this.symm⟩ - -theorem finite_approx_of_totallyBounded {s : Set α} (hs : TotallyBounded s) : - ∀ ε > 0, ∃ t, t ⊆ s ∧ Set.Finite t ∧ s ⊆ ⋃ y ∈ t, ball y ε := by - intro ε ε_pos - rw [totallyBounded_iff_subset] at hs - exact hs _ (dist_mem_uniformity ε_pos) - -/-- Expressing uniform convergence using `dist` -/ -theorem tendstoUniformlyOnFilter_iff {F : ι → β → α} {f : β → α} {p : Filter ι} {p' : Filter β} : - TendstoUniformlyOnFilter F f p p' ↔ - ∀ ε > 0, ∀ᶠ n : ι × β in p ×ˢ p', dist (f n.snd) (F n.fst n.snd) < ε := by - refine ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu => ?_⟩ - rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ - exact (H ε εpos).mono fun n hn => hε hn - -/-- Expressing locally uniform convergence on a set using `dist`. -/ -theorem tendstoLocallyUniformlyOn_iff [TopologicalSpace β] {F : ι → β → α} {f : β → α} - {p : Filter ι} {s : Set β} : - TendstoLocallyUniformlyOn F f p s ↔ - ∀ ε > 0, ∀ x ∈ s, ∃ t ∈ 𝓝[s] x, ∀ᶠ n in p, ∀ y ∈ t, dist (f y) (F n y) < ε := by - refine ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu x hx => ?_⟩ - rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ - rcases H ε εpos x hx with ⟨t, ht, Ht⟩ - exact ⟨t, ht, Ht.mono fun n hs x hx => hε (hs x hx)⟩ - -/-- Expressing uniform convergence on a set using `dist`. -/ -theorem tendstoUniformlyOn_iff {F : ι → β → α} {f : β → α} {p : Filter ι} {s : Set β} : - TendstoUniformlyOn F f p s ↔ ∀ ε > 0, ∀ᶠ n in p, ∀ x ∈ s, dist (f x) (F n x) < ε := by - refine ⟨fun H ε hε => H _ (dist_mem_uniformity hε), fun H u hu => ?_⟩ - rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩ - exact (H ε εpos).mono fun n hs x hx => hε (hs x hx) - -/-- Expressing locally uniform convergence using `dist`. -/ -theorem tendstoLocallyUniformly_iff [TopologicalSpace β] {F : ι → β → α} {f : β → α} - {p : Filter ι} : - TendstoLocallyUniformly F f p ↔ - ∀ ε > 0, ∀ x : β, ∃ t ∈ 𝓝 x, ∀ᶠ n in p, ∀ y ∈ t, dist (f y) (F n y) < ε := by - simp only [← tendstoLocallyUniformlyOn_univ, tendstoLocallyUniformlyOn_iff, nhdsWithin_univ, - mem_univ, forall_const, exists_prop] - -/-- Expressing uniform convergence using `dist`. -/ -theorem tendstoUniformly_iff {F : ι → β → α} {f : β → α} {p : Filter ι} : - TendstoUniformly F f p ↔ ∀ ε > 0, ∀ᶠ n in p, ∀ x, dist (f x) (F n x) < ε := by - rw [← tendstoUniformlyOn_univ, tendstoUniformlyOn_iff] - simp - -protected theorem cauchy_iff {f : Filter α} : - Cauchy f ↔ NeBot f ∧ ∀ ε > 0, ∃ t ∈ f, ∀ x ∈ t, ∀ y ∈ t, dist x y < ε := - uniformity_basis_dist.cauchy_iff - theorem nhds_basis_ball : (𝓝 x).HasBasis (0 < ·) (ball x) := nhds_basis_uniformity uniformity_basis_dist @@ -949,18 +826,6 @@ theorem isOpen_singleton_iff {α : Type*} [PseudoMetricSpace α] {x : α} : IsOpen ({x} : Set α) ↔ ∃ ε > 0, ∀ y, dist y x < ε → y = x := by simp [isOpen_iff, subset_singleton_iff, mem_ball] -/-- Given a point `x` in a discrete subset `s` of a pseudometric space, there is an open ball -centered at `x` and intersecting `s` only at `x`. -/ -theorem exists_ball_inter_eq_singleton_of_mem_discrete [DiscreteTopology s] {x : α} (hx : x ∈ s) : - ∃ ε > 0, Metric.ball x ε ∩ s = {x} := - nhds_basis_ball.exists_inter_eq_singleton_of_mem_discrete hx - -/-- Given a point `x` in a discrete subset `s` of a pseudometric space, there is a closed ball -of positive radius centered at `x` and intersecting `s` only at `x`. -/ -theorem exists_closedBall_inter_eq_singleton_of_discrete [DiscreteTopology s] {x : α} (hx : x ∈ s) : - ∃ ε > 0, Metric.closedBall x ε ∩ s = {x} := - nhds_basis_closedBall.exists_inter_eq_singleton_of_mem_discrete hx - theorem _root_.Dense.exists_dist_lt {s : Set α} (hs : Dense s) (x : α) {ε : ℝ} (hε : 0 < ε) : ∃ y ∈ s, dist x y < ε := by have : (ball x ε).Nonempty := by simp [hε] @@ -1052,9 +917,6 @@ theorem Metric.emetric_closedBall_nnreal {x : α} {ε : ℝ≥0} : theorem Metric.emetric_ball_top (x : α) : EMetric.ball x ⊤ = univ := eq_univ_of_forall fun _ => edist_lt_top _ _ -theorem Metric.inseparable_iff {x y : α} : Inseparable x y ↔ dist x y = 0 := by - rw [EMetric.inseparable_iff, edist_nndist, dist_nndist, ENNReal.coe_eq_zero, NNReal.coe_eq_zero] - /-- Build a new pseudometric space from an old one where the bundled uniform structure is provably (but typically non-definitionaly) equal to some given uniform structure. See Note [forgetful inheritance]. @@ -1186,12 +1048,6 @@ theorem Metric.uniformity_eq_comap_nhds_zero : simp only [mem_uniformity_dist, (nhds_basis_ball.comap _).mem_iff] simp [subset_def, Real.dist_0_eq_abs] -theorem cauchySeq_iff_tendsto_dist_atTop_0 [Nonempty β] [SemilatticeSup β] {u : β → α} : - CauchySeq u ↔ Tendsto (fun n : β × β => dist (u n.1) (u n.2)) atTop (𝓝 0) := by - rw [cauchySeq_iff_tendsto, Metric.uniformity_eq_comap_nhds_zero, tendsto_comap_iff, - Function.comp_def] - simp_rw [Prod.map_fst, Prod.map_snd] - theorem tendsto_uniformity_iff_dist_tendsto_zero {f : ι → α × α} {p : Filter ι} : Tendsto f p (𝓤 α) ↔ Tendsto (fun x => dist (f x).1 (f x).2) p (𝓝 0) := by rw [Metric.uniformity_eq_comap_nhds_zero, tendsto_comap_iff, Function.comp_def] @@ -1259,22 +1115,6 @@ theorem dense_iff {s : Set α} : Dense s ↔ ∀ x, ∀ r > 0, (ball x r ∩ s). theorem denseRange_iff {f : β → α} : DenseRange f ↔ ∀ x, ∀ r > 0, ∃ y, dist x (f y) < r := forall_congr' fun x => by simp only [mem_closure_iff, exists_range_iff] --- Porting note: `TopologicalSpace.IsSeparable.separableSpace` moved to `EMetricSpace` - -/-- The preimage of a separable set by an inducing map is separable. -/ -protected theorem _root_.Inducing.isSeparable_preimage {f : β → α} [TopologicalSpace β] - (hf : Inducing f) {s : Set α} (hs : IsSeparable s) : IsSeparable (f ⁻¹' s) := by - have : SeparableSpace s := hs.separableSpace - have : SecondCountableTopology s := UniformSpace.secondCountable_of_separable _ - have : Inducing ((mapsTo_preimage f s).restrict _ _ _) := - (hf.comp inducing_subtype_val).codRestrict _ - have := this.secondCountableTopology - exact .of_subtype _ - -protected theorem _root_.Embedding.isSeparable_preimage {f : β → α} [TopologicalSpace β] - (hf : Embedding f) {s : Set α} (hs : IsSeparable s) : IsSeparable (f ⁻¹' s) := - hf.toInducing.isSeparable_preimage hs - /-- If a map is continuous on a separable set `s`, then the image of `s` is also separable. -/ theorem _root_.ContinuousOn.isSeparable_image [TopologicalSpace β] {f : α → β} {s : Set α} (hf : ContinuousOn f s) (hs : IsSeparable s) : IsSeparable (f '' s) := by @@ -1283,11 +1123,6 @@ theorem _root_.ContinuousOn.isSeparable_image [TopologicalSpace β] {f : α → end Metric -/-- A compact set is separable. -/ -theorem IsCompact.isSeparable {s : Set α} (hs : IsCompact s) : IsSeparable s := - haveI : CompactSpace s := isCompact_iff_compactSpace.mp hs - .of_subtype s - section Compact /-- Any compact set in a pseudometric space can be covered by finitely many balls of a given @@ -1302,27 +1137,6 @@ alias IsCompact.finite_cover_balls := finite_cover_balls_of_compact end Compact -namespace Metric - -section SecondCountable - -open TopologicalSpace - -/-- A pseudometric space is second countable if, for every `ε > 0`, there is a countable set which -is `ε`-dense. -/ -theorem secondCountable_of_almost_dense_set - (H : ∀ ε > (0 : ℝ), ∃ s : Set α, s.Countable ∧ ∀ x, ∃ y ∈ s, dist x y ≤ ε) : - SecondCountableTopology α := by - refine EMetric.secondCountable_of_almost_dense_set fun ε ε0 => ?_ - rcases ENNReal.lt_iff_exists_nnreal_btwn.1 ε0 with ⟨ε', ε'0, ε'ε⟩ - choose s hsc y hys hyx using H ε' (mod_cast ε'0) - refine ⟨s, hsc, iUnion₂_eq_univ_iff.2 fun x => ⟨y x, hys _, le_trans ?_ ε'ε.le⟩⟩ - exact mod_cast hyx x - -end SecondCountable - -end Metric - theorem lebesgue_number_lemma_of_metric {s : Set α} {ι : Sort*} {c : ι → Set α} (hs : IsCompact s) (hc₁ : ∀ i, IsOpen (c i)) (hc₂ : s ⊆ ⋃ i, c i) : ∃ δ > 0, ∀ x ∈ s, ∃ i, ball x δ ⊆ c i := by simpa only [ball, UniformSpace.ball, preimage_setOf_eq, dist_comm] @@ -1331,3 +1145,7 @@ theorem lebesgue_number_lemma_of_metric {s : Set α} {ι : Sort*} {c : ι → Se theorem lebesgue_number_lemma_of_metric_sUnion {s : Set α} {c : Set (Set α)} (hs : IsCompact s) (hc₁ : ∀ t ∈ c, IsOpen t) (hc₂ : s ⊆ ⋃₀ c) : ∃ δ > 0, ∀ x ∈ s, ∃ t ∈ c, ball x δ ⊆ t := by rw [sUnion_eq_iUnion] at hc₂; simpa using lebesgue_number_lemma_of_metric hs (by simpa) hc₂ + +instance : PseudoMetricSpace (Additive α) := ‹_› +instance : PseudoMetricSpace (Multiplicative α) := ‹_› +instance : PseudoMetricSpace αᵒᵈ := ‹_› diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean b/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean index 04b7530b136fa..7bde82bf0007b 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean @@ -3,9 +3,8 @@ 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.Topology.MetricSpace.Pseudo.Constructions import Mathlib.Topology.Order.DenselyOrdered +import Mathlib.Topology.MetricSpace.Pseudo.Constructions /-! # Extra lemmas about pseudo-metric spaces @@ -16,24 +15,6 @@ open scoped NNReal Topology variable {ι α : Type*} [PseudoMetricSpace α] -lemma Real.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 - -lemma Real.dist_right_le_of_mem_uIcc {x y z : ℝ} (h : y ∈ uIcc x z) : dist y z ≤ dist x z := by - simpa only [dist_comm _ z] using abs_sub_right_of_mem_uIcc h - -lemma Real.dist_le_of_mem_uIcc {x y x' y' : ℝ} (hx : x ∈ uIcc x' y') (hy : y ∈ uIcc x' y') : - dist x y ≤ dist x' y' := - abs_sub_le_of_uIcc_subset_uIcc <| uIcc_subset_uIcc (by rwa [uIcc_comm]) (by rwa [uIcc_comm]) - -lemma Real.dist_le_of_mem_Icc {x y x' y' : ℝ} (hx : x ∈ Icc x' y') (hy : y ∈ Icc x' y') : - dist x y ≤ y' - x' := by - simpa only [Real.dist_eq, abs_of_nonpos (sub_nonpos.2 <| hx.1.trans hx.2), neg_sub] using - Real.dist_le_of_mem_uIcc (Icc_subset_uIcc hx) (Icc_subset_uIcc hy) - -lemma Real.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 - instance : OrderTopology ℝ := orderTopology_of_nhds_abs fun x => by simp only [nhds_basis_ball.eq_biInf, ball, Real.dist_eq, abs_sub_comm] @@ -109,13 +90,3 @@ lemma exists_isCompact_closedBall [WeaklyLocallyCompactSpace α] (x : α) : simpa only [and_comm] using (this.and self_mem_nhdsWithin).exists end Metric - -namespace Real -variable {π : ι → Type*} [Fintype ι] [∀ i, PseudoMetricSpace (π i)] {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 => - (Real.dist_le_of_mem_uIcc ?_ ?_).trans (dist_le_pi_dist x' y' b) <;> refine Icc_subset_uIcc ?_ - exacts [⟨hx.1 _, hx.2 _⟩, ⟨hy.1 _, hy.2 _⟩] - -end Real diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Pi.lean b/Mathlib/Topology/MetricSpace/Pseudo/Pi.lean new file mode 100644 index 0000000000000..3ace96bd69eb9 --- /dev/null +++ b/Mathlib/Topology/MetricSpace/Pseudo/Pi.lean @@ -0,0 +1,160 @@ +/- +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.Topology.Bornology.Constructions +import Mathlib.Topology.EMetricSpace.Basic +import Mathlib.Topology.MetricSpace.Pseudo.Defs + +/-! +# Product of pseudometric spaces + +This file constructs the infinity distance on finite products of pseudometric spaces. +-/ + +open Bornology Filter Metric Set +open scoped NNReal Topology + +variable {α β : Type*} [PseudoMetricSpace α] + +open Finset + +variable {π : β → Type*} [Fintype β] [∀ b, PseudoMetricSpace (π b)] + +/-- A finite product of pseudometric spaces is a pseudometric space, with the sup distance. -/ +instance pseudoMetricSpacePi : PseudoMetricSpace (∀ b, π b) := by + /- we construct the instance from the pseudoemetric space instance to avoid checking again that + the uniformity is the same as the product uniformity, but we register nevertheless a nice + formula for the distance -/ + let i := PseudoEMetricSpace.toPseudoMetricSpaceOfDist + (fun f g : ∀ b, π b => ((sup univ fun b => nndist (f b) (g b) : ℝ≥0) : ℝ)) + (fun f g => ((Finset.sup_lt_iff bot_lt_top).2 fun b _ => edist_lt_top _ _).ne) + (fun f g => by + simp only [edist_pi_def, edist_nndist, ← ENNReal.coe_finset_sup, ENNReal.coe_toReal]) + refine i.replaceBornology fun s => ?_ + simp only [← isBounded_def, isBounded_iff_eventually, ← forall_isBounded_image_eval_iff, + forall_mem_image, ← Filter.eventually_all, Function.eval_apply, @dist_nndist (π _)] + refine eventually_congr ((eventually_ge_atTop 0).mono fun C hC ↦ ?_) + lift C to ℝ≥0 using hC + refine ⟨fun H x hx y hy ↦ NNReal.coe_le_coe.2 <| Finset.sup_le fun b _ ↦ H b hx hy, + fun H b x hx y hy ↦ NNReal.coe_le_coe.2 ?_⟩ + simpa only using Finset.sup_le_iff.1 (NNReal.coe_le_coe.1 <| H hx hy) b (Finset.mem_univ b) + +lemma nndist_pi_def (f g : ∀ b, π b) : nndist f g = sup univ fun b => nndist (f b) (g b) := rfl + +lemma dist_pi_def (f g : ∀ b, π b) : dist f g = (sup univ fun b => nndist (f b) (g b) : ℝ≥0) := rfl + +lemma nndist_pi_le_iff {f g : ∀ b, π b} {r : ℝ≥0} : + nndist f g ≤ r ↔ ∀ b, nndist (f b) (g b) ≤ r := by simp [nndist_pi_def] + +lemma nndist_pi_lt_iff {f g : ∀ b, π b} {r : ℝ≥0} (hr : 0 < r) : + nndist f g < r ↔ ∀ b, nndist (f b) (g b) < r := by + rw [← bot_eq_zero'] at hr + simp [nndist_pi_def, Finset.sup_lt_iff hr] + +lemma nndist_pi_eq_iff {f g : ∀ b, π b} {r : ℝ≥0} (hr : 0 < r) : + nndist f g = r ↔ (∃ i, nndist (f i) (g i) = r) ∧ ∀ b, nndist (f b) (g b) ≤ r := by + rw [eq_iff_le_not_lt, nndist_pi_lt_iff hr, nndist_pi_le_iff, not_forall, and_comm] + simp_rw [not_lt, and_congr_left_iff, le_antisymm_iff] + intro h + refine exists_congr fun b => ?_ + apply (and_iff_right <| h _).symm + +lemma dist_pi_lt_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 < r) : + dist f g < r ↔ ∀ b, dist (f b) (g b) < r := by + lift r to ℝ≥0 using hr.le + exact nndist_pi_lt_iff hr + +lemma dist_pi_le_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 ≤ r) : + dist f g ≤ r ↔ ∀ b, dist (f b) (g b) ≤ r := by + lift r to ℝ≥0 using hr + exact nndist_pi_le_iff + +lemma dist_pi_eq_iff {f g : ∀ b, π b} {r : ℝ} (hr : 0 < r) : + dist f g = r ↔ (∃ i, dist (f i) (g i) = r) ∧ ∀ b, dist (f b) (g b) ≤ r := by + lift r to ℝ≥0 using hr.le + simp_rw [← coe_nndist, NNReal.coe_inj, nndist_pi_eq_iff hr, NNReal.coe_le_coe] + +lemma dist_pi_le_iff' [Nonempty β] {f g : ∀ b, π b} {r : ℝ} : + dist f g ≤ r ↔ ∀ b, dist (f b) (g b) ≤ r := by + by_cases hr : 0 ≤ r + · exact dist_pi_le_iff hr + · exact iff_of_false (fun h => hr <| dist_nonneg.trans h) fun h => + hr <| dist_nonneg.trans <| h <| Classical.arbitrary _ + +lemma dist_pi_const_le (a b : α) : (dist (fun _ : β => a) fun _ => b) ≤ dist a b := + (dist_pi_le_iff dist_nonneg).2 fun _ => le_rfl + +lemma nndist_pi_const_le (a b : α) : (nndist (fun _ : β => a) fun _ => b) ≤ nndist a b := + nndist_pi_le_iff.2 fun _ => le_rfl + +@[simp] +lemma dist_pi_const [Nonempty β] (a b : α) : (dist (fun _ : β => a) fun _ => b) = dist a b := by + simpa only [dist_edist] using congr_arg ENNReal.toReal (edist_pi_const a b) + +@[simp] +lemma nndist_pi_const [Nonempty β] (a b : α) : (nndist (fun _ : β => a) fun _ => b) = nndist a b := + NNReal.eq <| dist_pi_const a b + +lemma nndist_le_pi_nndist (f g : ∀ b, π b) (b : β) : nndist (f b) (g b) ≤ nndist f g := by + rw [← ENNReal.coe_le_coe, ← edist_nndist, ← edist_nndist] + exact edist_le_pi_edist f g b + +lemma dist_le_pi_dist (f g : ∀ b, π b) (b : β) : dist (f b) (g b) ≤ dist f g := by + simp only [dist_nndist, NNReal.coe_le_coe, nndist_le_pi_nndist f g b] + +/-- An open ball in a product space is a product of open balls. See also `ball_pi'` +for a version assuming `Nonempty β` instead of `0 < r`. -/ +lemma ball_pi (x : ∀ b, π b) {r : ℝ} (hr : 0 < r) : + ball x r = Set.pi univ fun b => ball (x b) r := by + ext p + simp [dist_pi_lt_iff hr] + +/-- An open ball in a product space is a product of open balls. See also `ball_pi` +for a version assuming `0 < r` instead of `Nonempty β`. -/ +lemma ball_pi' [Nonempty β] (x : ∀ b, π b) (r : ℝ) : + ball x r = Set.pi univ fun b => ball (x b) r := + (lt_or_le 0 r).elim (ball_pi x) fun hr => by simp [ball_eq_empty.2 hr] + +/-- A closed ball in a product space is a product of closed balls. See also `closedBall_pi'` +for a version assuming `Nonempty β` instead of `0 ≤ r`. -/ +lemma closedBall_pi (x : ∀ b, π b) {r : ℝ} (hr : 0 ≤ r) : + closedBall x r = Set.pi univ fun b => closedBall (x b) r := by + ext p + simp [dist_pi_le_iff hr] + +/-- A closed ball in a product space is a product of closed balls. See also `closedBall_pi` +for a version assuming `0 ≤ r` instead of `Nonempty β`. -/ +lemma closedBall_pi' [Nonempty β] (x : ∀ b, π b) (r : ℝ) : + closedBall x r = Set.pi univ fun b => closedBall (x b) r := + (le_or_lt 0 r).elim (closedBall_pi x) fun hr => by simp [closedBall_eq_empty.2 hr] + +/-- A sphere in a product space is a union of spheres on each component restricted to the closed +ball. -/ +lemma sphere_pi (x : ∀ b, π b) {r : ℝ} (h : 0 < r ∨ Nonempty β) : + sphere x r = (⋃ i : β, Function.eval i ⁻¹' sphere (x i) r) ∩ closedBall x r := by + obtain hr | rfl | hr := lt_trichotomy r 0 + · simp [hr] + · rw [closedBall_eq_sphere_of_nonpos le_rfl, eq_comm, Set.inter_eq_right] + letI := h.resolve_left (lt_irrefl _) + inhabit β + refine subset_iUnion_of_subset default ?_ + intro x hx + replace hx := hx.le + rw [dist_pi_le_iff le_rfl] at hx + exact le_antisymm (hx default) dist_nonneg + · ext + simp [dist_pi_eq_iff hr, dist_pi_le_iff hr.le] + +@[simp] +lemma Fin.nndist_insertNth_insertNth {n : ℕ} {α : Fin (n + 1) → Type*} + [∀ i, PseudoMetricSpace (α i)] (i : Fin (n + 1)) (x y : α i) (f g : ∀ j, α (i.succAbove j)) : + nndist (i.insertNth x f) (i.insertNth y g) = max (nndist x y) (nndist f g) := + eq_of_forall_ge_iff fun c => by simp [nndist_pi_le_iff, i.forall_iff_succAbove] + +@[simp] +lemma Fin.dist_insertNth_insertNth {n : ℕ} {α : Fin (n + 1) → Type*} + [∀ i, PseudoMetricSpace (α i)] (i : Fin (n + 1)) (x y : α i) (f g : ∀ j, α (i.succAbove j)) : + dist (i.insertNth x f) (i.insertNth y g) = max (dist x y) (dist f g) := by + simp only [dist_nndist, Fin.nndist_insertNth_insertNth, NNReal.coe_max] diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Real.lean b/Mathlib/Topology/MetricSpace/Pseudo/Real.lean new file mode 100644 index 0000000000000..ae477178577c5 --- /dev/null +++ b/Mathlib/Topology/MetricSpace/Pseudo/Real.lean @@ -0,0 +1,45 @@ +/- +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.Topology.MetricSpace.Pseudo.Pi + +/-! +# Lemmas about distances between points in intervals in `ℝ`. +-/ + +open Bornology Filter Metric Set +open scoped NNReal Topology + +namespace Real + +variable {ι α : Type*} [PseudoMetricSpace α] + +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 + +lemma dist_right_le_of_mem_uIcc {x y z : ℝ} (h : y ∈ uIcc x z) : dist y z ≤ dist x z := by + simpa only [dist_comm _ z] using abs_sub_right_of_mem_uIcc h + +lemma dist_le_of_mem_uIcc {x y x' y' : ℝ} (hx : x ∈ uIcc x' y') (hy : y ∈ uIcc x' y') : + dist x y ≤ dist x' y' := + abs_sub_le_of_uIcc_subset_uIcc <| uIcc_subset_uIcc (by rwa [uIcc_comm]) (by rwa [uIcc_comm]) + +lemma dist_le_of_mem_Icc {x y x' y' : ℝ} (hx : x ∈ Icc x' y') (hy : y ∈ Icc x' y') : + dist x y ≤ y' - x' := by + simpa only [Real.dist_eq, abs_of_nonpos (sub_nonpos.2 <| hx.1.trans hx.2), neg_sub] using + Real.dist_le_of_mem_uIcc (Icc_subset_uIcc hx) (Icc_subset_uIcc hy) + +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' : ι → ℝ} + +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 => + (Real.dist_le_of_mem_uIcc ?_ ?_).trans (dist_le_pi_dist x' y' b) <;> refine Icc_subset_uIcc ?_ + exacts [⟨hx.1 _, hx.2 _⟩, ⟨hy.1 _, hy.2 _⟩] + +end Real diff --git a/Mathlib/Topology/Metrizable/Uniformity.lean b/Mathlib/Topology/Metrizable/Uniformity.lean index 84aa5a4ff45ae..5a7ca2208776e 100644 --- a/Mathlib/Topology/Metrizable/Uniformity.lean +++ b/Mathlib/Topology/Metrizable/Uniformity.lean @@ -282,3 +282,10 @@ lemma TotallyBounded.isSeparable [UniformSpace X] [i : IsCountablyGenerated ( exact EMetric.ball_subset_closedBall obtain ⟨t, _, htc, hts⟩ := EMetric.subset_countable_closure_of_almost_dense_set s h' exact ⟨t, htc, hts⟩ + +open TopologicalSpace in +instance (priority := 100) DiscreteTopology.metrizableSpace + {α} [TopologicalSpace α] [DiscreteTopology α] : + MetrizableSpace α := by + obtain rfl := DiscreteTopology.eq_bot (α := α) + exact @UniformSpace.metrizableSpace α ⊥ (isCountablyGenerated_principal _) _ From ab9032907fcf786dcada23be512f904443aee780 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Aug 2024 01:08:22 +0000 Subject: [PATCH 316/359] chore: Fix implicitness of `floor_le_floor` (#15856) --- Mathlib/Algebra/Order/Floor.lean | 16 ++++++---------- 1 file changed, 6 insertions(+), 10 deletions(-) diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index d12b1da461bd2..fd12802ad8138 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -132,7 +132,7 @@ end OrderedSemiring section LinearOrderedSemiring -variable [LinearOrderedSemiring α] [FloorSemiring α] {a : α} {n : ℕ} +variable [LinearOrderedSemiring α] [FloorSemiring α] {a b : α} {n : ℕ} theorem floor_lt (ha : 0 ≤ a) : ⌊a⌋₊ < n ↔ a < n := lt_iff_lt_of_le_iff_le <| le_floor_iff ha @@ -183,8 +183,7 @@ theorem floor_mono : Monotone (floor : α → ℕ) := fun a b h => by exact Nat.zero_le _ · exact le_floor ((floor_le ha).trans h) -@[gcongr] -theorem floor_le_floor : ∀ x y : α, x ≤ y → ⌊x⌋₊ ≤ ⌊y⌋₊ := floor_mono +@[gcongr] lemma floor_le_floor (hab : a ≤ b) : ⌊a⌋₊ ≤ ⌊b⌋₊ := floor_mono hab theorem le_floor_iff' (hn : n ≠ 0) : n ≤ ⌊a⌋₊ ↔ (n : α) ≤ a := by obtain ha | ha := le_total a 0 @@ -281,8 +280,7 @@ theorem ceil_natCast (n : ℕ) : ⌈(n : α)⌉₊ = n := theorem ceil_mono : Monotone (ceil : α → ℕ) := gc_ceil_coe.monotone_l -@[gcongr] -theorem ceil_le_ceil : ∀ x y : α, x ≤ y → ⌈x⌉₊ ≤ ⌈y⌉₊ := ceil_mono +@[gcongr] lemma ceil_le_ceil (hab : a ≤ b) : ⌈a⌉₊ ≤ ⌈b⌉₊ := ceil_mono hab @[simp] theorem ceil_zero : ⌈(0 : α)⌉₊ = 0 := by rw [← Nat.cast_zero, ceil_natCast] @@ -556,7 +554,7 @@ def FloorRing.ofCeil (α) [LinearOrderedRing α] (ceil : α → ℤ) namespace Int -variable [LinearOrderedRing α] [FloorRing α] {z : ℤ} {a : α} +variable [LinearOrderedRing α] [FloorRing α] {z : ℤ} {a b : α} /-- `Int.floor a` is the greatest integer `z` such that `z ≤ a`. It is denoted with `⌊a⌋`. -/ def floor : α → ℤ := @@ -658,8 +656,7 @@ theorem floor_one : ⌊(1 : α)⌋ = 1 := by rw [← cast_one, floor_intCast] theorem floor_mono : Monotone (floor : α → ℤ) := gc_coe_floor.monotone_u -@[gcongr] -theorem floor_le_floor : ∀ x y : α, x ≤ y → ⌊x⌋ ≤ ⌊y⌋ := floor_mono +@[gcongr] lemma floor_le_floor (hab : a ≤ b) : ⌊a⌋ ≤ ⌊b⌋ := floor_mono hab theorem floor_pos : 0 < ⌊a⌋ ↔ 1 ≤ a := by -- Porting note: broken `convert le_floor` @@ -1078,8 +1075,7 @@ theorem ceil_ofNat (n : ℕ) [n.AtLeastTwo] : ⌈(no_index (OfNat.ofNat n : α)) theorem ceil_mono : Monotone (ceil : α → ℤ) := gc_ceil_coe.monotone_l -@[gcongr] -theorem ceil_le_ceil : ∀ x y : α, x ≤ y → ⌈x⌉ ≤ ⌈y⌉ := ceil_mono +@[gcongr] lemma ceil_le_ceil (hab : a ≤ b) : ⌈a⌉ ≤ ⌈b⌉ := ceil_mono hab @[simp] theorem ceil_add_int (a : α) (z : ℤ) : ⌈a + z⌉ = ⌈a⌉ + z := by From dc5804b64c1d397a30ef524d1a85a4557944f3fd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Aug 2024 01:08:23 +0000 Subject: [PATCH 317/359] feat: `ZMod.forall`, `ZMod.exists` (#15861) From LeanAPAP --- Mathlib/Data/ZMod/Basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index fed33c8dea188..6737b8715dbd6 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -231,6 +231,9 @@ theorem intCast_surjective : Function.Surjective ((↑) : ℤ → ZMod n) := @[deprecated (since := "2024-04-17")] alias int_cast_surjective := intCast_surjective +lemma «forall» {P : ZMod n → Prop} : (∀ x, P x) ↔ ∀ x : ℤ, P x := intCast_surjective.forall +lemma «exists» {P : ZMod n → Prop} : (∃ x, P x) ↔ ∃ x : ℤ, P x := intCast_surjective.exists + theorem cast_id : ∀ (n) (i : ZMod n), (ZMod.cast i : ZMod n) = i | 0, _ => Int.cast_id | _ + 1, i => natCast_zmod_val i From 76ccc39d64d8f3b4c934d8fb47b477acf0a8a3ef Mon Sep 17 00:00:00 2001 From: Jeremy Tan Jie Rui Date: Fri, 16 Aug 2024 05:39:21 +0000 Subject: [PATCH 318/359] feat: Dyck words up to semilength (#14162) First part of #9781. --- Mathlib.lean | 1 + .../Combinatorics/Enumerative/DyckWord.lean | 199 ++++++++++++++++++ 2 files changed, 200 insertions(+) create mode 100644 Mathlib/Combinatorics/Enumerative/DyckWord.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3d99e987052c1..d7cb00827b04c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1877,6 +1877,7 @@ import Mathlib.Combinatorics.Derangements.Finite import Mathlib.Combinatorics.Enumerative.Catalan import Mathlib.Combinatorics.Enumerative.Composition import Mathlib.Combinatorics.Enumerative.DoubleCounting +import Mathlib.Combinatorics.Enumerative.DyckWord import Mathlib.Combinatorics.Enumerative.Partition import Mathlib.Combinatorics.HalesJewett import Mathlib.Combinatorics.Hall.Basic diff --git a/Mathlib/Combinatorics/Enumerative/DyckWord.lean b/Mathlib/Combinatorics/Enumerative/DyckWord.lean new file mode 100644 index 0000000000000..c7ccaf3c8c5ca --- /dev/null +++ b/Mathlib/Combinatorics/Enumerative/DyckWord.lean @@ -0,0 +1,199 @@ +/- +Copyright (c) 2024 Jeremy Tan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Tan +-/ +import Mathlib.Algebra.Order.Ring.Nat +import Mathlib.Data.List.Nodup + +/-! +# Dyck words + +A Dyck word is a sequence consisting of an equal number `n` of symbols of two types such that +for all prefixes one symbol occurs at least as many times as the other. +If the symbols are `(` and `)` the latter restriction is equivalent to balanced brackets; +if they are `U = (1, 1)` and `D = (1, -1)` the sequence is a lattice path from `(0, 0)` to `(0, 2n)` +and the restriction requires the path to never go below the x-axis. + +## Main definitions + +* `DyckWord`: a list of `U`s and `D`s with as many `U`s as `D`s and with every prefix having +at least as many `U`s as `D`s. +* `DyckWord.semilength`: semilength (half the length) of a Dyck word. + +## Implementation notes + +While any two-valued type could have been used for `DyckStep`, a new enumerated type is used here +to emphasise that the definition of a Dyck word does not depend on that underlying type. + +## TODO + +* Prove the bijection between Dyck words and rooted binary trees +(https://github.com/leanprover-community/mathlib4/pull/9781). +-/ + +open List + +/-- A `DyckStep` is either `U` or `D`, corresponding to `(` and `)` respectively. -/ +inductive DyckStep + | U : DyckStep + | D : DyckStep + deriving Inhabited, DecidableEq + +/-- Named in analogy to `Bool.dichotomy`. -/ +lemma DyckStep.dichotomy (s : DyckStep) : s = U ∨ s = D := by cases s <;> tauto + +open DyckStep + +/-- A Dyck word is a list of `DyckStep`s with as many `U`s as `D`s and with every prefix having +at least as many `U`s as `D`s. -/ +@[ext] +structure DyckWord where + /-- The underlying list -/ + toList : List DyckStep + /-- There are as many `U`s as `D`s -/ + count_U_eq_count_D : toList.count U = toList.count D + /-- Each prefix has as least as many `U`s as `D`s -/ + count_D_le_count_U i : (toList.take i).count D ≤ (toList.take i).count U + deriving DecidableEq + +attribute [coe] DyckWord.toList +instance : Coe DyckWord (List DyckStep) := ⟨DyckWord.toList⟩ + +instance : Add DyckWord where + add p q := ⟨p ++ q, by + simp only [count_append, p.count_U_eq_count_D, q.count_U_eq_count_D], by + simp only [take_append_eq_append_take, count_append] + exact fun _ ↦ add_le_add (p.count_D_le_count_U _) (q.count_D_le_count_U _)⟩ + +instance : Zero DyckWord := ⟨[], by simp, by simp⟩ + +/-- Dyck words form an additive monoid under concatenation, with the empty word as 0. -/ +instance : AddMonoid DyckWord where + add_zero p := by ext1; exact append_nil _ + zero_add p := by ext1; rfl + add_assoc p q r := by ext1; apply append_assoc + nsmul := nsmulRec + +namespace DyckWord + +variable (p q : DyckWord) + +lemma toList_eq_nil : p.toList = [] ↔ p = 0 := by rw [DyckWord.ext_iff]; rfl + +/-- The first element of a nonempty Dyck word is `U`. -/ +lemma head_eq_U (h : p.toList ≠ []) : p.toList.head h = U := by + rcases p with - | s; · tauto + rw [head_cons] + by_contra f + rename_i _ nonneg + simpa [s.dichotomy.resolve_left f] using nonneg 1 + +/-- The last element of a nonempty Dyck word is `D`. -/ +lemma getLast_eq_D (h : p.toList ≠ []) : p.toList.getLast h = D := by + by_contra f; have s := p.count_U_eq_count_D + rw [← dropLast_append_getLast h, (dichotomy _).resolve_right f] at s + simp_rw [dropLast_eq_take, count_append, count_singleton', ite_true, ite_false] at s + have := p.count_D_le_count_U (p.toList.length - 1); omega + +lemma cons_tail_dropLast_concat (h : p.toList ≠ []) : + U :: p.toList.dropLast.tail ++ [D] = p := by + have : p.toList.dropLast.take 1 = [p.toList.head h] := by + rcases p with - | ⟨s, ⟨- | ⟨t, r⟩⟩⟩ + · tauto + · rename_i bal _ + cases s <;> simp at bal + · tauto + nth_rw 2 [← p.toList.dropLast_append_getLast h, ← p.toList.dropLast.take_append_drop 1] + rw [p.getLast_eq_D, drop_one, this, p.head_eq_U] + rfl + +/-- Prefix of a Dyck word as a Dyck word, given that the count of `U`s and `D`s in it are equal. -/ +def take (i : ℕ) (hi : (p.toList.take i).count U = (p.toList.take i).count D) : DyckWord where + toList := p.toList.take i + count_U_eq_count_D := hi + count_D_le_count_U k := by rw [take_take]; exact p.count_D_le_count_U (min k i) + +/-- Suffix of a Dyck word as a Dyck word, given that the count of `U`s and `D`s in the prefix +are equal. -/ +def drop (i : ℕ) (hi : (p.toList.take i).count U = (p.toList.take i).count D) : DyckWord where + toList := p.toList.drop i + count_U_eq_count_D := by + have := p.count_U_eq_count_D + rw [← take_append_drop i p.toList, count_append, count_append] at this + omega + count_D_le_count_U k := by + rw [show i = min i (i + k) by omega, ← take_take] at hi + rw [take_drop, ← add_le_add_iff_left (((p.toList.take (i + k)).take i).count U), + ← count_append, hi, ← count_append, take_append_drop] + exact p.count_D_le_count_U _ + +/-- Nest `p` in one pair of brackets, i.e. `x` becomes `(x)`. -/ +def nest : DyckWord where + toList := [U] ++ p ++ [D] + count_U_eq_count_D := by simp [p.count_U_eq_count_D] + count_D_le_count_U i := by + simp only [take_append_eq_append_take, count_append] + rw [← add_rotate (count D _), ← add_rotate (count U _)] + apply add_le_add _ (p.count_D_le_count_U _) + rcases i.eq_zero_or_pos with hi | hi; · simp [hi] + rw [take_of_length_le (show [U].length ≤ i by rwa [length_singleton]), count_singleton'] + simp only [ite_true, ite_false] + rw [add_comm] + exact add_le_add (zero_le _) ((count_le_length _ _).trans (by simp)) + +/-- Denest `p`, i.e. `(x)` becomes `x`, given that `p` is strictly positive in its interior +(this ensures that `x` is a Dyck word). -/ +def denest (h : p.toList ≠ []) (pos : ∀ i, 0 < i → i < p.toList.length → + (p.toList.take i).count D < (p.toList.take i).count U) : DyckWord where + toList := p.toList.dropLast.tail + count_U_eq_count_D := by + have := p.count_U_eq_count_D + rw [← p.cons_tail_dropLast_concat h, count_append, count_cons] at this + simpa using this + count_D_le_count_U i := by + have l1 : p.toList.take 1 = [p.toList.head h] := by rcases p with - | - <;> tauto + have l3 : p.toList.length - 1 = p.toList.length - 1 - 1 + 1 := by + rcases p with - | ⟨s, ⟨- | ⟨t, r⟩⟩⟩ + · tauto + · rename_i bal _ + cases s <;> simp at bal + · tauto + rw [← drop_one, take_drop, dropLast_eq_take, take_take] + have ub : min (1 + i) (p.toList.length - 1) < p.toList.length := + (min_le_right _ p.toList.length.pred).trans_lt (Nat.pred_lt ((length_pos.mpr h).ne')) + have lb : 0 < min (1 + i) (p.toList.length - 1) := by + rw [l3, add_comm, min_add_add_right]; omega + have eq := pos _ lb ub + set j := min (1 + i) (p.toList.length - 1) + rw [← (p.toList.take j).take_append_drop 1, count_append, count_append, take_take, + min_eq_left (by omega), l1, p.head_eq_U] at eq + simp only [count_singleton', ite_true, ite_false] at eq + omega + +lemma denest_nest (h : p.toList ≠ []) (pos : ∀ i, 0 < i → i < p.toList.length → + (p.toList.take i).count D < (p.toList.take i).count U) : (p.denest h pos).nest = p := by + simpa [DyckWord.ext_iff] using p.cons_tail_dropLast_concat h + +section Semilength + +/-- The semilength of a Dyck word is half of the number of `DyckStep`s in it, or equivalently +its number of `U`s. -/ +def semilength : ℕ := p.toList.count U + +@[simp] lemma semilength_zero : semilength 0 = 0 := rfl +@[simp] lemma semilength_add : (p + q).semilength = p.semilength + q.semilength := count_append .. +@[simp] lemma semilength_nest : p.nest.semilength = p.semilength + 1 := by simp [semilength, nest] + +lemma semilength_eq_count_D : p.semilength = p.toList.count D := by + rw [← count_U_eq_count_D]; rfl + +@[simp] +lemma two_mul_semilength_eq_length : 2 * p.semilength = p.toList.length := by + nth_rw 1 [two_mul, semilength, p.count_U_eq_count_D, semilength] + convert (p.toList.length_eq_countP_add_countP (· == D)).symm + rw [count]; congr!; rename_i s; cases s <;> tauto + +end Semilength + +end DyckWord From 9045ac05908da3a71d2ee4e52abe4925154b284a Mon Sep 17 00:00:00 2001 From: Gabin Date: Fri, 16 Aug 2024 05:48:26 +0000 Subject: [PATCH 319/359] =?UTF-8?q?feat(ModelTheory):=20Fra=C3=AFss=C3=A9?= =?UTF-8?q?=20limits=20are=20unique=20(#9967)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove that Fraïsse limits are unique up to equivalence. Co-authored-by: Aaron Anderson Co-authored-by: Aaron Anderson <65780815+awainverse@users.noreply.github.com> Co-authored-by: Gabin Kolly Co-authored-by: Gabin <68381468+QuinnLesquimau@users.noreply.github.com> --- Mathlib/ModelTheory/DirectLimit.lean | 2 +- Mathlib/ModelTheory/Fraisse.lean | 103 ++++++++++++++++++++++++- Mathlib/ModelTheory/Substructures.lean | 37 +++++++-- 3 files changed, 132 insertions(+), 10 deletions(-) diff --git a/Mathlib/ModelTheory/DirectLimit.lean b/Mathlib/ModelTheory/DirectLimit.lean index 5f0d653d1499c..09c8d5f5ddb8e 100644 --- a/Mathlib/ModelTheory/DirectLimit.lean +++ b/Mathlib/ModelTheory/DirectLimit.lean @@ -417,7 +417,7 @@ theorem cg {ι : Type*} [Countable ι] [Preorder ι] [IsDirected ι (· ≤ ·)] Structure.CG L (DirectLimit G f) := by refine ⟨⟨⋃ i, DirectLimit.of L ι G f i '' Classical.choose (h i).out, ?_, ?_⟩⟩ · exact Set.countable_iUnion fun i => Set.Countable.image (Classical.choose_spec (h i).out).1 _ - · rw [eq_top_iff, Substructure.closure_unionᵢ] + · rw [eq_top_iff, Substructure.closure_iUnion] simp_rw [← Embedding.coe_toHom, Substructure.closure_image] rw [le_iSup_iff] intro S hS x _ diff --git a/Mathlib/ModelTheory/Fraisse.lean b/Mathlib/ModelTheory/Fraisse.lean index ca502e44d67e8..c0f806f87aaa1 100644 --- a/Mathlib/ModelTheory/Fraisse.lean +++ b/Mathlib/ModelTheory/Fraisse.lean @@ -1,10 +1,10 @@ /- Copyright (c) 2022 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Aaron Anderson +Authors: Aaron Anderson, Gabin Kolly -/ import Mathlib.ModelTheory.FinitelyGenerated -import Mathlib.ModelTheory.DirectLimit +import Mathlib.ModelTheory.PartialEquiv import Mathlib.ModelTheory.Bundled import Mathlib.Algebra.Order.Archimedean.Basic @@ -42,6 +42,8 @@ Fraïssé limit - the countable ultrahomogeneous structure with that age. essentially countable. - `FirstOrder.Language.exists_countable_is_age_of_iff` gives necessary and sufficient conditions for a class to be the age of a countable structure in a language with countably many functions. +- `FirstOrder.Language.IsFraisseLimit.nonempty_equiv` shows that any class which is Fraïsse has + at most one Fraïsse limit up to equivalence. ## Implementation Notes @@ -155,6 +157,22 @@ theorem age.jointEmbedding : JointEmbedding (L.age M) := fun _ hN _ hP => ⟨Embedding.comp (inclusion le_sup_left) hN.2.some.equivRange.toEmbedding⟩, ⟨Embedding.comp (inclusion le_sup_right) hP.2.some.equivRange.toEmbedding⟩⟩ +variable {M} + +theorem age.fg_substructure {S : L.Substructure M} (fg : S.FG) : Bundled.mk S ∈ L.age M := by + exact ⟨(Substructure.fg_iff_structure_fg _).1 fg, ⟨subtype _⟩⟩ + +variable (M) + +/-- Any class in the age of a structure has a representative which is a finitely generated +substructure. -/ +theorem age.has_representative_as_substructure : + ∀ C ∈ Quotient.mk' '' L.age M, ∃ V : {V : L.Substructure M // FG V}, + ⟦Bundled.mk V⟧ = C := by + rintro _ ⟨N, ⟨N_fg, ⟨N_incl⟩⟩, N_eq⟩ + refine N_eq.symm ▸ ⟨⟨N_incl.toHom.range, ?_⟩, Quotient.sound ⟨N_incl.equivRange.symm⟩⟩ + exact FG.range N_fg (Embedding.toHom N_incl) + /-- The age of a countable structure is essentially countable (has countably many isomorphism classes). -/ theorem age.countable_quotient [h : Countable M] : (Quotient.mk' '' L.age M).Countable := by @@ -265,6 +283,50 @@ structure IsFraisseLimit [Countable (Σ l, L.Functions l)] [Countable M] : Prop variable {M} +/-- Any embedding from a finitely generated `S` to an ultrahomogeneous structure `M` +can be extended to an embedding from any structure with an embedding to `M`. -/ +theorem IsUltrahomogeneous.extend_embedding (M_homog : L.IsUltrahomogeneous M) {S : Type*} + [L.Structure S] (S_FG : FG L S) {T : Type*} [L.Structure T] [h : Nonempty (T ↪[L] M)] + (f : S ↪[L] M) (g : S ↪[L] T) : + ∃ f' : T ↪[L] M, f = f'.comp g := by + let ⟨r⟩ := h + let s := r.comp g + let ⟨t, eq⟩ := M_homog s.toHom.range (S_FG.range s.toHom) (f.comp s.equivRange.symm.toEmbedding) + use t.toEmbedding.comp r + change _ = t.toEmbedding.comp s + ext x + have eq' := congr_fun (congr_arg DFunLike.coe eq) ⟨s x, Hom.mem_range.2 ⟨x, rfl⟩⟩ + simp only [Embedding.comp_apply, Hom.comp_apply, + Equiv.coe_toHom, Embedding.coe_toHom, coeSubtype] at eq' + simp only [Embedding.comp_apply, ← eq', Equiv.coe_toEmbedding, EmbeddingLike.apply_eq_iff_eq] + apply (Embedding.equivRange (Embedding.comp r g)).injective + simp only [Equiv.apply_symm_apply] + rfl + +/-- A countably generated structure is ultrahomogeneous if and only if any equivalence between +finitely generated substructures can be extended to any element in the domain.-/ +theorem isUltrahomogeneous_iff_IsExtensionPair (M_CG : CG L M) : L.IsUltrahomogeneous M ↔ + L.IsExtensionPair M M := by + constructor + · intro M_homog ⟨f, f_FG⟩ m + let S := f.dom ⊔ closure L {m} + have dom_le_S : f.dom ≤ S := le_sup_left + let ⟨f', eq_f'⟩ := M_homog.extend_embedding (f.dom.fg_iff_structure_fg.1 f_FG) + ((subtype _).comp f.toEquiv.toEmbedding) (inclusion dom_le_S) (h := ⟨subtype _⟩) + refine ⟨⟨⟨S, f'.toHom.range, f'.equivRange⟩, f_FG.sup (fg_closure_singleton _)⟩, + subset_closure.trans (le_sup_right : _ ≤ S) (mem_singleton m), ⟨dom_le_S, ?_⟩⟩ + ext + simp only [Embedding.comp_apply, Equiv.coe_toEmbedding, coeSubtype, eq_f', + Embedding.equivRange_apply, Substructure.coe_inclusion, EmbeddingLike.apply_eq_iff_eq] + · intro h S S_FG f + let ⟨g, ⟨dom_le_dom, eq⟩⟩ := + equiv_between_cg M_CG M_CG ⟨⟨S, f.toHom.range, f.equivRange⟩, S_FG⟩ h h + use g + simp only [Embedding.subtype_equivRange] at eq + rw [← eq] + ext + rfl + theorem IsUltrahomogeneous.amalgamation_age (h : L.IsUltrahomogeneous M) : Amalgamation (L.age M) := by rintro N P Q NP NQ ⟨Nfg, ⟨-⟩⟩ ⟨Pfg, ⟨PM⟩⟩ ⟨Qfg, ⟨QM⟩⟩ @@ -302,6 +364,43 @@ theorem isFraisse [Countable (Σ l, L.Functions l)] [Countable M] (h : IsFraisse IsFraisse K := (congr rfl h.age).mp h.ultrahomogeneous.age_isFraisse +variable {K} {N : Type w} [L.Structure N] +variable [Countable (Σ l, L.Functions l)] [Countable M] [Countable N] +variable (hM : IsFraisseLimit K M) (hN : IsFraisseLimit K N) + +include hM hN + +protected theorem isExtensionPair : L.IsExtensionPair M N := by + intro ⟨f, f_FG⟩ m + let S := f.dom ⊔ closure L {m} + have S_FG : S.FG := f_FG.sup (Substructure.fg_closure_singleton _) + have S_in_age_N : ⟨S, inferInstance⟩ ∈ L.age N := by + rw [hN.age, ← hM.age] + exact ⟨(fg_iff_structure_fg S).1 S_FG, ⟨subtype _⟩⟩ + haveI nonempty_S_N : Nonempty (S ↪[L] N) := S_in_age_N.2 + let ⟨g, g_eq⟩ := hN.ultrahomogeneous.extend_embedding (f.dom.fg_iff_structure_fg.1 f_FG) + ((subtype f.cod).comp f.toEquiv.toEmbedding) (inclusion (le_sup_left : _ ≤ S)) + refine ⟨⟨⟨S, g.toHom.range, g.equivRange⟩, S_FG⟩, + subset_closure.trans (le_sup_right : _ ≤ S) (mem_singleton m), ⟨le_sup_left, ?_⟩⟩ + simp only [Subtype.mk_le_mk, PartialEquiv.le_def, g_eq] + rfl + +/-- The Fraïssé limit of a class is unique, in that any two Fraïssé limits are isomorphic. -/ +theorem nonempty_equiv : Nonempty (M ≃[L] N) := by + let S : L.Substructure M := ⊥ + have S_fg : FG L S := (fg_iff_structure_fg _).1 Substructure.fg_bot + obtain ⟨_, ⟨emb_S : S ↪[L] N⟩⟩ : ⟨S, inferInstance⟩ ∈ L.age N := by + rw [hN.age, ← hM.age] + exact ⟨S_fg, ⟨subtype _⟩⟩ + let v : M ≃ₚ[L] N := { + dom := S + cod := emb_S.toHom.range + toEquiv := emb_S.equivRange + } + exact ⟨Exists.choose (equiv_between_cg cg_of_countable cg_of_countable + ⟨v, ((Substructure.fg_iff_structure_fg _).2 S_fg)⟩ (hM.isExtensionPair hN) + (hN.isExtensionPair hM))⟩ + end IsFraisseLimit end Language diff --git a/Mathlib/ModelTheory/Substructures.lean b/Mathlib/ModelTheory/Substructures.lean index a4ac5cf414f99..a91af2f676773 100644 --- a/Mathlib/ModelTheory/Substructures.lean +++ b/Mathlib/ModelTheory/Substructures.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Gabin Kolly -/ +import Mathlib.Data.Fintype.Order import Mathlib.Order.Closure import Mathlib.ModelTheory.Semantics import Mathlib.ModelTheory.Encoding @@ -341,7 +342,7 @@ theorem closure_univ : closure L (univ : Set M) = ⊤ := theorem closure_union (s t : Set M) : closure L (s ∪ t) = closure L s ⊔ closure L t := (Substructure.gi L M).gc.l_sup -theorem closure_unionᵢ {ι} (s : ι → Set M) : closure L (⋃ i, s i) = ⨆ i, closure L (s i) := +theorem closure_iUnion {ι} (s : ι → Set M) : closure L (⋃ i, s i) = ⨆ i, closure L (s i) := (Substructure.gi L M).gc.l_iSup instance small_bot : Small.{u} (⊥ : L.Substructure M) := by @@ -349,6 +350,28 @@ instance small_bot : Small.{u} (⊥ : L.Substructure M) := by haveI : Small.{u} (∅ : Set M) := small_subsingleton _ exact Substructure.small_closure +theorem iSup_eq_closure {ι : Sort*} (S : ι → L.Substructure M) : + ⨆ i, S i = closure L (⋃ i, (S i : Set M)) := by simp_rw [closure_iUnion, closure_eq] + +-- This proof uses the fact that `Substructure.closure` is finitary. +theorem mem_iSup_of_directed {ι : Type*} [hι : Nonempty ι] {S : ι → L.Substructure M} + (hS : Directed (· ≤ ·) S) {x : M} : + x ∈ ⨆ i, S i ↔ ∃ i, x ∈ S i := by + refine ⟨?_, fun ⟨i, hi⟩ ↦ le_iSup S i hi⟩ + suffices x ∈ closure L (⋃ i, (S i : Set M)) → ∃ i, x ∈ S i by + simpa only [closure_iUnion, closure_eq (S _)] using this + refine fun hx ↦ closure_induction hx (fun _ ↦ mem_iUnion.1) (fun f v hC ↦ ?_) + simp_rw [Set.mem_setOf] at * + have ⟨i, hi⟩ := hS.finite_le (fun i ↦ Classical.choose (hC i)) + refine ⟨i, (S i).fun_mem f v (fun j ↦ hi j (Classical.choose_spec (hC j)))⟩ + +-- This proof uses the fact that `Substructure.closure` is finitary. +theorem mem_sSup_of_directedOn {S : Set (L.Substructure M)} (Sne : S.Nonempty) + (hS : DirectedOn (· ≤ ·) S) {x : M} : + x ∈ sSup S ↔ ∃ s ∈ S, x ∈ s := by + haveI : Nonempty S := Sne.to_subtype + simp only [sSup_eq_iSup', mem_iSup_of_directed hS.directed_val, Subtype.exists, exists_prop] + /-! ### `comap` and `map` -/ @@ -437,7 +460,7 @@ theorem map_sup (S T : L.Substructure M) (f : M →[L] N) : (S ⊔ T).map f = S. (gc_map_comap f).l_sup theorem map_iSup {ι : Sort*} (f : M →[L] N) (s : ι → L.Substructure M) : - (iSup s).map f = ⨆ i, (s i).map f := + (⨆ i, s i).map f = ⨆ i, (s i).map f := (gc_map_comap f).l_iSup theorem comap_inf (S T : L.Substructure N) (f : M →[L] N) : @@ -445,7 +468,7 @@ theorem comap_inf (S T : L.Substructure N) (f : M →[L] N) : (gc_map_comap f).u_inf theorem comap_iInf {ι : Sort*} (f : M →[L] N) (s : ι → L.Substructure N) : - (iInf s).comap f = ⨅ i, (s i).comap f := + (⨅ i, s i).comap f = ⨅ i, (s i).comap f := (gc_map_comap f).u_iInf @[simp] @@ -493,14 +516,14 @@ theorem comap_inf_map_of_injective (S T : L.Substructure M) : (S.map f ⊓ T.map (gciMapComap hf).u_inf_l _ _ theorem comap_iInf_map_of_injective (S : ι → L.Substructure M) : - (⨅ i, (S i).map f).comap f = iInf S := + (⨅ i, (S i).map f).comap f = ⨅ i, S i := (gciMapComap hf).u_iInf_l _ theorem comap_sup_map_of_injective (S T : L.Substructure M) : (S.map f ⊔ T.map f).comap f = S ⊔ T := (gciMapComap hf).u_sup_l _ _ theorem comap_iSup_map_of_injective (S : ι → L.Substructure M) : - (⨆ i, (S i).map f).comap f = iSup S := + (⨆ i, (S i).map f).comap f = ⨆ i, S i := (gciMapComap hf).u_iSup_l _ theorem map_le_map_iff_of_injective {S T : L.Substructure M} : S.map f ≤ T.map f ↔ S ≤ T := @@ -536,7 +559,7 @@ theorem map_inf_comap_of_surjective (S T : L.Substructure N) : (giMapComap hf).l_inf_u _ _ theorem map_iInf_comap_of_surjective (S : ι → L.Substructure N) : - (⨅ i, (S i).comap f).map f = iInf S := + (⨅ i, (S i).comap f).map f = ⨅ i, S i := (giMapComap hf).l_iInf_u _ theorem map_sup_comap_of_surjective (S T : L.Substructure N) : @@ -544,7 +567,7 @@ theorem map_sup_comap_of_surjective (S T : L.Substructure N) : (giMapComap hf).l_sup_u _ _ theorem map_iSup_comap_of_surjective (S : ι → L.Substructure N) : - (⨆ i, (S i).comap f).map f = iSup S := + (⨆ i, (S i).comap f).map f = ⨆ i, S i := (giMapComap hf).l_iSup_u _ theorem comap_le_comap_iff_of_surjective {S T : L.Substructure N} : S.comap f ≤ T.comap f ↔ S ≤ T := From 07a0ae9126dff1d2838a59b03ac6763e22ed6cb4 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 16 Aug 2024 07:51:37 +0000 Subject: [PATCH 320/359] feat(Measure/Restrict): add `QuasiMeasurePreserving.restrict` (#15661) --- Mathlib/MeasureTheory/Measure/Restrict.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index 0acae1a750f6c..515994e89c22a 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -370,6 +370,20 @@ theorem exists_mem_of_measure_ne_zero_of_ae (hs : μ s ≠ 0) {p : α → Prop} rw [← μ.restrict_apply_self, ← frequently_ae_mem_iff] at hs exact (hs.and_eventually hp).exists +/-- If a quasi measure preserving map `f` maps a set `s` to a set `t`, +then it is quasi measure preserving with respect to the restrictions of the measures. -/ +theorem QuasiMeasurePreserving.restrict {ν : Measure β} {f : α → β} + (hf : QuasiMeasurePreserving f μ ν) {t : Set β} (hmaps : MapsTo f s t) : + QuasiMeasurePreserving f (μ.restrict s) (ν.restrict t) where + measurable := hf.measurable + absolutelyContinuous := by + refine AbsolutelyContinuous.mk fun u hum ↦ ?_ + suffices ν (u ∩ t) = 0 → μ (f ⁻¹' u ∩ s) = 0 by simpa [hum, hf.measurable, hf.measurable hum] + refine fun hu ↦ measure_mono_null ?_ (hf.preimage_null hu) + rw [preimage_inter] + gcongr + assumption + /-! ### Extensionality results -/ /-- Two measures are equal if they have equal restrictions on a spanning collection of sets From c9122d8d32fb4410f5d0e503be249d2f1984c554 Mon Sep 17 00:00:00 2001 From: kkytola Date: Fri, 16 Aug 2024 08:13:01 +0000 Subject: [PATCH 321/359] chore: a clean-up of files on finite measures, probability measures, and portmanteau theorem (#12822) --- .../MeasureTheory/Measure/FiniteMeasure.lean | 172 ++++++++---------- .../MeasureTheory/Measure/Portmanteau.lean | 147 +++++++-------- .../Measure/ProbabilityMeasure.lean | 81 ++++----- 3 files changed, 175 insertions(+), 225 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean index 788da2c597187..ac8e1902e55e3 100644 --- a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean @@ -81,14 +81,7 @@ weak convergence of measures, finite measure noncomputable section -open MeasureTheory - -open Set - -open Filter - -open BoundedContinuousFunction - +open MeasureTheory Set Filter BoundedContinuousFunction open scoped Topology ENNReal NNReal BoundedContinuousFunction namespace MeasureTheory @@ -125,15 +118,12 @@ def _root_.MeasureTheory.FiniteMeasure (Ω : Type*) [MeasurableSpace Ω] : Type def toMeasure : FiniteMeasure Ω → Measure Ω := Subtype.val /-- A finite measure can be interpreted as a measure. -/ -instance instCoe : Coe (FiniteMeasure Ω) (MeasureTheory.Measure Ω) where - coe := toMeasure +instance instCoe : Coe (FiniteMeasure Ω) (MeasureTheory.Measure Ω) := { coe := toMeasure } -instance isFiniteMeasure (μ : FiniteMeasure Ω) : IsFiniteMeasure (μ : Measure Ω) := - μ.prop +instance isFiniteMeasure (μ : FiniteMeasure Ω) : IsFiniteMeasure (μ : Measure Ω) := μ.prop @[simp] -theorem val_eq_toMeasure (ν : FiniteMeasure Ω) : ν.val = (ν : Measure Ω) := - rfl +theorem val_eq_toMeasure (ν : FiniteMeasure Ω) : ν.val = (ν : Measure Ω) := rfl theorem toMeasure_injective : Function.Injective ((↑) : FiniteMeasure Ω → Measure Ω) := Subtype.coe_injective @@ -157,15 +147,18 @@ theorem ennreal_coeFn_eq_coeFn_toMeasure (ν : FiniteMeasure Ω) (s : Set Ω) : (ν s : ℝ≥0∞) = (ν : Measure Ω) s := ENNReal.coe_toNNReal (measure_lt_top (↑ν) s).ne -theorem apply_mono (μ : FiniteMeasure Ω) {s₁ s₂ : Set Ω} (h : s₁ ⊆ s₂) : μ s₁ ≤ μ s₂ := by - change ((μ : Measure Ω) s₁).toNNReal ≤ ((μ : Measure Ω) s₂).toNNReal - have key : (μ : Measure Ω) s₁ ≤ (μ : Measure Ω) s₂ := (μ : Measure Ω).mono h - apply (ENNReal.toNNReal_le_toNNReal (measure_ne_top _ s₁) (measure_ne_top _ s₂)).mpr key +@[simp] +theorem null_iff_toMeasure_null (ν : FiniteMeasure Ω) (s : Set Ω) : + ν s = 0 ↔ (ν : Measure Ω) s = 0 := + ⟨fun h ↦ by rw [← ennreal_coeFn_eq_coeFn_toMeasure, h, ENNReal.coe_zero], + fun h ↦ congrArg ENNReal.toNNReal h⟩ + +theorem apply_mono (μ : FiniteMeasure Ω) {s₁ s₂ : Set Ω} (h : s₁ ⊆ s₂) : μ s₁ ≤ μ s₂ := + ENNReal.toNNReal_mono (measure_ne_top _ s₂) ((μ : Measure Ω).mono h) /-- The (total) mass of a finite measure `μ` is `μ univ`, i.e., the cast to `NNReal` of `(μ : measure Ω) univ`. -/ -def mass (μ : FiniteMeasure Ω) : ℝ≥0 := - μ univ +def mass (μ : FiniteMeasure Ω) : ℝ≥0 := μ univ @[simp] theorem apply_le_mass (μ : FiniteMeasure Ω) (s : Set Ω) : μ s ≤ μ.mass := by simpa using apply_mono μ (subset_univ s) @@ -179,8 +172,7 @@ instance instZero : Zero (FiniteMeasure Ω) where zero := ⟨0, MeasureTheory.is @[simp, norm_cast] lemma coeFn_zero : ⇑(0 : FiniteMeasure Ω) = 0 := rfl @[simp] -theorem zero_mass : (0 : FiniteMeasure Ω).mass = 0 := - rfl +theorem zero_mass : (0 : FiniteMeasure Ω).mass = 0 := rfl @[simp] theorem mass_zero_iff (μ : FiniteMeasure Ω) : μ.mass = 0 ↔ μ = 0 := by @@ -189,9 +181,8 @@ theorem mass_zero_iff (μ : FiniteMeasure Ω) : μ.mass = 0 ↔ μ = 0 := by apply Measure.measure_univ_eq_zero.mp rwa [← ennreal_mass, ENNReal.coe_eq_zero] -theorem mass_nonzero_iff (μ : FiniteMeasure Ω) : μ.mass ≠ 0 ↔ μ ≠ 0 := by - rw [not_iff_not] - exact FiniteMeasure.mass_zero_iff μ +theorem mass_nonzero_iff (μ : FiniteMeasure Ω) : μ.mass ≠ 0 ↔ μ ≠ 0 := + not_iff_not.mpr <| FiniteMeasure.mass_zero_iff μ @[ext] theorem eq_of_forall_toMeasure_apply_eq (μ ν : FiniteMeasure Ω) @@ -205,8 +196,7 @@ theorem eq_of_forall_apply_eq (μ ν : FiniteMeasure Ω) ext1 s s_mble simpa [ennreal_coeFn_eq_coeFn_toMeasure] using congr_arg ((↑) : ℝ≥0 → ℝ≥0∞) (h s s_mble) -instance instInhabited : Inhabited (FiniteMeasure Ω) := - ⟨0⟩ +instance instInhabited : Inhabited (FiniteMeasure Ω) := ⟨0⟩ instance instAdd : Add (FiniteMeasure Ω) where add μ ν := ⟨μ + ν, MeasureTheory.isFiniteMeasureAdd⟩ @@ -217,13 +207,11 @@ instance instSMul : SMul R (FiniteMeasure Ω) where smul (c : R) μ := ⟨c • (μ : Measure Ω), MeasureTheory.isFiniteMeasureSMulOfNNRealTower⟩ @[simp, norm_cast] -theorem toMeasure_zero : ((↑) : FiniteMeasure Ω → Measure Ω) 0 = 0 := - rfl +theorem toMeasure_zero : ((↑) : FiniteMeasure Ω → Measure Ω) 0 = 0 := rfl -- Porting note: with `simp` here the `coeFn` lemmas below fall prey to `simpNF`: the LHS simplifies @[norm_cast] -theorem toMeasure_add (μ ν : FiniteMeasure Ω) : ↑(μ + ν) = (↑μ + ↑ν : Measure Ω) := - rfl +theorem toMeasure_add (μ ν : FiniteMeasure Ω) : ↑(μ + ν) = (↑μ + ↑ν : Measure Ω) := rfl @[simp, norm_cast] theorem toMeasure_smul (c : R) (μ : FiniteMeasure Ω) : ↑(c • μ) = c • (μ : Measure Ω) := @@ -242,7 +230,7 @@ theorem coeFn_smul [IsScalarTower R ℝ≥0 ℝ≥0] (c : R) (μ : FiniteMeasure funext; simp [← ENNReal.coe_inj, ENNReal.coe_smul] instance instAddCommMonoid : AddCommMonoid (FiniteMeasure Ω) := - toMeasure_injective.addCommMonoid (↑) toMeasure_zero toMeasure_add fun _ _ => toMeasure_smul _ _ + toMeasure_injective.addCommMonoid (↑) toMeasure_zero toMeasure_add fun _ _ ↦ toMeasure_smul _ _ /-- Coercion is an `AddMonoidHom`. -/ @[simps] @@ -265,8 +253,7 @@ def restrict (μ : FiniteMeasure Ω) (A : Set Ω) : FiniteMeasure Ω where property := MeasureTheory.isFiniteMeasureRestrict (μ : Measure Ω) A theorem restrict_measure_eq (μ : FiniteMeasure Ω) (A : Set Ω) : - (μ.restrict A : Measure Ω) = (μ : Measure Ω).restrict A := - rfl + (μ.restrict A : Measure Ω) = (μ : Measure Ω).restrict A := rfl theorem restrict_apply_measure (μ : FiniteMeasure Ω) (A : Set Ω) {s : Set Ω} (s_mble : MeasurableSet s) : (μ.restrict A : Measure Ω) s = (μ : Measure Ω) (s ∩ A) := @@ -358,8 +345,7 @@ theorem testAgainstNN_smul [IsScalarTower R ℝ≥0 ℝ≥0] [PseudoMetricSpace ENNReal.coe_smul] simp_rw [← smul_one_smul ℝ≥0∞ c (f _ : ℝ≥0∞), ← smul_one_smul ℝ≥0∞ c (lintegral _ _ : ℝ≥0∞), smul_eq_mul] - exact - @lintegral_const_mul _ _ (μ : Measure Ω) (c • (1 : ℝ≥0∞)) _ f.measurable_coe_ennreal_comp + exact lintegral_const_mul (c • (1 : ℝ≥0∞)) f.measurable_coe_ennreal_comp theorem testAgainstNN_lipschitz_estimate (μ : FiniteMeasure Ω) (f g : Ω →ᵇ ℝ≥0) : μ.testAgainstNN f ≤ μ.testAgainstNN g + nndist f g * μ.mass := by @@ -370,32 +356,29 @@ theorem testAgainstNN_lipschitz_estimate (μ : FiniteMeasure Ω) (f g : Ω → have le_dist : ∀ ω, dist (f ω) (g ω) ≤ nndist f g := BoundedContinuousFunction.dist_coe_le_dist intro ω have le' : f ω ≤ g ω + nndist f g := by - apply (NNReal.le_add_nndist (f ω) (g ω)).trans - rw [add_le_add_iff_left] - exact dist_le_coe.mp (le_dist ω) + calc f ω + _ ≤ g ω + nndist (f ω) (g ω) := NNReal.le_add_nndist (f ω) (g ω) + _ ≤ g ω + nndist f g := (add_le_add_iff_left (g ω)).mpr (le_dist ω) have le : (f ω : ℝ≥0∞) ≤ (g ω : ℝ≥0∞) + nndist f g := by - rw [← ENNReal.coe_add] - exact ENNReal.coe_mono le' + simpa only [← ENNReal.coe_add] using (by exact_mod_cast le') rwa [coe_nnreal_ennreal_nndist] at le theorem testAgainstNN_lipschitz (μ : FiniteMeasure Ω) : - LipschitzWith μ.mass fun f : Ω →ᵇ ℝ≥0 => μ.testAgainstNN f := by + LipschitzWith μ.mass fun f : Ω →ᵇ ℝ≥0 ↦ μ.testAgainstNN f := by rw [lipschitzWith_iff_dist_le_mul] intro f₁ f₂ suffices abs (μ.testAgainstNN f₁ - μ.testAgainstNN f₂ : ℝ) ≤ μ.mass * dist f₁ f₂ by rwa [NNReal.dist_eq] apply abs_le.mpr constructor - · have key' := μ.testAgainstNN_lipschitz_estimate f₂ f₁ - rw [mul_comm] at key' + · have key := μ.testAgainstNN_lipschitz_estimate f₂ f₁ + rw [mul_comm] at key suffices ↑(μ.testAgainstNN f₂) ≤ ↑(μ.testAgainstNN f₁) + ↑μ.mass * dist f₁ f₂ by linarith - have key := NNReal.coe_mono key' - rwa [NNReal.coe_add, NNReal.coe_mul, nndist_comm] at key - · have key' := μ.testAgainstNN_lipschitz_estimate f₁ f₂ - rw [mul_comm] at key' + simpa [nndist_comm] using NNReal.coe_mono key + · have key := μ.testAgainstNN_lipschitz_estimate f₁ f₂ + rw [mul_comm] at key suffices ↑(μ.testAgainstNN f₁) ≤ ↑(μ.testAgainstNN f₂) + ↑μ.mass * dist f₁ f₂ by linarith - have key := NNReal.coe_mono key' - rwa [NNReal.coe_add, NNReal.coe_mul] at key + simpa using NNReal.coe_mono key /-- Finite measures yield elements of the `WeakDual` of bounded continuous nonnegative functions via `MeasureTheory.FiniteMeasure.testAgainstNN`, i.e., integration. -/ @@ -411,8 +394,7 @@ theorem coe_toWeakDualBCNN (μ : FiniteMeasure Ω) : ⇑μ.toWeakDualBCNN = μ.t @[simp] theorem toWeakDualBCNN_apply (μ : FiniteMeasure Ω) (f : Ω →ᵇ ℝ≥0) : - μ.toWeakDualBCNN f = (∫⁻ x, f x ∂(μ : Measure Ω)).toNNReal := - rfl + μ.toWeakDualBCNN f = (∫⁻ x, f x ∂(μ : Measure Ω)).toNNReal := rfl /-- The topology of weak convergence on `MeasureTheory.FiniteMeasure Ω` is inherited (induced) from the weak-* topology on `WeakDual ℝ≥0 (Ω →ᵇ ℝ≥0)` via the function @@ -426,8 +408,8 @@ theorem toWeakDualBCNN_continuous : Continuous (@toWeakDualBCNN Ω _ _ _) := /-- Integration of (nonnegative bounded continuous) test functions against finite Borel measures depends continuously on the measure. -/ theorem continuous_testAgainstNN_eval (f : Ω →ᵇ ℝ≥0) : - Continuous fun μ : FiniteMeasure Ω => μ.testAgainstNN f := by - show Continuous ((fun φ : WeakDual ℝ≥0 (Ω →ᵇ ℝ≥0) => φ f) ∘ toWeakDualBCNN) + Continuous fun μ : FiniteMeasure Ω ↦ μ.testAgainstNN f := by + show Continuous ((fun φ : WeakDual ℝ≥0 (Ω →ᵇ ℝ≥0) ↦ φ f) ∘ toWeakDualBCNN) refine Continuous.comp ?_ (toWeakDualBCNN_continuous (Ω := Ω)) exact WeakBilin.eval_continuous (𝕜 := ℝ≥0) (E := (Ω →ᵇ ℝ≥0) →L[ℝ≥0] ℝ≥0) _ _ /- porting note: without explicitly providing `𝕜` and `E` TC synthesis times @@ -435,29 +417,29 @@ theorem continuous_testAgainstNN_eval (f : Ω →ᵇ ℝ≥0) : `set_option synthInstance.maxHeartbeats 47000` was sufficient. -/ /-- The total mass of a finite measure depends continuously on the measure. -/ -theorem continuous_mass : Continuous fun μ : FiniteMeasure Ω => μ.mass := by +theorem continuous_mass : Continuous fun μ : FiniteMeasure Ω ↦ μ.mass := by simp_rw [← testAgainstNN_one]; exact continuous_testAgainstNN_eval 1 /-- Convergence of finite measures implies the convergence of their total masses. -/ theorem _root_.Filter.Tendsto.mass {γ : Type*} {F : Filter γ} {μs : γ → FiniteMeasure Ω} - {μ : FiniteMeasure Ω} (h : Tendsto μs F (𝓝 μ)) : Tendsto (fun i => (μs i).mass) F (𝓝 μ.mass) := + {μ : FiniteMeasure Ω} (h : Tendsto μs F (𝓝 μ)) : Tendsto (fun i ↦ (μs i).mass) F (𝓝 μ.mass) := (continuous_mass.tendsto μ).comp h -theorem tendsto_iff_weak_star_tendsto {γ : Type*} {F : Filter γ} {μs : γ → FiniteMeasure Ω} +theorem tendsto_iff_weakDual_tendsto {γ : Type*} {F : Filter γ} {μs : γ → FiniteMeasure Ω} {μ : FiniteMeasure Ω} : - Tendsto μs F (𝓝 μ) ↔ Tendsto (fun i => (μs i).toWeakDualBCNN) F (𝓝 μ.toWeakDualBCNN) := + Tendsto μs F (𝓝 μ) ↔ Tendsto (fun i ↦ (μs i).toWeakDualBCNN) F (𝓝 μ.toWeakDualBCNN) := Inducing.tendsto_nhds_iff ⟨rfl⟩ theorem tendsto_iff_forall_toWeakDualBCNN_tendsto {γ : Type*} {F : Filter γ} {μs : γ → FiniteMeasure Ω} {μ : FiniteMeasure Ω} : Tendsto μs F (𝓝 μ) ↔ - ∀ f : Ω →ᵇ ℝ≥0, Tendsto (fun i => (μs i).toWeakDualBCNN f) F (𝓝 (μ.toWeakDualBCNN f)) := by - rw [tendsto_iff_weak_star_tendsto, tendsto_iff_forall_eval_tendsto_topDualPairing]; rfl + ∀ f : Ω →ᵇ ℝ≥0, Tendsto (fun i ↦ (μs i).toWeakDualBCNN f) F (𝓝 (μ.toWeakDualBCNN f)) := by + rw [tendsto_iff_weakDual_tendsto, tendsto_iff_forall_eval_tendsto_topDualPairing]; rfl theorem tendsto_iff_forall_testAgainstNN_tendsto {γ : Type*} {F : Filter γ} {μs : γ → FiniteMeasure Ω} {μ : FiniteMeasure Ω} : Tendsto μs F (𝓝 μ) ↔ - ∀ f : Ω →ᵇ ℝ≥0, Tendsto (fun i => (μs i).testAgainstNN f) F (𝓝 (μ.testAgainstNN f)) := by + ∀ f : Ω →ᵇ ℝ≥0, Tendsto (fun i ↦ (μs i).testAgainstNN f) F (𝓝 (μ.testAgainstNN f)) := by rw [FiniteMeasure.tendsto_iff_forall_toWeakDualBCNN_tendsto]; rfl /-- If the total masses of finite measures tend to zero, then the measures tend to @@ -465,23 +447,21 @@ zero. This formulation concerns the associated functionals on bounded continuous nonnegative test functions. See `MeasureTheory.FiniteMeasure.tendsto_zero_of_tendsto_zero_mass` for a formulation stating the weak convergence of measures. -/ theorem tendsto_zero_testAgainstNN_of_tendsto_zero_mass {γ : Type*} {F : Filter γ} - {μs : γ → FiniteMeasure Ω} (mass_lim : Tendsto (fun i => (μs i).mass) F (𝓝 0)) (f : Ω →ᵇ ℝ≥0) : - Tendsto (fun i => (μs i).testAgainstNN f) F (𝓝 0) := by + {μs : γ → FiniteMeasure Ω} (mass_lim : Tendsto (fun i ↦ (μs i).mass) F (𝓝 0)) (f : Ω →ᵇ ℝ≥0) : + Tendsto (fun i ↦ (μs i).testAgainstNN f) F (𝓝 0) := by apply tendsto_iff_dist_tendsto_zero.mpr - have obs := fun i => (μs i).testAgainstNN_lipschitz_estimate f 0 + have obs := fun i ↦ (μs i).testAgainstNN_lipschitz_estimate f 0 simp_rw [testAgainstNN_zero, zero_add] at obs simp_rw [show ∀ i, dist ((μs i).testAgainstNN f) 0 = (μs i).testAgainstNN f by simp only [dist_nndist, NNReal.nndist_zero_eq_val', eq_self_iff_true, imp_true_iff]] - refine squeeze_zero (fun i => NNReal.coe_nonneg _) obs ?_ - have lim_pair : Tendsto (fun i => (⟨nndist f 0, (μs i).mass⟩ : ℝ × ℝ)) F (𝓝 ⟨nndist f 0, 0⟩) := by - refine (Prod.tendsto_iff _ _).mpr ⟨tendsto_const_nhds, ?_⟩ - exact (NNReal.continuous_coe.tendsto 0).comp mass_lim - have key := tendsto_mul.comp lim_pair - rwa [mul_zero] at key + apply squeeze_zero (fun i ↦ NNReal.coe_nonneg _) obs + have lim_pair : Tendsto (fun i ↦ (⟨nndist f 0, (μs i).mass⟩ : ℝ × ℝ)) F (𝓝 ⟨nndist f 0, 0⟩) := + (Prod.tendsto_iff _ _).mpr ⟨tendsto_const_nhds, (NNReal.continuous_coe.tendsto 0).comp mass_lim⟩ + simpa using tendsto_mul.comp lim_pair /-- If the total masses of finite measures tend to zero, then the measures tend to zero. -/ theorem tendsto_zero_of_tendsto_zero_mass {γ : Type*} {F : Filter γ} {μs : γ → FiniteMeasure Ω} - (mass_lim : Tendsto (fun i => (μs i).mass) F (𝓝 0)) : Tendsto μs F (𝓝 0) := by + (mass_lim : Tendsto (fun i ↦ (μs i).mass) F (𝓝 0)) : Tendsto μs F (𝓝 0) := by rw [tendsto_iff_forall_testAgainstNN_tendsto] intro f convert tendsto_zero_testAgainstNN_of_tendsto_zero_mass mass_lim f @@ -493,7 +473,7 @@ theorem tendsto_iff_forall_lintegral_tendsto {γ : Type*} {F : Filter γ} {μs : {μ : FiniteMeasure Ω} : Tendsto μs F (𝓝 μ) ↔ ∀ f : Ω →ᵇ ℝ≥0, - Tendsto (fun i => ∫⁻ x, f x ∂(μs i : Measure Ω)) F (𝓝 (∫⁻ x, f x ∂(μ : Measure Ω))) := by + Tendsto (fun i ↦ ∫⁻ x, f x ∂(μs i : Measure Ω)) F (𝓝 (∫⁻ x, f x ∂(μ : Measure Ω))) := by rw [tendsto_iff_forall_toWeakDualBCNN_tendsto] simp_rw [toWeakDualBCNN_apply _ _, ← testAgainstNN_coe_eq, ENNReal.tendsto_coe, ENNReal.toNNReal_coe] @@ -557,10 +537,10 @@ A related result with more general assumptions is -/ theorem tendsto_lintegral_nn_of_le_const (μ : FiniteMeasure Ω) {fs : ℕ → Ω →ᵇ ℝ≥0} {c : ℝ≥0} (fs_le_const : ∀ n ω, fs n ω ≤ c) {f : Ω → ℝ≥0} - (fs_lim : ∀ ω, Tendsto (fun n => fs n ω) atTop (𝓝 (f ω))) : - Tendsto (fun n => ∫⁻ ω, fs n ω ∂(μ : Measure Ω)) atTop (𝓝 (∫⁻ ω, f ω ∂(μ : Measure Ω))) := + (fs_lim : ∀ ω, Tendsto (fun n ↦ fs n ω) atTop (𝓝 (f ω))) : + Tendsto (fun n ↦ ∫⁻ ω, fs n ω ∂(μ : Measure Ω)) atTop (𝓝 (∫⁻ ω, f ω ∂(μ : Measure Ω))) := tendsto_lintegral_nn_filter_of_le_const μ - (eventually_of_forall fun n => eventually_of_forall (fs_le_const n)) + (eventually_of_forall fun n ↦ eventually_of_forall (fs_le_const n)) (eventually_of_forall fs_lim) /-- A bounded convergence theorem for a finite measure: @@ -579,8 +559,8 @@ A related result using `MeasureTheory.lintegral` for integration is theorem tendsto_testAgainstNN_filter_of_le_const {ι : Type*} {L : Filter ι} [L.IsCountablyGenerated] {μ : FiniteMeasure Ω} {fs : ι → Ω →ᵇ ℝ≥0} {c : ℝ≥0} (fs_le_const : ∀ᶠ i in L, ∀ᵐ ω : Ω ∂(μ : Measure Ω), fs i ω ≤ c) {f : Ω →ᵇ ℝ≥0} - (fs_lim : ∀ᵐ ω : Ω ∂(μ : Measure Ω), Tendsto (fun i => fs i ω) L (𝓝 (f ω))) : - Tendsto (fun i => μ.testAgainstNN (fs i)) L (𝓝 (μ.testAgainstNN f)) := by + (fs_lim : ∀ᵐ ω : Ω ∂(μ : Measure Ω), Tendsto (fun i ↦ fs i ω) L (𝓝 (f ω))) : + Tendsto (fun i ↦ μ.testAgainstNN (fs i)) L (𝓝 (μ.testAgainstNN f)) := by apply (ENNReal.tendsto_toNNReal (f.lintegral_lt_top_of_nnreal (μ : Measure Ω)).ne).comp exact tendsto_lintegral_nn_filter_of_le_const μ fs_le_const fs_lim @@ -597,10 +577,10 @@ Related results: -/ theorem tendsto_testAgainstNN_of_le_const {μ : FiniteMeasure Ω} {fs : ℕ → Ω →ᵇ ℝ≥0} {c : ℝ≥0} (fs_le_const : ∀ n ω, fs n ω ≤ c) {f : Ω →ᵇ ℝ≥0} - (fs_lim : ∀ ω, Tendsto (fun n => fs n ω) atTop (𝓝 (f ω))) : - Tendsto (fun n => μ.testAgainstNN (fs n)) atTop (𝓝 (μ.testAgainstNN f)) := + (fs_lim : ∀ ω, Tendsto (fun n ↦ fs n ω) atTop (𝓝 (f ω))) : + Tendsto (fun n ↦ μ.testAgainstNN (fs n)) atTop (𝓝 (μ.testAgainstNN f)) := tendsto_testAgainstNN_filter_of_le_const - (eventually_of_forall fun n => eventually_of_forall (fs_le_const n)) + (eventually_of_forall fun n ↦ eventually_of_forall (fs_le_const n)) (eventually_of_forall fs_lim) end FiniteMeasureBoundedConvergence @@ -619,27 +599,23 @@ variable {Ω : Type*} [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurabl theorem tendsto_of_forall_integral_tendsto {γ : Type*} {F : Filter γ} {μs : γ → FiniteMeasure Ω} {μ : FiniteMeasure Ω} - (h : - ∀ f : Ω →ᵇ ℝ, - Tendsto (fun i => ∫ x, f x ∂(μs i : Measure Ω)) F (𝓝 (∫ x, f x ∂(μ : Measure Ω)))) : + (h : ∀ f : Ω →ᵇ ℝ, + Tendsto (fun i ↦ ∫ x, f x ∂(μs i : Measure Ω)) F (𝓝 (∫ x, f x ∂(μ : Measure Ω)))) : Tendsto μs F (𝓝 μ) := by - apply (@tendsto_iff_forall_lintegral_tendsto Ω _ _ _ γ F μs μ).mpr + apply tendsto_iff_forall_lintegral_tendsto.mpr intro f - have key := - @ENNReal.tendsto_toReal_iff _ F _ - (fun i => (f.lintegral_lt_top_of_nnreal (μs i)).ne) _ (f.lintegral_lt_top_of_nnreal μ).ne - simp only [ENNReal.ofReal_coe_nnreal] at key - apply key.mp + apply (ENNReal.tendsto_toReal_iff (fi := F) + (fun i ↦ (f.lintegral_lt_top_of_nnreal (μs i)).ne) (f.lintegral_lt_top_of_nnreal μ).ne).mp have lip : LipschitzWith 1 ((↑) : ℝ≥0 → ℝ) := isometry_subtype_coe.lipschitz set f₀ := BoundedContinuousFunction.comp _ lip f with _def_f₀ have f₀_eq : ⇑f₀ = ((↑) : ℝ≥0 → ℝ) ∘ ⇑f := rfl - have f₀_nn : 0 ≤ ⇑f₀ := fun _ => by + have f₀_nn : 0 ≤ ⇑f₀ := fun _ ↦ by simp only [f₀_eq, Pi.zero_apply, Function.comp_apply, NNReal.zero_le_coe] have f₀_ae_nn : 0 ≤ᵐ[(μ : Measure Ω)] ⇑f₀ := eventually_of_forall f₀_nn - have f₀_ae_nns : ∀ i, 0 ≤ᵐ[(μs i : Measure Ω)] ⇑f₀ := fun i => eventually_of_forall f₀_nn + have f₀_ae_nns : ∀ i, 0 ≤ᵐ[(μs i : Measure Ω)] ⇑f₀ := fun i ↦ eventually_of_forall f₀_nn have aux := integral_eq_lintegral_of_nonneg_ae f₀_ae_nn f₀.continuous.measurable.aestronglyMeasurable - have auxs := fun i => + have auxs := fun i ↦ integral_eq_lintegral_of_nonneg_ae (f₀_ae_nns i) f₀.continuous.measurable.aestronglyMeasurable simp_rw [f₀_eq, Function.comp_apply, ENNReal.ofReal_coe_nnreal] at aux auxs simpa only [← aux, ← auxs] using h f₀ @@ -650,7 +626,7 @@ theorem tendsto_iff_forall_integral_tendsto {γ : Type*} {F : Filter γ} {μs : {μ : FiniteMeasure Ω} : Tendsto μs F (𝓝 μ) ↔ ∀ f : Ω →ᵇ ℝ, - Tendsto (fun i => ∫ x, f x ∂(μs i : Measure Ω)) F (𝓝 (∫ x, f x ∂(μ : Measure Ω))) := by + Tendsto (fun i ↦ ∫ x, f x ∂(μs i : Measure Ω)) F (𝓝 (∫ x, f x ∂(μ : Measure Ω))) := by refine ⟨?_, tendsto_of_forall_integral_tendsto⟩ rw [tendsto_iff_forall_lintegral_tendsto] intro h f @@ -661,9 +637,9 @@ theorem tendsto_iff_forall_integral_tendsto {γ : Type*} {F : Filter γ} {μs : have tends_neg := (ENNReal.tendsto_toReal (f_neg.lintegral_lt_top_of_nnreal μ).ne).comp (h f_neg) have aux : ∀ g : Ω →ᵇ ℝ≥0, - (ENNReal.toReal ∘ fun i : γ => ∫⁻ x : Ω, ↑(g x) ∂(μs i : Measure Ω)) = fun i : γ => - (∫⁻ x : Ω, ↑(g x) ∂(μs i : Measure Ω)).toReal := - fun _ => rfl + (ENNReal.toReal ∘ fun i : γ ↦ ∫⁻ x : Ω, ↑(g x) ∂(μs i : Measure Ω)) = + fun i : γ ↦ (∫⁻ x : Ω, ↑(g x) ∂(μs i : Measure Ω)).toReal := + fun _ ↦ rfl simp_rw [aux, BoundedContinuousFunction.toReal_lintegral_coe_eq_integral] at tends_pos tends_neg exact Tendsto.sub tends_pos tends_neg @@ -695,8 +671,8 @@ lemma map_apply' (ν : FiniteMeasure Ω) {f : Ω → Ω'} (f_aemble : AEMeasurab lemma map_apply_of_aemeasurable (ν : FiniteMeasure Ω) {f : Ω → Ω'} (f_aemble : AEMeasurable f ν) {A : Set Ω'} (A_mble : MeasurableSet A) : ν.map f A = ν (f ⁻¹' A) := by - have := ν.map_apply' f_aemble A_mble - exact (ENNReal.toNNReal_eq_toNNReal_iff' (measure_ne_top _ _) (measure_ne_top _ _)).mpr this + have key := ν.map_apply' f_aemble A_mble + exact (ENNReal.toNNReal_eq_toNNReal_iff' (measure_ne_top _ _) (measure_ne_top _ _)).mpr key lemma map_apply (ν : FiniteMeasure Ω) {f : Ω → Ω'} (f_mble : Measurable f) {A : Set Ω'} (A_mble : MeasurableSet A) : diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 674c9e951c5a9..19386844f64da 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -51,12 +51,11 @@ theorem, however, is most convenient for probability measures on pseudo-emetriza their Borel sigma algebras. Some specific considerations on the assumptions in the different implications: - * `MeasureTheory.FiniteMeasure.limsup_measure_closed_le_of_tendsto` assumes - `PseudoEMetricSpace`. The only reason is to have bounded continuous pointwise approximations - to the indicator function of a closed set. Clearly for example metrizability or - pseudo-emetrizability would be sufficient assumptions. The typeclass assumptions should be later - adjusted in a way that takes into account use cases, but the proof will presumably remain - essentially the same. + * `MeasureTheory.FiniteMeasure.limsup_measure_closed_le_of_tendsto`, i.e., implication (T) → (C), + assumes that in the underlying topological space, indicator functions of closed sets have + decreasing bounded continuous pointwise approximating sequences. The assumption is in the form + of the type class `HasOuterApproxClosed`. Type class inference knows that for example the more + common assumptions of metrizability or pseudo-emetrizability suffice. * Where formulations are currently only provided for probability measures, one can obtain the finite measure formulations using the characterization of convergence of finite measures by their total masses and their probability-normalized versions, i.e., by @@ -99,11 +98,10 @@ of measures. variable {Ω : Type*} [MeasurableSpace Ω] -/-- **Portmanteau theorem** -/ theorem le_measure_compl_liminf_of_limsup_measure_le {ι : Type*} {L : Filter ι} {μ : Measure Ω} {μs : ι → Measure Ω} [IsProbabilityMeasure μ] [∀ i, IsProbabilityMeasure (μs i)] {E : Set Ω} - (E_mble : MeasurableSet E) (h : (L.limsup fun i => μs i E) ≤ μ E) : - μ Eᶜ ≤ L.liminf fun i => μs i Eᶜ := by + (E_mble : MeasurableSet E) (h : (L.limsup fun i ↦ μs i E) ≤ μ E) : + μ Eᶜ ≤ L.liminf fun i ↦ μs i Eᶜ := by rcases L.eq_or_neBot with rfl | hne · simp only [liminf_bot, le_top] have meas_Ec : μ Eᶜ = 1 - μ E := by @@ -112,24 +110,22 @@ theorem le_measure_compl_liminf_of_limsup_measure_le {ι : Type*} {L : Filter ι intro i simpa only [measure_univ] using measure_compl E_mble (measure_lt_top (μs i) E).ne simp_rw [meas_Ec, meas_i_Ec] - have obs : - (L.liminf fun i : ι => 1 - μs i E) = L.liminf ((fun x => 1 - x) ∘ fun i : ι => μs i E) := rfl - rw [obs] - have := antitone_const_tsub.map_limsup_of_continuousAt (F := L) - (fun i => μs i E) (ENNReal.continuous_sub_left ENNReal.one_ne_top).continuousAt - simp_rw [← this] - exact antitone_const_tsub h + rw [show (L.liminf fun i : ι ↦ 1 - μs i E) = L.liminf ((fun x ↦ 1 - x) ∘ fun i : ι ↦ μs i E) + from rfl] + have key := antitone_const_tsub.map_limsup_of_continuousAt (F := L) + (fun i ↦ μs i E) (ENNReal.continuous_sub_left ENNReal.one_ne_top).continuousAt + simpa [← key] using antitone_const_tsub h theorem le_measure_liminf_of_limsup_measure_compl_le {ι : Type*} {L : Filter ι} {μ : Measure Ω} {μs : ι → Measure Ω} [IsProbabilityMeasure μ] [∀ i, IsProbabilityMeasure (μs i)] {E : Set Ω} - (E_mble : MeasurableSet E) (h : (L.limsup fun i => μs i Eᶜ) ≤ μ Eᶜ) : - μ E ≤ L.liminf fun i => μs i E := + (E_mble : MeasurableSet E) (h : (L.limsup fun i ↦ μs i Eᶜ) ≤ μ Eᶜ) : + μ E ≤ L.liminf fun i ↦ μs i E := compl_compl E ▸ le_measure_compl_liminf_of_limsup_measure_le (MeasurableSet.compl E_mble) h theorem limsup_measure_compl_le_of_le_liminf_measure {ι : Type*} {L : Filter ι} {μ : Measure Ω} {μs : ι → Measure Ω} [IsProbabilityMeasure μ] [∀ i, IsProbabilityMeasure (μs i)] {E : Set Ω} - (E_mble : MeasurableSet E) (h : μ E ≤ L.liminf fun i => μs i E) : - (L.limsup fun i => μs i Eᶜ) ≤ μ Eᶜ := by + (E_mble : MeasurableSet E) (h : μ E ≤ L.liminf fun i ↦ μs i E) : + (L.limsup fun i ↦ μs i Eᶜ) ≤ μ Eᶜ := by rcases L.eq_or_neBot with rfl | hne · simp only [limsup_bot, bot_le] have meas_Ec : μ Eᶜ = 1 - μ E := by @@ -138,18 +134,16 @@ theorem limsup_measure_compl_le_of_le_liminf_measure {ι : Type*} {L : Filter ι intro i simpa only [measure_univ] using measure_compl E_mble (measure_lt_top (μs i) E).ne simp_rw [meas_Ec, meas_i_Ec] - have obs : - (L.limsup fun i : ι => 1 - μs i E) = L.limsup ((fun x => 1 - x) ∘ fun i : ι => μs i E) := rfl - rw [obs] - have := antitone_const_tsub.map_liminf_of_continuousAt (F := L) - (fun i => μs i E) (ENNReal.continuous_sub_left ENNReal.one_ne_top).continuousAt - simp_rw [← this] - exact antitone_const_tsub h + rw [show (L.limsup fun i : ι ↦ 1 - μs i E) = L.limsup ((fun x ↦ 1 - x) ∘ fun i : ι ↦ μs i E) + from rfl] + have key := antitone_const_tsub.map_liminf_of_continuousAt (F := L) + (fun i ↦ μs i E) (ENNReal.continuous_sub_left ENNReal.one_ne_top).continuousAt + simpa [← key] using antitone_const_tsub h theorem limsup_measure_le_of_le_liminf_measure_compl {ι : Type*} {L : Filter ι} {μ : Measure Ω} {μs : ι → Measure Ω} [IsProbabilityMeasure μ] [∀ i, IsProbabilityMeasure (μs i)] {E : Set Ω} - (E_mble : MeasurableSet E) (h : μ Eᶜ ≤ L.liminf fun i => μs i Eᶜ) : - (L.limsup fun i => μs i E) ≤ μ E := + (E_mble : MeasurableSet E) (h : μ Eᶜ ≤ L.liminf fun i ↦ μs i Eᶜ) : + (L.limsup fun i ↦ μs i E) ≤ μ E := compl_compl E ▸ limsup_measure_compl_le_of_le_liminf_measure (MeasurableSet.compl E_mble) h variable [TopologicalSpace Ω] [OpensMeasurableSpace Ω] @@ -166,8 +160,8 @@ under a candidate limit measure. theorem limsup_measure_closed_le_iff_liminf_measure_open_ge {ι : Type*} {L : Filter ι} {μ : Measure Ω} {μs : ι → Measure Ω} [IsProbabilityMeasure μ] [∀ i, IsProbabilityMeasure (μs i)] : - (∀ F, IsClosed F → (L.limsup fun i => μs i F) ≤ μ F) ↔ - ∀ G, IsOpen G → μ G ≤ L.liminf fun i => μs i G := by + (∀ F, IsClosed F → (L.limsup fun i ↦ μs i F) ≤ μ F) ↔ + ∀ G, IsOpen G → μ G ≤ L.liminf fun i ↦ μs i G := by constructor · intro h G G_open exact le_measure_liminf_of_limsup_measure_compl_le @@ -201,23 +195,23 @@ variable {Ω : Type*} [MeasurableSpace Ω] theorem tendsto_measure_of_le_liminf_measure_of_limsup_measure_le {ι : Type*} {L : Filter ι} {μ : Measure Ω} {μs : ι → Measure Ω} {E₀ E E₁ : Set Ω} (E₀_subset : E₀ ⊆ E) (subset_E₁ : E ⊆ E₁) - (nulldiff : μ (E₁ \ E₀) = 0) (h_E₀ : μ E₀ ≤ L.liminf fun i => μs i E₀) - (h_E₁ : (L.limsup fun i => μs i E₁) ≤ μ E₁) : L.Tendsto (fun i => μs i E) (𝓝 (μ E)) := by + (nulldiff : μ (E₁ \ E₀) = 0) (h_E₀ : μ E₀ ≤ L.liminf fun i ↦ μs i E₀) + (h_E₁ : (L.limsup fun i ↦ μs i E₁) ≤ μ E₁) : L.Tendsto (fun i ↦ μs i E) (𝓝 (μ E)) := by apply tendsto_of_le_liminf_of_limsup_le · have E₀_ae_eq_E : E₀ =ᵐ[μ] E := EventuallyLE.antisymm E₀_subset.eventuallyLE (subset_E₁.eventuallyLE.trans (ae_le_set.mpr nulldiff)) calc μ E = μ E₀ := measure_congr E₀_ae_eq_E.symm - _ ≤ L.liminf fun i => μs i E₀ := h_E₀ - _ ≤ L.liminf fun i => μs i E := - liminf_le_liminf (eventually_of_forall fun _ => measure_mono E₀_subset) + _ ≤ L.liminf fun i ↦ μs i E₀ := h_E₀ + _ ≤ L.liminf fun i ↦ μs i E := + liminf_le_liminf (eventually_of_forall fun _ ↦ measure_mono E₀_subset) · have E_ae_eq_E₁ : E =ᵐ[μ] E₁ := EventuallyLE.antisymm subset_E₁.eventuallyLE ((ae_le_set.mpr nulldiff).trans E₀_subset.eventuallyLE) calc - (L.limsup fun i => μs i E) ≤ L.limsup fun i => μs i E₁ := - limsup_le_limsup (eventually_of_forall fun _ => measure_mono subset_E₁) + (L.limsup fun i ↦ μs i E) ≤ L.limsup fun i ↦ μs i E₁ := + limsup_le_limsup (eventually_of_forall fun _ ↦ measure_mono subset_E₁) _ ≤ μ E₁ := h_E₁ _ = μ E := measure_congr E_ae_eq_E₁.symm · infer_param @@ -233,9 +227,9 @@ sequence converge to its measure under the candidate limit measure. -/ theorem tendsto_measure_of_null_frontier {ι : Type*} {L : Filter ι} {μ : Measure Ω} {μs : ι → Measure Ω} [IsProbabilityMeasure μ] [∀ i, IsProbabilityMeasure (μs i)] - (h_opens : ∀ G, IsOpen G → μ G ≤ L.liminf fun i => μs i G) {E : Set Ω} - (E_nullbdry : μ (frontier E) = 0) : L.Tendsto (fun i => μs i E) (𝓝 (μ E)) := - haveI h_closeds : ∀ F, IsClosed F → (L.limsup fun i => μs i F) ≤ μ F := + (h_opens : ∀ G, IsOpen G → μ G ≤ L.liminf fun i ↦ μs i G) {E : Set Ω} + (E_nullbdry : μ (frontier E) = 0) : L.Tendsto (fun i ↦ μs i E) (𝓝 (μ E)) := + haveI h_closeds : ∀ F, IsClosed F → (L.limsup fun i ↦ μs i F) ≤ μ F := limsup_measure_closed_le_iff_liminf_measure_open_ge.mpr h_opens tendsto_measure_of_le_liminf_measure_of_limsup_measure_le interior_subset subset_closure E_nullbdry (h_opens _ isOpen_interior) (h_closeds _ isClosed_closure) @@ -273,27 +267,24 @@ theorem FiniteMeasure.limsup_measure_closed_le_of_tendsto {Ω ι : Type*} {L : F [MeasurableSpace Ω] [TopologicalSpace Ω] [HasOuterApproxClosed Ω] [OpensMeasurableSpace Ω] {μ : FiniteMeasure Ω} {μs : ι → FiniteMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) {F : Set Ω} (F_closed : IsClosed F) : - (L.limsup fun i => (μs i : Measure Ω) F) ≤ (μ : Measure Ω) F := by + (L.limsup fun i ↦ (μs i : Measure Ω) F) ≤ (μ : Measure Ω) F := by rcases L.eq_or_neBot with rfl | hne · simp only [limsup_bot, bot_le] apply ENNReal.le_of_forall_pos_le_add intro ε ε_pos _ + have ε_pos' := (ENNReal.half_pos (ENNReal.coe_ne_zero.mpr ε_pos.ne.symm)).ne.symm let fs := F_closed.apprSeq have key₁ : Tendsto (fun n ↦ ∫⁻ ω, (fs n ω : ℝ≥0∞) ∂μ) atTop (𝓝 ((μ : Measure Ω) F)) := HasOuterApproxClosed.tendsto_lintegral_apprSeq F_closed (μ : Measure Ω) - have room₁ : (μ : Measure Ω) F < (μ : Measure Ω) F + ε / 2 := by - apply - ENNReal.lt_add_right (measure_lt_top (μ : Measure Ω) F).ne - (ENNReal.div_pos_iff.mpr ⟨(ENNReal.coe_pos.mpr ε_pos).ne.symm, ENNReal.two_ne_top⟩).ne.symm - rcases eventually_atTop.mp (eventually_lt_of_tendsto_lt room₁ key₁) with ⟨M, hM⟩ + 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₁ 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 := by - apply ENNReal.lt_add_right (ne_of_lt ?_) - (ENNReal.div_pos_iff.mpr ⟨(ENNReal.coe_pos.mpr ε_pos).ne.symm, ENNReal.two_ne_top⟩).ne.symm - apply BoundedContinuousFunction.lintegral_lt_top_of_nnreal - have ev_near := Eventually.mono (eventually_lt_of_tendsto_lt room₂ key₂) fun n => le_of_lt + (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 (fun n ↦ le_trans (HasOuterApproxClosed.measure_le_lintegral F_closed (μs n) M)) apply (Filter.limsup_le_limsup ev_near').trans @@ -309,9 +300,9 @@ theorem ProbabilityMeasure.limsup_measure_closed_le_of_tendsto {Ω ι : Type*} { [MeasurableSpace Ω] [TopologicalSpace Ω] [OpensMeasurableSpace Ω] [HasOuterApproxClosed Ω] {μ : ProbabilityMeasure Ω} {μs : ι → ProbabilityMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) {F : Set Ω} (F_closed : IsClosed F) : - (L.limsup fun i => (μs i : Measure Ω) F) ≤ (μ : Measure Ω) F := by + (L.limsup fun i ↦ (μs i : Measure Ω) F) ≤ (μ : Measure Ω) F := by apply FiniteMeasure.limsup_measure_closed_le_of_tendsto - ((ProbabilityMeasure.tendsto_nhds_iff_toFiniteMeasure_tendsto_nhds L).mp μs_lim) F_closed + ((tendsto_nhds_iff_toFiniteMeasure_tendsto_nhds L).mp μs_lim) F_closed /-- One implication of the portmanteau theorem: Weak convergence of probability measures implies that the liminf of the measures of any open set @@ -321,9 +312,9 @@ theorem ProbabilityMeasure.le_liminf_measure_open_of_tendsto {Ω ι : Type*} {L [MeasurableSpace Ω] [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] [HasOuterApproxClosed Ω] {μ : ProbabilityMeasure Ω} {μs : ι → ProbabilityMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) {G : Set Ω} (G_open : IsOpen G) : - (μ : Measure Ω) G ≤ L.liminf fun i => (μs i : Measure Ω) G := + (μ : Measure Ω) G ≤ L.liminf fun i ↦ (μs i : Measure Ω) G := haveI h_closeds : ∀ F, IsClosed F → (L.limsup fun i ↦ (μs i : Measure Ω) F) ≤ (μ : Measure Ω) F := - fun _ F_closed => ProbabilityMeasure.limsup_measure_closed_le_of_tendsto μs_lim F_closed + fun _ F_closed ↦ limsup_measure_closed_le_of_tendsto μs_lim F_closed le_measure_liminf_of_limsup_measure_compl_le G_open.measurableSet (h_closeds _ (isClosed_compl_iff.mpr G_open)) @@ -331,9 +322,9 @@ theorem ProbabilityMeasure.tendsto_measure_of_null_frontier_of_tendsto' {Ω ι : {L : Filter ι} [MeasurableSpace Ω] [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] [HasOuterApproxClosed Ω] {μ : ProbabilityMeasure Ω} {μs : ι → ProbabilityMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) {E : Set Ω} (E_nullbdry : (μ : Measure Ω) (frontier E) = 0) : - Tendsto (fun i => (μs i : Measure Ω) E) L (𝓝 ((μ : Measure Ω) E)) := - haveI h_opens : ∀ G, IsOpen G → (μ : Measure Ω) G ≤ L.liminf fun i => (μs i : Measure Ω) G := - fun _ G_open => ProbabilityMeasure.le_liminf_measure_open_of_tendsto μs_lim G_open + Tendsto (fun i ↦ (μs i : Measure Ω) E) L (𝓝 ((μ : Measure Ω) E)) := + haveI h_opens : ∀ G, IsOpen G → (μ : Measure Ω) G ≤ L.liminf fun i ↦ (μs i : Measure Ω) G := + fun _ G_open ↦ le_liminf_measure_open_of_tendsto μs_lim G_open tendsto_measure_of_null_frontier h_opens E_nullbdry /-- One implication of the portmanteau theorem: @@ -347,10 +338,8 @@ A version with coercions to ordinary `ℝ≥0∞`-valued measures is theorem ProbabilityMeasure.tendsto_measure_of_null_frontier_of_tendsto {Ω ι : Type*} {L : Filter ι} [MeasurableSpace Ω] [PseudoEMetricSpace Ω] [OpensMeasurableSpace Ω] [HasOuterApproxClosed Ω] {μ : ProbabilityMeasure Ω} {μs : ι → ProbabilityMeasure Ω} (μs_lim : Tendsto μs L (𝓝 μ)) - {E : Set Ω} (E_nullbdry : μ (frontier E) = 0) : Tendsto (fun i => μs i E) L (𝓝 (μ E)) := by - have E_nullbdry' : (μ : Measure Ω) (frontier E) = 0 := by - rw [← ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure, E_nullbdry, ENNReal.coe_zero] - have key := ProbabilityMeasure.tendsto_measure_of_null_frontier_of_tendsto' μs_lim E_nullbdry' + {E : Set Ω} (E_nullbdry : μ (frontier E) = 0) : Tendsto (fun i ↦ μs i E) L (𝓝 (μ E)) := by + have key := tendsto_measure_of_null_frontier_of_tendsto' μs_lim (by simpa using E_nullbdry) exact (ENNReal.tendsto_toNNReal (measure_ne_top (↑μ) E)).comp key end ConvergenceImpliesLimsupClosedLE --section @@ -385,7 +374,7 @@ variable {Ω : Type*} [PseudoEMetricSpace Ω] [MeasurableSpace Ω] [OpensMeasura theorem exists_null_frontier_thickening (μ : Measure Ω) [SFinite μ] (s : Set Ω) {a b : ℝ} (hab : a < b) : ∃ r ∈ Ioo a b, μ (frontier (Metric.thickening r s)) = 0 := by have mbles : ∀ r : ℝ, MeasurableSet (frontier (Metric.thickening r s)) := - fun r => isClosed_frontier.measurableSet + fun r ↦ isClosed_frontier.measurableSet have disjs := Metric.frontier_thickening_disjoint s have key := Measure.countable_meas_pos_of_disjoint_iUnion (μ := μ) mbles disjs have aux := measure_diff_null (s := Ioo a b) (Set.Countable.measure_zero key volume) @@ -402,8 +391,8 @@ theorem exists_null_frontiers_thickening (μ : Measure Ω) [SFinite μ] (s : Set have obs := fun n : ℕ => exists_null_frontier_thickening μ s (Rs_pos n) refine ⟨fun n : ℕ => (obs n).choose, ⟨?_, ?_⟩⟩ · exact tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds Rs_lim - (fun n => (obs n).choose_spec.1.1.le) fun n => (obs n).choose_spec.1.2.le - · exact fun n => ⟨(obs n).choose_spec.1.1, (obs n).choose_spec.2⟩ + (fun n ↦ (obs n).choose_spec.1.1.le) fun n ↦ (obs n).choose_spec.1.2.le + · exact fun n ↦ ⟨(obs n).choose_spec.1.1, (obs n).choose_spec.2⟩ /-- One implication of the portmanteau theorem: Assuming that for all Borel sets E whose boundary ∂E carries no probability mass under a @@ -430,15 +419,13 @@ lemma limsup_measure_closed_le_of_forall_tendsto_measure intros ε ε_pos μF_finite have keyB := tendsto_measure_cthickening_of_isClosed (μ := μ) (s := F) ⟨1, ⟨by simp only [gt_iff_lt, zero_lt_one], measure_ne_top _ _⟩⟩ F_closed - have nhd : Iio (μ F + ε) ∈ 𝓝 (μ F) := by - apply Iio_mem_nhds - exact ENNReal.lt_add_right μF_finite.ne (ENNReal.coe_pos.mpr ε_pos).ne' + have nhd : Iio (μ F + ε) ∈ 𝓝 (μ F) := + Iio_mem_nhds <| ENNReal.lt_add_right μF_finite.ne (ENNReal.coe_pos.mpr ε_pos).ne' specialize rs_lim (keyB nhd) - simp only [mem_map, mem_atTop_sets, mem_preimage, mem_Iio] at rs_lim + simp only [mem_map, mem_atTop_sets, ge_iff_le, mem_preimage, mem_Iio] at rs_lim obtain ⟨m, hm⟩ := rs_lim - have aux' := fun i ↦ measure_mono (μ := μs i) (Metric.self_subset_thickening (rs_pos m) F) have aux : (fun i ↦ (μs i F)) ≤ᶠ[L] (fun i ↦ μs i (Metric.thickening (rs m) F)) := - eventually_of_forall aux' + eventually_of_forall <| fun i ↦ measure_mono (Metric.self_subset_thickening (rs_pos m) F) refine (limsup_le_limsup aux).trans ?_ rw [Tendsto.limsup_eq (key m)] apply (measure_mono (Metric.thickening_subset_cthickening (rs m) F)).trans (hm m rfl.le).le @@ -506,7 +493,7 @@ lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure f.continuous.measurable.aestronglyMeasurable] let g := BoundedContinuousFunction.comp _ Real.lipschitzWith_toNNReal f have bound : ∀ i, ∫⁻ x, ENNReal.ofReal (f x) ∂(μs i) ≤ nndist 0 g := fun i ↦ by - simpa only [coe_nnreal_ennreal_nndist, measure_univ, mul_one] using + 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 @@ -515,7 +502,7 @@ lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure simp only [measure_univ, mul_one] at obs apply lt_of_le_of_lt _ (show (‖f‖₊ : ℝ≥0∞) < ∞ from ENNReal.coe_lt_top) apply liminf_le_of_le - · refine ⟨0, eventually_of_forall (by simp only [zero_le, forall_const])⟩ + · refine ⟨0, eventually_of_forall (by simp only [ge_iff_le, zero_le, forall_const])⟩ · intro x hx obtain ⟨i, hi⟩ := hx.exists apply le_trans hi @@ -542,13 +529,9 @@ theorem tendsto_of_forall_isOpen_le_liminf {μ : ProbabilityMeasure Ω} have aux : ENNReal.ofNNReal (liminf (fun i ↦ μs i G) atTop) = liminf (ENNReal.ofNNReal ∘ fun i ↦ μs i G) atTop := by refine Monotone.map_liminf_of_continuousAt (F := atTop) ENNReal.coe_mono (μs · G) ?_ ?_ ?_ - · apply ENNReal.continuous_coe.continuousAt - · apply IsBoundedUnder.isCoboundedUnder_ge ⟨1, ?_⟩ - simp only [eventually_map, ProbabilityMeasure.apply_le_one, eventually_atTop, ge_iff_le, - implies_true, forall_const, exists_const] - · use 0 - simp only [zero_le, eventually_map, eventually_atTop, implies_true, forall_const, - exists_const] + · exact ENNReal.continuous_coe.continuousAt + · exact IsBoundedUnder.isCoboundedUnder_ge ⟨1, by simp⟩ + · exact ⟨0, by simp⟩ have obs := ENNReal.coe_mono h_opens simp only [ne_eq, ProbabilityMeasure.ennreal_coeFn_eq_coeFn_toMeasure, aux] at obs convert obs diff --git a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean index e990eebe4d4d6..d908cc8f3afd6 100644 --- a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean @@ -74,14 +74,7 @@ convergence in distribution, convergence in law, weak convergence of measures, p noncomputable section -open MeasureTheory - -open Set - -open Filter - -open BoundedContinuousFunction - +open MeasureTheory Set Filter BoundedContinuousFunction open scoped Topology ENNReal NNReal BoundedContinuousFunction namespace MeasureTheory @@ -120,8 +113,7 @@ instance [Inhabited Ω] : Inhabited (ProbabilityMeasure Ω) := def toMeasure : ProbabilityMeasure Ω → Measure Ω := Subtype.val /-- A probability measure can be interpreted as a measure. -/ -instance : Coe (ProbabilityMeasure Ω) (MeasureTheory.Measure Ω) where - coe := toMeasure +instance : Coe (ProbabilityMeasure Ω) (MeasureTheory.Measure Ω) := { coe := toMeasure } instance (μ : ProbabilityMeasure Ω) : IsProbabilityMeasure (μ : Measure Ω) := μ.prop @@ -129,8 +121,7 @@ instance (μ : ProbabilityMeasure Ω) : IsProbabilityMeasure (μ : Measure Ω) : @[simp, norm_cast] lemma coe_mk (μ : Measure Ω) (hμ) : toMeasure ⟨μ, hμ⟩ = μ := rfl @[simp] -theorem val_eq_to_measure (ν : ProbabilityMeasure Ω) : ν.val = (ν : Measure Ω) := - rfl +theorem val_eq_to_measure (ν : ProbabilityMeasure Ω) : ν.val = (ν : Measure Ω) := rfl theorem toMeasure_injective : Function.Injective ((↑) : ProbabilityMeasure Ω → Measure Ω) := Subtype.coe_injective @@ -157,8 +148,7 @@ theorem coeFn_univ_ne_zero (ν : ProbabilityMeasure Ω) : ν univ ≠ 0 := by simp only [coeFn_univ, Ne, one_ne_zero, not_false_iff] /-- A probability measure can be interpreted as a finite measure. -/ -def toFiniteMeasure (μ : ProbabilityMeasure Ω) : FiniteMeasure Ω := - ⟨μ, inferInstance⟩ +def toFiniteMeasure (μ : ProbabilityMeasure Ω) : FiniteMeasure Ω := ⟨μ, inferInstance⟩ @[simp] lemma coeFn_toFiniteMeasure (μ : ProbabilityMeasure Ω) : ⇑μ.toFiniteMeasure = μ := rfl lemma toFiniteMeasure_apply (μ : ProbabilityMeasure Ω) (s : Set Ω) : @@ -166,13 +156,11 @@ lemma toFiniteMeasure_apply (μ : ProbabilityMeasure Ω) (s : Set Ω) : @[simp] theorem toMeasure_comp_toFiniteMeasure_eq_toMeasure (ν : ProbabilityMeasure Ω) : - (ν.toFiniteMeasure : Measure Ω) = (ν : Measure Ω) := - rfl + (ν.toFiniteMeasure : Measure Ω) = (ν : Measure Ω) := rfl @[simp] theorem coeFn_comp_toFiniteMeasure_eq_coeFn (ν : ProbabilityMeasure Ω) : - (ν.toFiniteMeasure : Set Ω → ℝ≥0) = (ν : Set Ω → ℝ≥0) := - rfl + (ν.toFiniteMeasure : Set Ω → ℝ≥0) = (ν : Set Ω → ℝ≥0) := rfl @[simp] theorem toFiniteMeasure_apply_eq_apply (ν : ProbabilityMeasure Ω) (s : Set Ω) : @@ -184,6 +172,12 @@ theorem ennreal_coeFn_eq_coeFn_toMeasure (ν : ProbabilityMeasure Ω) (s : Set rw [← coeFn_comp_toFiniteMeasure_eq_coeFn, FiniteMeasure.ennreal_coeFn_eq_coeFn_toMeasure, toMeasure_comp_toFiniteMeasure_eq_toMeasure] +@[simp] +theorem null_iff_toMeasure_null (ν : ProbabilityMeasure Ω) (s : Set Ω) : + ν s = 0 ↔ (ν : Measure Ω) s = 0 := + ⟨fun h ↦ by rw [← ennreal_coeFn_eq_coeFn_toMeasure, h, ENNReal.coe_zero], + fun h ↦ congrArg ENNReal.toNNReal h⟩ + theorem apply_mono (μ : ProbabilityMeasure Ω) {s₁ s₂ : Set Ω} (h : s₁ ⊆ s₂) : μ s₁ ≤ μ s₂ := by rw [← coeFn_comp_toFiniteMeasure_eq_coeFn] exact MeasureTheory.FiniteMeasure.apply_mono _ h @@ -215,15 +209,14 @@ theorem mass_toFiniteMeasure (μ : ProbabilityMeasure Ω) : μ.toFiniteMeasure.m μ.coeFn_univ theorem toFiniteMeasure_nonzero (μ : ProbabilityMeasure Ω) : μ.toFiniteMeasure ≠ 0 := by - rw [← FiniteMeasure.mass_nonzero_iff, μ.mass_toFiniteMeasure] - exact one_ne_zero + simp [← FiniteMeasure.mass_nonzero_iff] section convergence_in_distribution variable [TopologicalSpace Ω] [OpensMeasurableSpace Ω] theorem testAgainstNN_lipschitz (μ : ProbabilityMeasure Ω) : - LipschitzWith 1 fun f : Ω →ᵇ ℝ≥0 => μ.toFiniteMeasure.testAgainstNN f := + LipschitzWith 1 fun f : Ω →ᵇ ℝ≥0 ↦ μ.toFiniteMeasure.testAgainstNN f := μ.mass_toFiniteMeasure ▸ μ.toFiniteMeasure.testAgainstNN_lipschitz /-- The topology of weak convergence on `MeasureTheory.ProbabilityMeasure Ω`. This is inherited @@ -243,21 +236,19 @@ def toWeakDualBCNN : ProbabilityMeasure Ω → WeakDual ℝ≥0 (Ω →ᵇ ℝ @[simp] theorem coe_toWeakDualBCNN (μ : ProbabilityMeasure Ω) : - ⇑μ.toWeakDualBCNN = μ.toFiniteMeasure.testAgainstNN := - rfl + ⇑μ.toWeakDualBCNN = μ.toFiniteMeasure.testAgainstNN := rfl @[simp] theorem toWeakDualBCNN_apply (μ : ProbabilityMeasure Ω) (f : Ω →ᵇ ℝ≥0) : - μ.toWeakDualBCNN f = (∫⁻ ω, f ω ∂(μ : Measure Ω)).toNNReal := - rfl + μ.toWeakDualBCNN f = (∫⁻ ω, f ω ∂(μ : Measure Ω)).toNNReal := rfl -theorem toWeakDualBCNN_continuous : Continuous fun μ : ProbabilityMeasure Ω => μ.toWeakDualBCNN := +theorem toWeakDualBCNN_continuous : Continuous fun μ : ProbabilityMeasure Ω ↦ μ.toWeakDualBCNN := FiniteMeasure.toWeakDualBCNN_continuous.comp toFiniteMeasure_continuous /- Integration of (nonnegative bounded continuous) test functions against Borel probability measures depends continuously on the measure. -/ theorem continuous_testAgainstNN_eval (f : Ω →ᵇ ℝ≥0) : - Continuous fun μ : ProbabilityMeasure Ω => μ.toFiniteMeasure.testAgainstNN f := + Continuous fun μ : ProbabilityMeasure Ω ↦ μ.toFiniteMeasure.testAgainstNN f := (FiniteMeasure.continuous_testAgainstNN_eval f).comp toFiniteMeasure_continuous -- The canonical mapping from probability measures to finite measures is an embedding. @@ -265,7 +256,7 @@ theorem toFiniteMeasure_embedding (Ω : Type*) [MeasurableSpace Ω] [Topological [OpensMeasurableSpace Ω] : Embedding (toFiniteMeasure : ProbabilityMeasure Ω → FiniteMeasure Ω) := { induced := rfl - inj := fun _μ _ν h => Subtype.eq <| congr_arg FiniteMeasure.toMeasure h } + inj := fun _μ _ν h ↦ Subtype.eq <| congr_arg FiniteMeasure.toMeasure h } theorem tendsto_nhds_iff_toFiniteMeasure_tendsto_nhds {δ : Type*} (F : Filter δ) {μs : δ → ProbabilityMeasure Ω} {μ₀ : ProbabilityMeasure Ω} : @@ -279,7 +270,7 @@ theorem tendsto_iff_forall_lintegral_tendsto {γ : Type*} {F : Filter γ} {μs : γ → ProbabilityMeasure Ω} {μ : ProbabilityMeasure Ω} : Tendsto μs F (𝓝 μ) ↔ ∀ f : Ω →ᵇ ℝ≥0, - Tendsto (fun i => ∫⁻ ω, f ω ∂(μs i : Measure Ω)) F (𝓝 (∫⁻ ω, f ω ∂(μ : Measure Ω))) := by + Tendsto (fun i ↦ ∫⁻ ω, f ω ∂(μs i : Measure Ω)) F (𝓝 (∫⁻ ω, f ω ∂(μ : Measure Ω))) := by rw [tendsto_nhds_iff_toFiniteMeasure_tendsto_nhds] exact FiniteMeasure.tendsto_iff_forall_lintegral_tendsto @@ -290,7 +281,7 @@ theorem tendsto_iff_forall_integral_tendsto {γ : Type*} {F : Filter γ} {μs : γ → ProbabilityMeasure Ω} {μ : ProbabilityMeasure Ω} : Tendsto μs F (𝓝 μ) ↔ ∀ f : Ω →ᵇ ℝ, - Tendsto (fun i => ∫ ω, f ω ∂(μs i : Measure Ω)) F (𝓝 (∫ ω, f ω ∂(μ : Measure Ω))) := by + Tendsto (fun i ↦ ∫ ω, f ω ∂(μs i : Measure Ω)) F (𝓝 (∫ ω, f ω ∂(μ : Measure Ω))) := by rw [tendsto_nhds_iff_toFiniteMeasure_tendsto_nhds] rw [FiniteMeasure.tendsto_iff_forall_integral_tendsto] rfl @@ -411,25 +402,25 @@ variable {μ} theorem tendsto_testAgainstNN_of_tendsto_normalize_testAgainstNN_of_tendsto_mass {γ : Type*} {F : Filter γ} {μs : γ → FiniteMeasure Ω} - (μs_lim : Tendsto (fun i => (μs i).normalize) F (𝓝 μ.normalize)) - (mass_lim : Tendsto (fun i => (μs i).mass) F (𝓝 μ.mass)) (f : Ω →ᵇ ℝ≥0) : - Tendsto (fun i => (μs i).testAgainstNN f) F (𝓝 (μ.testAgainstNN f)) := by + (μs_lim : Tendsto (fun i ↦ (μs i).normalize) F (𝓝 μ.normalize)) + (mass_lim : Tendsto (fun i ↦ (μs i).mass) F (𝓝 μ.mass)) (f : Ω →ᵇ ℝ≥0) : + Tendsto (fun i ↦ (μs i).testAgainstNN f) F (𝓝 (μ.testAgainstNN f)) := by by_cases h_mass : μ.mass = 0 · simp only [μ.mass_zero_iff.mp h_mass, zero_testAgainstNN_apply, zero_mass, - eq_self_iff_true] at * + eq_self_iff_true] at mass_lim ⊢ exact tendsto_zero_testAgainstNN_of_tendsto_zero_mass mass_lim f - simp_rw [fun i => (μs i).testAgainstNN_eq_mass_mul f, μ.testAgainstNN_eq_mass_mul f] + simp_rw [fun i ↦ (μs i).testAgainstNN_eq_mass_mul f, μ.testAgainstNN_eq_mass_mul f] rw [ProbabilityMeasure.tendsto_nhds_iff_toFiniteMeasure_tendsto_nhds] at μs_lim rw [tendsto_iff_forall_testAgainstNN_tendsto] at μs_lim have lim_pair : - Tendsto (fun i => (⟨(μs i).mass, (μs i).normalize.toFiniteMeasure.testAgainstNN f⟩ : ℝ≥0 × ℝ≥0)) + Tendsto (fun i ↦ (⟨(μs i).mass, (μs i).normalize.toFiniteMeasure.testAgainstNN f⟩ : ℝ≥0 × ℝ≥0)) F (𝓝 ⟨μ.mass, μ.normalize.toFiniteMeasure.testAgainstNN f⟩) := (Prod.tendsto_iff _ _).mpr ⟨mass_lim, μs_lim f⟩ exact tendsto_mul.comp lim_pair theorem tendsto_normalize_testAgainstNN_of_tendsto {γ : Type*} {F : Filter γ} {μs : γ → FiniteMeasure Ω} (μs_lim : Tendsto μs F (𝓝 μ)) (nonzero : μ ≠ 0) (f : Ω →ᵇ ℝ≥0) : - Tendsto (fun i => (μs i).normalize.toFiniteMeasure.testAgainstNN f) F + Tendsto (fun i ↦ (μs i).normalize.toFiniteMeasure.testAgainstNN f) F (𝓝 (μ.normalize.toFiniteMeasure.testAgainstNN f)) := by have lim_mass := μs_lim.mass have aux : {(0 : ℝ≥0)}ᶜ ∈ 𝓝 μ.mass := @@ -445,7 +436,7 @@ theorem tendsto_normalize_testAgainstNN_of_tendsto {γ : Type*} {F : Filter γ} apply normalize_testAgainstNN _ hi simp_rw [tendsto_congr' eve, μ.normalize_testAgainstNN nonzero] have lim_pair : - Tendsto (fun i => (⟨(μs i).mass⁻¹, (μs i).testAgainstNN f⟩ : ℝ≥0 × ℝ≥0)) F + Tendsto (fun i ↦ (⟨(μs i).mass⁻¹, (μs i).testAgainstNN f⟩ : ℝ≥0 × ℝ≥0)) F (𝓝 ⟨μ.mass⁻¹, μ.testAgainstNN f⟩) := by refine (Prod.tendsto_iff _ _).mpr ⟨?_, ?_⟩ · exact (continuousOn_inv₀.continuousAt aux).tendsto.comp lim_mass @@ -455,27 +446,27 @@ theorem tendsto_normalize_testAgainstNN_of_tendsto {γ : Type*} {F : Filter γ} /-- If the normalized versions of finite measures converge weakly and their total masses also converge, then the finite measures themselves converge weakly. -/ theorem tendsto_of_tendsto_normalize_testAgainstNN_of_tendsto_mass {γ : Type*} {F : Filter γ} - {μs : γ → FiniteMeasure Ω} (μs_lim : Tendsto (fun i => (μs i).normalize) F (𝓝 μ.normalize)) - (mass_lim : Tendsto (fun i => (μs i).mass) F (𝓝 μ.mass)) : Tendsto μs F (𝓝 μ) := by + {μs : γ → FiniteMeasure Ω} (μs_lim : Tendsto (fun i ↦ (μs i).normalize) F (𝓝 μ.normalize)) + (mass_lim : Tendsto (fun i ↦ (μs i).mass) F (𝓝 μ.mass)) : Tendsto μs F (𝓝 μ) := by rw [tendsto_iff_forall_testAgainstNN_tendsto] - exact fun f => + exact fun f ↦ tendsto_testAgainstNN_of_tendsto_normalize_testAgainstNN_of_tendsto_mass μs_lim mass_lim f /-- If finite measures themselves converge weakly to a nonzero limit measure, then their normalized versions also converge weakly. -/ theorem tendsto_normalize_of_tendsto {γ : Type*} {F : Filter γ} {μs : γ → FiniteMeasure Ω} (μs_lim : Tendsto μs F (𝓝 μ)) (nonzero : μ ≠ 0) : - Tendsto (fun i => (μs i).normalize) F (𝓝 μ.normalize) := by + Tendsto (fun i ↦ (μs i).normalize) F (𝓝 μ.normalize) := by rw [ProbabilityMeasure.tendsto_nhds_iff_toFiniteMeasure_tendsto_nhds, tendsto_iff_forall_testAgainstNN_tendsto] - exact fun f => tendsto_normalize_testAgainstNN_of_tendsto μs_lim nonzero f + exact fun f ↦ tendsto_normalize_testAgainstNN_of_tendsto μs_lim nonzero f /-- The weak convergence of finite measures to a nonzero limit can be characterized by the weak convergence of both their normalized versions (probability measures) and their total masses. -/ theorem tendsto_normalize_iff_tendsto {γ : Type*} {F : Filter γ} {μs : γ → FiniteMeasure Ω} (nonzero : μ ≠ 0) : - Tendsto (fun i => (μs i).normalize) F (𝓝 μ.normalize) ∧ - Tendsto (fun i => (μs i).mass) F (𝓝 μ.mass) ↔ + Tendsto (fun i ↦ (μs i).normalize) F (𝓝 μ.normalize) ∧ + Tendsto (fun i ↦ (μs i).mass) F (𝓝 μ.mass) ↔ Tendsto μs F (𝓝 μ) := by constructor · rintro ⟨normalized_lim, mass_lim⟩ From 45ee234cc84e156ca6c42639b2c735993918210a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 16 Aug 2024 08:13:02 +0000 Subject: [PATCH 322/359] feat(Measure/ContinuousPreimage): new file (#15367) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit - move `tendsto_measure_symmDiff_preimage_nhds_zero` to a new file; - prove that `{z | f z ⁻¹' t =ᵐ[μ] s}` is a closed set, if `f` continuously depends on `z`. --- Mathlib.lean | 1 + .../ContinuousCompMeasurePreserving.lean | 66 +--------- .../Measure/ContinuousPreimage.lean | 121 ++++++++++++++++++ 3 files changed, 123 insertions(+), 65 deletions(-) create mode 100644 Mathlib/MeasureTheory/Measure/ContinuousPreimage.lean diff --git a/Mathlib.lean b/Mathlib.lean index d7cb00827b04c..15f573a3fa753 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3216,6 +3216,7 @@ import Mathlib.MeasureTheory.Measure.AEMeasurable import Mathlib.MeasureTheory.Measure.AddContent import Mathlib.MeasureTheory.Measure.Complex import Mathlib.MeasureTheory.Measure.Content +import Mathlib.MeasureTheory.Measure.ContinuousPreimage import Mathlib.MeasureTheory.Measure.Count import Mathlib.MeasureTheory.Measure.Dirac import Mathlib.MeasureTheory.Measure.DiracProba diff --git a/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean b/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean index 09380e0b12e50..74a783f2afe90 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.MeasureTheory.Function.SimpleFuncDenseLp -import Mathlib.MeasureTheory.Measure.Regular +import Mathlib.MeasureTheory.Measure.ContinuousPreimage /-! # Continuity of `MeasureTheory.Lp.compMeasurePreserving` @@ -30,70 +30,6 @@ variable {α X Y : Type*} {μ : Measure X} {ν : Measure Y} [μ.InnerRegularCompactLTTop] [IsLocallyFiniteMeasure ν] namespace MeasureTheory - -/-- Let `X` and `Y` be R₁ topological spaces -with Borel σ-algebras and measures `μ` and `ν`, respectively. -Suppose that `μ` is inner regular for finite measure sets with respect to compact sets -and `ν` is a locally finite measure. -Let `f : α → C(X, Y)` be a family of continuous maps -that converges to a continuous map `g : C(X, Y)` in the compact-open topology along a filter `l`. -Suppose that `g` is a measure preserving map -and `f a` is a measure preserving map eventually along `l`. -Then for any finite measure measurable set `s`, -the preimages `f a ⁻¹' s` tend to the preimage `g ⁻¹' s` in measure. -More precisely, the measure of the symmetric difference of these two sets tends to zero. -/ -theorem tendsto_measure_symmDiff_preimage_nhds_zero - {l : Filter α} {f : α → C(X, Y)} {g : C(X, Y)} {s : Set Y} (hfg : Tendsto f l (𝓝 g)) - (hf : ∀ᶠ a in l, MeasurePreserving (f a) μ ν) (hg : MeasurePreserving g μ ν) - (hs : NullMeasurableSet s ν) (hνs : ν s ≠ ∞) : - Tendsto (fun a ↦ μ ((f a ⁻¹' s) ∆ (g ⁻¹' s))) l (𝓝 0) := by - have : ν.InnerRegularCompactLTTop := by - rw [← hg.map_eq] - exact .map_of_continuous (map_continuous _) - rw [ENNReal.tendsto_nhds_zero] - intro ε hε - -- Without loss of generality, `s` is an open set. - wlog hso : IsOpen s generalizing s ε - · have H : 0 < ε / 3 := ENNReal.div_pos hε.ne' ENNReal.coe_ne_top - -- Indeed, we can choose an open set `U` such that `ν (U ∆ s) < ε / 3`, - -- apply the lemma to `U`, then use the triangle inequality for `μ (_ ∆ _)`. - rcases hs.exists_isOpen_symmDiff_lt hνs H.ne' with ⟨U, hUo, hU, hUs⟩ - have hmU : NullMeasurableSet U ν := hUo.measurableSet.nullMeasurableSet - replace hUs := hUs.le - filter_upwards [hf, this hmU hU.ne _ H hUo] with a hfa ha - calc - μ ((f a ⁻¹' s) ∆ (g ⁻¹' s)) - ≤ μ ((f a ⁻¹' s) ∆ (f a ⁻¹' U)) + μ ((f a ⁻¹' U) ∆ (g ⁻¹' U)) - + μ ((g ⁻¹' U) ∆ (g ⁻¹' s)) := by - refine (measure_symmDiff_le _ (g ⁻¹' U) _).trans ?_ - gcongr - apply measure_symmDiff_le - _ ≤ ε / 3 + ε / 3 + ε / 3 := by - gcongr - · rwa [← preimage_symmDiff, hfa.measure_preimage (hs.symmDiff hmU), symmDiff_comm] - · rwa [← preimage_symmDiff, hg.measure_preimage (hmU.symmDiff hs)] - _ = ε := by simp - -- Take a compact closed subset `K ⊆ g ⁻¹' s` of almost full measure, - -- `μ (g ⁻¹' s \ K) < ε / 2`. - have hνs' : μ (g ⁻¹' s) ≠ ∞ := by rwa [hg.measure_preimage hs] - obtain ⟨K, hKg, hKco, hKcl, hKμ⟩ : - ∃ K, MapsTo g K s ∧ IsCompact K ∧ IsClosed K ∧ μ (g ⁻¹' s \ K) < ε / 2 := - (hg.measurable hso.measurableSet).exists_isCompact_isClosed_diff_lt hνs' <| by simp [hε.ne'] - have hKm : MeasurableSet K := hKcl.measurableSet - -- Take `a` such that `f a` is measure preserving and maps `K` to `s`. - -- This is possible, because `K` is a compact set and `s` is an open set. - filter_upwards [hf, ContinuousMap.tendsto_nhds_compactOpen.mp hfg K hKco s hso hKg] with a hfa ha - -- Then each of the sets `g ⁻¹' s ∆ K = g ⁻¹' s \ K` and `f a ⁻¹' s ∆ K = f a ⁻¹' s \ K` - -- have measure at most `ε / 2`, thus `f a ⁻¹' s ∆ g ⁻¹' s` has measure at most `ε`. - rw [← ENNReal.add_halves ε] - refine (measure_symmDiff_le _ K _).trans ?_ - rw [symmDiff_of_ge ha.subset_preimage, symmDiff_of_le hKg.subset_preimage] - gcongr - have hK' : μ K ≠ ∞ := ne_top_of_le_ne_top hνs' <| measure_mono hKg.subset_preimage - rw [measure_diff_le_iff_le_add hKm ha.subset_preimage hK', hfa.measure_preimage hs, - ← hg.measure_preimage hs, ← measure_diff_le_iff_le_add hKm hKg.subset_preimage hK'] - exact hKμ.le - namespace Lp variable (μ ν) diff --git a/Mathlib/MeasureTheory/Measure/ContinuousPreimage.lean b/Mathlib/MeasureTheory/Measure/ContinuousPreimage.lean new file mode 100644 index 0000000000000..2de20eb3e8ac1 --- /dev/null +++ b/Mathlib/MeasureTheory/Measure/ContinuousPreimage.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2024 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Topology.CompactOpen +import Mathlib.Dynamics.Ergodic.MeasurePreserving +import Mathlib.MeasureTheory.Measure.Regular + +/-! +# Continuity of the preimage of a set under a measure preserving continuous function + +In this file we prove that the preimage of a null measurable set `s : Set Y` +under a measure preserving continuous function `f : C(X, Y)` is continuous in `f` +in the sense that `μ ((f a ⁻¹' s) ∆ (g ⁻¹' s))` tends to zero as `f a` tends to `g`. + +As a corollary, we show that +for a continuous family of continuous maps `f z : C(X, Y)`, +a null measurable set `s`, and a null measurable set `t` of finite measure, +the set of parameters `z` such that `f z ⁻¹' t` is a.e. equal to `s` is a closed set. +-/ + +open Filter Set +open scoped ENNReal symmDiff Topology + +namespace MeasureTheory + +variable {α X Y Z : Type*} + [TopologicalSpace X] [MeasurableSpace X] [BorelSpace X] [R1Space X] + [TopologicalSpace Y] [MeasurableSpace Y] [BorelSpace Y] [R1Space Y] + [TopologicalSpace Z] + {μ : Measure X} {ν : Measure Y} [μ.InnerRegularCompactLTTop] [IsLocallyFiniteMeasure ν] + +/-- Let `X` and `Y` be R₁ topological spaces +with Borel σ-algebras and measures `μ` and `ν`, respectively. +Suppose that `μ` is inner regular for finite measure sets with respect to compact sets +and `ν` is a locally finite measure. +Let `f : α → C(X, Y)` be a family of continuous maps +that converges to a continuous map `g : C(X, Y)` in the compact-open topology along a filter `l`. +Suppose that `g` is a measure preserving map +and `f a` is a measure preserving map eventually along `l`. +Then for any finite measure measurable set `s`, +the preimages `f a ⁻¹' s` tend to the preimage `g ⁻¹' s` in measure. +More precisely, the measure of the symmetric difference of these two sets tends to zero. -/ +theorem tendsto_measure_symmDiff_preimage_nhds_zero + {l : Filter α} {f : α → C(X, Y)} {g : C(X, Y)} {s : Set Y} (hfg : Tendsto f l (𝓝 g)) + (hf : ∀ᶠ a in l, MeasurePreserving (f a) μ ν) (hg : MeasurePreserving g μ ν) + (hs : NullMeasurableSet s ν) (hνs : ν s ≠ ∞) : + Tendsto (fun a ↦ μ ((f a ⁻¹' s) ∆ (g ⁻¹' s))) l (𝓝 0) := by + have : ν.InnerRegularCompactLTTop := by + rw [← hg.map_eq] + exact .map_of_continuous (map_continuous _) + rw [ENNReal.tendsto_nhds_zero] + intro ε hε + -- Without loss of generality, `s` is an open set. + wlog hso : IsOpen s generalizing s ε + · have H : 0 < ε / 3 := ENNReal.div_pos hε.ne' ENNReal.coe_ne_top + -- Indeed, we can choose an open set `U` such that `ν (U ∆ s) < ε / 3`, + -- apply the lemma to `U`, then use the triangle inequality for `μ (_ ∆ _)`. + rcases hs.exists_isOpen_symmDiff_lt hνs H.ne' with ⟨U, hUo, hU, hUs⟩ + have hmU : NullMeasurableSet U ν := hUo.measurableSet.nullMeasurableSet + replace hUs := hUs.le + filter_upwards [hf, this hmU hU.ne _ H hUo] with a hfa ha + calc + μ ((f a ⁻¹' s) ∆ (g ⁻¹' s)) + ≤ μ ((f a ⁻¹' s) ∆ (f a ⁻¹' U)) + μ ((f a ⁻¹' U) ∆ (g ⁻¹' U)) + + μ ((g ⁻¹' U) ∆ (g ⁻¹' s)) := by + refine (measure_symmDiff_le _ (g ⁻¹' U) _).trans ?_ + gcongr + apply measure_symmDiff_le + _ ≤ ε / 3 + ε / 3 + ε / 3 := by + gcongr + · rwa [← preimage_symmDiff, hfa.measure_preimage (hs.symmDiff hmU), symmDiff_comm] + · rwa [← preimage_symmDiff, hg.measure_preimage (hmU.symmDiff hs)] + _ = ε := by simp + -- Take a compact closed subset `K ⊆ g ⁻¹' s` of almost full measure, + -- `μ (g ⁻¹' s \ K) < ε / 2`. + have hνs' : μ (g ⁻¹' s) ≠ ∞ := by rwa [hg.measure_preimage hs] + obtain ⟨K, hKg, hKco, hKcl, hKμ⟩ : + ∃ K, MapsTo g K s ∧ IsCompact K ∧ IsClosed K ∧ μ (g ⁻¹' s \ K) < ε / 2 := + (hg.measurable hso.measurableSet).exists_isCompact_isClosed_diff_lt hνs' <| by simp [hε.ne'] + have hKm : MeasurableSet K := hKcl.measurableSet + -- Take `a` such that `f a` is measure preserving and maps `K` to `s`. + -- This is possible, because `K` is a compact set and `s` is an open set. + filter_upwards [hf, ContinuousMap.tendsto_nhds_compactOpen.mp hfg K hKco s hso hKg] with a hfa ha + -- Then each of the sets `g ⁻¹' s ∆ K = g ⁻¹' s \ K` and `f a ⁻¹' s ∆ K = f a ⁻¹' s \ K` + -- have measure at most `ε / 2`, thus `f a ⁻¹' s ∆ g ⁻¹' s` has measure at most `ε`. + rw [← ENNReal.add_halves ε] + refine (measure_symmDiff_le _ K _).trans ?_ + rw [symmDiff_of_ge ha.subset_preimage, symmDiff_of_le hKg.subset_preimage] + gcongr + have hK' : μ K ≠ ∞ := ne_top_of_le_ne_top hνs' <| measure_mono hKg.subset_preimage + rw [measure_diff_le_iff_le_add hKm ha.subset_preimage hK', hfa.measure_preimage hs, + ← hg.measure_preimage hs, ← measure_diff_le_iff_le_add hKm hKg.subset_preimage hK'] + exact hKμ.le + +/-- Let `f : Z → C(X, Y)` be a continuous (in the compact open topology) family +of continuous measure preserving maps. +Let `t : Set Y` be a null measurable set of finite measure. +Then for any `s`, the set of parameters `z` +such that the preimage of `t` under `f_z` is a.e. equal to `s` +is a closed set. + +In particular, if `X = Y` and `s = t`, +then we see that the a.e. stabilizer of a set is a closed set. -/ +theorem isClosed_setOf_preimage_ae_eq {f : Z → C(X, Y)} (hf : Continuous f) + (hfm : ∀ z, MeasurePreserving (f z) μ ν) (s : Set X) + {t : Set Y} (htm : NullMeasurableSet t ν) (ht : ν t ≠ ∞) : + IsClosed {z | f z ⁻¹' t =ᵐ[μ] s} := by + rw [← isOpen_compl_iff, isOpen_iff_mem_nhds] + intro z hz + replace hz : ∀ᶠ ε : ℝ≥0∞ in 𝓝 0, ε < μ ((f z ⁻¹' t) ∆ s) := by + apply gt_mem_nhds + rwa [pos_iff_ne_zero, ne_eq, measure_symmDiff_eq_zero_iff] + filter_upwards [(tendsto_measure_symmDiff_preimage_nhds_zero (hf.tendsto z) + (eventually_of_forall hfm) (hfm z) htm ht).eventually hz] with w hw + intro (hw' : f w ⁻¹' t =ᵐ[μ] s) + rw [measure_congr (hw'.symmDiff (ae_eq_refl _)), symmDiff_comm] at hw + exact hw.false + +end MeasureTheory From a6b7481a9abf30111d5efb2a336f2a1137857cc5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Aug 2024 08:13:03 +0000 Subject: [PATCH 323/359] feat: Essential supremum according to the counting measure (#15403) ... is just the usual supremum From LeanAPAP --- Mathlib/MeasureTheory/Function/EssSup.lean | 13 +++++++++++++ .../MeasureTheory/Function/LpSeminorm/Basic.lean | 6 ++++++ .../Function/StronglyMeasurable/Basic.lean | 3 +++ Mathlib/MeasureTheory/OuterMeasure/AE.lean | 14 ++++++++++++-- Mathlib/Order/Filter/Basic.lean | 2 ++ Mathlib/Order/LiminfLimsup.lean | 4 ++++ 6 files changed, 40 insertions(+), 2 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/EssSup.lean b/Mathlib/MeasureTheory/Function/EssSup.lean index ceac0322c5019..999cef7eab633 100644 --- a/Mathlib/MeasureTheory/Function/EssSup.lean +++ b/Mathlib/MeasureTheory/Function/EssSup.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.BorelSpace.Order +import Mathlib.MeasureTheory.Measure.Count import Mathlib.Order.Filter.ENNReal /-! @@ -175,6 +176,18 @@ theorem essSup_smul_measure {f : α → β} {c : ℝ≥0∞} (hc : c ≠ 0) : essSup f (c • μ) = essSup f μ := by simp_rw [essSup, Measure.ae_smul_measure_eq hc] +lemma essSup_eq_iSup (hμ : ∀ a, μ {a} ≠ 0) (f : α → β) : essSup f μ = ⨆ i, f i := by + rw [essSup, ae_eq_top.2 hμ, limsup_top] + +lemma essInf_eq_iInf (hμ : ∀ a, μ {a} ≠ 0) (f : α → β) : essInf f μ = ⨅ i, f i := by + rw [essInf, ae_eq_top.2 hμ, liminf_top] + +@[simp] lemma essSup_count [MeasurableSingletonClass α] (f : α → β) : essSup f .count = ⨆ i, f i := + essSup_eq_iSup (by simp) _ + +@[simp] lemma essInf_count [MeasurableSingletonClass α] (f : α → β) : essInf f .count = ⨅ i, f i := + essInf_eq_iInf (by simp) _ + section TopologicalSpace variable {γ : Type*} {mγ : MeasurableSpace γ} {f : α → γ} {g : γ → β} diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 7cdd71d229171..052668ca96ad7 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -878,6 +878,12 @@ theorem eLpNorm_le_add_measure_left (f : α → F) (μ ν : Measure α) {p : ℝ @[deprecated (since := "2024-07-27")] alias snorm_le_add_measure_left := eLpNorm_le_add_measure_left +lemma eLpNormEssSup_eq_iSup (hμ : ∀ a, μ {a} ≠ 0) (f : α → E) : eLpNormEssSup f μ = ⨆ a, ↑‖f a‖₊ := + essSup_eq_iSup hμ _ + +@[simp] lemma eLpNormEssSup_count [MeasurableSingletonClass α] (f : α → E) : + eLpNormEssSup f .count = ⨆ a, ↑‖f a‖₊ := essSup_count _ + theorem Memℒp.left_of_add_measure {f : α → E} (h : Memℒp f p (μ + ν)) : Memℒp f p μ := h.mono_measure <| Measure.le_add_right <| le_refl _ diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index e5e60b719ee57..4428be6393921 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -1107,6 +1107,9 @@ namespace AEStronglyMeasurable variable {m : MeasurableSpace α} {μ ν : Measure α} [TopologicalSpace β] [TopologicalSpace γ] {f g : α → β} +lemma of_finite [DiscreteMeasurableSpace α] [Finite α] : AEStronglyMeasurable f μ := + ⟨_, .of_finite _, ae_eq_rfl⟩ + section Mk /-- A `StronglyMeasurable` function such that `f =ᵐ[μ] hf.mk f`. See lemmas diff --git a/Mathlib/MeasureTheory/OuterMeasure/AE.lean b/Mathlib/MeasureTheory/OuterMeasure/AE.lean index ba7bd3fa9d612..9291c540a5bcf 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/AE.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/AE.lean @@ -106,8 +106,9 @@ theorem ae_ball_iff {ι : Type*} {S : Set ι} (hS : S.Countable) {p : α → ∀ (∀ᵐ x ∂μ, ∀ i (hi : i ∈ S), p x i hi) ↔ ∀ i (hi : i ∈ S), ∀ᵐ x ∂μ, p x i hi := eventually_countable_ball hS -theorem ae_eq_refl (f : α → β) : f =ᵐ[μ] f := - EventuallyEq.rfl +lemma ae_eq_refl (f : α → β) : f =ᵐ[μ] f := EventuallyEq.rfl +lemma ae_eq_rfl {f : α → β} : f =ᵐ[μ] f := EventuallyEq.rfl +lemma ae_eq_comm {f g : α → β} : f =ᵐ[μ] g ↔ g =ᵐ[μ] f := eventuallyEq_comm theorem ae_eq_symm {f g : α → β} (h : f =ᵐ[μ] g) : g =ᵐ[μ] f := h.symm @@ -115,6 +116,15 @@ theorem ae_eq_symm {f g : α → β} (h : f =ᵐ[μ] g) : g =ᵐ[μ] f := theorem ae_eq_trans {f g h : α → β} (h₁ : f =ᵐ[μ] g) (h₂ : g =ᵐ[μ] h) : f =ᵐ[μ] h := h₁.trans h₂ +@[simp] lemma ae_eq_top : ae μ = ⊤ ↔ ∀ a, μ {a} ≠ 0 := by + simp only [Filter.ext_iff, mem_ae_iff, mem_top, ne_eq] + refine ⟨fun h a ha ↦ by simpa [ha] using (h {a}ᶜ).1, fun h s ↦ ⟨fun hs ↦ ?_, ?_⟩⟩ + · rw [← compl_empty_iff, ← not_nonempty_iff_eq_empty] + rintro ⟨a, ha⟩ + exact h _ $ measure_mono_null (singleton_subset_iff.2 ha) hs + · rintro rfl + simp + theorem ae_le_of_ae_lt {β : Type*} [Preorder β] {f g : α → β} (h : ∀ᵐ x ∂μ, f x < g x) : f ≤ᵐ[μ] g := h.mono fun _ ↦ le_of_lt diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index a4e53468a767d..9111ee66adb85 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -1309,6 +1309,8 @@ protected theorem EventuallyEq.rfl {l : Filter α} {f : α → β} : f =ᶠ[l] f theorem EventuallyEq.symm {f g : α → β} {l : Filter α} (H : f =ᶠ[l] g) : g =ᶠ[l] f := H.mono fun _ => Eq.symm +lemma eventuallyEq_comm {f g : α → β} {l : Filter α} : f =ᶠ[l] g ↔ g =ᶠ[l] f := ⟨.symm, .symm⟩ + @[trans] theorem EventuallyEq.trans {l : Filter α} {f g h : α → β} (H₁ : f =ᶠ[l] g) (H₂ : g =ᶠ[l] h) : f =ᶠ[l] h := diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index daf5f974aacdb..11b9e050b8436 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -773,6 +773,8 @@ theorem HasBasis.limsup_eq_iInf_iSup {p : ι → Prop} {s : ι → Set β} {f : (h : f.HasBasis p s) : limsup u f = ⨅ (i) (_ : p i), ⨆ a ∈ s i, u a := (h.map u).limsSup_eq_iInf_sSup.trans <| by simp only [sSup_image, id] +@[simp] lemma limsup_top (u : β → α) : limsup u ⊤ = ⨆ i, u i := by simp [limsup_eq_iInf_iSup] + theorem blimsup_congr' {f : Filter β} {p q : β → Prop} {u : β → α} (h : ∀ᶠ x in f, u x ≠ ⊥ → (p x ↔ q x)) : blimsup u f p = blimsup u f q := by simp only [blimsup_eq] @@ -810,6 +812,8 @@ theorem liminf_eq_iSup_iInf_of_nat {u : ℕ → α} : liminf u atTop = ⨆ n : theorem liminf_eq_iSup_iInf_of_nat' {u : ℕ → α} : liminf u atTop = ⨆ n : ℕ, ⨅ i : ℕ, u (i + n) := @limsup_eq_iInf_iSup_of_nat' αᵒᵈ _ _ +@[simp] lemma liminf_top (u : β → α) : liminf u ⊤ = ⨅ i, u i := by simp [liminf_eq_iSup_iInf] + theorem HasBasis.liminf_eq_iSup_iInf {p : ι → Prop} {s : ι → Set β} {f : Filter β} {u : β → α} (h : f.HasBasis p s) : liminf u f = ⨆ (i) (_ : p i), ⨅ a ∈ s i, u a := HasBasis.limsup_eq_iInf_iSup (α := αᵒᵈ) h From 3913a5eea1ac1b2287100db6c586f4cf1f92ee06 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Aug 2024 09:35:32 +0000 Subject: [PATCH 324/359] feat(NNRat): More basic lemmas (#15047) From LeanAPAP --- Mathlib/Algebra/Algebra/Rat.lean | 100 +++++++++++++--- .../Analysis/SpecialFunctions/Log/Basic.lean | 1 + Mathlib/Data/NNRat/Defs.lean | 17 ++- Mathlib/Data/Rat/Cast/CharZero.lean | 112 +++++++++++------- Mathlib/Data/Rat/Cast/Defs.lean | 50 ++++++-- Mathlib/Data/Rat/Cast/Lemmas.lean | 8 +- Mathlib/NumberTheory/Bernoulli.lean | 2 +- Mathlib/NumberTheory/PythagoreanTriples.lean | 2 +- test/TCSynth.lean | 2 +- test/instance_diamonds/algebra_rat.lean | 7 ++ 10 files changed, 223 insertions(+), 78 deletions(-) create mode 100644 test/instance_diamonds/algebra_rat.lean diff --git a/Mathlib/Algebra/Algebra/Rat.lean b/Mathlib/Algebra/Algebra/Rat.lean index 6ad708e5d3439..ed05e1a90ee28 100644 --- a/Mathlib/Algebra/Algebra/Rat.lean +++ b/Mathlib/Algebra/Algebra/Rat.lean @@ -3,8 +3,9 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Yury Kudryashov -/ -import Mathlib.Data.Rat.Cast.CharZero import Mathlib.Algebra.Algebra.Defs +import Mathlib.Algebra.GroupWithZero.Action.Basic +import Mathlib.Data.Rat.Cast.CharZero /-! # Further basic results about `Algebra`'s over `ℚ`. @@ -12,9 +13,9 @@ import Mathlib.Algebra.Algebra.Defs This file could usefully be split further. -/ -namespace RingHom +variable {F R S : Type*} -variable {R S : Type*} +namespace RingHom -- Porting note: changed `[Ring R] [Ring S]` to `[Semiring R] [Semiring S]` -- otherwise, Lean failed to find a `Subsingleton (ℚ →+* S)` instance @@ -25,26 +26,89 @@ theorem map_rat_algebraMap [Semiring R] [Semiring S] [Algebra ℚ R] [Algebra end RingHom -section Rat +namespace NNRat +variable [DivisionSemiring R] [CharZero R] + +section Semiring +variable [Semiring S] [Module ℚ≥0 S] + +variable (R) in +/-- `nnqsmul` is equal to any other module structure via a cast. -/ +lemma cast_smul_eq_nnqsmul [Module R S] (q : ℚ≥0) (a : S) : (q : R) • a = q • a := by + refine MulAction.injective₀ (G₀ := ℚ≥0) (Nat.cast_ne_zero.2 q.den_pos.ne') ?_ + dsimp + rw [← mul_smul, den_mul_eq_num, Nat.cast_smul_eq_nsmul, Nat.cast_smul_eq_nsmul, ← smul_assoc, + nsmul_eq_mul q.den, ← cast_natCast, ← cast_mul, den_mul_eq_num, cast_natCast, + Nat.cast_smul_eq_nsmul] + +end Semiring + +section DivisionSemiring +variable [DivisionSemiring S] [CharZero S] + +instance _root_.DivisionSemiring.toNNRatAlgebra : Algebra ℚ≥0 R where + smul_def' := smul_def + toRingHom := castHom _ + commutes' := cast_commute + +instance _root_.RingHomClass.toLinearMapClassNNRat [FunLike F R S] [RingHomClass F R S] : + LinearMapClass F ℚ≥0 R S where + map_smulₛₗ f q a := by simp [smul_def, cast_id] + +variable [SMul R S] + +instance instSMulCommClass [SMulCommClass R S S] : SMulCommClass ℚ≥0 R S where + smul_comm q a b := by simp [smul_def, mul_smul_comm] + +instance instSMulCommClass' [SMulCommClass S R S] : SMulCommClass R ℚ≥0 S := + have := SMulCommClass.symm S R S; SMulCommClass.symm _ _ _ + +end DivisionSemiring +end NNRat + +namespace Rat +variable [DivisionRing R] [CharZero R] + +section Ring +variable [Ring S] [Module ℚ S] + +variable (R) in +/-- `nnqsmul` is equal to any other module structure via a cast. -/ +lemma cast_smul_eq_qsmul [Module R S] (q : ℚ) (a : S) : (q : R) • a = q • a := by + refine MulAction.injective₀ (G₀ := ℚ) (Nat.cast_ne_zero.2 q.den_pos.ne') ?_ + dsimp + rw [← mul_smul, den_mul_eq_num, Nat.cast_smul_eq_nsmul, Int.cast_smul_eq_zsmul, ← smul_assoc, + nsmul_eq_mul q.den, ← cast_natCast, ← cast_mul, den_mul_eq_num, cast_intCast, + Int.cast_smul_eq_zsmul] + +end Ring + +section DivisionRing +variable [DivisionRing S] [CharZero S] + +instance _root_.DivisionRing.toRatAlgebra : Algebra ℚ R where + smul_def' := smul_def + toRingHom := castHom _ + commutes' := cast_commute + +instance _root_.RingHomClass.toLinearMapClassRat [FunLike F R S] [RingHomClass F R S] : + LinearMapClass F ℚ R S where + map_smulₛₗ f q a := by simp [smul_def, cast_id] + +variable [SMul R S] -/-- Every division ring of characteristic zero is an algebra over the rationals. -/ -instance DivisionRing.toRatAlgebra {α} [DivisionRing α] [CharZero α] : Algebra ℚ α where - smul := (· • ·) - smul_def' := Rat.smul_def - toRingHom := Rat.castHom α - commutes' := Rat.cast_commute +instance instSMulCommClass [SMulCommClass R S S] : SMulCommClass ℚ R S where + smul_comm q a b := by simp [smul_def, mul_smul_comm] -/-- The rational numbers are an algebra over the non-negative rationals. -/ -instance : Algebra NNRat ℚ := - NNRat.coeHom.toAlgebra +instance instSMulCommClass' [SMulCommClass S R S] : SMulCommClass R ℚ S := + have := SMulCommClass.symm S R S; SMulCommClass.symm _ _ _ -/-- The two `Algebra ℚ ℚ` instances should coincide. -/ -example : DivisionRing.toRatAlgebra = Algebra.id ℚ := - rfl +end DivisionRing -@[simp] theorem algebraMap_rat_rat : algebraMap ℚ ℚ = RingHom.id ℚ := rfl +@[deprecated Algebra.id.map_eq_id (since := "2024-07-30")] +lemma _root_.algebraMap_rat_rat : algebraMap ℚ ℚ = RingHom.id ℚ := rfl -instance algebra_rat_subsingleton {α} [Semiring α] : Subsingleton (Algebra ℚ α) := +instance algebra_rat_subsingleton {R} [Semiring R] : Subsingleton (Algebra ℚ R) := ⟨fun x y => Algebra.algebra_ext x y <| RingHom.congr_fun <| Subsingleton.elim _ _⟩ end Rat diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index 7b1feb85d5ed5..ffe83425fd545 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -6,6 +6,7 @@ Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne import Mathlib.Analysis.SpecialFunctions.Exp import Mathlib.Data.Nat.Factorization.Defs import Mathlib.Analysis.NormedSpace.Real +import Mathlib.Data.Rat.Cast.CharZero /-! # Real logarithm diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index db2875604c859..7273bddabfad4 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -30,6 +30,17 @@ Whenever you state a lemma about the coercion `ℚ≥0 → ℚ`, check that Lean `Subtype.val`. Else your lemma will never apply. -/ +library_note "specialised high priority simp lemma" /-- +It sometimes happens that a `@[simp]` lemma declared early in the library can be proved by `simp` +using later, more general simp lemmas. In that case, the following reasons might be arguments for +the early lemma to be tagged `@[simp high]` (rather than `@[simp, nolint simpNF]` or +un``@[simp]``ed): +1. There is a significant portion of the library which needs the early lemma to be available via + `simp` and which doesn't have access to the more general lemmas. +2. The more general lemmas have more complicated typeclass assumptions, causing rewrites with them + to be slower. +-/ + open Function deriving instance CanonicallyOrderedCommSemiring for NNRat @@ -56,7 +67,8 @@ theorem ext : (p : ℚ) = (q : ℚ) → p = q := protected theorem coe_injective : Injective ((↑) : ℚ≥0 → ℚ) := Subtype.coe_injective -@[simp, norm_cast] +-- See note [specialised high priority simp lemma] +@[simp high, norm_cast] theorem coe_inj : (p : ℚ) = q ↔ p = q := Subtype.coe_inj @@ -109,7 +121,8 @@ theorem coe_mul (p q : ℚ≥0) : ((p * q : ℚ≥0) : ℚ) = p * q := theorem coe_sub (h : q ≤ p) : ((p - q : ℚ≥0) : ℚ) = p - q := max_eq_left <| le_sub_comm.2 <| by rwa [sub_zero] -@[simp] +-- See note [specialised high priority simp lemma] +@[simp high] theorem coe_eq_zero : (q : ℚ) = 0 ↔ q = 0 := by norm_cast theorem coe_ne_zero : (q : ℚ) ≠ 0 ↔ q ≠ 0 := diff --git a/Mathlib/Data/Rat/Cast/CharZero.lean b/Mathlib/Data/Rat/Cast/CharZero.lean index a6ea0dbb7aa1b..e24e6bf537738 100644 --- a/Mathlib/Data/Rat/Cast/CharZero.lean +++ b/Mathlib/Data/Rat/Cast/CharZero.lean @@ -10,21 +10,15 @@ import Mathlib.Data.Rat.Cast.Defs # Casts of rational numbers into characteristic zero fields (or division rings). -/ +open Function variable {F ι α β : Type*} namespace Rat +variable [DivisionRing α] [CharZero α] {p q : ℚ} -open Rat - -section WithDivRing - -variable [DivisionRing α] - -@[simp, norm_cast] -theorem cast_inj [CharZero α] : ∀ {m n : ℚ}, (m : α) = n ↔ m = n - | ⟨n₁, d₁, d₁0, c₁⟩, ⟨n₂, d₂, d₂0, c₂⟩ => by - refine ⟨fun h => ?_, congr_arg _⟩ +lemma cast_injective : Injective ((↑) : ℚ → α) + | ⟨n₁, d₁, d₁0, c₁⟩, ⟨n₂, d₂, d₂0, c₂⟩, h => by have d₁a : (d₁ : α) ≠ 0 := Nat.cast_ne_zero.2 d₁0 have d₂a : (d₂ : α) ≠ 0 := Nat.cast_ne_zero.2 d₂0 rw [mk'_eq_divInt, mk'_eq_divInt] at h ⊢ @@ -34,30 +28,21 @@ theorem cast_inj [CharZero α] : ∀ {m n : ℚ}, (m : α) = n ↔ m = n Int.cast_mul, ← Int.cast_natCast d₂, ← Int.cast_mul, Int.cast_inj, ← mkRat_eq_iff d₁0 d₂0] at h -theorem cast_injective [CharZero α] : Function.Injective ((↑) : ℚ → α) - | _, _ => cast_inj.1 +@[simp, norm_cast] lemma cast_inj : (p : α) = q ↔ p = q := cast_injective.eq_iff -@[simp] -theorem cast_eq_zero [CharZero α] {n : ℚ} : (n : α) = 0 ↔ n = 0 := by rw [← cast_zero, cast_inj] - -theorem cast_ne_zero [CharZero α] {n : ℚ} : (n : α) ≠ 0 ↔ n ≠ 0 := - not_congr cast_eq_zero - -@[simp, norm_cast] -theorem cast_add [CharZero α] (m n) : ((m + n : ℚ) : α) = m + n := - cast_add_of_ne_zero (Nat.cast_ne_zero.2 <| ne_of_gt m.pos) (Nat.cast_ne_zero.2 <| ne_of_gt n.pos) +@[simp, norm_cast] lemma cast_eq_zero : (p : α) = 0 ↔ p = 0 := cast_injective.eq_iff' cast_zero +lemma cast_ne_zero : (p : α) ≠ 0 ↔ p ≠ 0 := cast_eq_zero.ne -@[simp, norm_cast] -theorem cast_sub [CharZero α] (m n) : ((m - n : ℚ) : α) = m - n := - cast_sub_of_ne_zero (Nat.cast_ne_zero.2 <| ne_of_gt m.pos) (Nat.cast_ne_zero.2 <| ne_of_gt n.pos) +@[simp, norm_cast] lemma cast_add (p q : ℚ) : ↑(p + q) = (p + q : α) := + cast_add_of_ne_zero (Nat.cast_ne_zero.2 p.pos.ne') (Nat.cast_ne_zero.2 q.pos.ne') -@[simp, norm_cast] -theorem cast_mul [CharZero α] (m n) : ((m * n : ℚ) : α) = m * n := - cast_mul_of_ne_zero (Nat.cast_ne_zero.2 <| ne_of_gt m.pos) (Nat.cast_ne_zero.2 <| ne_of_gt n.pos) +@[simp, norm_cast] lemma cast_sub (p q : ℚ) : ↑(p - q) = (p - q : α) := + cast_sub_of_ne_zero (Nat.cast_ne_zero.2 p.pos.ne') (Nat.cast_ne_zero.2 q.pos.ne') -variable (α) -variable [CharZero α] +@[simp, norm_cast] lemma cast_mul (p q : ℚ) : ↑(p * q) = (p * q : α) := + cast_mul_of_ne_zero (Nat.cast_ne_zero.2 p.pos.ne') (Nat.cast_ne_zero.2 q.pos.ne') +variable (α) in /-- Coercion `ℚ → α` as a `RingHom`. -/ def castHom : ℚ →+* α where toFun := (↑) @@ -66,32 +51,67 @@ def castHom : ℚ →+* α where map_zero' := cast_zero map_add' := cast_add -variable {α} - -@[simp] -theorem coe_cast_hom : ⇑(castHom α) = ((↑) : ℚ → α) := - rfl +@[simp] lemma coe_castHom : ⇑(castHom α) = ((↑) : ℚ → α) := rfl -@[simp, norm_cast] -theorem cast_inv (n) : ((n⁻¹ : ℚ) : α) = (n : α)⁻¹ := - map_inv₀ (castHom α) _ +@[deprecated (since := "2024-07-22")] alias coe_cast_hom := coe_castHom -@[simp, norm_cast] -theorem cast_div (m n) : ((m / n : ℚ) : α) = m / n := - map_div₀ (castHom α) _ _ +@[simp, norm_cast] lemma cast_inv (p : ℚ) : ↑(p⁻¹) = (p⁻¹ : α) := map_inv₀ (castHom α) _ +@[simp, norm_cast] lemma cast_div (p q : ℚ) : ↑(p / q) = (p / q : α) := map_div₀ (castHom α) .. @[simp, norm_cast] -theorem cast_zpow (q : ℚ) (n : ℤ) : ((q ^ n : ℚ) : α) = (q : α) ^ n := - map_zpow₀ (castHom α) q n +lemma cast_zpow (p : ℚ) (n : ℤ) : ↑(p ^ n) = (p ^ n : α) := map_zpow₀ (castHom α) .. @[norm_cast] theorem cast_mk (a b : ℤ) : (a /. b : α) = a / b := by simp only [divInt_eq_div, cast_div, cast_intCast] +end Rat + +namespace NNRat +variable [DivisionSemiring α] [CharZero α] {p q : ℚ≥0} + +lemma cast_injective : Injective ((↑) : ℚ≥0 → α) := by + rintro p q hpq + rw [NNRat.cast_def, NNRat.cast_def, Commute.div_eq_div_iff] at hpq + rw [← p.num_div_den, ← q.num_div_den, div_eq_div_iff] + norm_cast at hpq ⊢ + any_goals norm_cast + any_goals apply den_ne_zero + exact Nat.cast_commute .. + +@[simp, norm_cast] lemma cast_inj : (p : α) = q ↔ p = q := cast_injective.eq_iff + +@[simp, norm_cast] lemma cast_eq_zero : (q : α) = 0 ↔ q = 0 := by rw [← cast_zero, cast_inj] +lemma cast_ne_zero : (q : α) ≠ 0 ↔ q ≠ 0 := cast_eq_zero.not + +@[simp, norm_cast] lemma cast_add (p q : ℚ≥0) : ↑(p + q) = (p + q : α) := + cast_add_of_ne_zero (Nat.cast_ne_zero.2 p.den_pos.ne') (Nat.cast_ne_zero.2 q.den_pos.ne') + +@[simp, norm_cast] lemma cast_mul (p q) : (p * q : ℚ≥0) = (p * q : α) := + cast_mul_of_ne_zero (Nat.cast_ne_zero.2 p.den_pos.ne') (Nat.cast_ne_zero.2 q.den_pos.ne') + +variable (α) in +/-- Coercion `ℚ≥0 → α` as a `RingHom`. -/ +def castHom : ℚ≥0 →+* α where + toFun := (↑) + map_one' := cast_one + map_mul' := cast_mul + map_zero' := cast_zero + map_add' := cast_add + +@[simp, norm_cast] lemma coe_castHom : ⇑(castHom α) = (↑) := rfl + +@[simp, norm_cast] lemma cast_inv (p) : (p⁻¹ : ℚ≥0) = (p : α)⁻¹ := map_inv₀ (castHom α) _ +@[simp, norm_cast] lemma cast_div (p q) : (p / q : ℚ≥0) = (p / q : α) := map_div₀ (castHom α) .. + @[simp, norm_cast] -theorem cast_pow (q : ℚ) (k : ℕ) : ↑(q ^ k) = (q : α) ^ k := - (castHom α).map_pow q k +lemma cast_zpow (q : ℚ≥0) (p : ℤ) : ↑(q ^ p) = ((q : α) ^ p : α) := map_zpow₀ (castHom α) .. -end WithDivRing +@[simp] +lemma cast_divNat (a b : ℕ) : (divNat a b : α) = a / b := by + rw [← cast_natCast, ← cast_natCast b, ← cast_div] + congr + ext + apply Rat.mkRat_eq_div -end Rat +end NNRat diff --git a/Mathlib/Data/Rat/Cast/Defs.lean b/Mathlib/Data/Rat/Cast/Defs.lean index e9e86730c7e10..643f4426c88c1 100644 --- a/Mathlib/Data/Rat/Cast/Defs.lean +++ b/Mathlib/Data/Rat/Cast/Defs.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.Group.Commute.Basic import Mathlib.Algebra.GroupWithZero.Units.Lemmas import Mathlib.Algebra.Order.Field.Rat import Mathlib.Data.Int.Cast.Lemmas +import Mathlib.Data.NNRat.Lemmas import Mathlib.Data.Rat.Lemmas /-! @@ -222,19 +223,47 @@ theorem map_ratCast [DivisionRing α] [DivisionRing β] [RingHomClass F α β] ( @[simp] lemma eq_ratCast [DivisionRing α] [FunLike F ℚ α] [RingHomClass F ℚ α] (f : F) (q : ℚ) : f q = q := by rw [← map_ratCast f, Rat.cast_id] -namespace MonoidWithZeroHom +namespace MonoidWithZeroHomClass -variable {M₀ : Type*} [MonoidWithZero M₀] [FunLike F ℚ M₀] [MonoidWithZeroHomClass F ℚ M₀] -variable {f g : F} +variable {M₀ : Type*} [MonoidWithZero M₀] -/-- If `f` and `g` agree on the integers then they are equal `φ`. -/ +section NNRat +variable [FunLike F ℚ≥0 M₀] [MonoidWithZeroHomClass F ℚ≥0 M₀] {f g : F} + +/-- If monoid with zero homs `f` and `g` from `ℚ≥0` agree on the naturals then they are equal. -/ +lemma ext_nnrat' (h : ∀ n : ℕ, f n = g n) : f = g := + (DFunLike.ext f g) fun r => by + rw [← r.num_div_den, div_eq_mul_inv, map_mul, map_mul, h, eq_on_inv₀ f g] + apply h + +/-- If monoid with zero homs `f` and `g` from `ℚ≥0` agree on the naturals then they are equal. + +See note [partially-applied ext lemmas] for why `comp` is used here. -/ +@[ext] +lemma ext_nnrat {f g : ℚ≥0 →*₀ M₀} + (h : f.comp (Nat.castRingHom ℚ≥0 : ℕ →*₀ ℚ≥0) = g.comp (Nat.castRingHom ℚ≥0)) : f = g := + ext_nnrat' <| DFunLike.congr_fun h + +/-- If monoid with zero homs `f` and `g` from `ℚ≥0` agree on the positive naturals then they are +equal. -/ +lemma ext_nnrat_on_pnat (same_on_pnat : ∀ n : ℕ, 0 < n → f n = g n) : f = g := + ext_nnrat' <| DFunLike.congr_fun <| ext_nat'' + ((f : ℚ≥0 →*₀ M₀).comp (Nat.castRingHom ℚ≥0 : ℕ →*₀ ℚ≥0)) + ((g : ℚ≥0 →*₀ M₀).comp (Nat.castRingHom ℚ≥0 : ℕ →*₀ ℚ≥0)) (by simpa) + +end NNRat + +section Rat +variable [FunLike F ℚ M₀] [MonoidWithZeroHomClass F ℚ M₀] {f g : F} + +/-- If monoid with zero homs `f` and `g` from `ℚ` agree on the integers then they are equal. -/ theorem ext_rat' (h : ∀ m : ℤ, f m = g m) : f = g := (DFunLike.ext f g) fun r => by rw [← r.num_div_den, div_eq_mul_inv, map_mul, map_mul, h, ← Int.cast_natCast, eq_on_inv₀ f g] apply h -/-- If `f` and `g` agree on the integers then they are equal `φ`. +/-- If monoid with zero homs `f` and `g` from `ℚ` agree on the integers then they are equal. See note [partially-applied ext lemmas] for why `comp` is used here. -/ @[ext] @@ -242,7 +271,8 @@ theorem ext_rat {f g : ℚ →*₀ M₀} (h : f.comp (Int.castRingHom ℚ : ℤ →*₀ ℚ) = g.comp (Int.castRingHom ℚ)) : f = g := ext_rat' <| DFunLike.congr_fun h -/-- Positive integer values of a morphism `φ` and its value on `-1` completely determine `φ`. -/ +/-- If monoid with zero homs `f` and `g` from `ℚ` agree on the positive naturals and `-1` then +they are equal. -/ theorem ext_rat_on_pnat (same_on_neg_one : f (-1) = g (-1)) (same_on_pnat : ∀ n : ℕ, 0 < n → f n = g n) : f = g := ext_rat' <| @@ -252,16 +282,20 @@ theorem ext_rat_on_pnat (same_on_neg_one : f (-1) = g (-1)) (g : ℚ →*₀ M₀).comp (Int.castRingHom ℚ : ℤ →*₀ ℚ) from ext_int' (by simpa) (by simpa) -end MonoidWithZeroHom +end Rat +end MonoidWithZeroHomClass /-- Any two ring homomorphisms from `ℚ` to a semiring are equal. If the codomain is a division ring, then this lemma follows from `eq_ratCast`. -/ theorem RingHom.ext_rat {R : Type*} [Semiring R] [FunLike F ℚ R] [RingHomClass F ℚ R] (f g : F) : f = g := - MonoidWithZeroHom.ext_rat' <| + MonoidWithZeroHomClass.ext_rat' <| RingHom.congr_fun <| ((f : ℚ →+* R).comp (Int.castRingHom ℚ)).ext_int ((g : ℚ →+* R).comp (Int.castRingHom ℚ)) +instance NNRat.subsingleton_ringHom {R : Type*} [Semiring R] : Subsingleton (ℚ≥0 →+* R) where + allEq f g := MonoidWithZeroHomClass.ext_nnrat' <| by simp + instance Rat.subsingleton_ringHom {R : Type*} [Semiring R] : Subsingleton (ℚ →+* R) := ⟨RingHom.ext_rat⟩ diff --git a/Mathlib/Data/Rat/Cast/Lemmas.lean b/Mathlib/Data/Rat/Cast/Lemmas.lean index 5abedffa0bd2f..9dea343405326 100644 --- a/Mathlib/Data/Rat/Cast/Lemmas.lean +++ b/Mathlib/Data/Rat/Cast/Lemmas.lean @@ -3,8 +3,9 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Data.Rat.Cast.Defs import Mathlib.Algebra.Field.Basic +import Mathlib.Data.Rat.Cast.Defs +import Mathlib.Tactic.Positivity.Basic /-! # Some exiled lemmas about casting @@ -20,6 +21,11 @@ namespace Rat variable {α : Type*} [DivisionRing α] +@[simp, norm_cast] +lemma cast_pow (p : ℚ) (n : ℕ) : ↑(p ^ n) = (p ^ n : α) := by + rw [cast_def, cast_def, den_pow, num_pow, Nat.cast_pow, Int.cast_pow, div_eq_mul_inv, ← inv_pow, + ← (Int.cast_commute _ _).mul_pow, ← div_eq_mul_inv] + -- Porting note: rewrote proof @[simp] theorem cast_inv_nat (n : ℕ) : ((n⁻¹ : ℚ) : α) = (n : α)⁻¹ := by diff --git a/Mathlib/NumberTheory/Bernoulli.lean b/Mathlib/NumberTheory/Bernoulli.lean index e6c6d64dc34d4..fbd3725afff82 100644 --- a/Mathlib/NumberTheory/Bernoulli.lean +++ b/Mathlib/NumberTheory/Bernoulli.lean @@ -293,7 +293,7 @@ theorem sum_range_pow (n p : ℕ) : apply sum_congr rfl intros m h simp only [f, exp_pow_eq_rescale_exp, rescale, one_div, coeff_mk, RingHom.coe_mk, coeff_exp, - RingHom.id_apply, cast_mul, algebraMap_rat_rat] + RingHom.id_apply, cast_mul, Algebra.id.map_eq_id] -- manipulate factorials and binomial coefficients simp? at h says simp only [succ_eq_add_one, mem_range] at h rw [choose_eq_factorial_div_factorial h.le, eq_comm, div_eq_iff (hne q.succ), succ_eq_add_one, diff --git a/Mathlib/NumberTheory/PythagoreanTriples.lean b/Mathlib/NumberTheory/PythagoreanTriples.lean index a9e537bc9d8e7..b2482ab6382cc 100644 --- a/Mathlib/NumberTheory/PythagoreanTriples.lean +++ b/Mathlib/NumberTheory/PythagoreanTriples.lean @@ -528,7 +528,7 @@ theorem isPrimitiveClassified_of_coprime_of_odd_of_pos (hc : Int.gcd x y = 1) (h apply Rat.div_int_inj hzpos _ (h.coprime_of_coprime hc) h1.2.2.2 · show w = _ rw [← Rat.divInt_eq_div, ← Rat.divInt_mul_right (by norm_num : (2 : ℤ) ≠ 0)] - rw [Int.ediv_mul_cancel h1.1, Int.ediv_mul_cancel h1.2.1, hw2] + rw [Int.ediv_mul_cancel h1.1, Int.ediv_mul_cancel h1.2.1, hw2, Rat.divInt_eq_div] norm_cast · apply (mul_lt_mul_right (by norm_num : 0 < (2 : ℤ))).mp rw [Int.ediv_mul_cancel h1.1, zero_mul] diff --git a/test/TCSynth.lean b/test/TCSynth.lean index 391a3a3bd3101..f26afeaae9055 100644 --- a/test/TCSynth.lean +++ b/test/TCSynth.lean @@ -63,7 +63,7 @@ section -- Initial issue: https://github.com/leanprover-community/mathlib4/issues/12230 open Complex in -set_option synthInstance.maxHeartbeats 3000 in +set_option synthInstance.maxHeartbeats 3200 in example (x : ℝ) : abs (cos x + sin x * I) = 1 := by simp end diff --git a/test/instance_diamonds/algebra_rat.lean b/test/instance_diamonds/algebra_rat.lean new file mode 100644 index 0000000000000..ac1fa3612b061 --- /dev/null +++ b/test/instance_diamonds/algebra_rat.lean @@ -0,0 +1,7 @@ +import Mathlib.Algebra.Algebra.Rat + +/-- The two `Algebra ℚ≥0 ℚ≥0` instances should coincide. -/ +example : DivisionSemiring.toNNRatAlgebra = Algebra.id ℚ≥0 := rfl + +/-- The two `Algebra ℚ ℚ` instances should coincide. -/ +example : DivisionRing.toRatAlgebra = Algebra.id ℚ := rfl From f66c0d7d3e938a825479b94afc611e50c92bfd49 Mon Sep 17 00:00:00 2001 From: Lorenzo Luccioli Date: Fri, 16 Aug 2024 10:35:29 +0000 Subject: [PATCH 325/359] feat (MeasureTheory.Decomposition.RadonNicodym): `rnDeriv_mul_rnDeriv'` (#15548) Add `rnDeriv_mul_rnDeriv'`, an alternative version of `rnDeriv_mul_rnDeriv`. The difference is that the new version requires the absolute continuity of the second measure wrt the third (as opposed to the first wrt the second), but the thesis is true a.e. wrt the second measure (as opposed to wrt the third). --- Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean index 04bc751e124ce..aa62a9d723094 100644 --- a/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean +++ b/Mathlib/MeasureTheory/Decomposition/RadonNikodym.lean @@ -400,6 +400,16 @@ lemma rnDeriv_mul_rnDeriv {κ : Measure α} [SigmaFinite μ] [SigmaFinite ν] [S · exact rnDeriv_ne_top _ _ · rw [Measure.withDensity_rnDeriv_eq _ _ hμν] +lemma rnDeriv_mul_rnDeriv' {κ : Measure α} [SigmaFinite μ] [SigmaFinite ν] [SigmaFinite κ] + (hνκ : ν ≪ κ) : + μ.rnDeriv ν * ν.rnDeriv κ =ᵐ[ν] μ.rnDeriv κ := by + obtain ⟨h_meas, h_sing, hμν⟩ := Measure.haveLebesgueDecomposition_spec μ ν + filter_upwards [hνκ <| Measure.rnDeriv_add' (μ.singularPart ν) (ν.withDensity (μ.rnDeriv ν)) κ, + hνκ <| Measure.rnDeriv_withDensity_left_of_absolutelyContinuous hνκ h_meas.aemeasurable, + Measure.rnDeriv_eq_zero_of_mutuallySingular h_sing hνκ] with x hx1 hx2 hx3 + nth_rw 2 [hμν] + rw [hx1, Pi.add_apply, hx2, Pi.mul_apply, hx3, Pi.zero_apply, zero_add] + lemma rnDeriv_le_one_of_le (hμν : μ ≤ ν) [SigmaFinite ν] : μ.rnDeriv ν ≤ᵐ[ν] 1 := by refine ae_le_of_forall_setLIntegral_le_of_sigmaFinite (μ.measurable_rnDeriv ν) fun s _ _ ↦ ?_ simp only [Pi.one_apply, MeasureTheory.setLIntegral_one] From fc16569029ce5c4b661978ffc55c931ec37b7321 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Aug 2024 12:15:08 +0000 Subject: [PATCH 326/359] feat: A subgroup of a `ZMod`-module is a `ZMod`-module (#15866) From LeanAPAP --- Mathlib/Data/ZMod/Basic.lean | 24 ++++++++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 6737b8715dbd6..19c83df818076 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import Mathlib.Algebra.Ring.Prod +import Mathlib.GroupTheory.GroupAction.SubMulAction import Mathlib.GroupTheory.OrderOfElement import Mathlib.Tactic.FinCases import Mathlib.Tactic.Linarith @@ -1372,6 +1373,29 @@ end lift end ZMod +section Module +variable {S G : Type*} [AddCommGroup G] {n : ℕ} [Module (ZMod n) G] [SetLike S G] + [AddSubgroupClass S G] {K : S} {x : G} + +lemma zmod_smul_mem (hx : x ∈ K) : ∀ a : ZMod n, a • x ∈ K := by + simpa [ZMod.forall, Int.cast_smul_eq_zsmul] using zsmul_mem hx + +/-- This cannot be made an instance because of the `[Module (ZMod n) G]` argument and the fact that +`n` only appears in the second argument of `SMulMemClass`, which is an `OutParam`. -/ +lemma smulMemClass : SMulMemClass S (ZMod n) G where smul_mem _ _ {_x} hx := zmod_smul_mem hx _ + +namespace AddSubgroupClass + +instance instZModSMul : SMul (ZMod n) K where smul a x := ⟨a • x, zmod_smul_mem x.2 _⟩ + +@[simp, norm_cast] lemma coe_zmod_smul (a : ZMod n) (x : K) : ↑(a • x) = (a • x : G) := rfl + +instance instZModModule : Module (ZMod n) K := + Subtype.coe_injective.module _ (AddSubmonoidClass.subtype K) coe_zmod_smul + +end AddSubgroupClass +end Module + section AddGroup variable {α : Type*} [AddGroup α] {n : ℕ} From 5df12ecc31f10fe6ca3eb9c42c9016240acb93d7 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Fri, 16 Aug 2024 12:54:29 +0000 Subject: [PATCH 327/359] perf (Site.DenseSubsite): speed up file (#15350) `rintro` and `obtain` are misbehaving here in a way `rcases` is not. We cut the instructions by ~40% by prefacing some `rintro` calls with a `dsimp` and by replacing `obtain`s with `have; rcases`. --- Mathlib/CategoryTheory/Sites/DenseSubsite.lean | 17 ++++++++++------- 1 file changed, 10 insertions(+), 7 deletions(-) diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean index dba0ba9fb3362..f689c95f03d54 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean @@ -588,10 +588,8 @@ lemma isIso_ranCounit_app_of_isDenseSubsite (Y : Sheaf J A) (U X) : apply (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).hom_ext rintro ⟨⟨⟨⟩⟩, ⟨W⟩, g⟩ obtain ⟨g, rfl⟩ : ∃ g' : G.obj W ⟶ G.obj U, g = g'.op := ⟨g.unop, rfl⟩ - simp only [id_obj, comp_obj, StructuredArrow.proj_obj, RightExtension.coneAt_pt, - RightExtension.mk_left, RightExtension.coneAt_π_app, const_obj_obj, op_obj, - whiskeringLeft_obj_obj, RightExtension.mk_hom] apply (Y.2 X _ (IsDenseSubsite.imageSieve_mem J K G g)).isSeparatedFor.ext + 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, @@ -611,11 +609,16 @@ lemma isIso_ranCounit_app_of_isDenseSubsite (Y : Sheaf J A) (U X) : · simp only [const_obj_obj, op_obj, map_comp, hl] simp only [← map_comp_assoc, r.w] · simp [← map_comp, ← op_comp, hiUV] - · rintro ⟨⟨⟨⟩⟩, ⟨W₁⟩, g₁⟩ ⟨⟨⟨⟩⟩, ⟨W₂⟩, g₂⟩ ⟨⟨⟨⟨⟩⟩⟩, i, hi⟩ + · dsimp + rintro ⟨⟨⟨⟩⟩, ⟨W₁⟩, g₁⟩ ⟨⟨⟨⟩⟩, ⟨W₂⟩, g₂⟩ ⟨⟨⟨⟨⟩⟩⟩, i, hi⟩ dsimp at g₁ g₂ i hi - obtain rfl : g₂ = g₁ ≫ (G.map i.unop).op := by simpa only [Category.id_comp] using hi - obtain ⟨g, rfl⟩ : ∃ g' : G.obj W₁ ⟶ G.obj U, g₁ = g'.op := ⟨g₁.unop, rfl⟩ - obtain ⟨i, rfl⟩ : ∃ i' : W₂ ⟶ W₁, i = i'.op := ⟨i.unop, rfl⟩ + -- See issue #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)⟩ From 1fa16c895b501a376eedb6d926895836e6ba20b7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Aug 2024 12:54:30 +0000 Subject: [PATCH 328/359] chore: Fix precedence of `~r` infix for `IsRotated` (#15876) --- Mathlib/Data/List/Rotate.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/List/Rotate.lean b/Mathlib/Data/List/Rotate.lean index 854cb1f9533fe..c38e181f92a8d 100644 --- a/Mathlib/Data/List/Rotate.lean +++ b/Mathlib/Data/List/Rotate.lean @@ -374,7 +374,8 @@ def IsRotated : Prop := ∃ n, l.rotate n = l' @[inherit_doc List.IsRotated] -infixr:1000 " ~r " => IsRotated +-- This matches the precedence of the infix `~` for `List.Perm`, and of other relation infixes +infixr:50 " ~r " => IsRotated variable {l l'} From 9c20287c848ca85c2089199689f1e8a1bda100c1 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Fri, 16 Aug 2024 13:44:33 +0000 Subject: [PATCH 329/359] feat(Topology/Order/LowerUpperTopology): Scott and Upper topologies coincide on a complete linear order (#12234) This PR shows that the subbasis used to define the lower (upper) topology together with the universal set form a basis for the lower (upper) topology on a linear order. When the linear order is complete, the subbasis with the universal set is the topology. A consequence of this result is that the Scott and upper topologies coincide on a complete linear order. Co-authored-by: Christopher Hoskin Co-authored-by: Christopher Hoskin --- .../Topology/Order/LowerUpperTopology.lean | 84 +++++++++++++++++-- Mathlib/Topology/Order/ScottTopology.lean | 59 ++++++++----- 2 files changed, 117 insertions(+), 26 deletions(-) diff --git a/Mathlib/Topology/Order/LowerUpperTopology.lean b/Mathlib/Topology/Order/LowerUpperTopology.lean index 0f131e2cfc173..a1a516349e3ba 100644 --- a/Mathlib/Topology/Order/LowerUpperTopology.lean +++ b/Mathlib/Topology/Order/LowerUpperTopology.lean @@ -68,14 +68,14 @@ def upper (α : Type*) [Preorder α] : TopologicalSpace α := generateFrom {s | /-- Type synonym for a preorder equipped with the lower set topology. -/ def WithLower (α : Type*) := α -variable {α β} +variable {α β : Type*} namespace WithLower -/-- `toLower` is the identity function to the `WithLower` of a type. -/ +/-- `toLower` is the identity function to the `WithLower` of a type. -/ @[match_pattern] def toLower : α ≃ WithLower α := Equiv.refl _ -/-- `ofLower` is the identity function from the `WithLower` of a type. -/ +/-- `ofLower` is the identity function from the `WithLower` of a type. -/ @[match_pattern] def ofLower : WithLower α ≃ α := Equiv.refl _ @[simp] lemma to_WithLower_symm_eq : (@toLower α).symm = ofLower := rfl @@ -116,10 +116,10 @@ end WithLower def WithUpper (α : Type*) := α namespace WithUpper -/-- `toUpper` is the identity function to the `WithUpper` of a type. -/ +/-- `toUpper` is the identity function to the `WithUpper` of a type. -/ @[match_pattern] def toUpper : α ≃ WithUpper α := Equiv.refl _ -/-- `ofUpper` is the identity function from the `WithUpper` of a type. -/ +/-- `ofUpper` is the identity function from the `WithUpper` of a type. -/ @[match_pattern] def ofUpper : WithUpper α ≃ α := Equiv.refl _ @[simp] lemma to_WithUpper_symm_eq {α} : (@toUpper α).symm = ofUpper := rfl @@ -281,6 +281,61 @@ instance (priority := 90) t0Space : T0Space α := end PartialOrder +section LinearOrder + +variable [LinearOrder α] [TopologicalSpace α] [IsLower α] + +lemma isTopologicalBasis_insert_univ_subbasis : + IsTopologicalBasis (insert univ {s : Set α | ∃ a, (Ici a)ᶜ = s}) := + isTopologicalBasis_of_subbasis_of_inter (by rw [topology_eq α]; rfl) (by + rintro _ ⟨b, rfl⟩ _ ⟨c, rfl⟩ + use b ⊓ c + rw [compl_Ici, compl_Ici, compl_Ici, Iio_inter_Iio]) + +end LinearOrder + +section CompleteLinearOrder + +variable [CompleteLinearOrder α] [t : TopologicalSpace α] [IsLower α] + +lemma isTopologicalSpace_basis (U : Set α) : IsOpen U ↔ U = univ ∨ ∃ a, (Ici a)ᶜ = U := by + by_cases hU : U = univ + simp only [hU, isOpen_univ, compl_Ici, true_or] + refine ⟨?_, isTopologicalBasis_insert_univ_subbasis.isOpen⟩ + intro hO + apply Or.inr + convert IsTopologicalBasis.open_eq_sUnion isTopologicalBasis_insert_univ_subbasis hO + constructor + · intro ⟨a, ha⟩ + use {U} + constructor + · apply subset_trans (singleton_subset_iff.mpr _) (subset_insert _ _) + use a + · rw [sUnion_singleton] + · intro ⟨S, hS1, hS2⟩ + have hUS : univ ∉ S := by + by_contra hUS' + apply hU + rw [hS2] + exact sUnion_eq_univ_iff.mpr (fun a => ⟨univ, hUS', trivial⟩) + use sSup {a | (Ici a)ᶜ ∈ S} + rw [hS2, sUnion_eq_compl_sInter_compl, compl_inj_iff] + apply le_antisymm + · intro b hb + simp only [sInter_image, mem_iInter, mem_compl_iff] + intro s hs + obtain ⟨a,ha⟩ := (subset_insert_iff_of_not_mem hUS).mp hS1 hs + subst hS2 ha + simp_all only [compl_Ici, mem_Ici, sSup_le_iff, mem_setOf_eq, mem_Iio, not_lt] + · intro b hb + rw [mem_Ici, sSup_le_iff] + intro c hc + simp only [sInter_image, mem_iInter] at hb + rw [← not_lt, ← mem_Iio, ← compl_Ici] + exact hb _ hc + +end CompleteLinearOrder + end IsLower @@ -365,6 +420,25 @@ instance (priority := 90) t0Space : T0Space α := end PartialOrder +section LinearOrder + +variable [LinearOrder α] [TopologicalSpace α] [IsUpper α] + +lemma isTopologicalBasis_insert_univ_subbasis : + IsTopologicalBasis (insert univ {s : Set α | ∃ a, (Iic a)ᶜ = s}) := + IsLower.isTopologicalBasis_insert_univ_subbasis (α := αᵒᵈ) + +end LinearOrder + +section CompleteLinearOrder + +variable [CompleteLinearOrder α] [t : TopologicalSpace α] [IsUpper α] + +lemma isTopologicalSpace_basis (U : Set α) : IsOpen U ↔ U = univ ∨ ∃ a, (Iic a)ᶜ = U := + IsLower.isTopologicalSpace_basis (α := αᵒᵈ) U + +end CompleteLinearOrder + end IsUpper instance instIsLowerProd [Preorder α] [TopologicalSpace α] [IsLower α] diff --git a/Mathlib/Topology/Order/ScottTopology.lean b/Mathlib/Topology/Order/ScottTopology.lean index 2dc2f2655354e..8b5a06371adf8 100644 --- a/Mathlib/Topology/Order/ScottTopology.lean +++ b/Mathlib/Topology/Order/ScottTopology.lean @@ -165,12 +165,6 @@ lemma isOpen_iff : IsOpen s ↔ ∀ ⦃d : Set α⦄, d.Nonempty → DirectedOn (· ≤ ·) d → ∀ ⦃a : α⦄, IsLUB d a → a ∈ s → ∃ b ∈ d, Ici b ∩ d ⊆ s := by erw [topology_eq_scottHausdorff (α := α)]; rfl -lemma isOpen_of_isLowerSet (h : IsLowerSet s) : IsOpen s := - isOpen_iff.2 fun _d ⟨b, hb⟩ _ _ hda ha ↦ ⟨b, hb, fun _ hc ↦ h (mem_upperBounds.1 hda.1 _ hc.2) ha⟩ - -lemma isClosed_of_isUpperSet (h : IsUpperSet s) : IsClosed s := - isOpen_compl_iff.1 <| isOpen_of_isLowerSet h.compl - lemma dirSupInacc_of_isOpen (h : IsOpen s) : DirSupInacc s := fun d hd₁ hd₂ a hda hd₃ ↦ by obtain ⟨b, hbd, hb⟩ := isOpen_iff.1 h hd₁ hd₂ hda hd₃; exact ⟨b, hbd, hb ⟨le_rfl, hbd⟩⟩ @@ -181,6 +175,20 @@ lemma dirSupClosed_of_isClosed (h : IsClosed s) : DirSupClosed s := end IsScottHausdorff end ScottHausdorff +section ScottHausdorff +namespace IsScottHausdorff + +variable {s : Set α} [Preorder α] {t : TopologicalSpace α} [IsScottHausdorff α] + +lemma isOpen_of_isLowerSet (h : IsLowerSet s) : IsOpen s := + isOpen_iff.2 fun _d ⟨b, hb⟩ _ _ hda ha ↦ ⟨b, hb, fun _ hc ↦ h (mem_upperBounds.1 hda.1 _ hc.2) ha⟩ + +lemma isClosed_of_isUpperSet (h : IsUpperSet s) : IsClosed s := + isOpen_compl_iff.1 <| isOpen_of_isLowerSet h.compl + +end IsScottHausdorff +end ScottHausdorff + /-! ### Scott topology -/ section Scott @@ -297,24 +305,33 @@ end PartialOrder section CompleteLinearOrder -variable [CompleteLinearOrder α] [TopologicalSpace α] [Topology.IsScott α] +variable [CompleteLinearOrder α] -lemma isOpen_iff_Iic_compl_or_univ (U : Set α) : - IsOpen U ↔ (∃ (a : α), U = (Iic a)ᶜ) ∨ U = univ := by +lemma isOpen_iff_Iic_compl_or_univ [TopologicalSpace α] [Topology.IsScott α] (U : Set α) : + IsOpen U ↔ U = univ ∨ ∃ a, (Iic a)ᶜ = U := by constructor · intro hU rcases eq_empty_or_nonempty Uᶜ with eUc | neUc - · exact Or.inr (compl_empty_iff.mp eUc) - · apply Or.inl + · exact Or.inl (compl_empty_iff.mp eUc) + · apply Or.inr use sSup Uᶜ - rw [eq_compl_comm, le_antisymm_iff] - exact ⟨(isLowerSet_of_isClosed hU.isClosed_compl).Iic_subset - (dirSupClosed_iff_forall_sSup.mp (dirSupClosed_of_isClosed hU.isClosed_compl) - neUc (isChain_of_trichotomous Uᶜ).directedOn le_rfl), - fun _ ha ↦ le_sSup ha⟩ - · rintro (⟨a,rfl⟩ | rfl) - · exact isClosed_Iic.isOpen_compl + rw [compl_eq_comm, le_antisymm_iff] + exact ⟨fun _ ha ↦ le_sSup ha, (isLowerSet_of_isClosed hU.isClosed_compl).Iic_subset + (dirSupClosed_iff_forall_sSup.mp (dirSupClosed_of_isClosed hU.isClosed_compl) + neUc (isChain_of_trichotomous Uᶜ).directedOn le_rfl)⟩ + · rintro (rfl | ⟨a, rfl⟩) · exact isOpen_univ + · exact isClosed_Iic.isOpen_compl + +-- N.B. A number of conditions equivalent to `scott α = upper α` are given in Gierz _et al_, +-- Chapter III, Exercise 3.23. +lemma scott_eq_upper_of_completeLinearOrder : scott α = upper α := by + letI := upper α + ext U + rw [@Topology.IsUpper.isTopologicalSpace_basis _ _ (upper α) + ({ topology_eq_upperTopology := rfl }) U] + letI := scott α + rw [@isOpen_iff_Iic_compl_or_univ _ _ (scott α) ({ topology_eq_scott := rfl }) U] end CompleteLinearOrder @@ -367,8 +384,8 @@ end Scott variable [Preorder α] lemma scottHausdorff_le_lower : scottHausdorff α ≤ lower α := - fun s h => @IsScottHausdorff.isOpen_of_isLowerSet _ _ (scottHausdorff α) _ _ - <| (@IsLower.isLowerSet_of_isOpen (Topology.WithLower α) _ _ _ s h) + fun s h => IsScottHausdorff.isOpen_of_isLowerSet (t := scottHausdorff α) + <| (@IsLower.isLowerSet_of_isOpen (Topology.WithLower α) _ _ _ s h) variable [TopologicalSpace α] @@ -382,7 +399,7 @@ lemma IsScott.scottHausdorff_le [IsScott α] : scottHausdorff α ≤ ‹Topologi lemma IsLower.scottHausdorff_le [IsLower α] : scottHausdorff α ≤ ‹TopologicalSpace α› := fun _ h ↦ - @IsScottHausdorff.isOpen_of_isLowerSet _ _ (scottHausdorff α) _ _ + IsScottHausdorff.isOpen_of_isLowerSet (t := scottHausdorff α) <| IsLower.isLowerSet_of_isOpen h end Topology From e610482553afd2ebc7e072a8be1db795620e1d50 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Fri, 16 Aug 2024 13:44:34 +0000 Subject: [PATCH 330/359] chore(SetTheory/Ordinal/NaturalOps): cleanup termination_by (#15745) We simplify some of the proofs and definitions that use `termination_by` by moving the arguments before the colon. The diff might seem large, but all I really did besides that was reindent proofs. --- Mathlib/SetTheory/Ordinal/NaturalOps.lean | 155 +++++++++++----------- 1 file changed, 75 insertions(+), 80 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/NaturalOps.lean b/Mathlib/SetTheory/Ordinal/NaturalOps.lean index 7742f066f545b..e61e3fd7ddbe2 100644 --- a/Mathlib/SetTheory/Ordinal/NaturalOps.lean +++ b/Mathlib/SetTheory/Ordinal/NaturalOps.lean @@ -173,10 +173,9 @@ to normal ordinal addition, it is commutative. Natural addition can equivalently be characterized as the ordinal resulting from adding up corresponding coefficients in the Cantor normal forms of `a` and `b`. -/ -noncomputable def nadd : Ordinal → Ordinal → Ordinal - | a, b => - max (blsub.{u, u} a fun a' _ => nadd a' b) (blsub.{u, u} b fun b' _ => nadd a b') - termination_by o₁ o₂ => (o₁, o₂) +noncomputable def nadd (a b : Ordinal) : Ordinal := + max (blsub.{u, u} a fun a' _ => nadd a' b) (blsub.{u, u} b fun b' _ => nadd a b') +termination_by (a, b) @[inherit_doc] scoped[NaturalOps] infixl:65 " ♯ " => Ordinal.nadd @@ -191,9 +190,9 @@ distributive (over natural addition). Natural multiplication can equivalently be characterized as the ordinal resulting from multiplying the Cantor normal forms of `a` and `b` as if they were polynomials in `ω`. Addition of exponents is done via natural addition. -/ -noncomputable def nmul : Ordinal.{u} → Ordinal.{u} → Ordinal.{u} - | a, b => sInf {c | ∀ a' < a, ∀ b' < b, nmul a' b ♯ nmul a b' < c ♯ nmul a' b'} -termination_by a b => (a, b) +noncomputable def nmul (a b : Ordinal.{u}) : Ordinal.{u} := + sInf {c | ∀ a' < a, ∀ b' < b, nmul a' b ♯ nmul a b' < c ♯ nmul a' b'} +termination_by (a, b) @[inherit_doc] scoped[NaturalOps] infixl:70 " ⨳ " => Ordinal.nmul @@ -230,11 +229,10 @@ theorem nadd_le_nadd_right (h : b ≤ c) (a) : b ♯ a ≤ c ♯ a := by variable (a b) -theorem nadd_comm : ∀ a b, a ♯ b = b ♯ a - | a, b => by - rw [nadd_def, nadd_def, max_comm] - congr <;> ext <;> apply nadd_comm - termination_by a b => (a,b) +theorem nadd_comm (a b) : a ♯ b = b ♯ a := by + rw [nadd_def, nadd_def, max_comm] + congr <;> ext <;> apply nadd_comm +termination_by (a,b) theorem blsub_nadd_of_mono {f : ∀ c < a ♯ b, Ordinal.{max u v}} (hf : ∀ {i j} (hi hj), i ≤ j → f i hi ≤ f j hj) : @@ -465,17 +463,16 @@ theorem lt_nmul_iff : c < a ⨳ b ↔ ∃ a' < a, ∃ b' < b, c ♯ a' ⨳ b' theorem nmul_le_iff : a ⨳ b ≤ c ↔ ∀ a' < a, ∀ b' < b, a' ⨳ b ♯ a ⨳ b' < c ♯ a' ⨳ b' := by rw [← not_iff_not]; simp [lt_nmul_iff] -theorem nmul_comm : ∀ a b, a ⨳ b = b ⨳ a - | a, b => by - rw [nmul, nmul] - congr; ext x; constructor <;> intro H c hc d hd - -- Porting note: had to add additional arguments to `nmul_comm` here - -- for the termination checker. - · rw [nadd_comm, ← nmul_comm d b, ← nmul_comm a c, ← nmul_comm d] - exact H _ hd _ hc - · rw [nadd_comm, nmul_comm a d, nmul_comm c, nmul_comm c] - exact H _ hd _ hc -termination_by a b => (a, b) +theorem nmul_comm (a b) : a ⨳ b = b ⨳ a := by + rw [nmul, nmul] + congr; ext x; constructor <;> intro H c hc d hd + -- Porting note: had to add additional arguments to `nmul_comm` here + -- for the termination checker. + · rw [nadd_comm, ← nmul_comm d b, ← nmul_comm a c, ← nmul_comm d] + exact H _ hd _ hc + · rw [nadd_comm, nmul_comm a d, nmul_comm c, nmul_comm c] + exact H _ hd _ hc +termination_by (a, b) @[simp] theorem nmul_zero (a) : a ⨳ 0 = 0 := by @@ -519,47 +516,46 @@ theorem nmul_le_nmul_of_nonneg_right (h₁ : a ≤ b) (h₂ : 0 ≤ c) : a ⨳ c rw [nmul_comm, nmul_comm b] exact nmul_le_nmul_of_nonneg_left h₁ h₂ -theorem nmul_nadd : ∀ a b c, a ⨳ (b ♯ c) = a ⨳ b ♯ a ⨳ c - | a, b, c => by - refine le_antisymm (nmul_le_iff.2 fun a' ha d hd => ?_) - (nadd_le_iff.2 ⟨fun d hd => ?_, fun d hd => ?_⟩) - · -- Porting note: adding arguments to `nmul_nadd` for the termination checker. - rw [nmul_nadd a' b c] - rcases lt_nadd_iff.1 hd with (⟨b', hb, hd⟩ | ⟨c', hc, hd⟩) - · have := nadd_lt_nadd_of_lt_of_le (nmul_nadd_lt ha hb) (nmul_nadd_le ha.le hd) - -- Porting note: adding arguments to `nmul_nadd` for the termination checker. - rw [nmul_nadd a' b' c, nmul_nadd a b' c] at this - simp only [nadd_assoc] at this - rwa [nadd_left_comm, nadd_left_comm _ (a ⨳ b'), nadd_left_comm (a ⨳ b), - nadd_lt_nadd_iff_left, nadd_left_comm (a' ⨳ b), nadd_left_comm (a ⨳ b), - nadd_lt_nadd_iff_left, ← nadd_assoc, ← nadd_assoc] at this - · have := nadd_lt_nadd_of_le_of_lt (nmul_nadd_le ha.le hd) (nmul_nadd_lt ha hc) - -- Porting note: adding arguments to `nmul_nadd` for the termination checker. - rw [nmul_nadd a' b c', nmul_nadd a b c'] at this - simp only [nadd_assoc] at this - rwa [nadd_left_comm, nadd_comm (a ⨳ c), nadd_left_comm (a' ⨳ d), nadd_left_comm (a ⨳ c'), - nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, nadd_comm (a' ⨳ c), nadd_left_comm (a ⨳ d), - nadd_left_comm (a' ⨳ b), nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, nadd_comm (a ⨳ d), - nadd_comm (a' ⨳ d), ← nadd_assoc, ← nadd_assoc] at this - · rcases lt_nmul_iff.1 hd with ⟨a', ha, b', hb, hd⟩ - have := nadd_lt_nadd_of_le_of_lt hd (nmul_nadd_lt ha (nadd_lt_nadd_right hb c)) +theorem nmul_nadd (a b c : Ordinal) : a ⨳ (b ♯ c) = a ⨳ b ♯ a ⨳ c := by + refine le_antisymm (nmul_le_iff.2 fun a' ha d hd => ?_) + (nadd_le_iff.2 ⟨fun d hd => ?_, fun d hd => ?_⟩) + · -- Porting note: adding arguments to `nmul_nadd` for the termination checker. + rw [nmul_nadd a' b c] + rcases lt_nadd_iff.1 hd with (⟨b', hb, hd⟩ | ⟨c', hc, hd⟩) + · have := nadd_lt_nadd_of_lt_of_le (nmul_nadd_lt ha hb) (nmul_nadd_le ha.le hd) -- Porting note: adding arguments to `nmul_nadd` for the termination checker. - rw [nmul_nadd a' b c, nmul_nadd a b' c, nmul_nadd a'] at this + rw [nmul_nadd a' b' c, nmul_nadd a b' c] at this simp only [nadd_assoc] at this - rwa [nadd_left_comm (a' ⨳ b'), nadd_left_comm, nadd_lt_nadd_iff_left, nadd_left_comm, - nadd_left_comm _ (a' ⨳ b'), nadd_left_comm (a ⨳ b'), nadd_lt_nadd_iff_left, - nadd_left_comm (a' ⨳ c), nadd_left_comm, nadd_lt_nadd_iff_left, nadd_left_comm, - nadd_comm _ (a' ⨳ c), nadd_lt_nadd_iff_left] at this - · rcases lt_nmul_iff.1 hd with ⟨a', ha, c', hc, hd⟩ - have := nadd_lt_nadd_of_lt_of_le (nmul_nadd_lt ha (nadd_lt_nadd_left hc b)) hd + rwa [nadd_left_comm, nadd_left_comm _ (a ⨳ b'), nadd_left_comm (a ⨳ b), + nadd_lt_nadd_iff_left, nadd_left_comm (a' ⨳ b), nadd_left_comm (a ⨳ b), + nadd_lt_nadd_iff_left, ← nadd_assoc, ← nadd_assoc] at this + · have := nadd_lt_nadd_of_le_of_lt (nmul_nadd_le ha.le hd) (nmul_nadd_lt ha hc) -- Porting note: adding arguments to `nmul_nadd` for the termination checker. - rw [nmul_nadd a' b c, nmul_nadd a b c', nmul_nadd a'] at this + rw [nmul_nadd a' b c', nmul_nadd a b c'] at this simp only [nadd_assoc] at this - rwa [nadd_left_comm _ (a' ⨳ b), nadd_lt_nadd_iff_left, nadd_left_comm (a' ⨳ c'), - nadd_left_comm _ (a' ⨳ c), nadd_lt_nadd_iff_left, nadd_left_comm, nadd_comm (a' ⨳ c'), - nadd_left_comm _ (a ⨳ c'), nadd_lt_nadd_iff_left, nadd_comm _ (a' ⨳ c'), - nadd_comm _ (a' ⨳ c'), nadd_left_comm, nadd_lt_nadd_iff_left] at this -termination_by a b c => (a, b, c) + rwa [nadd_left_comm, nadd_comm (a ⨳ c), nadd_left_comm (a' ⨳ d), nadd_left_comm (a ⨳ c'), + nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, nadd_comm (a' ⨳ c), nadd_left_comm (a ⨳ d), + nadd_left_comm (a' ⨳ b), nadd_left_comm (a ⨳ b), nadd_lt_nadd_iff_left, nadd_comm (a ⨳ d), + nadd_comm (a' ⨳ d), ← nadd_assoc, ← nadd_assoc] at this + · rcases lt_nmul_iff.1 hd with ⟨a', ha, b', hb, hd⟩ + have := nadd_lt_nadd_of_le_of_lt hd (nmul_nadd_lt ha (nadd_lt_nadd_right hb c)) + -- Porting note: adding arguments to `nmul_nadd` for the termination checker. + rw [nmul_nadd a' b c, nmul_nadd a b' c, nmul_nadd a'] at this + simp only [nadd_assoc] at this + rwa [nadd_left_comm (a' ⨳ b'), nadd_left_comm, nadd_lt_nadd_iff_left, nadd_left_comm, + nadd_left_comm _ (a' ⨳ b'), nadd_left_comm (a ⨳ b'), nadd_lt_nadd_iff_left, + nadd_left_comm (a' ⨳ c), nadd_left_comm, nadd_lt_nadd_iff_left, nadd_left_comm, + nadd_comm _ (a' ⨳ c), nadd_lt_nadd_iff_left] at this + · rcases lt_nmul_iff.1 hd with ⟨a', ha, c', hc, hd⟩ + have := nadd_lt_nadd_of_lt_of_le (nmul_nadd_lt ha (nadd_lt_nadd_left hc b)) hd + -- Porting note: adding arguments to `nmul_nadd` for the termination checker. + rw [nmul_nadd a' b c, nmul_nadd a b c', nmul_nadd a'] at this + simp only [nadd_assoc] at this + rwa [nadd_left_comm _ (a' ⨳ b), nadd_lt_nadd_iff_left, nadd_left_comm (a' ⨳ c'), + nadd_left_comm _ (a' ⨳ c), nadd_lt_nadd_iff_left, nadd_left_comm, nadd_comm (a' ⨳ c'), + nadd_left_comm _ (a ⨳ c'), nadd_lt_nadd_iff_left, nadd_comm _ (a' ⨳ c'), + nadd_comm _ (a' ⨳ c'), nadd_left_comm, nadd_lt_nadd_iff_left] at this +termination_by (a, b, c) theorem nadd_nmul (a b c) : (a ♯ b) ⨳ c = a ⨳ c ♯ b ⨳ c := by rw [nmul_comm, nmul_nadd, nmul_comm, nmul_comm c] @@ -632,26 +628,25 @@ theorem nmul_le_iff₃' : d ♯ a' ⨳ (b' ⨳ c) ♯ a' ⨳ (b ⨳ c') ♯ a ⨳ (b' ⨳ c') := by rw [← not_iff_not]; simp [lt_nmul_iff₃'] -theorem nmul_assoc : ∀ a b c, a ⨳ b ⨳ c = a ⨳ (b ⨳ c) - | a, b, c => by - apply le_antisymm - · rw [nmul_le_iff₃] - intro a' ha b' hb c' hc - -- Porting note: the next line was just - -- repeat' rw [nmul_assoc] - -- but we need to spell out the arguments for the termination checker. - rw [nmul_assoc a' b c, nmul_assoc a b' c, nmul_assoc a b c', nmul_assoc a' b' c', - nmul_assoc a' b' c, nmul_assoc a' b c', nmul_assoc a b' c'] - exact nmul_nadd_lt₃' ha hb hc - · rw [nmul_le_iff₃'] - intro a' ha b' hb c' hc - -- Porting note: the next line was just - -- repeat' rw [← nmul_assoc] - -- but we need to spell out the arguments for the termination checker. - rw [← nmul_assoc a' b c, ← nmul_assoc a b' c, ← nmul_assoc a b c', ← nmul_assoc a' b' c', - ← nmul_assoc a' b' c, ← nmul_assoc a' b c', ← nmul_assoc a b' c'] - exact nmul_nadd_lt₃ ha hb hc -termination_by a b c => (a, b, c) +theorem nmul_assoc (a b c : Ordinal) : a ⨳ b ⨳ c = a ⨳ (b ⨳ c) := by + apply le_antisymm + · rw [nmul_le_iff₃] + intro a' ha b' hb c' hc + -- Porting note: the next line was just + -- repeat' rw [nmul_assoc] + -- but we need to spell out the arguments for the termination checker. + rw [nmul_assoc a' b c, nmul_assoc a b' c, nmul_assoc a b c', nmul_assoc a' b' c', + nmul_assoc a' b' c, nmul_assoc a' b c', nmul_assoc a b' c'] + exact nmul_nadd_lt₃' ha hb hc + · rw [nmul_le_iff₃'] + intro a' ha b' hb c' hc + -- Porting note: the next line was just + -- repeat' rw [← nmul_assoc] + -- but we need to spell out the arguments for the termination checker. + rw [← nmul_assoc a' b c, ← nmul_assoc a b' c, ← nmul_assoc a b c', ← nmul_assoc a' b' c', + ← nmul_assoc a' b' c, ← nmul_assoc a' b c', ← nmul_assoc a b' c'] + exact nmul_nadd_lt₃ ha hb hc +termination_by (a, b, c) end Ordinal From b899e710a91319b968b8572436a04bb6dc931463 Mon Sep 17 00:00:00 2001 From: Seewoo Lee Date: Fri, 16 Aug 2024 13:44:35 +0000 Subject: [PATCH 331/359] refactor: Restrict scope of `radical` to `UniqueFactorizationDomain` (#15826) Restrict scope of radical as suggested in [this comment](https://github.com/leanprover-community/mathlib4/pull/14873#discussion_r1716948063) by @YaelDillies. --- Mathlib/RingTheory/Radical.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/Mathlib/RingTheory/Radical.lean b/Mathlib/RingTheory/Radical.lean index 8de194f20477d..c86cee4872a46 100644 --- a/Mathlib/RingTheory/Radical.lean +++ b/Mathlib/RingTheory/Radical.lean @@ -36,7 +36,7 @@ noncomputable section open scoped Classical -open UniqueFactorizationMonoid +namespace UniqueFactorizationMonoid -- `CancelCommMonoidWithZero` is required by `UniqueFactorizationMonoid` variable {M : Type*} [CancelCommMonoidWithZero M] [NormalizationMonoid M] @@ -103,3 +103,5 @@ theorem radical_pow_of_prime {a : M} (ha : Prime a) {n : ℕ} (hn : 0 < n) : radical (a ^ n) = normalize a := by rw [radical_pow a hn] exact radical_of_prime ha + +end UniqueFactorizationMonoid From 4ecf882317aa62ad6b6499eb6544c2675e5850d0 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Fri, 16 Aug 2024 13:44:36 +0000 Subject: [PATCH 332/359] chore: fix some to_additive technical debt (#15852) --- Mathlib/Algebra/Group/Basic.lean | 2 +- Mathlib/Algebra/Group/Defs.lean | 6 +- Mathlib/Algebra/Group/InjSurj.lean | 4 +- Mathlib/Algebra/Group/Opposite.lean | 2 +- Mathlib/Algebra/Group/Pi/Basic.lean | 4 +- Mathlib/Algebra/Order/Group/Synonym.lean | 4 +- Mathlib/Algebra/Order/Hom/Basic.lean | 2 - .../Algebra/Order/Monoid/Unbundled/Basic.lean | 3 - .../Combinatorics/Additive/RuzsaCovering.lean | 2 - Mathlib/Data/Finset/Pointwise/Basic.lean | 2 +- Mathlib/Data/Set/Pointwise/Basic.lean | 11 +- Mathlib/Logic/Small/Group.lean | 141 +++++------------- .../Group/FundamentalDomain.lean | 1 - Mathlib/MeasureTheory/Group/Measure.lean | 4 +- Mathlib/Order/Filter/Pointwise.lean | 3 +- Mathlib/Tactic/ToAdditive/Frontend.lean | 4 + Mathlib/Topology/Support.lean | 11 +- scripts/style-exceptions.txt | 1 + 18 files changed, 64 insertions(+), 143 deletions(-) diff --git a/Mathlib/Algebra/Group/Basic.lean b/Mathlib/Algebra/Group/Basic.lean index e6d74154fcb84..c487707b09042 100644 --- a/Mathlib/Algebra/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Basic.lean @@ -409,7 +409,7 @@ theorem one_div_one_div : 1 / (1 / a) = a := by simp theorem div_eq_div_iff_comm : a / b = c / d ↔ b / a = d / c := inv_inj.symm.trans <| by simp only [inv_div] -@[to_additive SubtractionMonoid.toSubNegZeroMonoid] +@[to_additive] instance (priority := 100) DivisionMonoid.toDivInvOneMonoid : DivInvOneMonoid α := { DivisionMonoid.toDivInvMonoid with inv_one := by simpa only [one_div, inv_inv] using (inv_div (1 : α) 1).symm } diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 078d8fe493dd1..e166aa22f3016 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -957,7 +957,7 @@ class InvOneClass (G : Type*) extends One G, Inv G where protected inv_one : (1 : G)⁻¹ = 1 /-- A `DivInvMonoid` where `1⁻¹ = 1`. -/ -@[to_additive SubNegZeroMonoid] +@[to_additive] class DivInvOneMonoid (G : Type*) extends DivInvMonoid G, InvOneClass G -- FIXME: `to_additive` is not operating on the second parent. (#660) @@ -983,7 +983,7 @@ class SubtractionMonoid (G : Type u) extends SubNegMonoid G, InvolutiveNeg G whe `(a * b)⁻¹ = b⁻¹ * a⁻¹` and `a * b = 1 → a⁻¹ = b`. This is the immediate common ancestor of `Group` and `GroupWithZero`. -/ -@[to_additive SubtractionMonoid] +@[to_additive] class DivisionMonoid (G : Type u) extends DivInvMonoid G, InvolutiveInv G where protected mul_inv_rev (a b : G) : (a * b)⁻¹ = b⁻¹ * a⁻¹ /-- Despite the asymmetry of `inv_eq_of_mul`, the symmetric version is true thanks to the @@ -1090,7 +1090,7 @@ theorem mul_inv_cancel_right (a b : G) : a * b * b⁻¹ = a := by theorem inv_mul_cancel_right (a b : G) : a * b⁻¹ * b = a := by rw [mul_assoc, inv_mul_cancel, mul_one] -@[to_additive AddGroup.toSubtractionMonoid] +@[to_additive] instance (priority := 100) Group.toDivisionMonoid : DivisionMonoid G := { inv_inv := fun a ↦ inv_eq_of_mul (inv_mul_cancel a) mul_inv_rev := diff --git a/Mathlib/Algebra/Group/InjSurj.lean b/Mathlib/Algebra/Group/InjSurj.lean index 4f2c17838f4dc..41d0dd27c0297 100644 --- a/Mathlib/Algebra/Group/InjSurj.lean +++ b/Mathlib/Algebra/Group/InjSurj.lean @@ -240,7 +240,7 @@ protected abbrev divInvMonoid [DivInvMonoid M₂] (f : M₁ → M₂) (hf : Inje /-- A type endowed with `1`, `*`, `⁻¹`, and `/` is a `DivInvOneMonoid` if it admits an injective map that preserves `1`, `*`, `⁻¹`, and `/` to a `DivInvOneMonoid`. See note [reducible non-instances]. -/ -@[to_additive subNegZeroMonoid +@[to_additive "A type endowed with `0`, `+`, unary `-`, and binary `-` is a `SubNegZeroMonoid` if it admits an injective map that preserves `0`, `+`, unary `-`, and binary `-` to a `SubNegZeroMonoid`. This version takes custom `nsmul` and `zsmul` as `[SMul ℕ M₁]` and @@ -253,7 +253,7 @@ protected abbrev divInvOneMonoid [DivInvOneMonoid M₂] (f : M₁ → M₂) (hf /-- A type endowed with `1`, `*`, `⁻¹`, and `/` is a `DivisionMonoid` if it admits an injective map that preserves `1`, `*`, `⁻¹`, and `/` to a `DivisionMonoid`. See note [reducible non-instances] -/ -@[to_additive subtractionMonoid +@[to_additive "A type endowed with `0`, `+`, unary `-`, and binary `-` is a `SubtractionMonoid` if it admits an injective map that preserves `0`, `+`, unary `-`, and binary `-` to a `SubtractionMonoid`. This version takes custom `nsmul` and `zsmul` as `[SMul ℕ M₁]` diff --git a/Mathlib/Algebra/Group/Opposite.lean b/Mathlib/Algebra/Group/Opposite.lean index 1f1389337c3a4..07ef939bb18eb 100644 --- a/Mathlib/Algebra/Group/Opposite.lean +++ b/Mathlib/Algebra/Group/Opposite.lean @@ -165,7 +165,7 @@ instance instDivInvMonoid [DivInvMonoid α] : DivInvMonoid αᵐᵒᵖ where rw [unop_op, zpow_natCast, pow_succ', unop_mul, unop_op, zpow_natCast] zpow_neg' _ _ := unop_injective <| DivInvMonoid.zpow_neg' _ _ -@[to_additive AddOpposite.instSubtractionMonoid] +@[to_additive] instance instDivisionMonoid [DivisionMonoid α] : DivisionMonoid αᵐᵒᵖ where toDivInvMonoid := instDivInvMonoid __ := instInvolutiveInv diff --git a/Mathlib/Algebra/Group/Pi/Basic.lean b/Mathlib/Algebra/Group/Pi/Basic.lean index 92f9824806c5f..480759b083256 100644 --- a/Mathlib/Algebra/Group/Pi/Basic.lean +++ b/Mathlib/Algebra/Group/Pi/Basic.lean @@ -193,7 +193,7 @@ instance divInvMonoid [∀ i, DivInvMonoid (f i)] : DivInvMonoid (∀ i, f i) wh zpow_succ' := by intros; ext; exact DivInvMonoid.zpow_succ' _ _ zpow_neg' := by intros; ext; exact DivInvMonoid.zpow_neg' _ _ -@[to_additive Pi.subNegZeroMonoid] +@[to_additive] instance divInvOneMonoid [∀ i, DivInvOneMonoid (f i)] : DivInvOneMonoid (∀ i, f i) where inv_one := by ext; exact inv_one @@ -201,7 +201,7 @@ instance divInvOneMonoid [∀ i, DivInvOneMonoid (f i)] : DivInvOneMonoid (∀ i instance involutiveInv [∀ i, InvolutiveInv (f i)] : InvolutiveInv (∀ i, f i) where inv_inv := by intros; ext; exact inv_inv _ -@[to_additive Pi.subtractionMonoid] +@[to_additive] instance divisionMonoid [∀ i, DivisionMonoid (f i)] : DivisionMonoid (∀ i, f i) where __ := divInvMonoid __ := involutiveInv diff --git a/Mathlib/Algebra/Order/Group/Synonym.lean b/Mathlib/Algebra/Order/Group/Synonym.lean index 9a7dede0c96ec..03750731b0ab9 100644 --- a/Mathlib/Algebra/Order/Group/Synonym.lean +++ b/Mathlib/Algebra/Order/Group/Synonym.lean @@ -86,7 +86,7 @@ instance [h : InvolutiveInv α] : InvolutiveInv αᵒᵈ := h @[to_additive] instance [h : DivInvMonoid α] : DivInvMonoid αᵒᵈ := h -@[to_additive OrderDual.subtractionMonoid] +@[to_additive] instance [h : DivisionMonoid α] : DivisionMonoid αᵒᵈ := h @[to_additive OrderDual.subtractionCommMonoid] @@ -194,7 +194,7 @@ instance [h : InvolutiveInv α] : InvolutiveInv (Lex α) := h @[to_additive] instance [h : DivInvMonoid α] : DivInvMonoid (Lex α) := h -@[to_additive existing OrderDual.subtractionMonoid] +@[to_additive] instance [h : DivisionMonoid α] : DivisionMonoid (Lex α) := h @[to_additive existing OrderDual.subtractionCommMonoid] diff --git a/Mathlib/Algebra/Order/Hom/Basic.lean b/Mathlib/Algebra/Order/Hom/Basic.lean index 983b87c6868e5..584713f30f610 100644 --- a/Mathlib/Algebra/Order/Hom/Basic.lean +++ b/Mathlib/Algebra/Order/Hom/Basic.lean @@ -122,7 +122,6 @@ theorem le_map_mul_map_div [Group α] [CommSemigroup β] [LE β] [Submultiplicat theorem le_map_add_map_div [Group α] [AddCommSemigroup β] [LE β] [MulLEAddHomClass F α β] (f : F) (a b : α) : f a ≤ f b + f (a / b) := by simpa only [add_comm, div_mul_cancel] using map_mul_le_add f (a / b) b --- Porting note (#11215): TODO: `to_additive` clashes @[to_additive] theorem le_map_div_mul_map_div [Group α] [CommSemigroup β] [LE β] [SubmultiplicativeHomClass F α β] @@ -133,7 +132,6 @@ theorem le_map_div_mul_map_div [Group α] [CommSemigroup β] [LE β] [Submultipl theorem le_map_div_add_map_div [Group α] [AddCommSemigroup β] [LE β] [MulLEAddHomClass F α β] (f : F) (a b c : α) : f (a / c) ≤ f (a / b) + f (b / c) := by simpa only [div_mul_div_cancel'] using map_mul_le_add f (a / b) (b / c) --- Porting note (#11215): TODO: `to_additive` clashes namespace Mathlib.Meta.Positivity diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean index 05333081038c7..ec0256bb10e09 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Basic.lean @@ -965,9 +965,6 @@ theorem mul_eq_one_iff_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] have : b = 1 := le_antisymm this hb And.intro ‹a = 1› ‹b = 1›) (by rintro ⟨rfl, rfl⟩; rw [mul_one]) - -- Porting note: original proof of the second implication, - -- `fun ⟨ha', hb'⟩ => by rw [ha', hb', mul_one]`, - -- had its `to_additive`-ization fail due to some bug @[deprecated (since := "2024-07-24")] alias mul_eq_one_iff' := mul_eq_one_iff_of_one_le @[deprecated (since := "2024-07-24")] alias add_eq_zero_iff' := add_eq_zero_iff_of_nonneg diff --git a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean index 06670fc204bb7..522ba740a3844 100644 --- a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean +++ b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean @@ -65,8 +65,6 @@ lemma exists_subset_mul_div (hs : s.Finite) (ht' : t.Finite) (ht : t.Nonempty) : classical obtain ⟨u, hu, hsut⟩ := Finset.exists_subset_mul_div s ht refine ⟨u, ?_⟩ - -- `norm_cast` would find these automatically, but breaks `to_additive` when it does so - rw [← Finset.coe_mul, ← Finset.coe_mul, ← Finset.coe_div] norm_cast simp [*] diff --git a/Mathlib/Data/Finset/Pointwise/Basic.lean b/Mathlib/Data/Finset/Pointwise/Basic.lean index a20237588d38e..ffbddd2367556 100644 --- a/Mathlib/Data/Finset/Pointwise/Basic.lean +++ b/Mathlib/Data/Finset/Pointwise/Basic.lean @@ -873,7 +873,7 @@ protected theorem mul_eq_one_iff : s * t = 1 ↔ ∃ a b, s = {a} ∧ t = {b} simp_rw [← coe_inj, coe_mul, coe_one, Set.mul_eq_one_iff, coe_singleton] /-- `Finset α` is a division monoid under pointwise operations if `α` is. -/ -@[to_additive subtractionMonoid +@[to_additive "`Finset α` is a subtraction monoid under pointwise operations if `α` is."] protected def divisionMonoid : DivisionMonoid (Finset α) := coe_injective.divisionMonoid _ coe_one coe_mul coe_inv coe_div coe_pow coe_zpow diff --git a/Mathlib/Data/Set/Pointwise/Basic.lean b/Mathlib/Data/Set/Pointwise/Basic.lean index bc6cb8dcea2ee..81be1a2ed1658 100644 --- a/Mathlib/Data/Set/Pointwise/Basic.lean +++ b/Mathlib/Data/Set/Pointwise/Basic.lean @@ -749,14 +749,7 @@ theorem univ_mul_of_one_mem (ht : (1 : α) ∈ t) : univ * t = univ := theorem univ_mul_univ : (univ : Set α) * univ = univ := mul_univ_of_one_mem <| mem_univ _ ---TODO: `to_additive` trips up on the `1 : ℕ` used in the pattern-matching. -@[simp] -theorem nsmul_univ {α : Type*} [AddMonoid α] : ∀ {n : ℕ}, n ≠ 0 → n • (univ : Set α) = univ - | 0 => fun h => (h rfl).elim - | 1 => fun _ => one_nsmul _ - | n + 2 => fun _ => by rw [succ_nsmul, nsmul_univ n.succ_ne_zero, univ_add_univ] - -@[to_additive existing (attr := simp) nsmul_univ] +@[to_additive (attr := simp) nsmul_univ] theorem univ_pow : ∀ {n : ℕ}, n ≠ 0 → (univ : Set α) ^ n = univ | 0 => fun h => (h rfl).elim | 1 => fun _ => pow_one _ @@ -796,7 +789,7 @@ protected theorem mul_eq_one_iff : s * t = 1 ↔ ∃ a b, s = {a} ∧ t = {b} rw [singleton_mul_singleton, h, singleton_one] /-- `Set α` is a division monoid under pointwise operations if `α` is. -/ -@[to_additive subtractionMonoid +@[to_additive "`Set α` is a subtraction monoid under pointwise operations if `α` is."] protected noncomputable def divisionMonoid : DivisionMonoid (Set α) := { Set.monoid, Set.involutiveInv, Set.div, @Set.ZPow α _ _ _ with diff --git a/Mathlib/Logic/Small/Group.lean b/Mathlib/Logic/Small/Group.lean index fa16ff5f09481..3ec5b0486b284 100644 --- a/Mathlib/Logic/Small/Group.lean +++ b/Mathlib/Logic/Small/Group.lean @@ -14,161 +14,102 @@ noncomputable section variable {α : Type*} --- FIXME: here and below, why doesn't `to_additive` work? --- We're waiting on the fix for https://github.com/leanprover/lean4/issues/2077 to arrive. +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) +@[to_additive] +noncomputable instance [One α] [Small α] : One (Shrink α) := (equivShrink _).symm.one -instance [Zero α] [Small α] : Zero (Shrink α) := (equivShrink _).symm.zero - -@[to_additive existing] -instance [One α] [Small α] : One (Shrink α) := (equivShrink _).symm.one - -@[simp] -lemma equivShrink_symm_zero [Zero α] [Small α] : (equivShrink α).symm 0 = 0 := - (equivShrink α).symm_apply_apply 0 - -@[to_additive existing (attr := simp)] +@[to_additive (attr := simp)] lemma equivShrink_symm_one [One α] [Small α] : (equivShrink α).symm 1 = 1 := (equivShrink α).symm_apply_apply 1 -instance [Add α] [Small α] : Add (Shrink α) := (equivShrink _).symm.add - -@[to_additive existing] -instance [Mul α] [Small α] : Mul (Shrink α) := (equivShrink _).symm.mul +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) +@[to_additive] +noncomputable instance [Mul α] [Small α] : Mul (Shrink α) := (equivShrink _).symm.mul -@[simp] -lemma equivShrink_symm_add [Add α] [Small α] (x y : Shrink α) : - (equivShrink α).symm (x + y) = (equivShrink α).symm x + (equivShrink α).symm y := by - rw [Equiv.add_def] - simp - -@[simp] -lemma equivShrink_add [Add α] [Small α] (x y : α) : - equivShrink α (x + y) = equivShrink α x + equivShrink α y := by - rw [Equiv.add_def] - simp - -@[to_additive existing (attr := simp)] +@[to_additive (attr := simp)] lemma equivShrink_symm_mul [Mul α] [Small α] (x y : Shrink α) : (equivShrink α).symm (x * y) = (equivShrink α).symm x * (equivShrink α).symm y := by rw [Equiv.mul_def] simp -@[to_additive existing (attr := simp)] +@[to_additive (attr := simp)] lemma equivShrink_mul [Mul α] [Small α] (x y : α) : equivShrink α (x * y) = equivShrink α x * equivShrink α y := by rw [Equiv.mul_def] simp -instance [Sub α] [Small α] : Sub (Shrink α) := (equivShrink _).symm.sub - -@[to_additive existing] -instance [Div α] [Small α] : Div (Shrink α) := (equivShrink _).symm.div - -@[simp] -lemma equivShrink_symm_sub [Sub α] [Small α] (x y : Shrink α) : - (equivShrink α).symm (x - y) = (equivShrink α).symm x - (equivShrink α).symm y := by - rw [Equiv.sub_def] - simp - -@[simp] -lemma equivShrink_sub [Sub α] [Small α] (x y : α) : - equivShrink α (x - y) = equivShrink α x - equivShrink α y := by - rw [Equiv.sub_def] - simp +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) +@[to_additive] +noncomputable instance [Div α] [Small α] : Div (Shrink α) := (equivShrink _).symm.div -@[to_additive existing (attr := simp)] +@[to_additive (attr := simp)] lemma equivShrink_symm_div [Div α] [Small α] (x y : Shrink α) : (equivShrink α).symm (x / y) = (equivShrink α).symm x / (equivShrink α).symm y := by rw [Equiv.div_def] simp -@[to_additive existing (attr := simp)] +@[to_additive (attr := simp)] lemma equivShrink_div [Div α] [Small α] (x y : α) : equivShrink α (x / y) = equivShrink α x / equivShrink α y := by rw [Equiv.div_def] simp -instance [Neg α] [Small α] : Neg (Shrink α) := (equivShrink _).symm.Neg +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) +@[to_additive] +noncomputable instance [Inv α] [Small α] : Inv (Shrink α) := (equivShrink _).symm.Inv -@[to_additive existing] -instance [Inv α] [Small α] : Inv (Shrink α) := (equivShrink _).symm.Inv - -@[simp] -lemma equivShrink_symm_neg [Neg α] [Small α] (x : Shrink α) : - (equivShrink α).symm (-x) = -(equivShrink α).symm x := by - rw [Equiv.neg_def] - simp - -@[simp] -lemma equivShrink_neg [Neg α] [Small α] (x : α) : - equivShrink α (-x) = -equivShrink α x := by - rw [Equiv.neg_def] - simp - -@[to_additive existing (attr := simp)] +@[to_additive (attr := simp)] lemma equivShrink_symm_inv [Inv α] [Small α] (x : Shrink α) : (equivShrink α).symm x⁻¹ = ((equivShrink α).symm x)⁻¹ := by rw [Equiv.inv_def] simp -@[to_additive existing (attr := simp)] +@[to_additive (attr := simp)] lemma equivShrink_inv [Inv α] [Small α] (x : α) : equivShrink α x⁻¹ = (equivShrink α x)⁻¹ := by rw [Equiv.inv_def] simp -instance [AddSemigroup α] [Small α] : AddSemigroup (Shrink α) := (equivShrink _).symm.addSemigroup - -@[to_additive existing] -instance [Semigroup α] [Small α] : Semigroup (Shrink α) := (equivShrink _).symm.semigroup +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) +@[to_additive] +noncomputable instance [Semigroup α] [Small α] : Semigroup (Shrink α) := + (equivShrink _).symm.semigroup instance [SemigroupWithZero α] [Small α] : SemigroupWithZero (Shrink α) := (equivShrink _).symm.semigroupWithZero -instance [AddCommSemigroup α] [Small α] : AddCommSemigroup (Shrink α) := - (equivShrink _).symm.addCommSemigroup - -@[to_additive existing] -instance [CommSemigroup α] [Small α] : CommSemigroup (Shrink α) := +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) +@[to_additive] +noncomputable instance [CommSemigroup α] [Small α] : CommSemigroup (Shrink α) := (equivShrink _).symm.commSemigroup instance [MulZeroClass α] [Small α] : MulZeroClass (Shrink α) := (equivShrink _).symm.mulZeroClass -instance [AddZeroClass α] [Small α] : AddZeroClass (Shrink α) := - (equivShrink _).symm.addZeroClass - -@[to_additive existing] -instance [MulOneClass α] [Small α] : MulOneClass (Shrink α) := +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) +@[to_additive] +noncomputable instance [MulOneClass α] [Small α] : MulOneClass (Shrink α) := (equivShrink _).symm.mulOneClass instance [MulZeroOneClass α] [Small α] : MulZeroOneClass (Shrink α) := (equivShrink _).symm.mulZeroOneClass -instance [AddMonoid α] [Small α] : AddMonoid (Shrink α) := - (equivShrink _).symm.addMonoid - -@[to_additive existing] -instance [Monoid α] [Small α] : Monoid (Shrink α) := +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) +@[to_additive] +noncomputable instance [Monoid α] [Small α] : Monoid (Shrink α) := (equivShrink _).symm.monoid -instance [AddCommMonoid α] [Small α] : AddCommMonoid (Shrink α) := - (equivShrink _).symm.addCommMonoid - -@[to_additive existing] -instance [CommMonoid α] [Small α] : CommMonoid (Shrink α) := +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) +@[to_additive] +noncomputable instance [CommMonoid α] [Small α] : CommMonoid (Shrink α) := (equivShrink _).symm.commMonoid -instance [AddGroup α] [Small α] : AddGroup (Shrink α) := - (equivShrink _).symm.addGroup - -@[to_additive existing] -instance [Group α] [Small α] : Group (Shrink α) := +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) +@[to_additive] +noncomputable instance [Group α] [Small α] : Group (Shrink α) := (equivShrink _).symm.group -instance [AddCommGroup α] [Small α] : AddCommGroup (Shrink α) := - (equivShrink _).symm.addCommGroup - -@[to_additive existing] -instance [CommGroup α] [Small α] : CommGroup (Shrink α) := +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) +@[to_additive] +noncomputable instance [CommGroup α] [Small α] : CommGroup (Shrink α) := (equivShrink _).symm.commGroup diff --git a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean index 5b5f6208fd138..37cbb63632c50 100644 --- a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean +++ b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean @@ -540,7 +540,6 @@ theorem fundamentalInterior_union_fundamentalFrontier : theorem fundamentalFrontier_union_fundamentalInterior : fundamentalFrontier G s ∪ fundamentalInterior G s = s := inter_union_diff _ _ --- Porting note: there is a typo in `to_additive` in mathlib3, so there is no additive version @[to_additive (attr := simp) MeasureTheory.sdiff_addFundamentalInterior] theorem sdiff_fundamentalInterior : s \ fundamentalInterior G s = fundamentalFrontier G s := diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 4c2e837ace6f7..f1af9d26dac18 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -277,9 +277,7 @@ end Group namespace Measure --- Porting note: Even in `noncomputable section`, a definition with `to_additive` require --- `noncomputable` to generate an additive definition. --- Please refer to leanprover/lean4#2077. +-- TODO: noncomputable has to be specified explicitly. #1074 (item 8) /-- The measure `A ↦ μ (A⁻¹)`, where `A⁻¹` is the pointwise inverse of `A`. -/ @[to_additive "The measure `A ↦ μ (- A)`, where `- A` is the pointwise negation of `A`."] diff --git a/Mathlib/Order/Filter/Pointwise.lean b/Mathlib/Order/Filter/Pointwise.lean index b8be6d1b0884c..4e0a809056265 100644 --- a/Mathlib/Order/Filter/Pointwise.lean +++ b/Mathlib/Order/Filter/Pointwise.lean @@ -598,9 +598,8 @@ protected theorem mul_eq_one_iff : f * g = 1 ↔ ∃ a b, f = pure a ∧ g = pur rw [pure_mul_pure, h, pure_one] /-- `Filter α` is a division monoid under pointwise operations if `α` is. -/ -@[to_additive subtractionMonoid "`Filter α` is a subtraction monoid under pointwise operations if +@[to_additive "`Filter α` is a subtraction monoid under pointwise operations if `α` is."] --- Porting note: `to_additive` guessed `divisionAddMonoid` protected def divisionMonoid : DivisionMonoid (Filter α) := { Filter.monoid, Filter.instInvolutiveInv, Filter.instDiv, Filter.instZPow (α := α) with mul_inv_rev := fun s t => map_map₂_antidistrib mul_inv_rev diff --git a/Mathlib/Tactic/ToAdditive/Frontend.lean b/Mathlib/Tactic/ToAdditive/Frontend.lean index f3d84bce1ec87..f1201fe61ce4c 100644 --- a/Mathlib/Tactic/ToAdditive/Frontend.lean +++ b/Mathlib/Tactic/ToAdditive/Frontend.lean @@ -1044,6 +1044,10 @@ def fixAbbreviation : List String → List String | "le" :: "Zero" :: "Part" :: s => "negPart" :: fixAbbreviation s | "three" :: "GPFree" :: s => "three" :: "APFree" :: fixAbbreviation s | "Three" :: "GPFree" :: s => "Three" :: "APFree" :: fixAbbreviation s + | "Division" :: "Add" :: "Monoid" :: s => "SubtractionMonoid" :: fixAbbreviation s + | "division" :: "Add" :: "Monoid" :: s => "subtractionMonoid" :: fixAbbreviation s + | "Sub" :: "Neg" :: "Zero" :: "Add" :: "Monoid" :: s => "SubNegZeroMonoid" :: fixAbbreviation s + | "sub" :: "Neg" :: "Zero" :: "Add" :: "Monoid" :: s => "subNegZeroMonoid" :: fixAbbreviation s | x :: s => x :: fixAbbreviation s | [] => [] diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index fbccca79fa36f..3ddebbe31bfd6 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -212,15 +212,8 @@ theorem comp₂_left (hf : HasCompactMulSupport f) (hf₂ : HasCompactMulSupport f₂) (hm : m 1 1 = 1) : HasCompactMulSupport fun x => m (f x) (f₂ x) := by rw [hasCompactMulSupport_iff_eventuallyEq] at hf hf₂ ⊢ - #adaptation_note /-- `nightly-2024-03-11` - If we *either* (1) remove the type annotations on the - binders in the following `fun` or (2) revert `simp only` to `simp_rw`, `to_additive` fails - because an `OfNat.ofNat 1` is not replaced with `0`. Notably, as of this nightly, what used to - look like `OfNat.ofNat (nat_lit 1) x` in the proof term now looks like - `OfNat.ofNat (OfNat.ofNat (α := ℕ) (nat_lit 1)) x`, and this seems to trip up `to_additive`. - -/ - filter_upwards [hf, hf₂] using fun x (hx : f x = (1 : α → β) x) (hx₂ : f₂ x = (1 : α → γ) x) => by - simp only [hx, hx₂, Pi.one_apply, hm] + filter_upwards [hf, hf₂] with x hx hx₂ + simp_rw [hx, hx₂, Pi.one_apply, hm] @[to_additive] lemma isCompact_preimage [TopologicalSpace β] diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index 9edb3fdcd7751..1510a1e80cc85 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -55,6 +55,7 @@ Mathlib/SetTheory/Game/PGame.lean : line 1 : ERR_NUM_LIN : 1900 file contains 17 Mathlib/SetTheory/Ordinal/Arithmetic.lean : line 1 : ERR_NUM_LIN : 2400 file contains 2269 lines, try to split it up Mathlib/SetTheory/ZFC/Basic.lean : line 1 : ERR_NUM_LIN : 1700 file contains 1549 lines, try to split it up Mathlib/Tactic/CC/Addition.lean : line 1 : ERR_NUM_LIN : 2300 file contains 2100 lines, try to split it up +Mathlib/Tactic/ToAdditive/Frontend.lean : line 1 : ERR_NUM_LIN : 1700 file contains 1502 lines, try to split it up Mathlib/Topology/Algebra/Group/Basic.lean : line 1 : ERR_NUM_LIN : 2100 file contains 1932 lines, try to split it up Mathlib/Topology/Algebra/Module/Basic.lean : line 1 : ERR_NUM_LIN : 2600 file contains 2405 lines, try to split it up Mathlib/Topology/Basic.lean : line 1 : ERR_NUM_LIN : 1900 file contains 1731 lines, try to split it up From 9b07aee13e84b0d2fee3689c3614e55b8b2db1b5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Aug 2024 13:44:38 +0000 Subject: [PATCH 333/359] chore: Rename `nonzero_of_invertible` to `Invertible.ne_zero` (#15860) From LeanAPAP --- .../Wiedijk100Theorems/SolutionOfCubic.lean | 10 ++++----- Mathlib/Algebra/CharP/Invertible.lean | 2 +- Mathlib/Algebra/GroupWithZero/Invertible.lean | 22 ++++++++++--------- Mathlib/Algebra/Ring/Invertible.lean | 2 +- .../AffineSpace/Combination.lean | 2 +- .../RingTheory/WittVector/WittPolynomial.lean | 2 +- Mathlib/Tactic/NormNum/Result.lean | 2 +- 7 files changed, 22 insertions(+), 20 deletions(-) diff --git a/Archive/Wiedijk100Theorems/SolutionOfCubic.lean b/Archive/Wiedijk100Theorems/SolutionOfCubic.lean index b780b1926bc5e..f61bda51e27dd 100644 --- a/Archive/Wiedijk100Theorems/SolutionOfCubic.lean +++ b/Archive/Wiedijk100Theorems/SolutionOfCubic.lean @@ -73,8 +73,8 @@ theorem cubic_monic_eq_zero_iff (hω : IsPrimitiveRoot ω 3) (hp : p = (3 * c - x ^ 3 + b * x ^ 2 + c * x + d = 0 ↔ x = s - t - b / 3 ∨ x = s * ω - t * ω ^ 2 - b / 3 ∨ x = s * ω ^ 2 - t * ω - b / 3 := by let y := x + b / 3 - have hi2 : (2 : K) ≠ 0 := nonzero_of_invertible _ - have hi3 : (3 : K) ≠ 0 := nonzero_of_invertible _ + have hi2 : (2 : K) ≠ 0 := Invertible.ne_zero _ + have hi3 : (3 : K) ≠ 0 := Invertible.ne_zero _ have h9 : (9 : K) = 3 ^ 2 := by norm_num have h54 : (54 : K) = 2 * 3 ^ 3 := by norm_num have h₁ : x ^ 3 + b * x ^ 2 + c * x + d = y ^ 3 + 3 * p * y - 2 * q := by @@ -92,7 +92,7 @@ theorem cubic_eq_zero_iff (ha : a ≠ 0) (hω : IsPrimitiveRoot ω 3) a * x ^ 3 + b * x ^ 2 + c * x + d = 0 ↔ x = s - t - b / (3 * a) ∨ x = s * ω - t * ω ^ 2 - b / (3 * a) ∨ x = s * ω ^ 2 - t * ω - b / (3 * a) := by - have hi3 : (3 : K) ≠ 0 := nonzero_of_invertible _ + have hi3 : (3 : K) ≠ 0 := Invertible.ne_zero _ have h9 : (9 : K) = 3 ^ 2 := by norm_num have h54 : (54 : K) = 2 * 3 ^ 3 := by norm_num have h₁ : a * x ^ 3 + b * x ^ 2 + c * x + d @@ -119,8 +119,8 @@ theorem cubic_eq_zero_iff_of_p_eq_zero (ha : a ≠ 0) (hω : IsPrimitiveRoot ω x = s - b / (3 * a) ∨ x = s * ω - b / (3 * a) ∨ x = s * ω ^ 2 - b / (3 * a) := by have h₁ : ∀ x a₁ a₂ a₃ : K, x = a₁ ∨ x = a₂ ∨ x = a₃ ↔ (x - a₁) * (x - a₂) * (x - a₃) = 0 := by intros; simp only [mul_eq_zero, sub_eq_zero, or_assoc] - have hi2 : (2 : K) ≠ 0 := nonzero_of_invertible _ - have hi3 : (3 : K) ≠ 0 := nonzero_of_invertible _ + have hi2 : (2 : K) ≠ 0 := Invertible.ne_zero _ + have hi3 : (3 : K) ≠ 0 := Invertible.ne_zero _ have h54 : (54 : K) = 2 * 3 ^ 3 := by norm_num have hb2 : b ^ 2 = 3 * a * c := by rw [sub_eq_zero] at hpz; rw [hpz] have hb3 : b ^ 3 = 3 * a * b * c := by rw [pow_succ, hb2]; ring diff --git a/Mathlib/Algebra/CharP/Invertible.lean b/Mathlib/Algebra/CharP/Invertible.lean index ecff3a748a171..bdf5259b6cf73 100644 --- a/Mathlib/Algebra/CharP/Invertible.lean +++ b/Mathlib/Algebra/CharP/Invertible.lean @@ -28,7 +28,7 @@ def invertibleOfRingCharNotDvd {t : ℕ} (not_dvd : ¬ringChar K ∣ t) : Invert theorem not_ringChar_dvd_of_invertible {t : ℕ} [Invertible (t : K)] : ¬ringChar K ∣ t := by rw [← ringChar.spec, ← Ne] - exact nonzero_of_invertible (t : K) + exact Invertible.ne_zero (t : K) /-- A natural number `t` is invertible in a field `K` of characteristic `p` if `p` does not divide `t`. -/ diff --git a/Mathlib/Algebra/GroupWithZero/Invertible.lean b/Mathlib/Algebra/GroupWithZero/Invertible.lean index 36f4cf1c801f9..6895aebeb4888 100644 --- a/Mathlib/Algebra/GroupWithZero/Invertible.lean +++ b/Mathlib/Algebra/GroupWithZero/Invertible.lean @@ -18,16 +18,18 @@ universe u variable {α : Type u} -theorem nonzero_of_invertible [MulZeroOneClass α] (a : α) [Nontrivial α] [Invertible a] : a ≠ 0 := +theorem Invertible.ne_zero [MulZeroOneClass α] (a : α) [Nontrivial α] [Invertible a] : a ≠ 0 := fun ha => zero_ne_one <| calc 0 = ⅟ a * a := by simp [ha] - _ = 1 := invOf_mul_self a + _ = 1 := invOf_mul_self -instance (priority := 100) Invertible.ne_zero [MulZeroOneClass α] [Nontrivial α] (a : α) +@[deprecated (since := "2024-08-15")] alias nonzero_of_invertible := Invertible.ne_zero + +instance (priority := 100) Invertible.toNeZero [MulZeroOneClass α] [Nontrivial α] (a : α) [Invertible a] : NeZero a := - ⟨nonzero_of_invertible a⟩ + ⟨Invertible.ne_zero a⟩ section MonoidWithZero variable [MonoidWithZero α] @@ -48,15 +50,15 @@ def invertibleOfNonzero {a : α} (h : a ≠ 0) : Invertible a := @[simp] theorem invOf_eq_inv (a : α) [Invertible a] : ⅟ a = a⁻¹ := - invOf_eq_right_inv (mul_inv_cancel₀ (nonzero_of_invertible a)) + invOf_eq_right_inv (mul_inv_cancel₀ (Invertible.ne_zero a)) @[simp] theorem inv_mul_cancel_of_invertible (a : α) [Invertible a] : a⁻¹ * a = 1 := - inv_mul_cancel₀ (nonzero_of_invertible a) + inv_mul_cancel₀ (Invertible.ne_zero a) @[simp] theorem mul_inv_cancel_of_invertible (a : α) [Invertible a] : a * a⁻¹ = 1 := - mul_inv_cancel₀ (nonzero_of_invertible a) + mul_inv_cancel₀ (Invertible.ne_zero a) /-- `a` is the inverse of `a⁻¹` -/ def invertibleInv {a : α} [Invertible a] : Invertible a⁻¹ := @@ -64,15 +66,15 @@ def invertibleInv {a : α} [Invertible a] : Invertible a⁻¹ := @[simp] theorem div_mul_cancel_of_invertible (a b : α) [Invertible b] : a / b * b = a := - div_mul_cancel₀ a (nonzero_of_invertible b) + div_mul_cancel₀ a (Invertible.ne_zero b) @[simp] theorem mul_div_cancel_of_invertible (a b : α) [Invertible b] : a * b / b = a := - mul_div_cancel_right₀ a (nonzero_of_invertible b) + mul_div_cancel_right₀ a (Invertible.ne_zero b) @[simp] theorem div_self_of_invertible (a : α) [Invertible a] : a / a = 1 := - div_self (nonzero_of_invertible a) + div_self (Invertible.ne_zero a) /-- `b / a` is the inverse of `a / b` -/ def invertibleDiv (a b : α) [Invertible a] [Invertible b] : Invertible (a / b) := diff --git a/Mathlib/Algebra/Ring/Invertible.lean b/Mathlib/Algebra/Ring/Invertible.lean index a976fefc4ff53..67c2be22dc0da 100644 --- a/Mathlib/Algebra/Ring/Invertible.lean +++ b/Mathlib/Algebra/Ring/Invertible.lean @@ -33,7 +33,7 @@ theorem invOf_two_add_invOf_two [NonAssocSemiring α] [Invertible (2 : α)] : (⅟ 2 : α) + (⅟ 2 : α) = 1 := by rw [← two_mul, mul_invOf_self] theorem pos_of_invertible_cast [Semiring α] [Nontrivial α] (n : ℕ) [Invertible (n : α)] : 0 < n := - Nat.zero_lt_of_ne_zero fun h => nonzero_of_invertible (n : α) (h ▸ Nat.cast_zero) + Nat.zero_lt_of_ne_zero fun h => Invertible.ne_zero (n : α) (h ▸ Nat.cast_zero) theorem invOf_add_invOf [Semiring α] (a b : α) [Invertible a] [Invertible b] : ⅟a + ⅟b = ⅟a * (a + b) * ⅟b := by diff --git a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean index e1d31428561b4..ae32ab4bf5651 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Combination.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Combination.lean @@ -783,7 +783,7 @@ theorem centroid_pair [DecidableEq ι] [Invertible (2 : k)] (p : ι → P) (i₁ · have hc : (card ({i₁, i₂} : Finset ι) : k) ≠ 0 := by rw [card_insert_of_not_mem (not_mem_singleton.2 h), card_singleton] norm_num - exact nonzero_of_invertible _ + exact Invertible.ne_zero _ rw [centroid_def, affineCombination_eq_weightedVSubOfPoint_vadd_of_sum_eq_one _ _ _ (sum_centroidWeights_eq_one_of_cast_card_ne_zero _ hc) (p i₁)] diff --git a/Mathlib/RingTheory/WittVector/WittPolynomial.lean b/Mathlib/RingTheory/WittVector/WittPolynomial.lean index ff8d6143ceae0..de01b15cb0a5c 100644 --- a/Mathlib/RingTheory/WittVector/WittPolynomial.lean +++ b/Mathlib/RingTheory/WittVector/WittPolynomial.lean @@ -229,7 +229,7 @@ theorem xInTermsOfW_vars_aux (n : ℕ) : n ∈ (xInTermsOfW p ℚ n).vars ∧ (xInTermsOfW p ℚ n).vars ⊆ range (n + 1) := by apply Nat.strongInductionOn n; clear n intro n ih - rw [xInTermsOfW_eq, mul_comm, vars_C_mul _ (nonzero_of_invertible _), + rw [xInTermsOfW_eq, mul_comm, vars_C_mul _ (Invertible.ne_zero _), vars_sub_of_disjoint, vars_X, range_succ, insert_eq] on_goal 1 => simp only [true_and_iff, true_or_iff, eq_self_iff_true, mem_union, mem_singleton] diff --git a/Mathlib/Tactic/NormNum/Result.lean b/Mathlib/Tactic/NormNum/Result.lean index 2edf11a135ca5..82313df13e8d7 100644 --- a/Mathlib/Tactic/NormNum/Result.lean +++ b/Mathlib/Tactic/NormNum/Result.lean @@ -238,7 +238,7 @@ theorem IsRat.of_raw (α) [DivisionRing α] (n : ℤ) (d : ℕ) ⟨this, by simp [div_eq_mul_inv]⟩ theorem IsRat.den_nz {α} [DivisionRing α] {a n d} : IsRat (a : α) n d → (d : α) ≠ 0 - | ⟨_, _⟩ => nonzero_of_invertible (d : α) + | ⟨_, _⟩ => Invertible.ne_zero (d : α) /-- The result of `norm_num` running on an expression `x` of type `α`. Untyped version of `Result`. -/ From 7e532b5232b22bddd9d2122a6f27862c1fb81046 Mon Sep 17 00:00:00 2001 From: grunweg Date: Fri, 16 Aug 2024 13:44:39 +0000 Subject: [PATCH 334/359] chore: add missing end (#15868) Fix mathlib for running the missingEnd linter everywhere. Prerequiste to #15845. --- Mathlib/Algebra/Group/Hom/CompTypeclasses.lean | 2 ++ Mathlib/Algebra/Order/Monoid/OrderDual.lean | 1 - Mathlib/Algebra/Order/Monoid/Unbundled/OrderDual.lean | 3 +-- Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean | 2 ++ Mathlib/Data/Array/Defs.lean | 2 ++ Mathlib/Data/Int/Defs.lean | 2 ++ Mathlib/Data/List/GetD.lean | 2 ++ Mathlib/Data/MLList/IO.lean | 2 ++ Mathlib/Data/MLList/Split.lean | 2 ++ Mathlib/Data/Nat/Find.lean | 2 ++ Mathlib/Data/Option/Defs.lean | 2 ++ Mathlib/Init/Data/Int/Order.lean | 2 ++ Mathlib/Lean/Elab/Term.lean | 2 ++ Mathlib/Lean/Exception.lean | 4 ++++ Mathlib/Lean/GoalsLocation.lean | 2 ++ Mathlib/Lean/Json.lean | 2 ++ Mathlib/Lean/Meta/KAbstractPositions.lean | 2 ++ Mathlib/Lean/PrettyPrinter/Delaborator.lean | 2 ++ Mathlib/Logic/Function/CompTypeclasses.lean | 2 ++ Mathlib/Tactic/Abel.lean | 6 +----- Mathlib/Tactic/ApplyAt.lean | 2 ++ Mathlib/Tactic/ApplyFun.lean | 2 ++ Mathlib/Tactic/ApplyWith.lean | 2 ++ Mathlib/Tactic/ArithMult.lean | 2 ++ Mathlib/Tactic/Basic.lean | 2 ++ Mathlib/Tactic/Bound/Attribute.lean | 2 ++ Mathlib/Tactic/Cases.lean | 2 ++ Mathlib/Tactic/CasesM.lean | 2 ++ Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean | 6 +----- Mathlib/Tactic/CategoryTheory/Coherence.lean | 4 +--- Mathlib/Tactic/Check.lean | 2 ++ Mathlib/Tactic/Choose.lean | 2 ++ Mathlib/Tactic/ClearExcept.lean | 2 ++ Mathlib/Tactic/ClearExclamation.lean | 2 ++ Mathlib/Tactic/Clear_.lean | 2 ++ Mathlib/Tactic/Coe.lean | 2 ++ Mathlib/Tactic/CongrM.lean | 2 ++ Mathlib/Tactic/Conv.lean | 2 ++ Mathlib/Tactic/Core.lean | 2 ++ Mathlib/Tactic/ExistsI.lean | 2 ++ Mathlib/Tactic/Explode.lean | 4 ++++ Mathlib/Tactic/Explode/Datatypes.lean | 4 ++++ Mathlib/Tactic/Explode/Pretty.lean | 4 ++++ Mathlib/Tactic/ExtractGoal.lean | 2 ++ Mathlib/Tactic/FailIfNoProgress.lean | 2 ++ Mathlib/Tactic/FieldSimp.lean | 6 +----- Mathlib/Tactic/Find.lean | 2 ++ Mathlib/Tactic/FunProp/Attr.lean | 4 ++++ Mathlib/Tactic/FunProp/Core.lean | 4 ++++ Mathlib/Tactic/FunProp/Decl.lean | 4 ++++ Mathlib/Tactic/FunProp/Elab.lean | 4 ++++ Mathlib/Tactic/FunProp/FunctionData.lean | 4 ++++ Mathlib/Tactic/FunProp/Mor.lean | 6 ++++++ Mathlib/Tactic/FunProp/RefinedDiscrTree.lean | 2 ++ Mathlib/Tactic/FunProp/StateList.lean | 2 ++ Mathlib/Tactic/FunProp/Theorems.lean | 4 ++++ Mathlib/Tactic/FunProp/ToBatteries.lean | 4 ++++ Mathlib/Tactic/FunProp/Types.lean | 4 ++++ Mathlib/Tactic/GCongr/Core.lean | 4 ++++ Mathlib/Tactic/GCongr/ForwardAttr.lean | 4 ++++ Mathlib/Tactic/GeneralizeProofs.lean | 2 ++ Mathlib/Tactic/Group.lean | 6 +----- Mathlib/Tactic/Have.lean | 2 ++ Mathlib/Tactic/HaveI.lean | 2 ++ Mathlib/Tactic/HelpCmd.lean | 2 ++ Mathlib/Tactic/ITauto.lean | 2 ++ Mathlib/Tactic/InferParam.lean | 2 ++ Mathlib/Tactic/Inhabit.lean | 2 ++ Mathlib/Tactic/IntervalCases.lean | 4 +--- Mathlib/Tactic/IrreducibleDef.lean | 2 ++ .../Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean | 4 ++++ Mathlib/Tactic/LinearCombination.lean | 6 +----- Mathlib/Tactic/Linter/HaveLetLinter.lean | 4 ++++ Mathlib/Tactic/MkIffOfInductiveProp.lean | 2 ++ Mathlib/Tactic/ModCases.lean | 6 +----- Mathlib/Tactic/Monotonicity/Attr.lean | 6 ++++++ Mathlib/Tactic/Monotonicity/Basic.lean | 4 ++++ Mathlib/Tactic/MoveAdd.lean | 4 ++++ Mathlib/Tactic/Nontriviality/Core.lean | 4 ++++ Mathlib/Tactic/NormNum/Core.lean | 2 ++ Mathlib/Tactic/NormNum/Result.lean | 2 ++ Mathlib/Tactic/NthRewrite.lean | 2 ++ Mathlib/Tactic/PPWithUniv.lean | 4 ++++ Mathlib/Tactic/Peel.lean | 6 +----- Mathlib/Tactic/Polyrith.lean | 6 +----- Mathlib/Tactic/Positivity/Core.lean | 4 +--- Mathlib/Tactic/Propose.lean | 2 ++ Mathlib/Tactic/PushNeg.lean | 2 ++ Mathlib/Tactic/Qify.lean | 4 +--- Mathlib/Tactic/Recall.lean | 2 ++ Mathlib/Tactic/Recover.lean | 2 ++ Mathlib/Tactic/Relation/Rfl.lean | 2 ++ Mathlib/Tactic/Relation/Symm.lean | 2 ++ Mathlib/Tactic/Relation/Trans.lean | 2 ++ Mathlib/Tactic/Rename.lean | 2 ++ Mathlib/Tactic/RenameBVar.lean | 2 ++ Mathlib/Tactic/Replace.lean | 2 ++ Mathlib/Tactic/RewriteSearch.lean | 4 +--- Mathlib/Tactic/Rify.lean | 6 +----- Mathlib/Tactic/Ring/Basic.lean | 4 +--- Mathlib/Tactic/Ring/PNat.lean | 4 +--- Mathlib/Tactic/Ring/RingNF.lean | 4 +--- Mathlib/Tactic/Sat/FromLRAT.lean | 4 ++++ Mathlib/Tactic/Says.lean | 4 ++++ Mathlib/Tactic/ScopedNS.lean | 2 ++ Mathlib/Tactic/Set.lean | 2 ++ Mathlib/Tactic/SimpIntro.lean | 2 ++ Mathlib/Tactic/SimpRw.lean | 2 ++ Mathlib/Tactic/SplitIfs.lean | 2 ++ Mathlib/Tactic/Substs.lean | 4 ++++ Mathlib/Tactic/SuccessIfFailWithMsg.lean | 2 ++ Mathlib/Tactic/SwapVar.lean | 2 ++ Mathlib/Tactic/TFAE.lean | 4 ++++ Mathlib/Tactic/Tauto.lean | 2 ++ Mathlib/Tactic/TermCongr.lean | 4 ++++ Mathlib/Tactic/TryThis.lean | 2 ++ Mathlib/Tactic/Use.lean | 2 ++ Mathlib/Tactic/Variable.lean | 6 ++++++ Mathlib/Tactic/WLOG.lean | 2 ++ Mathlib/Tactic/Widget/Calc.lean | 2 ++ Mathlib/Tactic/Widget/CommDiag.lean | 4 +--- Mathlib/Tactic/Widget/InteractiveUnfold.lean | 4 ++++ Mathlib/Tactic/Zify.lean | 4 ++++ Mathlib/Util/AddRelatedDecl.lean | 2 ++ Mathlib/Util/AssertExists.lean | 2 ++ Mathlib/Util/AtomM.lean | 2 ++ Mathlib/Util/CountHeartbeats.lean | 4 ++++ Mathlib/Util/IncludeStr.lean | 2 ++ Mathlib/Util/Notation3.lean | 4 ++++ Mathlib/Util/Superscript.lean | 2 ++ Mathlib/Util/Tactic.lean | 2 ++ Mathlib/Util/TermBeta.lean | 2 ++ Mathlib/Util/WhatsNew.lean | 2 ++ Mathlib/Util/WithWeakNamespace.lean | 2 ++ 134 files changed, 317 insertions(+), 75 deletions(-) diff --git a/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean b/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean index fa2cc235e3d24..5f9f1cea90e7a 100644 --- a/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean +++ b/Mathlib/Algebra/Group/Hom/CompTypeclasses.lean @@ -106,3 +106,5 @@ theorem comp_assoc {Q : Type*} [Monoid Q] exact ⟨by simp only [← κ.comp_eq, ← h, ← κ'.comp_eq, MonoidHom.comp_assoc]⟩ end MonoidHom.CompTriple + +end MonoidHomCompTriple diff --git a/Mathlib/Algebra/Order/Monoid/OrderDual.lean b/Mathlib/Algebra/Order/Monoid/OrderDual.lean index 065c6ee564a2d..6c9ac68a97316 100644 --- a/Mathlib/Algebra/Order/Monoid/OrderDual.lean +++ b/Mathlib/Algebra/Order/Monoid/OrderDual.lean @@ -9,7 +9,6 @@ import Mathlib.Algebra.Order.Monoid.Defs /-! # Ordered monoid structures on the order dual. -/ - universe u variable {α : Type u} diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/OrderDual.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/OrderDual.lean index c1dc58f5a86f8..7a9819c7230f1 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/OrderDual.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/OrderDual.lean @@ -8,7 +8,6 @@ import Mathlib.Algebra.Order.Monoid.Unbundled.Defs /-! # Unbundled ordered monoid structures on the order dual. -/ - universe u variable {α : Type u} @@ -61,4 +60,4 @@ instance covariantClass_swap_mul_lt [LT α] [Mul α] CovariantClass αᵒᵈ αᵒᵈ (swap (· * ·)) (· < ·) := ⟨c.1.flip⟩ - +end OrderDual diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean index 738f12558d80d..66f0f84ca3589 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean @@ -1004,3 +1004,5 @@ end LinearOrderedRing @[deprecated (since := "2023-12-23")] alias zero_le_mul_right := mul_nonneg_iff_of_pos_right @[deprecated (since := "2023-12-23")] alias zero_lt_mul_left := mul_pos_iff_of_pos_left @[deprecated (since := "2023-12-23")] alias zero_lt_mul_right := mul_pos_iff_of_pos_right + +end OrderedCommRing diff --git a/Mathlib/Data/Array/Defs.lean b/Mathlib/Data/Array/Defs.lean index d6497446d3a52..2e86c181ba220 100644 --- a/Mathlib/Data/Array/Defs.lean +++ b/Mathlib/Data/Array/Defs.lean @@ -32,3 +32,5 @@ where cyclicPermuteAux : Array α → List Nat → α → Nat → Array α /-- Permute the array using a list of cycles. -/ def permute! [Inhabited α] (a : Array α) (ls : List (List Nat)) : Array α := ls.foldl (init := a) (·.cyclicPermute! ·) + +end Array diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index 313c5e9be0e6e..662009abfdf46 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -604,3 +604,5 @@ attribute [simp] natCast_pow @[deprecated (since := "2024-04-02")] alias sign_coe_nat_of_nonzero := sign_natCast_of_ne_zero @[deprecated (since := "2024-04-02")] alias toNat_coe_nat := toNat_natCast @[deprecated (since := "2024-04-02")] alias toNat_coe_nat_add_one := toNat_natCast_add_one + +end Int diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index d006f872973fa..d11003beb0ec4 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -135,3 +135,5 @@ theorem getI_eq_iget_get? (n : ℕ) : l.getI n = (l.get? n).iget := by theorem getI_zero_eq_headI : l.getI 0 = l.headI := by cases l <;> rfl end getI + +end List diff --git a/Mathlib/Data/MLList/IO.lean b/Mathlib/Data/MLList/IO.lean index a5c8a8ed25243..8f0a5da58d9d2 100644 --- a/Mathlib/Data/MLList/IO.lean +++ b/Mathlib/Data/MLList/IO.lean @@ -50,3 +50,5 @@ where put stdin.putStr input stdin.flush return child + +end MLList diff --git a/Mathlib/Data/MLList/Split.lean b/Mathlib/Data/MLList/Split.lean index 84dc19b7dbc81..1aaf2fbbd9927 100644 --- a/Mathlib/Data/MLList/Split.lean +++ b/Mathlib/Data/MLList/Split.lean @@ -108,3 +108,5 @@ starting a new sublist each time a predicate changes from `false` to `true`. -/ def splitAtBecomesTrue (L : MLList m α) (p : α → Bool) : MLList m (List α) := L.splitAtBecomesTrueM fun a => pure (.up (p a)) + +end MLList diff --git a/Mathlib/Data/Nat/Find.lean b/Mathlib/Data/Nat/Find.lean index 2f06a77d4afe8..30b550bb0fc2d 100644 --- a/Mathlib/Data/Nat/Find.lean +++ b/Mathlib/Data/Nat/Find.lean @@ -223,3 +223,5 @@ theorem findGreatest_of_ne_zero (h : Nat.findGreatest P n = m) (h0 : m ≠ 0) : (findGreatest_eq_iff.1 h).2.1 h0 end FindGreatest + +end Nat diff --git a/Mathlib/Data/Option/Defs.lean b/Mathlib/Data/Option/Defs.lean index df47456532460..e140c43e996c9 100644 --- a/Mathlib/Data/Option/Defs.lean +++ b/Mathlib/Data/Option/Defs.lean @@ -99,3 +99,5 @@ instance liftOrGet_isId (f : α → α → α) : Std.LawfulIdentity (liftOrGet f def _root_.Lean.LOption.toOption {α} : Lean.LOption α → Option α | .some a => some a | _ => none + +end Option diff --git a/Mathlib/Init/Data/Int/Order.lean b/Mathlib/Init/Data/Int/Order.lean index bae8f42395a7a..62b66941747ef 100644 --- a/Mathlib/Init/Data/Int/Order.lean +++ b/Mathlib/Init/Data/Int/Order.lean @@ -52,3 +52,5 @@ instance instLinearOrder : LinearOrder ℤ where protected theorem eq_zero_or_eq_zero_of_mul_eq_zero {a b : ℤ} (h : a * b = 0) : a = 0 ∨ b = 0 := Int.mul_eq_zero.mp h + +end Int diff --git a/Mathlib/Lean/Elab/Term.lean b/Mathlib/Lean/Elab/Term.lean index 30aa1819157a2..d536cad0358b2 100644 --- a/Mathlib/Lean/Elab/Term.lean +++ b/Mathlib/Lean/Elab/Term.lean @@ -22,3 +22,5 @@ def elabPattern (patt : Term) (expectedType? : Option Expr) : TermElabM Expr := let t ← elabTerm patt expectedType? synthesizeSyntheticMVars (postpone := .no) (ignoreStuckTC := true) instantiateMVars t + +end Lean.Elab.Term diff --git a/Mathlib/Lean/Exception.lean b/Mathlib/Lean/Exception.lean index ab87792720158..c15fad88d8c23 100644 --- a/Mathlib/Lean/Exception.lean +++ b/Mathlib/Lean/Exception.lean @@ -38,3 +38,7 @@ and the only commonality is the prefix of the string, so that's what we look for -/ def isFailedToSynthesize (e : Exception) : IO Bool := do pure <| (← e.toMessageData.toString).startsWith "failed to synthesize" + +end Exception + +end Lean diff --git a/Mathlib/Lean/GoalsLocation.lean b/Mathlib/Lean/GoalsLocation.lean index d6f34b2a228fe..b9efb1a01e8fe 100644 --- a/Mathlib/Lean/GoalsLocation.lean +++ b/Mathlib/Lean/GoalsLocation.lean @@ -30,3 +30,5 @@ def location : GoalsLocation → MetaM (Option Name) | ⟨_, .hypType fvarId _⟩ => some <$> fvarId.getUserName | ⟨_, .hypValue fvarId _⟩ => some <$> fvarId.getUserName | ⟨_, .target _⟩ => return none + +end Lean.SubExpr.GoalsLocation diff --git a/Mathlib/Lean/Json.lean b/Mathlib/Lean/Json.lean index 020937e439b01..c5ab7ec132470 100644 --- a/Mathlib/Lean/Json.lean +++ b/Mathlib/Lean/Json.lean @@ -39,3 +39,5 @@ instance {α : Type u} [FromJson α] (p : α → Prop) [DecidablePred p] : FromJ instance {α : Type u} [ToJson α] (p : α → Prop) : ToJson (Subtype p) where toJson x := toJson x.val + +end Lean diff --git a/Mathlib/Lean/Meta/KAbstractPositions.lean b/Mathlib/Lean/Meta/KAbstractPositions.lean index 2e3f801d542e0..60d3ded29a824 100644 --- a/Mathlib/Lean/Meta/KAbstractPositions.lean +++ b/Mathlib/Lean/Meta/KAbstractPositions.lean @@ -87,3 +87,5 @@ example (h : [5] ≠ []) : List.getLast [5] h = 5 := by def kabstractIsTypeCorrect (e subExpr : Expr) (pos : SubExpr.Pos) : MetaM Bool := do withLocalDeclD `_a (← inferType subExpr) fun fvar => do isTypeCorrect (← replaceSubexpr (fun _ => pure fvar) pos e) + +end Lean.Meta diff --git a/Mathlib/Lean/PrettyPrinter/Delaborator.lean b/Mathlib/Lean/PrettyPrinter/Delaborator.lean index 6ee7e6719093f..bcc547afa2239 100644 --- a/Mathlib/Lean/PrettyPrinter/Delaborator.lean +++ b/Mathlib/Lean/PrettyPrinter/Delaborator.lean @@ -37,3 +37,5 @@ def OptionsPerPos.setBool (opts : OptionsPerPos) (p : SubExpr.Pos) (n : Name) (v OptionsPerPos := let e := opts.findD p {} |>.setBool n v opts.insert p e + +end Lean.PrettyPrinter.Delaborator diff --git a/Mathlib/Logic/Function/CompTypeclasses.lean b/Mathlib/Logic/Function/CompTypeclasses.lean index 202f3bf2502f2..7fcd8ca32286c 100644 --- a/Mathlib/Logic/Function/CompTypeclasses.lean +++ b/Mathlib/Logic/Function/CompTypeclasses.lean @@ -64,3 +64,5 @@ lemma comp_apply {M N P : Type*} rw [← h.comp_eq, Function.comp_apply] end CompTriple + +end CompTriple diff --git a/Mathlib/Tactic/Abel.lean b/Mathlib/Tactic/Abel.lean index 105a5e5f2c9bb..a0c5c9ec31aff 100644 --- a/Mathlib/Tactic/Abel.lean +++ b/Mathlib/Tactic/Abel.lean @@ -563,8 +563,4 @@ macro (name := abelConv) "abel" : conv => @[inherit_doc abelConv] macro "abel!" : conv => `(conv| first | discharge => abel1! | try_this abel_nf!) -end Abel - -end Tactic - -end Mathlib +end Mathlib.Tactic.Abel diff --git a/Mathlib/Tactic/ApplyAt.lean b/Mathlib/Tactic/ApplyAt.lean index 4f3e2750087e7..1bf1be33a22a4 100644 --- a/Mathlib/Tactic/ApplyAt.lean +++ b/Mathlib/Tactic/ApplyAt.lean @@ -38,3 +38,5 @@ elab "apply" t:term "at" i:ident : tactic => withSynthesize <| withMainContext d (← mkAppOptM' f (mvs.pop.push ldecl.toExpr |>.map fun e => some e)) let (_, mainGoal) ← mainGoal.intro1P replaceMainGoal <| [mainGoal] ++ mvs.pop.toList.map fun e => e.mvarId! + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/ApplyFun.lean b/Mathlib/Tactic/ApplyFun.lean index 09d1d12840f90..e9aa791b2eb59 100644 --- a/Mathlib/Tactic/ApplyFun.lean +++ b/Mathlib/Tactic/ApplyFun.lean @@ -214,3 +214,5 @@ elab_rules : tactic | `(tactic| apply_fun $f $[$loc]? $[using $P]?) => do (atTarget := withMainContext do replaceMainGoal <| ← applyFunTarget f P (← getMainGoal)) (failed := fun _ ↦ throwError "apply_fun failed") + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/ApplyWith.lean b/Mathlib/Tactic/ApplyWith.lean index abd1fb386c292..c4ddb7097accc 100644 --- a/Mathlib/Tactic/ApplyWith.lean +++ b/Mathlib/Tactic/ApplyWith.lean @@ -23,3 +23,5 @@ open Lean Meta Elab Tactic Term elab (name := applyWith) "apply" " (" &"config" " := " cfg:term ") " e:term : tactic => do let cfg ← unsafe evalTerm ApplyConfig (mkConst ``ApplyConfig) cfg evalApplyLikeTactic (·.apply · cfg) e + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/ArithMult.lean b/Mathlib/Tactic/ArithMult.lean index f8bc0c339d012..c2e5f0530ec2d 100644 --- a/Mathlib/Tactic/ArithMult.lean +++ b/Mathlib/Tactic/ArithMult.lean @@ -41,3 +41,5 @@ proof term. -/ macro (name := arith_mult?) "arith_mult?" c:Aesop.tactic_clause* : tactic => `(tactic| { show_term { arith_mult $c* } }) + +end ArithmeticFunction diff --git a/Mathlib/Tactic/Basic.lean b/Mathlib/Tactic/Basic.lean index 45d7eec835f00..3d427acacd7df 100644 --- a/Mathlib/Tactic/Basic.lean +++ b/Mathlib/Tactic/Basic.lean @@ -157,3 +157,5 @@ elab (name := clearValue) "clear_value" hs:(ppSpace colGt term:max)+ : tactic => replaceMainGoal [mvarId] attribute [pp_with_univ] ULift PUnit PEmpty + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Bound/Attribute.lean b/Mathlib/Tactic/Bound/Attribute.lean index f2633e11e5b0c..a33e962df9870 100644 --- a/Mathlib/Tactic/Bound/Attribute.lean +++ b/Mathlib/Tactic/Bound/Attribute.lean @@ -137,3 +137,5 @@ context. A typical example is exposing an inequality field of a structure, such `HasPowerSeriesOnBall.r_pos`. -/ macro "bound_forward" : attr => `(attr|aesop safe forward (rule_sets := [$(Lean.mkIdent `Bound):ident])) + +end Mathlib.Tactic.Bound diff --git a/Mathlib/Tactic/Cases.lean b/Mathlib/Tactic/Cases.lean index 3f949e67bc8ef..b803e05583a44 100644 --- a/Mathlib/Tactic/Cases.lean +++ b/Mathlib/Tactic/Cases.lean @@ -164,3 +164,5 @@ elab (name := cases') "cases' " tgts:(Parser.Tactic.casesTarget,+) usingArg:((" let subgoals ← ElimApp.evalNames elimInfo result.alts withArg (numEqs := targets.size) (toClear := targetsNew) (toTag := toTag) setGoals <| subgoals.toList ++ gs + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/CasesM.lean b/Mathlib/Tactic/CasesM.lean index c8fce2ba7a258..a1b012c5b4baf 100644 --- a/Mathlib/Tactic/CasesM.lean +++ b/Mathlib/Tactic/CasesM.lean @@ -164,3 +164,5 @@ constructorm* _ ∨ _, _ ∧ _, True elab (name := constructorM) "constructorm" recursive:"*"? ppSpace pats:term,+ : tactic => do let pats ← elabPatterns pats.getElems liftMetaTactic (constructorMatching · (matchPatterns pats) recursive.isSome) + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean b/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean index cdae67c853745..7f68ebf77c516 100644 --- a/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean +++ b/Mathlib/Tactic/CategoryTheory/BicategoryCoherence.lean @@ -159,8 +159,4 @@ theorem assoc_liftHom₂ {f g h i : a ⟶ b} [LiftHom f] [LiftHom g] [LiftHom h] (η : f ⟶ g) (θ : g ⟶ h) (ι : h ⟶ i) [LiftHom₂ η] [LiftHom₂ θ] : η ≫ θ ≫ ι = (η ≫ θ) ≫ ι := (Category.assoc _ _ _).symm -end BicategoryCoherence - -end Tactic - -end Mathlib +end Mathlib.Tactic.BicategoryCoherence diff --git a/Mathlib/Tactic/CategoryTheory/Coherence.lean b/Mathlib/Tactic/CategoryTheory/Coherence.lean index 53555e21afdd6..9aff001798481 100644 --- a/Mathlib/Tactic/CategoryTheory/Coherence.lean +++ b/Mathlib/Tactic/CategoryTheory/Coherence.lean @@ -293,6 +293,4 @@ elab_rules : tactic end Coherence -end Tactic - -end Mathlib +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Check.lean b/Mathlib/Tactic/Check.lean index 2579740e185f8..f2014e8970861 100644 --- a/Mathlib/Tactic/Check.lean +++ b/Mathlib/Tactic/Check.lean @@ -55,3 +55,5 @@ Like the `#check` command, the `#check` tactic allows stuck typeclass instance p These become metavariables in the output. -/ elab tk:"#check " colGt term:term : tactic => elabCheckTactic tk true term + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Choose.lean b/Mathlib/Tactic/Choose.lean index 8e5fd8cbf8861..910c9d6fe5919 100644 --- a/Mathlib/Tactic/Choose.lean +++ b/Mathlib/Tactic/Choose.lean @@ -220,3 +220,5 @@ elab_rules : tactic syntax "choose!" (ppSpace colGt binderIdent)+ (" using " term)? : tactic macro_rules | `(tactic| choose! $[$ids]* $[using $h]?) => `(tactic| choose ! $[$ids]* $[using $h]?) + +end Mathlib.Tactic.Choose diff --git a/Mathlib/Tactic/ClearExcept.lean b/Mathlib/Tactic/ClearExcept.lean index b2791172f5776..1e4296b36ed44 100644 --- a/Mathlib/Tactic/ClearExcept.lean +++ b/Mathlib/Tactic/ClearExcept.lean @@ -34,3 +34,5 @@ elab_rules : tactic if let none ← isClass? decl.type then toClear := toClear.push decl.fvarId goal.tryClearMany toClear + +end Lean.Elab.Tactic diff --git a/Mathlib/Tactic/ClearExclamation.lean b/Mathlib/Tactic/ClearExclamation.lean index 2f34db728bb70..5935d32048c32 100644 --- a/Mathlib/Tactic/ClearExclamation.lean +++ b/Mathlib/Tactic/ClearExclamation.lean @@ -17,3 +17,5 @@ elab (name := clear!) "clear!" hs:(ppSpace colGt ident)* : tactic => do let fvarIds ← getFVarIds hs liftMetaTactic1 fun goal ↦ do goal.tryClearMany <| (← collectForwardDeps (fvarIds.map .fvar) true).map (·.fvarId!) + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Clear_.lean b/Mathlib/Tactic/Clear_.lean index 5d0c8245c8c97..0e2d325b31be1 100644 --- a/Mathlib/Tactic/Clear_.lean +++ b/Mathlib/Tactic/Clear_.lean @@ -22,3 +22,5 @@ elab (name := clear_) "clear_" : tactic => if let none ← isClass? decl.type then toClear := toClear.push decl.fvarId goal.tryClearMany toClear + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Coe.lean b/Mathlib/Tactic/Coe.lean index 88a98f81116a5..ce4de64cb81eb 100644 --- a/Mathlib/Tactic/Coe.lean +++ b/Mathlib/Tactic/Coe.lean @@ -61,3 +61,5 @@ elab "(" "↥" ")" : term <= expectedType => ensureHasType b ty else throwError "cannot coerce to sort{indentExpr x}" + +end Lean.Elab.Term.CoeImpl diff --git a/Mathlib/Tactic/CongrM.lean b/Mathlib/Tactic/CongrM.lean index de9a449def47e..c41789aec94d7 100644 --- a/Mathlib/Tactic/CongrM.lean +++ b/Mathlib/Tactic/CongrM.lean @@ -76,3 +76,5 @@ elab_rules : tactic let gStx ← Term.exprToSyntax (← getMainTarget) -- Gives the expected type to `refine` as a workaround for its elaboration order. evalTactic <| ← `(tactic| refine (congr($(⟨pattern⟩)) : $gStx)) + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Conv.lean b/Mathlib/Tactic/Conv.lean index 44dfd2db367fd..c3936735e3032 100644 --- a/Mathlib/Tactic/Conv.lean +++ b/Mathlib/Tactic/Conv.lean @@ -134,3 +134,5 @@ syntax "#simp" (&" only")? (simpArgs)? " =>"? ppSpace term : command macro_rules | `(#simp%$tk $[only%$o]? $[[$args,*]]? $[=>]? $e) => `(#conv%$tk simp $[only%$o]? $[[$args,*]]? => $e) + +end Mathlib.Tactic.Conv diff --git a/Mathlib/Tactic/Core.lean b/Mathlib/Tactic/Core.lean index bc9892963dcee..1aeb2717c5047 100644 --- a/Mathlib/Tactic/Core.lean +++ b/Mathlib/Tactic/Core.lean @@ -258,3 +258,5 @@ def getPackageDir (pkg : String) : IO System.FilePath := do /-- Returns the mathlib root directory. -/ def getMathlibDir := getPackageDir "Mathlib" + +end Mathlib diff --git a/Mathlib/Tactic/ExistsI.lean b/Mathlib/Tactic/ExistsI.lean index 17f49821f39a9..636ee485e54b7 100644 --- a/Mathlib/Tactic/ExistsI.lean +++ b/Mathlib/Tactic/ExistsI.lean @@ -31,3 +31,5 @@ example : ∃ x : Nat, ∃ y : Nat, x = y := by -/ macro "existsi " es:term,+ : tactic => `(tactic| refine ⟨$es,*, ?_⟩) + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Explode.lean b/Mathlib/Tactic/Explode.lean index da1139fbe0012..ee85b52734aea 100644 --- a/Mathlib/Tactic/Explode.lean +++ b/Mathlib/Tactic/Explode.lean @@ -271,3 +271,7 @@ elab "#explode " stx:term : command => withoutModifyingEnv <| Command.runTermEla let entries ← explode e let fitchTable : MessageData ← entriesToMessageData entries logInfo <|← addMessageContext m!"{heading}\n\n{fitchTable}\n" + +end Explode + +end Mathlib diff --git a/Mathlib/Tactic/Explode/Datatypes.lean b/Mathlib/Tactic/Explode/Datatypes.lean index d15f255f8b316..2572429fcd5e5 100644 --- a/Mathlib/Tactic/Explode/Datatypes.lean +++ b/Mathlib/Tactic/Explode/Datatypes.lean @@ -85,3 +85,7 @@ def Entries.add (entries : Entries) (expr : Expr) (entry : Entry) : Entry × Ent This is used by `let` bindings where `expr` is an fvar. -/ def Entries.addSynonym (entries : Entries) (expr : Expr) (entry : Entry) : Entries := ⟨entries.s.insert expr entry, entries.l⟩ + +end Explode + +end Mathlib diff --git a/Mathlib/Tactic/Explode/Pretty.lean b/Mathlib/Tactic/Explode/Pretty.lean index cb2b463d13943..79d34b772d0ec 100644 --- a/Mathlib/Tactic/Explode/Pretty.lean +++ b/Mathlib/Tactic/Explode/Pretty.lean @@ -63,3 +63,7 @@ def entriesToMessageData (entries : Entries) : MetaM MessageData := do let paddedThms ← padRight <| entries.l.map (·.thm) rowToMessageData paddedLines paddedDeps paddedThms entries.l + +end Explode + +end Mathlib diff --git a/Mathlib/Tactic/ExtractGoal.lean b/Mathlib/Tactic/ExtractGoal.lean index f32afd77cc52d..46ebf8c95a989 100644 --- a/Mathlib/Tactic/ExtractGoal.lean +++ b/Mathlib/Tactic/ExtractGoal.lean @@ -168,3 +168,5 @@ elab_rules : tactic let cmd := if ← Meta.isProp ty then "theorem" else "def" pure m!"{cmd} {sig} := sorry" logInfo msg + +end Mathlib.Tactic.ExtractGoal diff --git a/Mathlib/Tactic/FailIfNoProgress.lean b/Mathlib/Tactic/FailIfNoProgress.lean index d8e4dbedbbdd1..00a6545c4d97c 100644 --- a/Mathlib/Tactic/FailIfNoProgress.lean +++ b/Mathlib/Tactic/FailIfNoProgress.lean @@ -81,3 +81,5 @@ elab_rules : tactic let goal ← getMainGoal let l ← runAndFailIfNoProgress goal (evalTactic tacs) replaceMainGoal l + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/FieldSimp.lean b/Mathlib/Tactic/FieldSimp.lean index 8db488b31477d..b02abb4f47983 100644 --- a/Mathlib/Tactic/FieldSimp.lean +++ b/Mathlib/Tactic/FieldSimp.lean @@ -194,8 +194,4 @@ elab_rules : tactic _ ← simpLocation r.ctx {} dis loc -end FieldSimp - -end Tactic - -end Mathlib +end Mathlib.Tactic.FieldSimp diff --git a/Mathlib/Tactic/Find.lean b/Mathlib/Tactic/Find.lean index b95098e141720..6d91737566f29 100644 --- a/Mathlib/Tactic/Find.lean +++ b/Mathlib/Tactic/Find.lean @@ -135,3 +135,5 @@ elab "#find " t:term : tactic => do let t ← Term.elabTerm t none Term.synthesizeSyntheticMVars (postpone := .no) (ignoreStuckTC := true) findType t + +end Mathlib.Tactic.Find diff --git a/Mathlib/Tactic/FunProp/Attr.lean b/Mathlib/Tactic/FunProp/Attr.lean index 314b0c2f80965..b77b4b9cf2b0d 100644 --- a/Mathlib/Tactic/FunProp/Attr.lean +++ b/Mathlib/Tactic/FunProp/Attr.lean @@ -38,3 +38,7 @@ initialize funPropAttr : Unit ← erase := fun _declName => throwError "can't remove `funProp` attribute (not implemented yet)" } + +end Meta.FunProp + +end Mathlib diff --git a/Mathlib/Tactic/FunProp/Core.lean b/Mathlib/Tactic/FunProp/Core.lean index 1459b86032271..be950365215c5 100644 --- a/Mathlib/Tactic/FunProp/Core.lean +++ b/Mathlib/Tactic/FunProp/Core.lean @@ -690,3 +690,7 @@ mutual return none end + +end Meta.FunProp + +end Mathlib diff --git a/Mathlib/Tactic/FunProp/Decl.lean b/Mathlib/Tactic/FunProp/Decl.lean index 3e3e3d147570f..c744ad6ea026b 100644 --- a/Mathlib/Tactic/FunProp/Decl.lean +++ b/Mathlib/Tactic/FunProp/Decl.lean @@ -152,3 +152,7 @@ def tacticToDischarge (tacticCode : TSyntax `tactic) : Expr → MetaM (Option Ex let (result?, _) ← runTac?.run {} {} return result? + +end Meta.FunProp + +end Mathlib diff --git a/Mathlib/Tactic/FunProp/Elab.lean b/Mathlib/Tactic/FunProp/Elab.lean index 37a1f37478417..afa3b31d7cf78 100644 --- a/Mathlib/Tactic/FunProp/Elab.lean +++ b/Mathlib/Tactic/FunProp/Elab.lean @@ -84,3 +84,7 @@ def funPropTac : Tactic throwError msg | _ => throwUnsupportedSyntax + +end Meta.FunProp + +end Mathlib diff --git a/Mathlib/Tactic/FunProp/FunctionData.lean b/Mathlib/Tactic/FunProp/FunctionData.lean index 843b3963b807e..158fcddfe23b1 100644 --- a/Mathlib/Tactic/FunProp/FunctionData.lean +++ b/Mathlib/Tactic/FunProp/FunctionData.lean @@ -292,3 +292,7 @@ def FunctionData.decompositionOverArgs (fData : FunctionData) (args : Array Nat) return (f,g) catch _ => return none + +end Meta.FunProp + +end Mathlib diff --git a/Mathlib/Tactic/FunProp/Mor.lean b/Mathlib/Tactic/FunProp/Mor.lean index afe269094f0de..b9dd93ed43ec6 100644 --- a/Mathlib/Tactic/FunProp/Mor.lean +++ b/Mathlib/Tactic/FunProp/Mor.lean @@ -150,3 +150,9 @@ def mkAppN (f : Expr) (xs : Array Arg) : Expr := match x with | ⟨x, .none⟩ => (f.app x) | ⟨x, .some coe⟩ => (coe.app f).app x) + +end Mor + +end Meta.FunProp + +end Mathlib diff --git a/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean b/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean index 64e4a3bbc7639..c8bab48d68682 100644 --- a/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean +++ b/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean @@ -1152,3 +1152,5 @@ def mapArraysM (d : RefinedDiscrTree α) (f : Array α → m (Array β)) : m (Re /-- Apply a function to the array of values at each node in a `RefinedDiscrTree`. -/ def mapArrays (d : RefinedDiscrTree α) (f : Array α → Array β) : RefinedDiscrTree β := d.mapArraysM (m := Id) f + +end Mathlib.Meta.FunProp.RefinedDiscrTree diff --git a/Mathlib/Tactic/FunProp/StateList.lean b/Mathlib/Tactic/FunProp/StateList.lean index 2174cf4748411..57c9ac6bb2129 100644 --- a/Mathlib/Tactic/FunProp/StateList.lean +++ b/Mathlib/Tactic/FunProp/StateList.lean @@ -169,3 +169,5 @@ instance StateListT.monadControl : MonadControl m (StateListT σ m) where stM := StateList σ liftWith := fun f => do let s ← get; liftM (f (fun x => x s)) restoreM := fun x _ => x + +end Mathlib.Meta.FunProp diff --git a/Mathlib/Tactic/FunProp/Theorems.lean b/Mathlib/Tactic/FunProp/Theorems.lean index 26571a80035b4..1ab9b3850fec5 100644 --- a/Mathlib/Tactic/FunProp/Theorems.lean +++ b/Mathlib/Tactic/FunProp/Theorems.lean @@ -389,3 +389,7 @@ function property: {thm.funPropName}" transition theorem: {thm.thmName} function property: {thm.funPropName}" transitionTheoremsExt.add thm attrKind + +end Meta.FunProp + +end Mathlib diff --git a/Mathlib/Tactic/FunProp/ToBatteries.lean b/Mathlib/Tactic/FunProp/ToBatteries.lean index a673dc50a4f3c..f553da7935be3 100644 --- a/Mathlib/Tactic/FunProp/ToBatteries.lean +++ b/Mathlib/Tactic/FunProp/ToBatteries.lean @@ -130,3 +130,7 @@ def etaExpand1 (f : Expr) : MetaM Expr := do else withDefault do forallBoundedTelescope (← inferType f) (.some 1) fun xs _ => do mkLambdaFVars xs (mkAppN f xs) + +end Meta.FunProp + +end Mathlib diff --git a/Mathlib/Tactic/FunProp/Types.lean b/Mathlib/Tactic/FunProp/Types.lean index 398048a9db06a..3d8330020fa34 100644 --- a/Mathlib/Tactic/FunProp/Types.lean +++ b/Mathlib/Tactic/FunProp/Types.lean @@ -162,3 +162,7 @@ def logError (msg : String) : FunPropM Unit := do s.msgLog else msg::s.msgLog} + +end Meta.FunProp + +end Mathlib diff --git a/Mathlib/Tactic/GCongr/Core.lean b/Mathlib/Tactic/GCongr/Core.lean index bcb5eeb0a05d9..67a7dc94d4836 100644 --- a/Mathlib/Tactic/GCongr/Core.lean +++ b/Mathlib/Tactic/GCongr/Core.lean @@ -527,3 +527,7 @@ elab_rules : tactic let g := Lean.MessageData.joinSep (unsolvedGoals.map Lean.MessageData.ofExpr) Format.line throwError "rel failed, cannot prove goal by 'substituting' the listed relationships. \ The steps which could not be automatically justified were:\n{g}" + +end GCongr + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/GCongr/ForwardAttr.lean b/Mathlib/Tactic/GCongr/ForwardAttr.lean index ad1561c067f03..b2223676bf279 100644 --- a/Mathlib/Tactic/GCongr/ForwardAttr.lean +++ b/Mathlib/Tactic/GCongr/ForwardAttr.lean @@ -52,3 +52,7 @@ initialize registerBuiltinAttribute { setEnv <| forwardExt.addEntry env (declName, ext) | _ => throwUnsupportedSyntax } + +end GCongr + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/GeneralizeProofs.lean b/Mathlib/Tactic/GeneralizeProofs.lean index ecd3ab4028c91..8bb825b5909f9 100644 --- a/Mathlib/Tactic/GeneralizeProofs.lean +++ b/Mathlib/Tactic/GeneralizeProofs.lean @@ -520,3 +520,5 @@ elab (name := generalizeProofsElab) "generalize_proofs" config?:(Parser.Tactic.c let g' ← mkFreshExprSyntheticOpaqueMVar (← g.getType) (← g.getTag) g.assign g' return g'.mvarId! + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Group.lean b/Mathlib/Tactic/Group.lean index 050f79785fc0a..b20838a434f79 100644 --- a/Mathlib/Tactic/Group.lean +++ b/Mathlib/Tactic/Group.lean @@ -86,8 +86,4 @@ macro_rules | `(tactic| group $[$loc]?) => `(tactic| repeat (fail_if_no_progress (aux_group₁ $[$loc]? <;> aux_group₂ $[$loc]?))) -end Group - -end Tactic - -end Mathlib +end Mathlib.Tactic.Group diff --git a/Mathlib/Tactic/Have.lean b/Mathlib/Tactic/Have.lean index 1c96ee457da4f..c8240ce05b77f 100644 --- a/Mathlib/Tactic/Have.lean +++ b/Mathlib/Tactic/Have.lean @@ -101,3 +101,5 @@ elab_rules : tactic | `(tactic| let $n:optBinderIdent $bs* $[: $t:term]?) => withMainContext do let (goal1, goal2) ← haveLetCore (← getMainGoal) n bs t true replaceMainGoal [goal1, goal2] + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/HaveI.lean b/Mathlib/Tactic/HaveI.lean index e08ca9e7301a4..f0dce3425b640 100644 --- a/Mathlib/Tactic/HaveI.lean +++ b/Mathlib/Tactic/HaveI.lean @@ -41,3 +41,5 @@ macro_rules -/ macro "letI' " hd:haveDecl : doElem => `(doElem| assert! letIDummy $hd:haveDecl) + +end Mathlib.Tactic.HaveI diff --git a/Mathlib/Tactic/HelpCmd.lean b/Mathlib/Tactic/HelpCmd.lean index a46cf4763e507..5587ad008455a 100644 --- a/Mathlib/Tactic/HelpCmd.lean +++ b/Mathlib/Tactic/HelpCmd.lean @@ -316,3 +316,5 @@ syntax withPosition("#help " colGt &"command" "+"? macro_rules | `(#help command%$tk $[+%$more]? $(id)?) => `(#help cat$[+%$more]? $(mkIdentFrom tk `command) $(id)?) + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/ITauto.lean b/Mathlib/Tactic/ITauto.lean index fdb027224a4f9..11fea885b2ddf 100644 --- a/Mathlib/Tactic/ITauto.lean +++ b/Mathlib/Tactic/ITauto.lean @@ -732,3 +732,5 @@ macro_rules -- category := DocCategory.tactic -- declNames := [`tactic.interactive.itauto] -- tags := ["logic", "propositional logic", "intuitionistic logic", "decision procedure"] } + +end Mathlib.Tactic.ITauto diff --git a/Mathlib/Tactic/InferParam.lean b/Mathlib/Tactic/InferParam.lean index 568fcb26965c2..61823d34ed51a 100644 --- a/Mathlib/Tactic/InferParam.lean +++ b/Mathlib/Tactic/InferParam.lean @@ -32,3 +32,5 @@ elab (name := inferOptParam) "infer_param" : tactic => do evalTactic tacticSyntax else throwError "`infer_param` only solves goals of the form `optParam _ _` or `autoParam _ _`, not {tgt}" + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Inhabit.lean b/Mathlib/Tactic/Inhabit.lean index daa5b484d6c1e..cda2beb10fd0b 100644 --- a/Mathlib/Tactic/Inhabit.lean +++ b/Mathlib/Tactic/Inhabit.lean @@ -54,3 +54,5 @@ elab_rules : tactic | `(tactic| inhabit $[$h_name:ident :]? $term) => do let goal ← evalInhabit (← getMainGoal) h_name term replaceMainGoal [goal] + +end Lean.Elab.Tactic diff --git a/Mathlib/Tactic/IntervalCases.lean b/Mathlib/Tactic/IntervalCases.lean index d683390c9cec7..621a718ffea2c 100644 --- a/Mathlib/Tactic/IntervalCases.lean +++ b/Mathlib/Tactic/IntervalCases.lean @@ -397,6 +397,4 @@ elab_rules : tactic cont x xs[1]? subst g e lbs ubs (mustUseBounds := false) | _, _, _ => throwUnsupportedSyntax -end Tactic - -end Mathlib +end Mathlib.Tactic diff --git a/Mathlib/Tactic/IrreducibleDef.lean b/Mathlib/Tactic/IrreducibleDef.lean index e233d30979e21..2d3987da6409c 100644 --- a/Mathlib/Tactic/IrreducibleDef.lean +++ b/Mathlib/Tactic/IrreducibleDef.lean @@ -113,3 +113,5 @@ elab mods:declModifiers "irreducible_def" n_id:declId n_def:(irredDefLemma)? attribute [$attrs:attrInstance,*] $n) if prot.isSome then modifyEnv (addProtected · ((← getCurrNamespace) ++ n.getId)) + +end Lean.Elab.Command diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean index 8198ec0a53234..3c23e2eb30aa9 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/PositiveVector.lean @@ -98,3 +98,7 @@ def findPositiveVector {n m : Nat} {matType : Nat → Nat → Type} [UsableInSim return extractSolution res.snd else throwError "Simplex Algorithm failed" + +end SimplexAlgorithm + +end Linarith diff --git a/Mathlib/Tactic/LinearCombination.lean b/Mathlib/Tactic/LinearCombination.lean index 7b2fdbd850460..4677400bdced1 100644 --- a/Mathlib/Tactic/LinearCombination.lean +++ b/Mathlib/Tactic/LinearCombination.lean @@ -246,8 +246,4 @@ syntax "linear_combination2" (normStx)? (ppSpace colGt term)? : tactic elab_rules : tactic | `(tactic| linear_combination2 $[(norm := $tac)]? $(e)?) => elabLinearCombination tac none e true -end LinearCombination - -end Tactic - -end Mathlib +end Mathlib.Tactic.LinearCombination diff --git a/Mathlib/Tactic/Linter/HaveLetLinter.lean b/Mathlib/Tactic/Linter/HaveLetLinter.lean index 9e93804663ceb..c8f07ad46c4cc 100644 --- a/Mathlib/Tactic/Linter/HaveLetLinter.lean +++ b/Mathlib/Tactic/Linter/HaveLetLinter.lean @@ -129,3 +129,7 @@ def haveLetLinter : Linter where run := withSetOptionIn fun _stx => do You can disable this linter using `set_option linter.haveLet 0`" initialize addLinter haveLetLinter + +end haveLet + +end Mathlib.Linter diff --git a/Mathlib/Tactic/MkIffOfInductiveProp.lean b/Mathlib/Tactic/MkIffOfInductiveProp.lean index baf6206f73b80..151383da0f1f0 100644 --- a/Mathlib/Tactic/MkIffOfInductiveProp.lean +++ b/Mathlib/Tactic/MkIffOfInductiveProp.lean @@ -411,3 +411,5 @@ initialize Lean.registerBuiltinAttribute { | _ => throwError "unrecognized syntax" mkIffOfInductivePropImpl decl tgt idStx } + +end Mathlib.Tactic.MkIff diff --git a/Mathlib/Tactic/ModCases.lean b/Mathlib/Tactic/ModCases.lean index 760aaf8275f0b..9846cca299457 100644 --- a/Mathlib/Tactic/ModCases.lean +++ b/Mathlib/Tactic/ModCases.lean @@ -191,8 +191,4 @@ elab_rules : tactic | ~q(ℕ) => NatMod.modCases h e n | _ => throwError "mod_cases only works with Int and Nat" -end ModCases - -end Tactic - -end Mathlib +end Mathlib.Tactic.ModCases diff --git a/Mathlib/Tactic/Monotonicity/Attr.lean b/Mathlib/Tactic/Monotonicity/Attr.lean index 101f06b95807c..ce61e3c78dcda 100644 --- a/Mathlib/Tactic/Monotonicity/Attr.lean +++ b/Mathlib/Tactic/Monotonicity/Attr.lean @@ -30,3 +30,9 @@ initialize ext : LabelExtension ← ( relations on its domain and range, and possibly with side conditions." let mono := `mono registerLabelAttr mono descr mono) + +end Attr + +end Monotonicity + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Monotonicity/Basic.lean b/Mathlib/Tactic/Monotonicity/Basic.lean index 0b0ba3b09f108..c7c981caabe3c 100644 --- a/Mathlib/Tactic/Monotonicity/Basic.lean +++ b/Mathlib/Tactic/Monotonicity/Basic.lean @@ -53,3 +53,7 @@ elab_rules : tactic transparency := .reducible exfalso := false } liftMetaTactic fun g => do processSyntax cfg false false [] [] #[mkIdent `mono] [g] + +end Monotonicity + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/MoveAdd.lean b/Mathlib/Tactic/MoveAdd.lean index 0e14071ab836d..1f76421725627 100644 --- a/Mathlib/Tactic/MoveAdd.lean +++ b/Mathlib/Tactic/MoveAdd.lean @@ -459,3 +459,7 @@ elab "move_mul" rws:rwRuleSeq : tactic => do evalTactic (← `(tactic| move_oper $hmul $rws)) end parsing + +end MoveAdd + +end Mathlib diff --git a/Mathlib/Tactic/Nontriviality/Core.lean b/Mathlib/Tactic/Nontriviality/Core.lean index b3f566a5629c5..2f6edcb097560 100644 --- a/Mathlib/Tactic/Nontriviality/Core.lean +++ b/Mathlib/Tactic/Nontriviality/Core.lean @@ -123,3 +123,7 @@ syntax (name := nontriviality) "nontriviality" (ppSpace colGt term)? g.assert `inst ty m let g ← liftM <| tac <|> nontrivialityByElim α g stx[2][1].getSepArgs replaceMainGoal [(← g.intro1).2] + +end Nontriviality + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/NormNum/Core.lean b/Mathlib/Tactic/NormNum/Core.lean index 505eaca4e984f..2c3834acb555f 100644 --- a/Mathlib/Tactic/NormNum/Core.lean +++ b/Mathlib/Tactic/NormNum/Core.lean @@ -357,3 +357,5 @@ Unlike `norm_num`, this command does not fail when no simplifications are made. macro (name := normNumCmd) "#norm_num" cfg:(config)? o:(&" only")? args:(Parser.Tactic.simpArgs)? " :"? ppSpace e:term : command => `(command| #conv norm_num $(cfg)? $[only%$o]? $(args)? => $e) + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/NormNum/Result.lean b/Mathlib/Tactic/NormNum/Result.lean index 82313df13e8d7..a8a58dcf2cb7a 100644 --- a/Mathlib/Tactic/NormNum/Result.lean +++ b/Mathlib/Tactic/NormNum/Result.lean @@ -468,3 +468,5 @@ def Result.eqTrans {α : Q(Type u)} {a b : Q($α)} (eq : Q($a = $b)) : Result b | .isRat inst q n d proof => Result.isRat inst q n d q($eq ▸ $proof) end Meta.NormNum + +end Mathlib diff --git a/Mathlib/Tactic/NthRewrite.lean b/Mathlib/Tactic/NthRewrite.lean index b0b2c62bc73f9..2b57f362e7b19 100644 --- a/Mathlib/Tactic/NthRewrite.lean +++ b/Mathlib/Tactic/NthRewrite.lean @@ -146,3 +146,5 @@ macro (name := nthRwSeq) "nth_rw" c:(config)? ppSpace n:num+ s:rwRuleSeq l:(loca `(tactic| (nth_rewrite $(c)? $[$n]* [$rs,*] $(l)?; with_annotate_state $rbrak (try (with_reducible rfl)))) | _ => Macro.throwUnsupported + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/PPWithUniv.lean b/Mathlib/Tactic/PPWithUniv.lean index c4afe64ed1fe6..ad26e61447be9 100644 --- a/Mathlib/Tactic/PPWithUniv.lean +++ b/Mathlib/Tactic/PPWithUniv.lean @@ -47,3 +47,7 @@ initialize registerBuiltinAttribute { let attr ← Elab.elabAttr <| ← `(Term.attrInstance| delab $(mkIdent <| `app ++ src)) liftTermElabM <| Term.applyAttributes ``delabWithUniv #[{attr with kind}] | _ => throwUnsupportedSyntax } + +end PPWithUniv + +end Mathlib diff --git a/Mathlib/Tactic/Peel.lean b/Mathlib/Tactic/Peel.lean index 05de63976ca4e..89c1c50bfa47a 100644 --- a/Mathlib/Tactic/Peel.lean +++ b/Mathlib/Tactic/Peel.lean @@ -244,8 +244,4 @@ macro_rules | `(tactic| peel $[$n:num]? $[$e:term]? $[with $h*]? using $u:term) => `(tactic| peel $[$n:num]? $[$e:term]? $[with $h*]?; exact $u) -end Peel - -end Tactic - -end Mathlib +end Mathlib.Tactic.Peel diff --git a/Mathlib/Tactic/Polyrith.lean b/Mathlib/Tactic/Polyrith.lean index c2a538579b186..f20d12e722d99 100644 --- a/Mathlib/Tactic/Polyrith.lean +++ b/Mathlib/Tactic/Polyrith.lean @@ -426,8 +426,4 @@ elab_rules : tactic if !traceMe then Lean.Meta.Tactic.TryThis.addSuggestion tk stx | .error g => replaceMainGoal [g] -end Polyrith - -end Tactic - -end Mathlib +end Mathlib.Tactic.Polyrith diff --git a/Mathlib/Tactic/Positivity/Core.lean b/Mathlib/Tactic/Positivity/Core.lean index b7778a848a069..f6f04104e28db 100644 --- a/Mathlib/Tactic/Positivity/Core.lean +++ b/Mathlib/Tactic/Positivity/Core.lean @@ -414,6 +414,4 @@ elab (name := positivity) "positivity" : tactic => do end Positivity -end Tactic - -end Mathlib +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Propose.lean b/Mathlib/Tactic/Propose.lean index 1df6661535b19..9ca99b07d8d03 100644 --- a/Mathlib/Tactic/Propose.lean +++ b/Mathlib/Tactic/Propose.lean @@ -143,3 +143,5 @@ macro_rules `(tactic| have?%$tk ! $[: $type]? using $terms,*) | `(tactic| have!?%$tk $[: $type]? using $terms,*) => `(tactic| have?%$tk ! $[: $type]? using $terms,*) + +end Mathlib.Tactic.Propose diff --git a/Mathlib/Tactic/PushNeg.lean b/Mathlib/Tactic/PushNeg.lean index 2fe1ef4ab42da..32ceeec86884d 100644 --- a/Mathlib/Tactic/PushNeg.lean +++ b/Mathlib/Tactic/PushNeg.lean @@ -235,3 +235,5 @@ elab "push_neg" loc:(location)? : tactic => pushNegLocalDecl pushNegTarget (fun _ ↦ logInfo "push_neg couldn't find a negation to push") + +end Mathlib.Tactic.PushNeg diff --git a/Mathlib/Tactic/Qify.lean b/Mathlib/Tactic/Qify.lean index 6a264ce96ee33..f5b370e93f538 100644 --- a/Mathlib/Tactic/Qify.lean +++ b/Mathlib/Tactic/Qify.lean @@ -75,6 +75,4 @@ alias int_cast_ne := intCast_ne end Qify -end Tactic - -end Mathlib +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Recall.lean b/Mathlib/Tactic/Recall.lean index 2b242f8b21fd4..29e44ac4dc6fb 100644 --- a/Mathlib/Tactic/Recall.lean +++ b/Mathlib/Tactic/Recall.lean @@ -77,3 +77,5 @@ elab_rules : command else unless binders.getNumArgs == 0 do throwError "expected type after ':'" + +end Mathlib.Tactic.Recall diff --git a/Mathlib/Tactic/Recover.lean b/Mathlib/Tactic/Recover.lean index a56f0ddf7eaa1..b330002cecfd5 100644 --- a/Mathlib/Tactic/Recover.lean +++ b/Mathlib/Tactic/Recover.lean @@ -67,3 +67,5 @@ elab "recover " tacs:tacticSeq : tactic => do let unassignedMVarDependencies ← getUnassignedGoalMVarDependencies mvarId unassigned := unassigned.insertMany unassignedMVarDependencies.toList setGoals <| ((← getGoals) ++ unassigned.toList).eraseDups + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Relation/Rfl.lean b/Mathlib/Tactic/Relation/Rfl.lean index e114a5b15bd31..b35a56bd7c5c9 100644 --- a/Mathlib/Tactic/Relation/Rfl.lean +++ b/Mathlib/Tactic/Relation/Rfl.lean @@ -41,3 +41,5 @@ def _root_.Lean.Expr.relSidesIfRefl? (e : Expr) : MetaM (Option (Name × Expr × | some n => return some (n, lhs, rhs) | none => return none return none + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Relation/Symm.lean b/Mathlib/Tactic/Relation/Symm.lean index 273ff04b3d667..6057927cfb8b5 100644 --- a/Mathlib/Tactic/Relation/Symm.lean +++ b/Mathlib/Tactic/Relation/Symm.lean @@ -32,3 +32,5 @@ def _root_.Lean.Expr.relSidesIfSymm? (e : Expr) : MetaM (Option (Name × Expr × | some n => return some (n, lhs, rhs) | none => return none return none + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Relation/Trans.lean b/Mathlib/Tactic/Relation/Trans.lean index 80e8b89318c54..290d23e9aeb70 100644 --- a/Mathlib/Tactic/Relation/Trans.lean +++ b/Mathlib/Tactic/Relation/Trans.lean @@ -218,3 +218,5 @@ set_option hygiene false in macro_rules | `(tactic| transitivity) => `(tactic| trans) | `(tactic| transitivity $e) => `(tactic| trans $e) + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Rename.lean b/Mathlib/Tactic/Rename.lean index 34a4fa944e7a8..207f568849012 100644 --- a/Mathlib/Tactic/Rename.lean +++ b/Mathlib/Tactic/Rename.lean @@ -38,3 +38,5 @@ elab_rules : tactic withMainContext do for fvar in ids, tgt in bs do Elab.Term.addTermInfo' tgt (mkFVar fvar) + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/RenameBVar.lean b/Mathlib/Tactic/RenameBVar.lean index 2f4ded3799ee4..fe1706f2eacfc 100644 --- a/Mathlib/Tactic/RenameBVar.lean +++ b/Mathlib/Tactic/RenameBVar.lean @@ -49,3 +49,5 @@ elab "rename_bvar " old:ident " → " new:ident loc?:(location)? : tactic => do (fun fvarId ↦ renameBVarHyp mvarId fvarId old.getId new.getId) (renameBVarTarget mvarId old.getId new.getId) fun _ ↦ throwError "unexpected location syntax" + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Replace.lean b/Mathlib/Tactic/Replace.lean index 7a4da71738e92..150bb27671e9f 100644 --- a/Mathlib/Tactic/Replace.lean +++ b/Mathlib/Tactic/Replace.lean @@ -60,3 +60,5 @@ elab_rules : tactic match hId? with | some hId => replaceMainGoal [goal1, (← observing? <| goal2.clear hId).getD goal2] | none => replaceMainGoal [goal1, goal2] + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/RewriteSearch.lean b/Mathlib/Tactic/RewriteSearch.lean index d232d90ba1581..befd4ab7e0ee1 100644 --- a/Mathlib/Tactic/RewriteSearch.lean +++ b/Mathlib/Tactic/RewriteSearch.lean @@ -328,6 +328,4 @@ elab_rules : tactic | end RewriteSearch -end Tactic - -end Mathlib +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Rify.lean b/Mathlib/Tactic/Rify.lean index 8fe355b461dd2..cf51d4c2f58e1 100644 --- a/Mathlib/Tactic/Rify.lean +++ b/Mathlib/Tactic/Rify.lean @@ -86,8 +86,4 @@ alias rat_cast_ne := ratCast_ne @[rify_simps] lemma ofNat_rat_real (a : ℕ) [a.AtLeastTwo] : ((no_index (OfNat.ofNat a : ℚ)) : ℝ) = (OfNat.ofNat a : ℝ) := rfl -end Rify - -end Tactic - -end Mathlib +end Mathlib.Tactic.Rify diff --git a/Mathlib/Tactic/Ring/Basic.lean b/Mathlib/Tactic/Ring/Basic.lean index 49b7a2931f4f8..bda62f7eef640 100644 --- a/Mathlib/Tactic/Ring/Basic.lean +++ b/Mathlib/Tactic/Ring/Basic.lean @@ -1219,6 +1219,4 @@ elab (name := ring1) "ring1" tk:"!"? : tactic => liftMetaMAtMain fun g ↦ do end Ring -end Tactic - -end Mathlib +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Ring/PNat.lean b/Mathlib/Tactic/Ring/PNat.lean index 7a8519746d309..d5e2883e3d672 100644 --- a/Mathlib/Tactic/Ring/PNat.lean +++ b/Mathlib/Tactic/Ring/PNat.lean @@ -42,6 +42,4 @@ instance {n n' k} [h1 : CSLiftVal (n : ℕ+) n'] : end Ring -end Tactic - -end Mathlib +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Ring/RingNF.lean b/Mathlib/Tactic/Ring/RingNF.lean index 1b88fd1b13a71..4ac168ff6ec41 100644 --- a/Mathlib/Tactic/Ring/RingNF.lean +++ b/Mathlib/Tactic/Ring/RingNF.lean @@ -275,6 +275,4 @@ macro (name := ringConv) "ring" : conv => end RingNF -end Tactic - -end Mathlib +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Sat/FromLRAT.lean b/Mathlib/Tactic/Sat/FromLRAT.lean index a97a50a16c9b3..3bbdaa9c67d0c 100644 --- a/Mathlib/Tactic/Sat/FromLRAT.lean +++ b/Mathlib/Tactic/Sat/FromLRAT.lean @@ -669,3 +669,7 @@ elab "from_lrat " cnf:term:max ppSpace lrat:term:max : term => do example : ∀ (a b : Prop), (¬a ∧ ¬b ∨ a ∧ ¬b) ∨ ¬a ∧ b ∨ a ∧ b := from_lrat "p cnf 2 4 1 2 0 -1 2 0 1 -2 0 -1 -2 0" "5 -2 0 4 3 0 5 d 3 4 0 6 1 0 5 1 0 6 d 1 0 7 0 5 2 6 0" + +end Sat + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Says.lean b/Mathlib/Tactic/Says.lean index 146b5d3aa230e..72b195ef15967 100644 --- a/Mathlib/Tactic/Says.lean +++ b/Mathlib/Tactic/Says.lean @@ -134,3 +134,7 @@ elab_rules : tactic evalTactic result initialize Batteries.Linter.UnreachableTactic.addIgnoreTacticKind `Mathlib.Tactic.Says.says + +end Says + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/ScopedNS.lean b/Mathlib/Tactic/ScopedNS.lean index c12423cbc1cd4..0057e00d2ffe4 100644 --- a/Mathlib/Tactic/ScopedNS.lean +++ b/Mathlib/Tactic/ScopedNS.lean @@ -53,3 +53,5 @@ macro_rules | `(scoped[$ns] attribute [$[$attr:attr],*] $ids*) => `(with_weak_namespace $(mkIdentFrom ns <| rootNamespace ++ ns.getId) attribute [$[scoped $attr:attr],*] $ids*) + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Set.lean b/Mathlib/Tactic/Set.lean index 7ab63a973e54d..742da43bb1561 100644 --- a/Mathlib/Tactic/Set.lean +++ b/Mathlib/Tactic/Set.lean @@ -75,3 +75,5 @@ elab_rules : tactic evalTactic (← `(tactic| have%$tk $h : ($(← Term.exprToSyntax vale) : $(← Term.exprToSyntax ty)) = $a := rfl)) | _, _ => pure () + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/SimpIntro.lean b/Mathlib/Tactic/SimpIntro.lean index 72af478dd156a..7012c998ee2ba 100644 --- a/Mathlib/Tactic/SimpIntro.lean +++ b/Mathlib/Tactic/SimpIntro.lean @@ -74,3 +74,5 @@ elab "simp_intro" cfg:(config)? disch:(discharger)? g.withContext do let g? ← simpIntroCore g ctx (simprocs := simprocs) discharge? more.isSome ids.toList replaceMainGoal <| if let some g := g? then [g] else [] + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/SimpRw.lean b/Mathlib/Tactic/SimpRw.lean index 5b2e05e4bcc3e..fc9e243d4cc87 100644 --- a/Mathlib/Tactic/SimpRw.lean +++ b/Mathlib/Tactic/SimpRw.lean @@ -78,3 +78,5 @@ elab s:"simp_rw " cfg:(config)? rws:rwRuleSeq g:(location)? : tactic => focus do `(tactic| simp%$e $[$cfg]? only [← $e:term] $g ?) else `(tactic| simp%$e $[$cfg]? only [$e:term] $g ?)) + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/SplitIfs.lean b/Mathlib/Tactic/SplitIfs.lean index 13d9944826b89..d648be56512c6 100644 --- a/Mathlib/Tactic/SplitIfs.lean +++ b/Mathlib/Tactic/SplitIfs.lean @@ -154,3 +154,5 @@ elab_rules : tactic splitIfsCore loc names [] for name in ← names.get do logWarningAt name m!"unused name: {name}" + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Substs.lean b/Mathlib/Tactic/Substs.lean index 1e9dab5a3e5aa..b5f3ea55b20d3 100644 --- a/Mathlib/Tactic/Substs.lean +++ b/Mathlib/Tactic/Substs.lean @@ -22,3 +22,7 @@ syntax (name := substs) "substs" (colGt ppSpace ident)* : tactic macro_rules | `(tactic| substs $xs:ident*) => `(tactic| ($[subst $xs]*)) + +end Substs + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/SuccessIfFailWithMsg.lean b/Mathlib/Tactic/SuccessIfFailWithMsg.lean index d74fee7623ba5..9d80c1c15e17a 100644 --- a/Mathlib/Tactic/SuccessIfFailWithMsg.lean +++ b/Mathlib/Tactic/SuccessIfFailWithMsg.lean @@ -52,3 +52,5 @@ elab_rules : tactic Term.withoutErrToSorry <| withoutRecover do let msg ← unsafe Term.evalTerm String (.const ``String []) msg successIfFailWithMessage msg (evalTacticSeq tacs) tacs + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/SwapVar.lean b/Mathlib/Tactic/SwapVar.lean index 7bc39caecd595..96c6c0987b180 100644 --- a/Mathlib/Tactic/SwapVar.lean +++ b/Mathlib/Tactic/SwapVar.lean @@ -46,3 +46,5 @@ elab "swap_var " swapRules:(colGt swapRule),+ : tactic => do return lctx.setUserName fvarId₁ n₂ |>.setUserName fvarId₂ n₁ let mdecl := { mdecl with lctx := lctx } modifyMCtx fun mctx ↦ { mctx with decls := mctx.decls.insert mvarId mdecl } + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/TFAE.lean b/Mathlib/Tactic/TFAE.lean index d5ef4b815aad8..a5508905dec20 100644 --- a/Mathlib/Tactic/TFAE.lean +++ b/Mathlib/Tactic/TFAE.lean @@ -236,3 +236,7 @@ elab_rules : tactic let q2 ← AtomM.addAtom ty.bindingBody! hyps := hyps.push (q1, q2, hyp) proveTFAE hyps (← get).atoms is tfaeListQ + +end TFAE + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Tauto.lean b/Mathlib/Tactic/Tauto.lean index 2ccb7671879e9..14285f3eb5428 100644 --- a/Mathlib/Tactic/Tauto.lean +++ b/Mathlib/Tactic/Tauto.lean @@ -220,3 +220,5 @@ syntax (name := tauto) "tauto" (config)? : tactic elab_rules : tactic | `(tactic| tauto $[$cfg:config]?) => do let _cfg ← elabConfig (mkOptionalNode cfg) tautology + +end Mathlib.Tactic.Tauto diff --git a/Mathlib/Tactic/TermCongr.lean b/Mathlib/Tactic/TermCongr.lean index adc9350e2ff19..39a1d386b6041 100644 --- a/Mathlib/Tactic/TermCongr.lean +++ b/Mathlib/Tactic/TermCongr.lean @@ -638,3 +638,7 @@ def elabTermCongr : Term.TermElab := fun stx expectedType? => do let ty ← mkEq res.lhs res.rhs mkExpectedTypeHint pf ty | _ => throwUnsupportedSyntax + +end TermCongr + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/TryThis.lean b/Mathlib/Tactic/TryThis.lean index 8a8940474d5ef..9e0c3ea6a3cc0 100644 --- a/Mathlib/Tactic/TryThis.lean +++ b/Mathlib/Tactic/TryThis.lean @@ -26,3 +26,5 @@ elab tk:"try_this" tac:tactic : tactic => do elab tk:"try_this" tac:conv : conv => do Elab.Tactic.evalTactic tac Meta.Tactic.TryThis.addSuggestion tk tac (origSpan? := ← getRef) + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Use.lean b/Mathlib/Tactic/Use.lean index 8c6bc4f992c9a..0f48c40b79e55 100644 --- a/Mathlib/Tactic/Use.lean +++ b/Mathlib/Tactic/Use.lean @@ -202,3 +202,5 @@ elab (name := useSyntax) @[inherit_doc useSyntax] elab "use!" discharger?:(Parser.Tactic.discharger)? ppSpace args:term,+ : tactic => do runUse true (← mkUseDischarger discharger?) args.getElems.toList + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Variable.lean b/Mathlib/Tactic/Variable.lean index 1b303098b1b52..ae6e4f39f4ddb 100644 --- a/Mathlib/Tactic/Variable.lean +++ b/Mathlib/Tactic/Variable.lean @@ -308,3 +308,9 @@ where def ignorevariable? : Lean.Linter.IgnoreFunction := fun _ stack _ => stack.matches [`null, none, `null, ``Mathlib.Command.Variable.variable?] || stack.matches [`null, none, `null, `null, ``Mathlib.Command.Variable.variable?] + +end Variable + +end Command + +end Mathlib diff --git a/Mathlib/Tactic/WLOG.lean b/Mathlib/Tactic/WLOG.lean index 54a4007761d67..c996e5552b6df 100644 --- a/Mathlib/Tactic/WLOG.lean +++ b/Mathlib/Tactic/WLOG.lean @@ -138,3 +138,5 @@ elab_rules : tactic let goal ← getMainGoal let { reductionGoal, hypothesisGoal .. } ← goal.wlog h P xs H replaceMainGoal [reductionGoal, hypothesisGoal] + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Widget/Calc.lean b/Mathlib/Tactic/Widget/Calc.lean index 34ea5f388d745..bbcba1e9996ae 100644 --- a/Mathlib/Tactic/Widget/Calc.lean +++ b/Mathlib/Tactic/Widget/Calc.lean @@ -142,3 +142,5 @@ elab_rules : tactic Widget.savePanelWidgetInfo CalcPanel.javascriptHash (pure json) proofTerm isFirst := false evalCalc (← `(tactic|calc%$calcstx $stx)) + +end Lean.Elab.Tactic diff --git a/Mathlib/Tactic/Widget/CommDiag.lean b/Mathlib/Tactic/Widget/CommDiag.lean index b4b2a8a73ae77..1e089b2114a61 100644 --- a/Mathlib/Tactic/Widget/CommDiag.lean +++ b/Mathlib/Tactic/Widget/CommDiag.lean @@ -137,6 +137,4 @@ def commutativeSquarePresenter : ExprPresenter where end Widget -end Tactic - -end Mathlib +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Widget/InteractiveUnfold.lean b/Mathlib/Tactic/Widget/InteractiveUnfold.lean index d6b02f009dd2a..06ed44cc330f3 100644 --- a/Mathlib/Tactic/Widget/InteractiveUnfold.lean +++ b/Mathlib/Tactic/Widget/InteractiveUnfold.lean @@ -242,3 +242,7 @@ def elabUnfoldCommand : Command.CommandElab := fun stx => let unfolds := unfolds.toList.map (m! "· {·}") logInfo (m! "Unfolds for {e}:\n" ++ .joinSep unfolds "\n") + +end InteractiveUnfold + +end Mathlib.Tactic diff --git a/Mathlib/Tactic/Zify.lean b/Mathlib/Tactic/Zify.lean index 8123df3c5de73..ff226759fce03 100644 --- a/Mathlib/Tactic/Zify.lean +++ b/Mathlib/Tactic/Zify.lean @@ -117,3 +117,7 @@ variable {R : Type*} [AddGroupWithOne R] @[norm_cast] theorem Nat.cast_sub_of_lt {m n} (h : m < n) : ((n - m : ℕ) : R) = n - m := Nat.cast_sub h.le + +end Zify + +end Mathlib.Tactic diff --git a/Mathlib/Util/AddRelatedDecl.lean b/Mathlib/Util/AddRelatedDecl.lean index 229fb6dd488b1..f415c55828e22 100644 --- a/Mathlib/Util/AddRelatedDecl.lean +++ b/Mathlib/Util/AddRelatedDecl.lean @@ -74,3 +74,5 @@ def addRelatedDecl (src : Name) (suffix : String) (ref : Syntax) let attrs ← elabAttrs attrs Term.applyAttributes src attrs Term.applyAttributes tgt attrs + +end Mathlib.Tactic diff --git a/Mathlib/Util/AssertExists.lean b/Mathlib/Util/AssertExists.lean index 40beb0cc91e32..910bc40c84fef 100644 --- a/Mathlib/Util/AssertExists.lean +++ b/Mathlib/Util/AssertExists.lean @@ -80,3 +80,5 @@ elab "assert_not_imported " ids:ident+ : command => do let mods := (← getEnv).allImportedModuleNames for id in ids do if mods.contains id.getId then logWarningAt id m!"the module '{id}' is (transitively) imported" + +end diff --git a/Mathlib/Util/AtomM.lean b/Mathlib/Util/AtomM.lean index 3eb7d0ffd1304..89f19cc7ea3ef 100644 --- a/Mathlib/Util/AtomM.lean +++ b/Mathlib/Util/AtomM.lean @@ -47,3 +47,5 @@ def AtomM.addAtom (e : Expr) : AtomM Nat := do if ← withTransparency (← read).red <| isDefEq e c.atoms[i] then return i modifyGet fun c ↦ (c.atoms.size, { c with atoms := c.atoms.push e }) + +end Mathlib.Tactic diff --git a/Mathlib/Util/CountHeartbeats.lean b/Mathlib/Util/CountHeartbeats.lean index f9890c218568b..795aac8b29aa1 100644 --- a/Mathlib/Util/CountHeartbeats.lean +++ b/Mathlib/Util/CountHeartbeats.lean @@ -147,3 +147,7 @@ elab "count_heartbeats! " n:(num)? "in" ppLine cmd:command : command => do -- Then run once more, keeping the state. let counts := (← elabForHeartbeats cmd (revert := false)) :: counts logVariation counts + +end CountHeartbeats + +end Mathlib diff --git a/Mathlib/Util/IncludeStr.lean b/Mathlib/Util/IncludeStr.lean index 3cc088bfc91bb..691268eaf5e17 100644 --- a/Mathlib/Util/IncludeStr.lean +++ b/Mathlib/Util/IncludeStr.lean @@ -19,3 +19,5 @@ elab (name := includeStr) "include_str " str:str : term => do let some srcDir := srcPath.parent | throwError "{srcPath} not in a valid directory" let path := srcDir / str Lean.mkStrLit <$> IO.FS.readFile path + +end Mathlib.Util diff --git a/Mathlib/Util/Notation3.lean b/Mathlib/Util/Notation3.lean index fd02609f63fab..e7f520311181c 100644 --- a/Mathlib/Util/Notation3.lean +++ b/Mathlib/Util/Notation3.lean @@ -620,3 +620,7 @@ macro_rules | `($[$doc]? $(attr)? scoped[$ns] notation3 $(prec)? $(n)? $(prio)? $(pp)? $items* => $t) => `(with_weak_namespace $(mkIdentFrom ns <| rootNamespace ++ ns.getId) $[$doc]? $(attr)? scoped notation3 $(prec)? $(n)? $(prio)? $(pp)? $items* => $t) + +end Notation3 + +end Mathlib diff --git a/Mathlib/Util/Superscript.lean b/Mathlib/Util/Superscript.lean index 1883098e6ed5e..0bf0bdb933aa2 100644 --- a/Mathlib/Util/Superscript.lean +++ b/Mathlib/Util/Superscript.lean @@ -281,3 +281,5 @@ initialize registerAlias `subscript ``subscript subscript registerAliasCore Formatter.formatterAliasesRef `subscript subscript.formatter registerAliasCore Parenthesizer.parenthesizerAliasesRef `subscript subscript.parenthesizer + +end Mathlib.Tactic diff --git a/Mathlib/Util/Tactic.lean b/Mathlib/Util/Tactic.lean index f5deed46e9463..41a28ecce696f 100644 --- a/Mathlib/Util/Tactic.lean +++ b/Mathlib/Util/Tactic.lean @@ -72,3 +72,5 @@ exist in the local context of `mvarId`, nothing happens. def modifyLocalDecl [MonadMCtx m] (mvarId : MVarId) (fvarId : FVarId) (f : LocalDecl → LocalDecl) : m Unit := modifyLocalContext mvarId fun lctx ↦ lctx.modifyLocalDecl fvarId f + +end Mathlib.Tactic diff --git a/Mathlib/Util/TermBeta.lean b/Mathlib/Util/TermBeta.lean index f0a8b2ce66a13..894ced6315fb4 100644 --- a/Mathlib/Util/TermBeta.lean +++ b/Mathlib/Util/TermBeta.lean @@ -37,3 +37,5 @@ def elabBeta : TermElab := fun stx expectedType? => let e ← elabTerm t expectedType? return (← instantiateMVars e).headBeta | _ => throwUnsupportedSyntax + +end Mathlib.Util.TermBeta diff --git a/Mathlib/Util/WhatsNew.lean b/Mathlib/Util/WhatsNew.lean index 43e4c7034f986..bf3515832cd40 100644 --- a/Mathlib/Util/WhatsNew.lean +++ b/Mathlib/Util/WhatsNew.lean @@ -117,3 +117,5 @@ elab "whatsnew " "in" ppLine cmd:command : command => do finally let newEnv ← getEnv logInfo (← liftCoreM <| whatsNew oldEnv newEnv) + +end Mathlib.WhatsNew diff --git a/Mathlib/Util/WithWeakNamespace.lean b/Mathlib/Util/WithWeakNamespace.lean index e03f6fa9d1a16..04fcad50d7093 100644 --- a/Mathlib/Util/WithWeakNamespace.lean +++ b/Mathlib/Util/WithWeakNamespace.lean @@ -38,3 +38,5 @@ def withWeakNamespace {α : Type} (ns : Name) (m : CommandElabM α) : CommandEla /-- Changes the current namespace without causing scoped things to go out of scope -/ elab "with_weak_namespace " ns:ident cmd:command : command => withWeakNamespace ns.getId (elabCommand cmd) + +end Lean.Elab.Command From d03203d6da0a936fb8aabe8618f38a35d983b5d1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Dupuis?= Date: Fri, 16 Aug 2024 14:45:32 +0000 Subject: [PATCH 335/359] feat: define `rpow` and `sqrt` based on the continuous functional calculus (#15377) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR defines `CFC.nnrpow`, `CFC.sqrt` and `CFC.rpow` based on the continuous functional calculus. We define two separate versions `CFC.nnrpow` (based on the non-unital calculus) and `CFC.rpow` (unital) due to what happens at 0. Since `NNReal.rpow 0 0 = 1`, this means that this function does not map zero to zero when the exponent is zero, and hence `CFC.nnrpow a 0 = 0` whereas `CFC.rpow a 0 = 1`. Co-authored-by: Frédéric Dupuis <31101893+dupuisf@users.noreply.github.com> --- Mathlib.lean | 1 + .../ContinuousFunctionalCalculus/Unital.lean | 17 +- .../ContinuousFunctionalCalculus/Rpow.lean | 377 ++++++++++++++++++ .../SpecialFunctions/Pow/Continuity.lean | 13 + .../Analysis/SpecialFunctions/Pow/NNReal.lean | 5 + 5 files changed, 412 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean diff --git a/Mathlib.lean b/Mathlib.lean index 15f573a3fa753..9edaa6f62d5a8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1264,6 +1264,7 @@ import Mathlib.Analysis.SpecialFunctions.Complex.Log import Mathlib.Analysis.SpecialFunctions.Complex.LogBounds import Mathlib.Analysis.SpecialFunctions.Complex.LogDeriv import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.ExpLog +import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow import Mathlib.Analysis.SpecialFunctions.Exp import Mathlib.Analysis.SpecialFunctions.ExpDeriv import Mathlib.Analysis.SpecialFunctions.Exponential diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index bb752066f3920..796456e6c6588 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -807,7 +807,7 @@ variable {R A : Type*} {p : A → Prop} [OrderedCommSemiring R] [StarRing R] variable [MetricSpace R] [TopologicalSemiring R] [ContinuousStar R] variable [∀ (α) [TopologicalSpace α], StarOrderedRing C(α, R)] variable [TopologicalSpace A] [Ring A] [StarRing A] [PartialOrder A] [StarOrderedRing A] -variable [Algebra R A] [ContinuousFunctionalCalculus R p] +variable [Algebra R A] [instCFC : ContinuousFunctionalCalculus R p] lemma cfcHom_mono {a : A} (ha : p a) {f g : C(spectrum R a, R)} (hfg : f ≤ g) : cfcHom ha f ≤ cfcHom ha g := @@ -884,6 +884,21 @@ lemma one_le_cfc (f : R → R) (a : A) (h : ∀ x ∈ spectrum R a, 1 ≤ f x) end Semiring +section NNReal + +open scoped NNReal + +variable {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] [PartialOrder A] + [Algebra ℝ≥0 A] [ContinuousFunctionalCalculus ℝ≥0 (fun (a : A) => 0 ≤ a)] + +lemma CFC.inv_nonneg_of_nonneg (a : Aˣ) (ha : (0 : A) ≤ a := by cfc_tac) : (0 : A) ≤ a⁻¹ := + cfc_inv_id (R := ℝ≥0) a ▸ cfc_predicate _ (a : A) + +lemma CFC.inv_nonneg (a : Aˣ) : (0 : A) ≤ a⁻¹ ↔ (0 : A) ≤ a := + ⟨fun _ ↦ inv_inv a ▸ inv_nonneg_of_nonneg a⁻¹, fun _ ↦ inv_nonneg_of_nonneg a⟩ + +end NNReal + section Ring variable {R A : Type*} {p : A → Prop} [OrderedCommRing R] [StarRing R] diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean new file mode 100644 index 0000000000000..fc706358df05d --- /dev/null +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/Rpow.lean @@ -0,0 +1,377 @@ +/- +Copyright (c) 2024 Frédéric Dupuis. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Frédéric Dupuis +-/ + +import Mathlib.Analysis.SpecialFunctions.Pow.Real +import Mathlib.Analysis.Normed.Algebra.Spectrum +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.NonUnital + +/-! +# Real powers defined via the continuous functional calculus + +This file defines real powers via the continuous functional calculus (CFC) and builds its API. +This allows one to take real powers of matrices, operators, elements of a C⋆-algebra, etc. The +square root is also defined via the non-unital CFC. + +## Main declarations + ++ `CFC.nnrpow`: the `ℝ≥0` power function based on the non-unital CFC, i.e. `cfcₙ NNReal.rpow` + composed with `(↑) : ℝ≥0 → ℝ`. ++ `CFC.sqrt`: the square root function based on the non-unital CFC, i.e. `cfcₙ NNReal.sqrt` ++ `CFC.rpow`: the real power function based on the unital CFC, i.e. `cfc NNReal.rpow` + +## Implementation notes + +We define two separate versions `CFC.nnrpow` and `CFC.rpow` due to what happens at 0. Since +`NNReal.rpow 0 0 = 1`, this means that this function does not map zero to zero when the exponent +is zero, and hence `CFC.nnrpow a 0 = 0` whereas `CFC.rpow a 0 = 1`. Note that the non-unital version +only makes sense for nonnegative exponents, and hence we define it such that the exponent is in +`ℝ≥0`. + +## Notation + ++ We define a `Pow A ℝ` instance for `CFC.rpow`, i.e `a ^ y` with `A` an operator and `y : ℝ` works + as expected. Likewise, we define a `Pow A ℝ≥0` instance for `CFC.nnrpow`. Note that these are + low-priority instances, in order to avoid overriding instances such as `Pow ℝ ℝ`. + +## TODO + ++ Relate these to the log and exp functions ++ Lemmas about how these functions interact with commuting `a` and `b`. ++ Prove the order properties (operator monotonicity and concavity/convexity) +-/ + +open scoped NNReal + +namespace NNReal + +/-- Taking a nonnegative power of a nonnegative number. This is defined as a standalone definition +in order to speed up automation such as `cfc_cont_tac`. -/ +noncomputable abbrev nnrpow (a : ℝ≥0) (b : ℝ≥0) : ℝ≥0 := a ^ (b : ℝ) + +@[simp] lemma nnrpow_def (a b : ℝ≥0) : nnrpow a b = a ^ (b : ℝ) := rfl + +@[fun_prop] +lemma continuous_nnrpow_const (y : ℝ≥0) : Continuous (nnrpow · y) := + continuous_rpow_const zero_le_coe + +/- This is a "redeclaration" of the attribute to speed up the proofs in this file. -/ +attribute [fun_prop] continuousOn_rpow_const + +end NNReal + +namespace CFC + +section NonUnital + +variable {A : Type*} [PartialOrder A] [NonUnitalNormedRing A] [StarRing A] + [Module ℝ≥0 A] [SMulCommClass ℝ≥0 A A] [IsScalarTower ℝ≥0 A A] + [NonUnitalContinuousFunctionalCalculus ℝ≥0 (fun (a : A) => 0 ≤ a)] + +/- ## `nnrpow` -/ + +/-- Real powers of operators, based on the non-unital continuous functional calculus. -/ +noncomputable def nnrpow (a : A) (y : ℝ≥0) : A := cfcₙ (NNReal.nnrpow · y) a + +/-- Enable `a ^ y` notation for `CFC.nnrpow`. This is a low-priority instance to make sure it does +not take priority over other instances when they are available. -/ +noncomputable instance (priority := 100) : Pow A ℝ≥0 where + pow a y := nnrpow a y + +@[simp] +lemma nnrpow_eq_pow {a : A} {y : ℝ≥0} : nnrpow a y = a ^ y := rfl + +@[simp] +lemma nnrpow_nonneg {a : A} {x : ℝ≥0} : 0 ≤ a ^ x := cfcₙ_predicate _ a + +lemma nnrpow_def {a : A} {y : ℝ≥0} : a ^ y = cfcₙ (NNReal.nnrpow · y) a := rfl + +lemma nnrpow_add {a : A} {x y : ℝ≥0} (hx : 0 < x) (hy : 0 < y) : + a ^ (x + y) = a ^ x * a ^ y := by + simp only [nnrpow_def] + rw [← cfcₙ_mul _ _ a] + congr! 2 with z + exact_mod_cast NNReal.rpow_add' z <| ne_of_gt (add_pos hx hy) + +@[simp] +lemma nnrpow_zero {a : A} : a ^ (0 : ℝ≥0) = 0 := by + simp [nnrpow_def, cfcₙ_apply_of_not_map_zero] + +lemma nnrpow_one (a : A) (ha : 0 ≤ a := by cfc_tac) : a ^ (1 : ℝ≥0) = a := by + simp only [nnrpow_def, NNReal.nnrpow_def, NNReal.coe_one, NNReal.rpow_one] + change cfcₙ (id : ℝ≥0 → ℝ≥0) a = a + rw [cfcₙ_id ℝ≥0 a] + +lemma nnrpow_two (a : A) (ha : 0 ≤ a := by cfc_tac) : a ^ (2 : ℝ≥0) = a * a := by + simp only [nnrpow_def, NNReal.nnrpow_def, NNReal.coe_ofNat, NNReal.rpow_ofNat, pow_two] + change cfcₙ (fun z : ℝ≥0 => id z * id z) a = a * a + rw [cfcₙ_mul id id a, cfcₙ_id ℝ≥0 a] + +lemma nnrpow_three (a : A) (ha : 0 ≤ a := by cfc_tac) : a ^ (3 : ℝ≥0) = a * a * a := by + simp only [nnrpow_def, NNReal.nnrpow_def, NNReal.coe_ofNat, NNReal.rpow_ofNat, pow_three] + change cfcₙ (fun z : ℝ≥0 => id z * (id z * id z)) a = a * a * a + rw [cfcₙ_mul id _ a, cfcₙ_mul id _ a, ← mul_assoc, cfcₙ_id ℝ≥0 a] + +@[simp] +lemma zero_nnrpow {x : ℝ≥0} : (0 : A) ^ x = 0 := by simp [nnrpow_def] + +@[simp] +lemma nnrpow_nnrpow [UniqueNonUnitalContinuousFunctionalCalculus ℝ≥0 A] + {a : A} {x y : ℝ≥0} : (a ^ x) ^ y = a ^ (x * y) := by + by_cases ha : 0 ≤ a + case pos => + obtain (rfl | hx) := eq_zero_or_pos x <;> obtain (rfl | hy) := eq_zero_or_pos y + all_goals try simp + simp only [nnrpow_def, NNReal.coe_mul] + rw [← cfcₙ_comp _ _ a] + congr! 2 with u + ext + simp [Real.rpow_mul] + case neg => + simp [nnrpow_def, cfcₙ_apply_of_not_predicate a ha] + +lemma nnrpow_nnrpow_inv [UniqueNonUnitalContinuousFunctionalCalculus ℝ≥0 A] + (a : A) {x : ℝ≥0} (hx : x ≠ 0) (ha : 0 ≤ a := by cfc_tac) : (a ^ x) ^ x⁻¹ = a := by + simp [mul_inv_cancel₀ hx, nnrpow_one _ ha] + +lemma nnrpow_inv_nnrpow [UniqueNonUnitalContinuousFunctionalCalculus ℝ≥0 A] + (a : A) {x : ℝ≥0} (hx : x ≠ 0) (ha : 0 ≤ a := by cfc_tac) : (a ^ x⁻¹) ^ x = a := by + simp [inv_mul_cancel₀ hx, nnrpow_one _ ha] + +lemma nnrpow_inv_eq [UniqueNonUnitalContinuousFunctionalCalculus ℝ≥0 A] + (a b : A) {x : ℝ≥0} (hx : x ≠ 0) (ha : 0 ≤ a := by cfc_tac) (hb : 0 ≤ b := by cfc_tac) : + a ^ x⁻¹ = b ↔ b ^ x = a := + ⟨fun h ↦ nnrpow_inv_nnrpow a hx ▸ congr($(h) ^ x).symm, + fun h ↦ nnrpow_nnrpow_inv b hx ▸ congr($(h) ^ x⁻¹).symm⟩ + +/- ## `sqrt` -/ + +section sqrt + +/-- Square roots of operators, based on the non-unital continuous functional calculus. -/ +noncomputable def sqrt (a : A) : A := cfcₙ NNReal.sqrt a + +@[simp] +lemma sqrt_nonneg {a : A} : 0 ≤ sqrt a := cfcₙ_predicate _ a + +lemma sqrt_eq_nnrpow {a : A} : sqrt a = a ^ (1 / 2 : ℝ≥0) := by + simp only [sqrt, nnrpow, NNReal.coe_inv, NNReal.coe_ofNat, NNReal.rpow_eq_pow] + congr + ext + exact_mod_cast NNReal.sqrt_eq_rpow _ + +@[simp] +lemma sqrt_zero : sqrt (0 : A) = 0 := by simp [sqrt] + +variable [UniqueNonUnitalContinuousFunctionalCalculus ℝ≥0 A] + +@[simp] +lemma nnrpow_sqrt {a : A} {x : ℝ≥0} : (sqrt a) ^ x = a ^ (x / 2) := by + rw [sqrt_eq_nnrpow, nnrpow_nnrpow, one_div_mul_eq_div 2 x] + +lemma nnrpow_sqrt_two (a : A) (ha : 0 ≤ a := by cfc_tac) : (sqrt a) ^ (2 : ℝ≥0) = a := by + simp only [nnrpow_sqrt, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, div_self] + rw [nnrpow_one a] + +lemma sqrt_mul_sqrt_self (a : A) (ha : 0 ≤ a := by cfc_tac) : sqrt a * sqrt a = a := by + rw [← nnrpow_two _, nnrpow_sqrt_two _] + +@[simp] +lemma sqrt_nnrpow {a : A} {x : ℝ≥0} : sqrt (a ^ x) = a ^ (x / 2) := by + simp [sqrt_eq_nnrpow, div_eq_mul_inv] + +lemma sqrt_nnrpow_two (a : A) (ha : 0 ≤ a := by cfc_tac) : sqrt (a ^ (2 : ℝ≥0)) = a := by + simp only [sqrt_nnrpow, ne_eq, OfNat.ofNat_ne_zero, not_false_eq_true, div_self] + rw [nnrpow_one _] + +lemma sqrt_mul_self (a : A) (ha : 0 ≤ a := by cfc_tac) : sqrt (a * a) = a := by + rw [← nnrpow_two _, sqrt_nnrpow_two _] + +lemma mul_self_eq {a b : A} (h : sqrt a = b) (ha : 0 ≤ a := by cfc_tac) : + b * b = a := + h ▸ sqrt_mul_sqrt_self _ ha + +lemma sqrt_unique {a b : A} (h : b * b = a) (hb : 0 ≤ b := by cfc_tac) : + sqrt a = b := + h ▸ sqrt_mul_self b + +lemma sqrt_eq_iff (a b : A) (ha : 0 ≤ a := by cfc_tac) (hb : 0 ≤ b := by cfc_tac) : + sqrt a = b ↔ b * b = a := + ⟨(mul_self_eq ·), (sqrt_unique ·)⟩ + +end sqrt + +end NonUnital + +section Unital + +variable {A : Type*} [PartialOrder A] [NormedRing A] [StarRing A] + [NormedAlgebra ℝ A] [ContinuousFunctionalCalculus ℝ≥0 (fun (a : A) => 0 ≤ a)] + +/- ## `rpow` -/ + +/-- Real powers of operators, based on the unital continuous functional calculus. -/ +noncomputable def rpow (a : A) (y : ℝ) : A := cfc (fun x : ℝ≥0 => x ^ y) a + +/-- Enable `a ^ y` notation for `CFC.rpow`. This is a low-priority instance to make sure it does +not take priority over other instances when they are available (such as `Pow ℝ ℝ`). -/ +noncomputable instance (priority := 100) : Pow A ℝ where + pow a y := rpow a y + +@[simp] +lemma rpow_eq_pow {a : A} {y : ℝ} : rpow a y = a ^ y := rfl + +@[simp] +lemma rpow_nonneg {a : A} {y : ℝ} : 0 ≤ a ^ y := cfc_predicate _ a + +lemma rpow_def {a : A} {y : ℝ} : a ^ y = cfc (fun x : ℝ≥0 => x ^ y) a := rfl + +lemma rpow_one (a : A) (ha : 0 ≤ a := by cfc_tac) : a ^ (1 : ℝ) = a := by + simp only [rpow_def, NNReal.coe_one, NNReal.rpow_eq_pow, NNReal.rpow_one, cfc_id' ℝ≥0 a] + +@[simp] +lemma one_rpow {x : ℝ} : (1 : A) ^ x = (1 : A) := by simp [rpow_def] + +lemma rpow_zero (a : A) (ha : 0 ≤ a := by cfc_tac) : a ^ (0 : ℝ) = 1 := by + simp [rpow_def, cfc_const_one ℝ≥0 a] + +lemma zero_rpow {x : ℝ} (hx : x ≠ 0) : rpow (0 : A) x = 0 := by simp [rpow, NNReal.zero_rpow hx] + +lemma rpow_natCast (a : A) (n : ℕ) (ha : 0 ≤ a := by cfc_tac) : a ^ (n : ℝ) = a ^ n := by + rw [← cfc_pow_id (R := ℝ≥0) a n, rpow_def] + congr + simp + +@[simp] +lemma rpow_algebraMap {x : ℝ≥0} {y : ℝ} : + (algebraMap ℝ≥0 A x) ^ y = algebraMap ℝ≥0 A (x ^ y) := by + rw [rpow_def, cfc_algebraMap ..] + +lemma rpow_add {a : A} {x y : ℝ} (ha : 0 ∉ spectrum ℝ≥0 a) : + a ^ (x + y) = a ^ x * a ^ y := by + simp only [rpow_def, NNReal.rpow_eq_pow] + rw [← cfc_mul _ _ a] + refine cfc_congr ?_ + intro z hz + have : z ≠ 0 := by aesop + simp [NNReal.rpow_add this _ _] + +-- TODO: relate to a strict positivity condition +lemma rpow_rpow [UniqueContinuousFunctionalCalculus ℝ≥0 A] + (a : A) (x y : ℝ) (ha₁ : 0 ∉ spectrum ℝ≥0 a) (hx : x ≠ 0) (ha₂ : 0 ≤ a := by cfc_tac) : + (a ^ x) ^ y = a ^ (x * y) := by + simp only [rpow_def] + rw [← cfc_comp _ _ a ha₂] + refine cfc_congr fun _ _ => ?_ + simp [NNReal.rpow_mul] + +lemma rpow_rpow_of_exponent_nonneg [UniqueContinuousFunctionalCalculus ℝ≥0 A] (a : A) (x y : ℝ) + (hx : 0 ≤ x) (hy : 0 ≤ y) (ha₂ : 0 ≤ a := by cfc_tac) : (a ^ x) ^ y = a ^ (x * y) := by + simp only [rpow_def] + rw [← cfc_comp _ _ a] + refine cfc_congr fun _ _ => ?_ + simp [NNReal.rpow_mul] + +lemma rpow_mul_rpow_neg {a : A} (x : ℝ) (ha : 0 ∉ spectrum ℝ≥0 a) + (ha' : 0 ≤ a := by cfc_tac) : a ^ x * a ^ (-x) = 1 := by + rw [← rpow_add ha, add_neg_cancel, rpow_zero a] + +lemma rpow_neg_mul_rpow {a : A} (x : ℝ) (ha : 0 ∉ spectrum ℝ≥0 a) + (ha' : 0 ≤ a := by cfc_tac) : a ^ (-x) * a ^ x = 1 := by + rw [← rpow_add ha, neg_add_cancel, rpow_zero a] + +lemma rpow_neg_one_eq_inv (a : Aˣ) (ha : (0 : A) ≤ a := by cfc_tac) : + a ^ (-1 : ℝ) = (↑a⁻¹ : A) := by + refine a.inv_eq_of_mul_eq_one_left ?_ |>.symm + simpa [rpow_one (a : A)] using rpow_neg_mul_rpow 1 (spectrum.zero_not_mem ℝ≥0 a.isUnit) + +lemma rpow_neg [UniqueContinuousFunctionalCalculus ℝ≥0 A] (a : Aˣ) (x : ℝ) + (ha' : (0 : A) ≤ a := by cfc_tac) : (a : A) ^ (-x) = (↑a⁻¹ : A) ^ x := by + suffices h₁ : ContinuousOn (fun z ↦ z ^ x) (Inv.inv '' (spectrum ℝ≥0 (a : A))) by + rw [← cfc_inv_id (R := ℝ≥0) a, rpow_def, rpow_def, + ← cfc_comp' (fun z => z ^ x) (Inv.inv : ℝ≥0 → ℝ≥0) (a : A) h₁] + refine cfc_congr fun _ _ => ?_ + simp [NNReal.rpow_neg, NNReal.inv_rpow] + refine NNReal.continuousOn_rpow_const (.inl ?_) + rintro ⟨z, hz, hz'⟩ + exact spectrum.zero_not_mem ℝ≥0 a.isUnit <| inv_eq_zero.mp hz' ▸ hz + +lemma rpow_intCast (a : Aˣ) (n : ℤ) (ha : (0 : A) ≤ a := by cfc_tac) : + (a : A) ^ (n : ℝ) = (↑(a ^ n) : A) := by + rw [← cfc_zpow (R := ℝ≥0) a n, rpow_def] + refine cfc_congr fun _ _ => ?_ + simp + +section unital_vs_nonunital + +variable [∀ (a : A), CompactSpace (spectrum ℝ a)] + [UniqueNonUnitalContinuousFunctionalCalculus ℝ≥0 A] + +lemma nnrpow_eq_rpow {a : A} {x : ℝ≥0} (hx : 0 < x) : a ^ x = a ^ (x : ℝ) := by + rw [nnrpow_def (A := A), rpow_def, cfcₙ_eq_cfc] + +lemma sqrt_eq_rpow {a : A} : sqrt a = a ^ (1 / 2 : ℝ) := by + have : a ^ (1 / 2 : ℝ) = a ^ ((1 / 2 : ℝ≥0) : ℝ) := rfl + rw [this, ← nnrpow_eq_rpow (by norm_num), sqrt_eq_nnrpow (A := A)] + +lemma sqrt_eq_cfc {a : A} : sqrt a = cfc NNReal.sqrt a := by + unfold sqrt + rw [cfcₙ_eq_cfc] + +lemma sqrt_sq (a : A) (ha : 0 ≤ a := by cfc_tac) : sqrt (a ^ 2) = a := by + rw [pow_two, sqrt_mul_self (A := A) a] + +lemma sq_sqrt (a : A) (ha : 0 ≤ a := by cfc_tac) : (sqrt a) ^ 2 = a := by + rw [pow_two, sqrt_mul_sqrt_self (A := A) a] + +@[simp] +lemma sqrt_algebraMap {r : ℝ≥0} : sqrt (algebraMap ℝ≥0 A r) = algebraMap ℝ≥0 A (NNReal.sqrt r) := by + rw [sqrt_eq_cfc, cfc_algebraMap] + +@[simp] +lemma sqrt_one : sqrt (1 : A) = 1 := by simp [sqrt_eq_cfc] + +-- TODO: relate to a strict positivity condition +lemma sqrt_rpow [UniqueContinuousFunctionalCalculus ℝ≥0 A] {a : A} {x : ℝ} (h : 0 ∉ spectrum ℝ≥0 a) + (hx : x ≠ 0) : sqrt (a ^ x) = a ^ (x / 2) := by + by_cases hnonneg : 0 ≤ a + case pos => + simp only [sqrt_eq_rpow, div_eq_mul_inv, one_mul, rpow_rpow _ _ _ h hx] + case neg => + simp [sqrt_eq_cfc, rpow_def, cfc_apply_of_not_predicate a hnonneg] + +-- TODO: relate to a strict positivity condition +lemma rpow_sqrt [UniqueContinuousFunctionalCalculus ℝ≥0 A] (a : A) (x : ℝ) (h : 0 ∉ spectrum ℝ≥0 a) + (ha : 0 ≤ a := by cfc_tac) : (sqrt a) ^ x = a ^ (x / 2) := by + rw [sqrt_eq_rpow, div_eq_mul_inv, one_mul, + rpow_rpow _ _ _ h (by norm_num), inv_mul_eq_div] + +lemma sqrt_rpow_nnreal {a : A} {x : ℝ≥0} : sqrt (a ^ (x : ℝ)) = a ^ (x / 2 : ℝ) := by + by_cases htriv : 0 ≤ a + case neg => simp [sqrt_eq_cfc, rpow_def, cfc_apply_of_not_predicate a htriv] + case pos => + by_cases hx : x = 0 + case pos => simp [hx, rpow_zero _ htriv] + case neg => + have h₁ : 0 < x := lt_of_le_of_ne (by aesop) (Ne.symm hx) + have h₂ : (x : ℝ) / 2 = NNReal.toReal (x / 2) := rfl + have h₃ : 0 < x / 2 := by positivity + rw [← nnrpow_eq_rpow h₁, h₂, ← nnrpow_eq_rpow h₃, sqrt_nnrpow (A := A)] + +lemma rpow_sqrt_nnreal [UniqueContinuousFunctionalCalculus ℝ≥0 A] {a : A} {x : ℝ≥0} + (ha : 0 ≤ a := by cfc_tac) : (sqrt a) ^ (x : ℝ) = a ^ (x / 2 : ℝ) := by + by_cases hx : x = 0 + case pos => + have ha' : 0 ≤ sqrt a := by exact sqrt_nonneg + simp [hx, rpow_zero _ ha', rpow_zero _ ha] + case neg => + have h₁ : 0 ≤ (x : ℝ) := by exact NNReal.zero_le_coe + rw [sqrt_eq_rpow, rpow_rpow_of_exponent_nonneg _ _ _ (by norm_num) h₁, one_div_mul_eq_div] + +end unital_vs_nonunital + +end Unital + +end CFC diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean index 03ef52c7d48f1..623700a7cbb97 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean @@ -417,6 +417,19 @@ theorem continuousAt_rpow_const {x : ℝ≥0} {y : ℝ} (h : x ≠ 0 ∨ 0 ≤ y theorem continuous_rpow_const {y : ℝ} (h : 0 ≤ y) : Continuous fun x : ℝ≥0 => x ^ y := continuous_iff_continuousAt.2 fun _ => continuousAt_rpow_const (Or.inr h) +@[fun_prop] +theorem continuousOn_rpow_const_compl_zero {r : ℝ} : + ContinuousOn (fun z : ℝ≥0 => z ^ r) {0}ᶜ := + fun _ h => ContinuousAt.continuousWithinAt <| NNReal.continuousAt_rpow_const (.inl h) + +-- even though this follows from `ContinuousOn.mono` and the previous lemma, we include it for +-- automation purposes with `fun_prop`, because the side goal `0 ∉ s ∨ 0 ≤ r` is often easy to check +@[fun_prop] +theorem continuousOn_rpow_const {r : ℝ} {s : Set ℝ≥0} + (h : 0 ∉ s ∨ 0 ≤ r) : ContinuousOn (fun z : ℝ≥0 => z ^ r) s := + h.elim (fun _ ↦ ContinuousOn.mono (s := {0}ᶜ) (by fun_prop) (by aesop)) + (NNReal.continuous_rpow_const · |>.continuousOn) + end NNReal /-! ## Continuity for `ℝ≥0∞` powers -/ diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index 957d08e87a3bb..eeb41c187e3b7 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -116,6 +116,11 @@ theorem rpow_natCast (x : ℝ≥0) (n : ℕ) : x ^ (n : ℝ) = x ^ n := @[deprecated (since := "2024-04-17")] alias rpow_nat_cast := rpow_natCast +@[simp, norm_cast] +lemma rpow_intCast (x : ℝ≥0) (n : ℤ) : x ^ (n : ℝ) = x ^ n := by + cases n <;> simp only [Int.ofNat_eq_coe, Int.cast_natCast, rpow_natCast, zpow_natCast, + Int.cast_negSucc, rpow_neg, zpow_negSucc] + @[simp] lemma rpow_ofNat (x : ℝ≥0) (n : ℕ) [n.AtLeastTwo] : x ^ (no_index (OfNat.ofNat n) : ℝ) = x ^ (OfNat.ofNat n : ℕ) := From 55e1bed5e7ef3b2350e89d9198ba0f13997ae8cf Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 16 Aug 2024 15:12:56 +0000 Subject: [PATCH 336/359] refactor(Topology/Maps): use `structure` for `QuotientMap` (#15887) --- Mathlib/Topology/Defs/Induced.lean | 8 +++++--- Mathlib/Topology/Maps/Basic.lean | 11 ++++------- 2 files changed, 9 insertions(+), 10 deletions(-) diff --git a/Mathlib/Topology/Defs/Induced.lean b/Mathlib/Topology/Defs/Induced.lean index 201785096951d..2b89c259917e8 100644 --- a/Mathlib/Topology/Defs/Induced.lean +++ b/Mathlib/Topology/Defs/Induced.lean @@ -124,6 +124,8 @@ structure ClosedEmbedding (f : X → Y) extends Embedding f : Prop where /-- A function between topological spaces is a quotient map if it is surjective, and for all `s : Set Y`, `s` is open iff its preimage is an open set. -/ -def QuotientMap {X : Type*} {Y : Type*} [tX : TopologicalSpace X] [tY : TopologicalSpace Y] - (f : X → Y) : Prop := - Function.Surjective f ∧ tY = tX.coinduced f +@[mk_iff quotientMap_iff'] +structure QuotientMap {X : Type*} {Y : Type*} [tX : TopologicalSpace X] [tY : TopologicalSpace Y] + (f : X → Y) : Prop where + surjective : Function.Surjective f + eq_coinduced : tY = tX.coinduced f diff --git a/Mathlib/Topology/Maps/Basic.lean b/Mathlib/Topology/Maps/Basic.lean index 38f3f36cbcec0..e50f6343087cb 100644 --- a/Mathlib/Topology/Maps/Basic.lean +++ b/Mathlib/Topology/Maps/Basic.lean @@ -233,7 +233,7 @@ section QuotientMap variable [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] theorem quotientMap_iff : QuotientMap f ↔ Surjective f ∧ ∀ s : Set Y, IsOpen s ↔ IsOpen (f ⁻¹' s) := - and_congr Iff.rfl TopologicalSpace.ext_iff + (quotientMap_iff' _).trans <| and_congr Iff.rfl TopologicalSpace.ext_iff theorem quotientMap_iff_closed : QuotientMap f ↔ Surjective f ∧ ∀ s : Set Y, IsClosed s ↔ IsClosed (f ⁻¹' s) := @@ -246,13 +246,13 @@ protected theorem id : QuotientMap (@id X) := ⟨fun x => ⟨x, rfl⟩, coinduced_id.symm⟩ protected theorem comp (hg : QuotientMap g) (hf : QuotientMap f) : QuotientMap (g ∘ f) := - ⟨hg.left.comp hf.left, by rw [hg.right, hf.right, coinduced_compose]⟩ + ⟨hg.surjective.comp hf.surjective, by rw [hg.eq_coinduced, hf.eq_coinduced, coinduced_compose]⟩ protected theorem of_quotientMap_compose (hf : Continuous f) (hg : Continuous g) (hgf : QuotientMap (g ∘ f)) : QuotientMap g := ⟨hgf.1.of_comp, le_antisymm - (by rw [hgf.right, ← coinduced_compose]; exact coinduced_mono hf.coinduced_le) + (by rw [hgf.eq_coinduced, ← coinduced_compose]; exact coinduced_mono hf.coinduced_le) hg.coinduced_le⟩ theorem of_inverse {g : Y → X} (hf : Continuous f) (hg : Continuous g) (h : LeftInverse g f) : @@ -260,14 +260,11 @@ theorem of_inverse {g : Y → X} (hf : Continuous f) (hg : Continuous g) (h : Le QuotientMap.of_quotientMap_compose hf hg <| h.comp_eq_id.symm ▸ QuotientMap.id protected theorem continuous_iff (hf : QuotientMap f) : Continuous g ↔ Continuous (g ∘ f) := by - rw [continuous_iff_coinduced_le, continuous_iff_coinduced_le, hf.right, coinduced_compose] + rw [continuous_iff_coinduced_le, continuous_iff_coinduced_le, hf.eq_coinduced, coinduced_compose] protected theorem continuous (hf : QuotientMap f) : Continuous f := hf.continuous_iff.mp continuous_id -protected theorem surjective (hf : QuotientMap f) : Surjective f := - hf.1 - protected theorem isOpen_preimage (hf : QuotientMap f) {s : Set Y} : IsOpen (f ⁻¹' s) ↔ IsOpen s := ((quotientMap_iff.1 hf).2 s).symm From 355dce788c48537e54f9413239abaf225efe4684 Mon Sep 17 00:00:00 2001 From: Rida Hamadani Date: Fri, 16 Aug 2024 15:43:25 +0000 Subject: [PATCH 337/359] feat: topology on `ENat` (#15380) this is ported from Peter Nelson's [Matroid](https://github.com/apnelson1/Matroid) project Co-authored-by: Peter Nelson Co-authored-by: Yury G. Kudryashov --- Mathlib.lean | 1 + Mathlib/Topology/Instances/ENat.lean | 64 ++++++++++++++++++++++++++++ 2 files changed, 65 insertions(+) create mode 100644 Mathlib/Topology/Instances/ENat.lean diff --git a/Mathlib.lean b/Mathlib.lean index 9edaa6f62d5a8..2e127b5d5f086 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4469,6 +4469,7 @@ import Mathlib.Topology.Instances.CantorSet import Mathlib.Topology.Instances.Complex import Mathlib.Topology.Instances.Discrete import Mathlib.Topology.Instances.ENNReal +import Mathlib.Topology.Instances.ENat import Mathlib.Topology.Instances.EReal import Mathlib.Topology.Instances.Int import Mathlib.Topology.Instances.Irrational diff --git a/Mathlib/Topology/Instances/ENat.lean b/Mathlib/Topology/Instances/ENat.lean new file mode 100644 index 0000000000000..004d68f482d65 --- /dev/null +++ b/Mathlib/Topology/Instances/ENat.lean @@ -0,0 +1,64 @@ +/- +Copyright (c) 2022 Peter Nelson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Peter Nelson +-/ +import Mathlib.Data.ENat.Basic +import Mathlib.Topology.Algebra.Monoid +import Mathlib.Topology.Instances.Discrete +import Mathlib.Order.Interval.Set.WithBotTop + +/-! +# Topology on extended natural numbers +-/ + +open Set + +open scoped Topology Filter + +namespace ENat + +/-- +Topology on `ℕ∞`. + +Note: this is different from the `EMetricSpace` topology. The `EMetricSpace` topology has +`IsOpen {∞}`, but all neighborhoods of `∞` in `ℕ∞` contain infinite intervals. +-/ +instance : TopologicalSpace ℕ∞ := Preorder.topology ℕ∞ + +instance : OrderTopology ℕ∞ := ⟨rfl⟩ + +@[simp] theorem range_natCast : range ((↑) : ℕ → ℕ∞) = Iio ⊤ := + WithTop.range_coe + +theorem embedding_natCast : Embedding ((↑) : ℕ → ℕ∞) := + Nat.strictMono_cast.embedding_of_ordConnected <| range_natCast ▸ ordConnected_Iio + +theorem openEmbedding_natCast : OpenEmbedding ((↑) : ℕ → ℕ∞) := + ⟨embedding_natCast, range_natCast ▸ isOpen_Iio⟩ + +theorem nhds_natCast (n : ℕ) : 𝓝 (n : ℕ∞) = pure (n : ℕ∞) := by + simp [← openEmbedding_natCast.map_nhds_eq] + +@[simp] +protected theorem nhds_eq_pure {n : ℕ∞} (h : n ≠ ⊤) : 𝓝 n = pure n := by + lift n to ℕ using h + simp [nhds_natCast] + +theorem isOpen_singleton {x : ℕ∞} (hx : x ≠ ⊤) : IsOpen {x} := by + rw [isOpen_singleton_iff_nhds_eq_pure, ENat.nhds_eq_pure hx] + +theorem mem_nhds_iff {x : ℕ∞} {s : Set ℕ∞} (hx : x ≠ ⊤) : s ∈ 𝓝 x ↔ x ∈ s := by + simp [hx] + +theorem mem_nhds_natCast_iff (n : ℕ) {s : Set ℕ∞} : s ∈ 𝓝 (n : ℕ∞) ↔ (n : ℕ∞) ∈ s := + mem_nhds_iff (coe_ne_top _) + +instance : ContinuousAdd ℕ∞ := by + refine ⟨continuous_iff_continuousAt.2 fun (a, b) ↦ ?_⟩ + match a, b with + | ⊤, _ => exact tendsto_nhds_top_mono' continuousAt_fst fun p ↦ le_add_right le_rfl + | (a : ℕ), ⊤ => exact tendsto_nhds_top_mono' continuousAt_snd fun p ↦ le_add_left le_rfl + | (a : ℕ), (b : ℕ) => simp [ContinuousAt, nhds_prod_eq, tendsto_pure_nhds] + +end ENat From f191832269f931b02885c356470056fcffa829eb Mon Sep 17 00:00:00 2001 From: james18lpc Date: Fri, 16 Aug 2024 15:43:26 +0000 Subject: [PATCH 338/359] chore: Rename `eventually_of_forall` to `Eventually.of_forall` (#15444) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit … and `frequently_of_forall` to `Frequently.of_forall` Co-authored-by: Yael Dillies Co-authored-by: Yury G. Kudryashov --- Counterexamples/Phillips.lean | 4 +-- Mathlib/Analysis/Analytic/Basic.lean | 4 +-- Mathlib/Analysis/Analytic/Constructions.lean | 2 +- .../Asymptotics/AsymptoticEquivalent.lean | 2 +- Mathlib/Analysis/Asymptotics/Asymptotics.lean | 6 ++-- .../Asymptotics/SuperpolynomialDecay.lean | 6 ++-- .../BoxIntegral/Box/SubboxInduction.lean | 4 +-- .../Analysis/BoxIntegral/Integrability.lean | 2 +- .../BoxIntegral/Partition/Filter.lean | 2 +- .../Calculus/BumpFunction/Convolution.lean | 8 ++--- .../Calculus/BumpFunction/Normed.lean | 2 +- .../Analysis/Calculus/ContDiff/RCLike.lean | 2 +- Mathlib/Analysis/Calculus/Deriv/Comp.lean | 2 +- Mathlib/Analysis/Calculus/Deriv/Inv.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Basic.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Comp.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Equiv.lean | 8 ++--- .../Analysis/Calculus/FDeriv/Symmetric.lean | 2 +- .../ApproximatesLinearOn.lean | 4 +-- Mathlib/Analysis/Calculus/LHopital.lean | 24 ++++++------- .../Analysis/Calculus/LocalExtr/Basic.lean | 2 +- Mathlib/Analysis/Calculus/Rademacher.lean | 2 +- Mathlib/Analysis/Calculus/TangentCone.lean | 2 +- Mathlib/Analysis/Complex/AbsMax.lean | 2 +- Mathlib/Analysis/Complex/Hadamard.lean | 2 +- .../Analysis/Complex/LocallyUniformLimit.lean | 4 +-- .../Analysis/Complex/PhragmenLindelof.lean | 2 +- Mathlib/Analysis/Convex/Integral.lean | 2 +- Mathlib/Analysis/Convolution.lean | 28 +++++++-------- .../Analysis/Distribution/SchwartzSpace.lean | 2 +- Mathlib/Analysis/Fourier/AddCircle.lean | 2 +- .../Fourier/FourierTransformDeriv.lean | 2 +- .../Fourier/RiemannLebesgueLemma.lean | 2 +- .../FunctionalSpaces/SobolevInequality.lean | 4 +-- .../Analysis/InnerProductSpace/l2Space.lean | 4 +-- Mathlib/Analysis/LocallyConvex/Bounded.lean | 6 ++-- .../Analysis/LocallyConvex/WithSeminorms.lean | 2 +- Mathlib/Analysis/MellinTransform.lean | 16 ++++----- Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 2 +- Mathlib/Analysis/Normed/Field/Basic.lean | 2 +- Mathlib/Analysis/Normed/Group/AddTorsor.lean | 2 +- Mathlib/Analysis/Normed/Group/Basic.lean | 4 +-- .../Analysis/Normed/Group/InfiniteSum.lean | 2 +- Mathlib/Analysis/Normed/Lp/lpSpace.lean | 4 +-- .../Normed/Module/FiniteDimension.lean | 2 +- .../Normed/Operator/BoundedLinearMaps.lean | 2 +- .../Analysis/NormedSpace/FunctionSeries.lean | 2 +- .../OperatorNorm/Completeness.lean | 2 +- Mathlib/Analysis/SpecialFunctions/Exp.lean | 4 +-- .../Analysis/SpecialFunctions/Gamma/Beta.lean | 10 +++--- .../SpecialFunctions/Gamma/BohrMollerup.lean | 2 +- .../Gaussian/FourierTransform.lean | 4 +-- .../Gaussian/GaussianIntegral.lean | 2 +- .../Gaussian/PoissonSummation.lean | 6 ++-- .../SpecialFunctions/ImproperIntegrals.lean | 4 +-- .../SpecialFunctions/JapaneseBracket.lean | 6 ++-- .../SpecialFunctions/Polynomials.lean | 2 +- .../SpecialFunctions/Pow/Asymptotics.lean | 2 +- .../SpecialFunctions/Pow/Continuity.lean | 4 +-- Mathlib/Analysis/SpecificLimits/Basic.lean | 6 ++-- Mathlib/Analysis/SpecificLimits/Normed.lean | 2 +- .../Computability/AkraBazzi/AkraBazzi.lean | 6 ++-- Mathlib/Dynamics/OmegaLimit.lean | 2 +- .../Constructions/BorelSpace/Metrizable.lean | 2 +- .../Constructions/Prod/Basic.lean | 4 +-- .../Constructions/Prod/Integral.lean | 14 ++++---- .../Covering/Differentiation.lean | 10 +++--- .../MeasureTheory/Covering/LiminfLimsup.lean | 18 +++++----- .../MeasureTheory/Decomposition/Lebesgue.lean | 4 +-- .../Decomposition/SignedHahn.lean | 2 +- .../Function/AEEqOfIntegral.lean | 6 ++-- .../ConditionalExpectation/AEMeasurable.lean | 2 +- .../ConditionalExpectation/Basic.lean | 6 ++-- .../ConditionalExpectation/CondexpL1.lean | 2 +- .../Function/ConditionalExpectation/Real.lean | 16 ++++----- .../Function/ConvergenceInMeasure.lean | 2 +- .../Function/Intersectivity.lean | 4 +-- Mathlib/MeasureTheory/Function/Jacobian.lean | 2 +- Mathlib/MeasureTheory/Function/L1Space.lean | 12 +++---- Mathlib/MeasureTheory/Function/L2Space.lean | 6 ++-- Mathlib/MeasureTheory/Function/LpOrder.lean | 4 +-- .../Function/LpSeminorm/Basic.lean | 24 ++++++------- .../Function/LpSeminorm/CompareExp.lean | 8 ++--- .../LpSeminorm/TriangleInequality.lean | 2 +- Mathlib/MeasureTheory/Function/LpSpace.lean | 10 +++--- .../ContinuousCompMeasurePreserving.lean | 2 +- .../Function/SimpleFuncDenseLp.lean | 6 ++-- .../Function/StronglyMeasurable/Basic.lean | 2 +- .../Function/UniformIntegrable.lean | 2 +- .../Group/FundamentalDomain.lean | 2 +- Mathlib/MeasureTheory/Group/Prod.lean | 6 ++-- Mathlib/MeasureTheory/Integral/Average.lean | 6 ++-- Mathlib/MeasureTheory/Integral/Bochner.lean | 32 ++++++++--------- .../Integral/BoundedContinuousFunction.lean | 2 +- .../Integral/CircleIntegral.lean | 4 +-- .../Integral/DivergenceTheorem.lean | 4 +-- .../Integral/DominatedConvergence.lean | 20 +++++------ .../MeasureTheory/Integral/IntegrableOn.lean | 8 ++--- .../Integral/IntegralEqImproper.lean | 4 +-- .../Integral/IntervalIntegral.lean | 8 ++--- Mathlib/MeasureTheory/Integral/Layercake.lean | 12 +++---- Mathlib/MeasureTheory/Integral/Lebesgue.lean | 12 +++---- .../MeasureTheory/Integral/PeakFunction.lean | 10 +++--- Mathlib/MeasureTheory/Integral/Periodic.lean | 2 +- .../MeasureTheory/Integral/SetIntegral.lean | 24 ++++++------- Mathlib/MeasureTheory/Integral/SetToL1.lean | 8 ++--- .../Integral/VitaliCaratheodory.lean | 10 +++--- .../MeasureTheory/Measure/AEMeasurable.lean | 2 +- .../Measure/ContinuousPreimage.lean | 2 +- Mathlib/MeasureTheory/Measure/DiracProba.lean | 2 +- .../MeasureTheory/Measure/EverywherePos.lean | 2 +- .../MeasureTheory/Measure/FiniteMeasure.lean | 12 +++---- .../Measure/Haar/Disintegration.lean | 2 +- .../MeasureTheory/Measure/Haar/Unique.lean | 4 +-- .../Measure/HasOuterApproxClosed.lean | 4 +-- Mathlib/MeasureTheory/Measure/Hausdorff.lean | 2 +- .../Measure/Lebesgue/EqHaar.lean | 4 +-- .../Measure/Lebesgue/Integral.lean | 2 +- .../Measure/LevyProkhorovMetric.lean | 22 ++++++------ .../MeasureTheory/Measure/MeasureSpace.lean | 2 +- .../MeasureTheory/Measure/Portmanteau.lean | 16 ++++----- Mathlib/MeasureTheory/Measure/Restrict.lean | 2 +- Mathlib/MeasureTheory/Measure/Stieltjes.lean | 2 +- Mathlib/MeasureTheory/Order/UpperLower.lean | 4 +-- Mathlib/MeasureTheory/OuterMeasure/AE.lean | 2 +- Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean | 2 +- .../EisensteinSeries/MDifferentiable.lean | 2 +- Mathlib/NumberTheory/WellApproximable.lean | 2 +- Mathlib/Order/Filter/AtTopBot.lean | 8 ++--- Mathlib/Order/Filter/Basic.lean | 22 +++++++----- Mathlib/Order/Filter/ENNReal.lean | 2 +- Mathlib/Order/Filter/FilterProduct.lean | 2 +- Mathlib/Order/Filter/Germ/Basic.lean | 22 ++++++------ Mathlib/Order/Filter/Germ/OrderedMonoid.lean | 2 +- Mathlib/Order/Filter/IndicatorFunction.lean | 2 +- Mathlib/Order/Filter/Interval.lean | 2 +- Mathlib/Order/Filter/ModEq.lean | 2 +- Mathlib/Order/Filter/Ring.lean | 2 +- Mathlib/Order/LiminfLimsup.lean | 34 +++++++++---------- Mathlib/Probability/BorelCantelli.lean | 2 +- .../Probability/ConditionalExpectation.lean | 2 +- .../Distributions/Exponential.lean | 2 +- Mathlib/Probability/Independence/Kernel.lean | 2 +- Mathlib/Probability/Kernel/Composition.lean | 4 +-- Mathlib/Probability/Kernel/CondDistrib.lean | 2 +- .../Kernel/Disintegration/CDFToKernel.lean | 8 ++--- .../Kernel/Disintegration/Density.lean | 8 ++--- .../Kernel/Disintegration/Integral.lean | 4 +-- .../Probability/Kernel/IntegralCompProd.lean | 8 ++--- .../Kernel/MeasurableIntegral.lean | 6 ++-- Mathlib/Probability/Martingale/Basic.lean | 6 ++-- .../Probability/Martingale/BorelCantelli.lean | 4 +-- .../Martingale/OptionalSampling.lean | 2 +- .../Probability/Martingale/Upcrossing.lean | 4 +-- Mathlib/Probability/StrongLaw.lean | 12 +++---- Mathlib/Probability/Variance.lean | 2 +- Mathlib/Tactic/Peel.lean | 8 ++--- .../Topology/Algebra/InfiniteSum/Group.lean | 2 +- .../Topology/Algebra/InfiniteSum/Order.lean | 2 +- Mathlib/Topology/Algebra/Module/Basic.lean | 2 +- Mathlib/Topology/Algebra/Monoid.lean | 2 +- Mathlib/Topology/Algebra/Order/Floor.lean | 4 +-- .../Topology/Algebra/Order/LiminfLimsup.lean | 12 +++---- Mathlib/Topology/Basic.lean | 2 +- Mathlib/Topology/Bornology/Absorbs.lean | 2 +- .../Topology/ContinuousFunction/Bounded.lean | 6 ++-- .../ContinuousFunction/Weierstrass.lean | 2 +- .../ContinuousFunction/ZeroAtInfty.lean | 2 +- Mathlib/Topology/ContinuousOn.lean | 4 +-- Mathlib/Topology/DenseEmbedding.lean | 2 +- Mathlib/Topology/Germ.lean | 2 +- Mathlib/Topology/Instances/ENNReal.lean | 6 ++-- Mathlib/Topology/Instances/EReal.lean | 2 +- Mathlib/Topology/LocallyConstant/Basic.lean | 2 +- Mathlib/Topology/Maps/Proper/Basic.lean | 2 +- .../Maps/Proper/UniversallyClosed.lean | 2 +- Mathlib/Topology/MetricSpace/Lipschitz.lean | 2 +- .../Topology/MetricSpace/Pseudo/Lemmas.lean | 2 +- Mathlib/Topology/MetricSpace/Sequences.lean | 2 +- Mathlib/Topology/Order/Basic.lean | 8 ++--- .../Topology/Order/MonotoneConvergence.lean | 2 +- Mathlib/Topology/Order/OrderClosed.lean | 6 ++-- Mathlib/Topology/RestrictGenTopology.lean | 2 +- Mathlib/Topology/Semicontinuous.lean | 22 ++++++------ Mathlib/Topology/Sequences.lean | 8 ++--- Mathlib/Topology/UniformSpace/Ascoli.lean | 2 +- .../Topology/UniformSpace/Equicontinuity.lean | 16 ++++----- 187 files changed, 512 insertions(+), 508 deletions(-) diff --git a/Counterexamples/Phillips.lean b/Counterexamples/Phillips.lean index d4fc4fc003f7e..0f73a13e843b5 100644 --- a/Counterexamples/Phillips.lean +++ b/Counterexamples/Phillips.lean @@ -132,7 +132,7 @@ def boundedIntegrableFunctionsIntegralCLM [MeasurableSpace α] (μ : Measure α) intro f rw [mul_comm] apply norm_integral_le_of_norm_le_const - apply Filter.eventually_of_forall + apply Filter.Eventually.of_forall intro x exact BoundedContinuousFunction.norm_coe_le_norm f.1 x) @@ -434,7 +434,7 @@ theorem toFunctions_toMeasure [MeasurableSpace α] (μ : Measure α) [IsFiniteMe have : Integrable (fun _ => (1 : ℝ)) μ := integrable_const (1 : ℝ) apply this.mono' (Measurable.indicator (@measurable_const _ _ _ _ (1 : ℝ)) hs).aestronglyMeasurable - apply Filter.eventually_of_forall + apply Filter.Eventually.of_forall exact norm_indicator_le_one _ theorem toFunctions_toMeasure_continuousPart [MeasurableSpace α] [MeasurableSingletonClass α] diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index 81b7d49ceba23..2321dfb012dfd 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -330,7 +330,7 @@ theorem radius_le_radius_continuousLinearMap_comp (p : FormalMultilinearSeries apply le_radius_of_isBigO apply (IsBigO.trans_isLittleO _ (p.isLittleO_one_of_lt_radius hr)).isBigO refine IsBigO.mul (@IsBigOWith.isBigO _ _ _ _ _ ‖f‖ _ _ _ ?_) (isBigO_refl _ _) - refine IsBigOWith.of_bound (eventually_of_forall fun n => ?_) + refine IsBigOWith.of_bound (Eventually.of_forall fun n => ?_) simpa only [norm_norm] using f.norm_compContinuousMultilinearMap_le (p n) end FormalMultilinearSeries @@ -829,7 +829,7 @@ theorem HasFPowerSeriesOnBall.tendstoLocallyUniformlyOn' (hf : HasFPowerSeriesOn protected theorem HasFPowerSeriesOnBall.continuousOn (hf : HasFPowerSeriesOnBall f p x r) : ContinuousOn f (EMetric.ball x r) := hf.tendstoLocallyUniformlyOn'.continuousOn <| - eventually_of_forall fun n => + Eventually.of_forall fun n => ((p.partialSum_continuous n).comp (continuous_id.sub continuous_const)).continuousOn protected theorem HasFPowerSeriesAt.continuousAt (hf : HasFPowerSeriesAt f p x) : diff --git a/Mathlib/Analysis/Analytic/Constructions.lean b/Mathlib/Analysis/Analytic/Constructions.lean index 7942e61582f38..9f74981f7a02e 100644 --- a/Mathlib/Analysis/Analytic/Constructions.lean +++ b/Mathlib/Analysis/Analytic/Constructions.lean @@ -221,7 +221,7 @@ lemma formalMultilinearSeries_geometric_radius (𝕜) [NontriviallyNormedField simp_rw [IsLittleO, IsBigOWith, not_forall, norm_one, mul_one, not_eventually] refine ⟨1, one_pos, ?_⟩ - refine ((eventually_ne_atTop 0).mp (eventually_of_forall ?_)).frequently + refine ((eventually_ne_atTop 0).mp (Eventually.of_forall ?_)).frequently intro n hn push_neg rwa [norm_pow, one_lt_pow_iff_of_nonneg (norm_nonneg _) hn, diff --git a/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean b/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean index 1689aea70b8ce..42aba8ff24d51 100644 --- a/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean +++ b/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean @@ -195,7 +195,7 @@ theorem isEquivalent_of_tendsto_one (hz : ∀ᶠ x in l, v x = 0 → u x = 0) theorem isEquivalent_of_tendsto_one' (hz : ∀ x, v x = 0 → u x = 0) (huv : Tendsto (u / v) l (𝓝 1)) : u ~[l] v := - isEquivalent_of_tendsto_one (eventually_of_forall hz) huv + isEquivalent_of_tendsto_one (Eventually.of_forall hz) huv theorem isEquivalent_iff_tendsto_one (hz : ∀ᶠ x in l, v x ≠ 0) : u ~[l] v ↔ Tendsto (u / v) l (𝓝 1) := by diff --git a/Mathlib/Analysis/Asymptotics/Asymptotics.lean b/Mathlib/Analysis/Asymptotics/Asymptotics.lean index 2e4f8556940f8..f97d8a1772783 100644 --- a/Mathlib/Analysis/Asymptotics/Asymptotics.lean +++ b/Mathlib/Analysis/Asymptotics/Asymptotics.lean @@ -1077,7 +1077,7 @@ theorem isLittleO_principal {s : Set α} : f'' =o[𝓟 s] g' ↔ ∀ x ∈ s, f' ((continuous_id.mul continuous_const).tendsto' _ _ (zero_mul _)).mono_left inf_le_left apply le_of_tendsto_of_tendsto tendsto_const_nhds this - apply eventually_nhdsWithin_iff.2 (eventually_of_forall (fun c hc ↦ ?_)) + apply eventually_nhdsWithin_iff.2 (Eventually.of_forall (fun c hc ↦ ?_)) exact eventually_principal.1 (h hc) x hx · apply (isLittleO_zero g' _).congr' ?_ EventuallyEq.rfl exact fun x hx ↦ (h x hx).symm @@ -1564,11 +1564,11 @@ theorem isLittleO_iff_tendsto' {f g : α → 𝕜} (hgf : ∀ᶠ x in l, g x = 0 f =o[l] g ↔ Tendsto (fun x => f x / g x) l (𝓝 0) := ⟨IsLittleO.tendsto_div_nhds_zero, fun h => (((isLittleO_one_iff _).mpr h).mul_isBigO (isBigO_refl g l)).congr' - (hgf.mono fun _x => div_mul_cancel_of_imp) (eventually_of_forall fun _x => one_mul _)⟩ + (hgf.mono fun _x => div_mul_cancel_of_imp) (Eventually.of_forall fun _x => one_mul _)⟩ theorem isLittleO_iff_tendsto {f g : α → 𝕜} (hgf : ∀ x, g x = 0 → f x = 0) : f =o[l] g ↔ Tendsto (fun x => f x / g x) l (𝓝 0) := - isLittleO_iff_tendsto' (eventually_of_forall hgf) + isLittleO_iff_tendsto' (Eventually.of_forall hgf) alias ⟨_, isLittleO_of_tendsto'⟩ := isLittleO_iff_tendsto' diff --git a/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean b/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean index f0990ce3adde0..6068e8fe71de4 100644 --- a/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean +++ b/Mathlib/Analysis/Asymptotics/SuperpolynomialDecay.lean @@ -159,7 +159,7 @@ theorem SuperpolynomialDecay.trans_eventually_abs_le (hf : SuperpolynomialDecay rw [superpolynomialDecay_iff_abs_tendsto_zero] at hf ⊢ refine fun z => tendsto_of_tendsto_of_tendsto_of_le_of_le' tendsto_const_nhds (hf z) - (eventually_of_forall fun x => abs_nonneg _) (hfg.mono fun x hx => ?_) + (Eventually.of_forall fun x => abs_nonneg _) (hfg.mono fun x hx => ?_) calc |k x ^ z * g x| = |k x ^ z| * |g x| := abs_mul (k x ^ z) (g x) _ ≤ |k x ^ z| * |f x| := by gcongr _ * ?_; exact hx @@ -167,7 +167,7 @@ theorem SuperpolynomialDecay.trans_eventually_abs_le (hf : SuperpolynomialDecay theorem SuperpolynomialDecay.trans_abs_le (hf : SuperpolynomialDecay l k f) (hfg : ∀ x, |g x| ≤ |f x|) : SuperpolynomialDecay l k g := - hf.trans_eventually_abs_le (eventually_of_forall hfg) + hf.trans_eventually_abs_le (Eventually.of_forall hfg) end LinearOrderedCommRing @@ -206,7 +206,7 @@ theorem superpolynomialDecay_iff_abs_isBoundedUnder (hk : Tendsto k l atTop) : zero_mul m ▸ Tendsto.mul_const m ((tendsto_zero_iff_abs_tendsto_zero _).1 hk.inv_tendsto_atTop) refine - tendsto_of_tendsto_of_tendsto_of_le_of_le' h1 h2 (eventually_of_forall fun x => abs_nonneg _) + tendsto_of_tendsto_of_tendsto_of_le_of_le' h1 h2 (Eventually.of_forall fun x => abs_nonneg _) ((eventually_map.1 hm).mp ?_) refine (hk.eventually_ne_atTop 0).mono fun x hk0 hx => ?_ refine Eq.trans_le ?_ (mul_le_mul_of_nonneg_left hx <| abs_nonneg (k x)⁻¹) diff --git a/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean b/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean index dd973734df63e..c516b55504b50 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/SubboxInduction.lean @@ -151,9 +151,9 @@ theorem subbox_induction_on' {p : Box ι → Prop} (I : Box ι) simpa [hJsub] using tendsto_const_nhds.div_atTop (tendsto_pow_atTop_atTop_of_one_lt _root_.one_lt_two) replace hJlz : Tendsto (fun m ↦ (J m).lower) atTop (𝓝[Icc I.lower I.upper] z) := - tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within _ hJlz (eventually_of_forall hJl_mem) + tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within _ hJlz (Eventually.of_forall hJl_mem) replace hJuz : Tendsto (fun m ↦ (J m).upper) atTop (𝓝[Icc I.lower I.upper] z) := - tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within _ hJuz (eventually_of_forall hJu_mem) + tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within _ hJuz (Eventually.of_forall hJu_mem) rcases H_nhds z (h0 ▸ hzJ 0) with ⟨U, hUz, hU⟩ rcases (tendsto_lift'.1 (hJlz.Icc hJuz) U hUz).exists with ⟨m, hUm⟩ exact hJp m (hU (J m) (hJle m) m (hzJ m) hUm (hJsub m)) diff --git a/Mathlib/Analysis/BoxIntegral/Integrability.lean b/Mathlib/Analysis/BoxIntegral/Integrability.lean index 8de8af41da9e9..12a2a1c87c00d 100644 --- a/Mathlib/Analysis/BoxIntegral/Integrability.lean +++ b/Mathlib/Analysis/BoxIntegral/Integrability.lean @@ -294,7 +294,7 @@ theorem IntegrableOn.hasBoxIntegral [CompleteSpace E] {f : (ι → ℝ) → E} { integral_finset_biUnion π.boxes (fun J _ => J.measurableSet_coe) π.pairwiseDisjoint (hfgi _)] refine dist_sum_sum_le_of_le _ fun J hJ => ?_ rw [dist_eq_norm, ← integral_sub (hfi _ J hJ) (hgi J hJ)] - refine norm_integral_le_of_norm_le (hfgi _ J hJ) (eventually_of_forall fun x => ?_) + refine norm_integral_le_of_norm_le (hfgi _ J hJ) (Eventually.of_forall fun x => ?_) exact hfg_mono x (hNx (π.tag J)) /-- If `f : ℝⁿ → E` is continuous on a rectangular box `I`, then it is Box integrable on `I` diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean index caa6545c2bab8..bd77ed4ea788e 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Filter.lean @@ -518,7 +518,7 @@ theorem eventually_isPartition (l : IntegrationParams) (I : Box ι) : ∀ᶠ π in l.toFilteriUnion I ⊤, TaggedPrepartition.IsPartition π := eventually_iSup.2 fun _ => eventually_inf_principal.2 <| - eventually_of_forall fun π h => + Eventually.of_forall fun π h => π.isPartition_iff_iUnion_eq.2 (h.trans Prepartition.iUnion_top) end IntegrationParams diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean b/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean index bc3b1582c4f18..6640b13a2e378 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean @@ -85,8 +85,8 @@ nonrec theorem convolution_tendsto_right {ι} {φ : ι → ContDiffBump (0 : G)} (hig : ∀ᶠ i in l, AEStronglyMeasurable (g i) μ) (hcg : Tendsto (uncurry g) (l ×ˢ 𝓝 x₀) (𝓝 z₀)) (hk : Tendsto k l (𝓝 x₀)) : Tendsto (fun i => ((φ i).normed μ ⋆[lsmul ℝ ℝ, μ] g i) (k i)) l (𝓝 z₀) := - convolution_tendsto_right (eventually_of_forall fun i => (φ i).nonneg_normed) - (eventually_of_forall fun i => (φ i).integral_normed) (tendsto_support_normed_smallSets hφ) hig + convolution_tendsto_right (Eventually.of_forall fun i => (φ i).nonneg_normed) + (Eventually.of_forall fun i => (φ i).integral_normed) (tendsto_support_normed_smallSets hφ) hig hcg hk /-- Special case of `ContDiffBump.convolution_tendsto_right` where `g` is continuous, @@ -94,7 +94,7 @@ nonrec theorem convolution_tendsto_right {ι} {φ : ι → ContDiffBump (0 : G)} theorem convolution_tendsto_right_of_continuous {ι} {φ : ι → ContDiffBump (0 : G)} {l : Filter ι} (hφ : Tendsto (fun i => (φ i).rOut) l (𝓝 0)) (hg : Continuous g) (x₀ : G) : Tendsto (fun i => ((φ i).normed μ ⋆[lsmul ℝ ℝ, μ] g) x₀) l (𝓝 (g x₀)) := - convolution_tendsto_right hφ (eventually_of_forall fun _ => hg.aestronglyMeasurable) + convolution_tendsto_right hφ (Eventually.of_forall fun _ => hg.aestronglyMeasurable) ((hg.tendsto x₀).comp tendsto_snd) tendsto_const_nhds /-- If a function `g` is locally integrable, then the convolution `φ i * g` converges almost @@ -112,7 +112,7 @@ theorem ae_convolution_tendsto_right_of_locallyIntegrable filter_upwards [(Besicovitch.vitaliFamily μ).ae_tendsto_average_norm_sub hg] with x₀ h₀ simp only [convolution_eq_swap, lsmul_apply] have hφ' : Tendsto (fun i ↦ (φ i).rOut) l (𝓝[>] 0) := - tendsto_nhdsWithin_iff.2 ⟨hφ, eventually_of_forall (fun i ↦ (φ i).rOut_pos)⟩ + tendsto_nhdsWithin_iff.2 ⟨hφ, Eventually.of_forall (fun i ↦ (φ i).rOut_pos)⟩ have := (h₀.comp (Besicovitch.tendsto_filterAt μ x₀)).comp hφ' simp only [Function.comp] at this apply tendsto_integral_smul_of_tendsto_average_norm_sub (K ^ (FiniteDimensional.finrank ℝ G)) this diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean index 81bf4037d473e..fb8c29c0c7405 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean @@ -99,7 +99,7 @@ theorem measure_closedBall_le_integral : (μ (closedBall c f.rIn)).toReal ≤ (μ (closedBall c f.rIn)).toReal = ∫ x in closedBall c f.rIn, 1 ∂μ := by simp _ = ∫ x in closedBall c f.rIn, f x ∂μ := setIntegral_congr measurableSet_closedBall (fun x hx ↦ (one_of_mem_closedBall f hx).symm) - _ ≤ ∫ x, f x ∂μ := setIntegral_le_integral f.integrable (eventually_of_forall (fun x ↦ f.nonneg)) + _ ≤ ∫ x, f x ∂μ := setIntegral_le_integral f.integrable (Eventually.of_forall (fun x ↦ f.nonneg)) theorem normed_le_div_measure_closedBall_rIn [μ.IsOpenPosMeasure] (x : E) : f.normed μ x ≤ 1 / (μ (closedBall c f.rIn)).toReal := by diff --git a/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean b/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean index 134a423eaeca7..d9f341edf671a 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean @@ -89,7 +89,7 @@ theorem HasFTaylorSeriesUpToOn.exists_lipschitzOnWith_of_nnnorm_lt {E F : Type*} replace hK : ‖f' x‖₊ < K := by simpa only [f', LinearIsometryEquiv.nnnorm_map] exact hs.exists_nhdsWithin_lipschitzOnWith_of_hasFDerivWithinAt_of_nnnorm_lt - (eventually_nhdsWithin_iff.2 <| eventually_of_forall hder) hcont K hK + (eventually_nhdsWithin_iff.2 <| Eventually.of_forall hder) hcont K hK /-- If `f` has a formal Taylor series `p` up to order `1` on `{x} ∪ s`, where `s` is a convex set, then `f` is Lipschitz in a neighborhood of `x` within `s`. -/ diff --git a/Mathlib/Analysis/Calculus/Deriv/Comp.lean b/Mathlib/Analysis/Calculus/Deriv/Comp.lean index 1dca57b89bb78..0bd8b9fd5f895 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Comp.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Comp.lean @@ -81,7 +81,7 @@ theorem HasDerivAtFilter.scomp_of_eq (hg : HasDerivAtFilter g₁ g₁' y L') theorem HasDerivWithinAt.scomp_hasDerivAt (hg : HasDerivWithinAt g₁ g₁' s' (h x)) (hh : HasDerivAt h h' x) (hs : ∀ x, h x ∈ s') : HasDerivAt (g₁ ∘ h) (h' • g₁') x := - hg.scomp x hh <| tendsto_inf.2 ⟨hh.continuousAt, tendsto_principal.2 <| eventually_of_forall hs⟩ + hg.scomp x hh <| tendsto_inf.2 ⟨hh.continuousAt, tendsto_principal.2 <| Eventually.of_forall hs⟩ theorem HasDerivWithinAt.scomp_hasDerivAt_of_eq (hg : HasDerivWithinAt g₁ g₁' s' y) (hh : HasDerivAt h h' x) (hs : ∀ x, h x ∈ s') (hy : y = h x) : diff --git a/Mathlib/Analysis/Calculus/Deriv/Inv.lean b/Mathlib/Analysis/Calculus/Deriv/Inv.lean index 0fda49102f336..0c5de84efda13 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Inv.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Inv.lean @@ -47,7 +47,7 @@ theorem hasStrictDerivAt_inv (hx : x ≠ 0) : HasStrictDerivAt Inv.inv (-(x ^ 2) suffices (fun p : 𝕜 × 𝕜 => (p.1 - p.2) * ((x * x)⁻¹ - (p.1 * p.2)⁻¹)) =o[𝓝 (x, x)] fun p => (p.1 - p.2) * 1 by - refine this.congr' ?_ (eventually_of_forall fun _ => mul_one _) + refine this.congr' ?_ (Eventually.of_forall fun _ => mul_one _) refine Eventually.mono ((isOpen_ne.prod isOpen_ne).mem_nhds ⟨hx, hx⟩) ?_ rintro ⟨y, z⟩ ⟨hy, hz⟩ simp only [mem_setOf_eq] at hy hz diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index 031c14bda50c7..7da598a12bd5b 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -732,7 +732,7 @@ theorem HasStrictFDerivAt.isBigO_sub_rev {f' : E ≃L[𝕜] F} theorem HasFDerivAtFilter.isBigO_sub_rev (hf : HasFDerivAtFilter f f' x L) {C} (hf' : AntilipschitzWith C f') : (fun x' => x' - x) =O[L] fun x' => f x' - f x := have : (fun x' => x' - x) =O[L] fun x' => f' (x' - x) := - isBigO_iff.2 ⟨C, eventually_of_forall fun _ => ZeroHomClass.bound_of_antilipschitz f' hf' _⟩ + isBigO_iff.2 ⟨C, Eventually.of_forall fun _ => ZeroHomClass.bound_of_antilipschitz f' hf' _⟩ (this.trans (hf.isLittleO.trans_isBigO this).right_isBigO_add).congr (fun _ => rfl) fun _ => sub_add_cancel _ _ diff --git a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean index 4aea3d3f12abc..beaf6fad6b9fe 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Comp.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Comp.lean @@ -67,7 +67,7 @@ example {g : F → G} {g' : F →L[𝕜] G} (hg : HasFDerivAtFilter g g' (f x) ( refine .of_isLittleO <| this.triangle ?_ calc (fun x' : E => g' (f x' - f x) - g'.comp f' (x' - x)) - _ =ᶠ[L] fun x' => g' (f x' - f x - f' (x' - x)) := eventually_of_forall fun x' => by simp + _ =ᶠ[L] fun x' => g' (f x' - f x - f' (x' - x)) := Eventually.of_forall fun x' => by simp _ =O[L] fun x' => f x' - f x - f' (x' - x) := g'.isBigO_comp _ _ _ =o[L] fun x' => x' - x := hf.isLittleO diff --git a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean index d05eb97118ca7..e2f34a6adc71a 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Equiv.lean @@ -349,10 +349,10 @@ theorem HasStrictFDerivAt.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] refine this.trans_isLittleO ?_ clear this refine ((hf.comp_tendsto hg).symm.congr' - (hfg.mono ?_) (eventually_of_forall fun _ => rfl)).trans_isBigO ?_ + (hfg.mono ?_) (Eventually.of_forall fun _ => rfl)).trans_isBigO ?_ · rintro p ⟨hp1, hp2⟩ simp [hp1, hp2] - · refine (hf.isBigO_sub_rev.comp_tendsto hg).congr' (eventually_of_forall fun _ => rfl) + · refine (hf.isBigO_sub_rev.comp_tendsto hg).congr' (Eventually.of_forall fun _ => rfl) (hfg.mono ?_) rintro p ⟨hp1, hp2⟩ simp only [(· ∘ ·), hp1, hp2] @@ -375,7 +375,7 @@ theorem HasFDerivAt.of_local_left_inverse {f : E → F} {f' : E ≃L[𝕜] F} {g · intro p hp simp [hp, hfg.self_of_nhds] · refine ((hf.isBigO_sub_rev f'.antilipschitz).comp_tendsto hg).congr' - (eventually_of_forall fun _ => rfl) (hfg.mono ?_) + (Eventually.of_forall fun _ => rfl) (hfg.mono ?_) rintro p hp simp only [(· ∘ ·), hp, hfg.self_of_nhds] @@ -404,7 +404,7 @@ theorem HasFDerivWithinAt.eventually_ne (h : HasFDerivWithinAt f f' s x) (hf' : ∃ C, ∀ z, ‖z‖ ≤ C * ‖f' z‖) : ∀ᶠ z in 𝓝[s \ {x}] x, f z ≠ f x := by rw [nhdsWithin, diff_eq, ← inf_principal, ← inf_assoc, eventually_inf_principal] have A : (fun z => z - x) =O[𝓝[s] x] fun z => f' (z - x) := - isBigO_iff.2 <| hf'.imp fun C hC => eventually_of_forall fun z => hC _ + isBigO_iff.2 <| hf'.imp fun C hC => Eventually.of_forall fun z => hC _ have : (fun z => f z - f x) ~[𝓝[s] x] fun z => f' (z - x) := h.isLittleO.trans_isBigO A simpa [not_imp_not, sub_eq_zero] using (A.trans this.isBigO_symm).eq_zero_imp diff --git a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean index fa52f1fbe6370..b24d6c6e85fe1 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean @@ -319,4 +319,4 @@ theorem second_derivative_symmetric_of_eventually {f : E → F} {f' : E → E derivative is symmetric. -/ theorem second_derivative_symmetric {f : E → F} {f' : E → E →L[ℝ] F} {f'' : E →L[ℝ] E →L[ℝ] F} (hf : ∀ y, HasFDerivAt f (f' y) y) (hx : HasFDerivAt f' f'' x) (v w : E) : f'' v w = f'' w v := - second_derivative_symmetric_of_eventually (Filter.eventually_of_forall hf) hx v w + second_derivative_symmetric_of_eventually (Filter.Eventually.of_forall hf) hx v w diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean index 7994a8712c5d2..454b00502e126 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ApproximatesLinearOn.lean @@ -252,13 +252,13 @@ theorem surjOn_closedBall_of_nonlinearRightInverse obtain ⟨x, hx⟩ : ∃ x, Tendsto u atTop (𝓝 x) := cauchySeq_tendsto_of_complete this -- As all the `uₙ` belong to the ball `closedBall b ε`, so does their limit `x`. have xmem : x ∈ closedBall b ε := - isClosed_ball.mem_of_tendsto hx (eventually_of_forall fun n => C n _ (D n).2) + isClosed_ball.mem_of_tendsto hx (Eventually.of_forall fun n => C n _ (D n).2) refine ⟨x, xmem, ?_⟩ -- It remains to check that `f x = y`. This follows from continuity of `f` on `closedBall b ε` -- and from the fact that `f uₙ` is converging to `y` by construction. have hx' : Tendsto u atTop (𝓝[closedBall b ε] x) := by simp only [nhdsWithin, tendsto_inf, hx, true_and_iff, tendsto_principal] - exact eventually_of_forall fun n => C n _ (D n).2 + exact Eventually.of_forall fun n => C n _ (D n).2 have T1 : Tendsto (f ∘ u) atTop (𝓝 (f x)) := (hf.continuousOn.mono hε x xmem).tendsto.comp hx' have T2 : Tendsto (f ∘ u) atTop (𝓝 y) := by diff --git a/Mathlib/Analysis/Calculus/LHopital.lean b/Mathlib/Analysis/Calculus/LHopital.lean index 0e18fd06df1c3..653749f47e528 100644 --- a/Mathlib/Analysis/Calculus/LHopital.lean +++ b/Mathlib/Analysis/Calculus/LHopital.lean @@ -364,12 +364,12 @@ theorem lhopital_zero_nhds_right (hdf : ∀ᶠ x in 𝓝[>] a, DifferentiableAt (hdiv : Tendsto (fun x => (deriv f) x / (deriv g) x) (𝓝[>] a) l) : Tendsto (fun x => f x / g x) (𝓝[>] a) l := by have hdg : ∀ᶠ x in 𝓝[>] a, DifferentiableAt ℝ g x := - hg'.mp (eventually_of_forall fun _ hg' => + hg'.mp (Eventually.of_forall fun _ hg' => by_contradiction fun h => hg' (deriv_zero_of_not_differentiableAt h)) have hdf' : ∀ᶠ x in 𝓝[>] a, HasDerivAt f (deriv f x) x := - hdf.mp (eventually_of_forall fun _ => DifferentiableAt.hasDerivAt) + hdf.mp (Eventually.of_forall fun _ => DifferentiableAt.hasDerivAt) have hdg' : ∀ᶠ x in 𝓝[>] a, HasDerivAt g (deriv g x) x := - hdg.mp (eventually_of_forall fun _ => DifferentiableAt.hasDerivAt) + hdg.mp (Eventually.of_forall fun _ => DifferentiableAt.hasDerivAt) exact HasDerivAt.lhopital_zero_nhds_right hdf' hdg' hg' hfa hga hdiv /-- **L'Hôpital's rule** for approaching a real from the left, `deriv` version -/ @@ -379,12 +379,12 @@ theorem lhopital_zero_nhds_left (hdf : ∀ᶠ x in 𝓝[<] a, DifferentiableAt (hdiv : Tendsto (fun x => (deriv f) x / (deriv g) x) (𝓝[<] a) l) : Tendsto (fun x => f x / g x) (𝓝[<] a) l := by have hdg : ∀ᶠ x in 𝓝[<] a, DifferentiableAt ℝ g x := - hg'.mp (eventually_of_forall fun _ hg' => + hg'.mp (Eventually.of_forall fun _ hg' => by_contradiction fun h => hg' (deriv_zero_of_not_differentiableAt h)) have hdf' : ∀ᶠ x in 𝓝[<] a, HasDerivAt f (deriv f x) x := - hdf.mp (eventually_of_forall fun _ => DifferentiableAt.hasDerivAt) + hdf.mp (Eventually.of_forall fun _ => DifferentiableAt.hasDerivAt) have hdg' : ∀ᶠ x in 𝓝[<] a, HasDerivAt g (deriv g x) x := - hdg.mp (eventually_of_forall fun _ => DifferentiableAt.hasDerivAt) + hdg.mp (Eventually.of_forall fun _ => DifferentiableAt.hasDerivAt) exact HasDerivAt.lhopital_zero_nhds_left hdf' hdg' hg' hfa hga hdiv /-- **L'Hôpital's rule** for approaching a real, `deriv` version. This @@ -413,12 +413,12 @@ theorem lhopital_zero_atTop (hdf : ∀ᶠ x : ℝ in atTop, DifferentiableAt ℝ (hgtop : Tendsto g atTop (𝓝 0)) (hdiv : Tendsto (fun x => (deriv f) x / (deriv g) x) atTop l) : Tendsto (fun x => f x / g x) atTop l := by have hdg : ∀ᶠ x in atTop, DifferentiableAt ℝ g x := hg'.mp - (eventually_of_forall fun _ hg' => + (Eventually.of_forall fun _ hg' => by_contradiction fun h => hg' (deriv_zero_of_not_differentiableAt h)) have hdf' : ∀ᶠ x in atTop, HasDerivAt f (deriv f x) x := - hdf.mp (eventually_of_forall fun _ => DifferentiableAt.hasDerivAt) + hdf.mp (Eventually.of_forall fun _ => DifferentiableAt.hasDerivAt) have hdg' : ∀ᶠ x in atTop, HasDerivAt g (deriv g x) x := - hdg.mp (eventually_of_forall fun _ => DifferentiableAt.hasDerivAt) + hdg.mp (Eventually.of_forall fun _ => DifferentiableAt.hasDerivAt) exact HasDerivAt.lhopital_zero_atTop hdf' hdg' hg' hftop hgtop hdiv /-- **L'Hôpital's rule** for approaching -∞, `deriv` version -/ @@ -427,12 +427,12 @@ theorem lhopital_zero_atBot (hdf : ∀ᶠ x : ℝ in atBot, DifferentiableAt ℝ (hgbot : Tendsto g atBot (𝓝 0)) (hdiv : Tendsto (fun x => (deriv f) x / (deriv g) x) atBot l) : Tendsto (fun x => f x / g x) atBot l := by have hdg : ∀ᶠ x in atBot, DifferentiableAt ℝ g x := - hg'.mp (eventually_of_forall fun _ hg' => + hg'.mp (Eventually.of_forall fun _ hg' => by_contradiction fun h => hg' (deriv_zero_of_not_differentiableAt h)) have hdf' : ∀ᶠ x in atBot, HasDerivAt f (deriv f x) x := - hdf.mp (eventually_of_forall fun _ => DifferentiableAt.hasDerivAt) + hdf.mp (Eventually.of_forall fun _ => DifferentiableAt.hasDerivAt) have hdg' : ∀ᶠ x in atBot, HasDerivAt g (deriv g x) x := - hdg.mp (eventually_of_forall fun _ => DifferentiableAt.hasDerivAt) + hdg.mp (Eventually.of_forall fun _ => DifferentiableAt.hasDerivAt) exact HasDerivAt.lhopital_zero_atBot hdf' hdg' hg' hfbot hgbot hdiv end deriv diff --git a/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean b/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean index bfa8f865b600c..946cbc383845a 100644 --- a/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean +++ b/Mathlib/Analysis/Calculus/LocalExtr/Basic.lean @@ -83,7 +83,7 @@ theorem posTangentConeAt_mono : Monotone fun s => posTangentConeAt s a := by theorem mem_posTangentConeAt_of_frequently_mem (h : ∃ᶠ t : ℝ in 𝓝[>] 0, x + t • y ∈ s) : y ∈ posTangentConeAt s x := by obtain ⟨a, ha, has⟩ := Filter.exists_seq_forall_of_frequently h - refine ⟨a⁻¹, (a · • y), eventually_of_forall has, tendsto_inv_zero_atTop.comp ha, ?_⟩ + refine ⟨a⁻¹, (a · • y), Eventually.of_forall has, tendsto_inv_zero_atTop.comp ha, ?_⟩ refine tendsto_const_nhds.congr' ?_ filter_upwards [(tendsto_nhdsWithin_iff.1 ha).2] with n (hn : 0 < a n) simp [ne_of_gt hn] diff --git a/Mathlib/Analysis/Calculus/Rademacher.lean b/Mathlib/Analysis/Calculus/Rademacher.lean index cbca99bb29e9a..b1924f51dde25 100644 --- a/Mathlib/Analysis/Calculus/Rademacher.lean +++ b/Mathlib/Analysis/Calculus/Rademacher.lean @@ -64,7 +64,7 @@ variation, and is therefore ae differentiable, together with a Fubini argument. theorem memℒp_lineDeriv (hf : LipschitzWith C f) (v : E) : Memℒp (fun x ↦ lineDeriv ℝ f x v) ∞ μ := memℒp_top_of_bound (aestronglyMeasurable_lineDeriv hf.continuous μ) - (C * ‖v‖) (eventually_of_forall (fun _x ↦ norm_lineDeriv_le_of_lipschitz ℝ hf)) + (C * ‖v‖) (.of_forall fun _x ↦ norm_lineDeriv_le_of_lipschitz ℝ hf) variable [FiniteDimensional ℝ E] [IsAddHaarMeasure μ] diff --git a/Mathlib/Analysis/Calculus/TangentCone.lean b/Mathlib/Analysis/Calculus/TangentCone.lean index e4c5b68d694c1..280e5356cb943 100644 --- a/Mathlib/Analysis/Calculus/TangentCone.lean +++ b/Mathlib/Analysis/Calculus/TangentCone.lean @@ -87,7 +87,7 @@ theorem mem_tangentConeAt_of_pow_smul {r : 𝕜} (hr₀ : r ≠ 0) (hr : ‖r‖ theorem tangentCone_univ : tangentConeAt 𝕜 univ x = univ := let ⟨_r, hr₀, hr⟩ := exists_norm_lt_one 𝕜 eq_univ_of_forall fun _ ↦ mem_tangentConeAt_of_pow_smul (norm_pos_iff.1 hr₀) hr <| - eventually_of_forall fun _ ↦ mem_univ _ + Eventually.of_forall fun _ ↦ mem_univ _ theorem tangentCone_mono (h : s ⊆ t) : tangentConeAt 𝕜 s x ⊆ tangentConeAt 𝕜 t x := by rintro y ⟨c, d, ds, ctop, clim⟩ diff --git a/Mathlib/Analysis/Complex/AbsMax.lean b/Mathlib/Analysis/Complex/AbsMax.lean index 9bb3354667e7c..8c266e5611fc5 100644 --- a/Mathlib/Analysis/Complex/AbsMax.lean +++ b/Mathlib/Analysis/Complex/AbsMax.lean @@ -330,7 +330,7 @@ theorem eventually_eq_or_eq_zero_of_isLocalMin_norm {f : E → ℂ} {c : E} refine or_iff_not_imp_right.mpr fun h => ?_ have h1 : ∀ᶠ z in 𝓝 c, f z ≠ 0 := hf.self_of_nhds.continuousAt.eventually_ne h have h2 : IsLocalMax (norm ∘ f)⁻¹ c := hc.inv (h1.mono fun z => norm_pos_iff.mpr) - have h3 : IsLocalMax (norm ∘ f⁻¹) c := by refine h2.congr (eventually_of_forall ?_); simp + have h3 : IsLocalMax (norm ∘ f⁻¹) c := by refine h2.congr (Eventually.of_forall ?_); simp have h4 : ∀ᶠ z in 𝓝 c, DifferentiableAt ℂ f⁻¹ z := by filter_upwards [hf, h1] with z h using h.inv filter_upwards [eventually_eq_of_isLocalMax_norm h4 h3] with z using inv_inj.mp diff --git a/Mathlib/Analysis/Complex/Hadamard.lean b/Mathlib/Analysis/Complex/Hadamard.lean index 663f0d998f29a..88d6c801d518c 100644 --- a/Mathlib/Analysis/Complex/Hadamard.lean +++ b/Mathlib/Analysis/Complex/Hadamard.lean @@ -219,7 +219,7 @@ theorem norm_mul_invInterpStrip_le_one_of_mem_verticalClosedStrip (f : ℂ → E rw [Asymptotics.isBigO_iff] use 1 rw [eventually_inf_principal] - apply eventually_of_forall + apply Eventually.of_forall intro x hx norm_num exact (hBF x ((preimage_mono Ioo_subset_Icc_self) hx)).trans diff --git a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean index 09792d0f7a4bd..14b603b86f48e 100644 --- a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean +++ b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean @@ -170,7 +170,7 @@ theorem differentiableOn_tsum_of_summable_norm {u : ι → ℝ} (hu : Summable u DifferentiableOn ℂ (fun w : ℂ => ∑' i : ι, F i w) U := by classical have hc := (tendstoUniformlyOn_tsum hu hF_le).tendstoLocallyUniformlyOn - refine hc.differentiableOn (eventually_of_forall fun s => ?_) hU + refine hc.differentiableOn (Eventually.of_forall fun s => ?_) hU exact DifferentiableOn.sum fun i _ => hf i /-- If the terms in the sum `∑' (i : ι), F i` are uniformly bounded on `U` by a @@ -182,7 +182,7 @@ theorem hasSum_deriv_of_summable_norm {u : ι → ℝ} (hu : Summable u) HasSum (fun i : ι => deriv (F i) z) (deriv (fun w : ℂ => ∑' i : ι, F i w) z) := by rw [HasSum] have hc := (tendstoUniformlyOn_tsum hu hF_le).tendstoLocallyUniformlyOn - convert (hc.deriv (eventually_of_forall fun s => + convert (hc.deriv (Eventually.of_forall fun s => DifferentiableOn.sum fun i _ => hf i) hU).tendsto_at hz using 1 ext1 s exact (deriv_sum fun i _ => (hf i).differentiableAt (hU.mem_nhds hz)).symm diff --git a/Mathlib/Analysis/Complex/PhragmenLindelof.lean b/Mathlib/Analysis/Complex/PhragmenLindelof.lean index a4a9e03590e0f..c08ca1fd6b87e 100644 --- a/Mathlib/Analysis/Complex/PhragmenLindelof.lean +++ b/Mathlib/Analysis/Complex/PhragmenLindelof.lean @@ -733,7 +733,7 @@ theorem right_half_plane_of_bounded_on_real (hd : DiffContOnCl ℂ f {z | 0 < z. refine right_half_plane_of_tendsto_zero_on_real hd ?_ ?_ (fun y => ?_) hz · rcases hexp with ⟨c, hc, B, hO⟩ refine ⟨c, hc, B, (IsBigO.of_bound 1 ?_).trans hO⟩ - refine eventually_inf_principal.2 <| eventually_of_forall fun z hz => ?_ + refine eventually_inf_principal.2 <| Eventually.of_forall fun z hz => ?_ rw [hgn, one_mul] refine mul_le_of_le_one_left (norm_nonneg _) (Real.exp_le_one_iff.2 ?_) exact mul_nonpos_of_nonpos_of_nonneg ε₀.le (le_of_lt hz) diff --git a/Mathlib/Analysis/Convex/Integral.lean b/Mathlib/Analysis/Convex/Integral.lean index 435c9f7ec4a19..12b11b68057db 100644 --- a/Mathlib/Analysis/Convex/Integral.lean +++ b/Mathlib/Analysis/Convex/Integral.lean @@ -68,7 +68,7 @@ theorem Convex.integral_mem [IsProbabilityMeasure μ] (hs : Convex ℝ s) (hsc : set G : ℕ → SimpleFunc α E := SimpleFunc.approxOn _ hgm.measurable (range g ∩ s) y₀ h₀ have : Tendsto (fun n => (G n).integral μ) atTop (𝓝 <| ∫ x, g x ∂μ) := tendsto_integral_approxOn_of_measurable hfi _ hg _ (integrable_const _) - refine hsc.mem_of_tendsto this (eventually_of_forall fun n => hs.sum_mem ?_ ?_ ?_) + refine hsc.mem_of_tendsto this (Eventually.of_forall fun n => hs.sum_mem ?_ ?_ ?_) · exact fun _ _ => ENNReal.toReal_nonneg · rw [← ENNReal.toReal_sum, (G n).sum_range_measure_preimage_singleton, measure_univ, ENNReal.one_toReal] diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index a1f0062306e40..7beb76c9e195a 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -227,7 +227,7 @@ theorem ConvolutionExistsAt.ofNorm' {x₀ : G} (hmf : AEStronglyMeasurable f μ) (hmg : AEStronglyMeasurable g <| map (fun t => x₀ - t) μ) : ConvolutionExistsAt f g x₀ L μ := by refine (h.const_mul ‖L‖).mono' - (hmf.convolution_integrand_snd' L hmg) (eventually_of_forall fun x => ?_) + (hmf.convolution_integrand_snd' L hmg) (Eventually.of_forall fun x => ?_) rw [mul_apply', ← mul_assoc] apply L.le_opNorm₂ @@ -278,15 +278,15 @@ theorem Integrable.convolution_integrand (hf : Integrable f ν) (hg : Integrable have h2_meas : AEStronglyMeasurable (fun y : G => ∫ x : G, ‖L (f y) (g (x - y))‖ ∂μ) ν := h_meas.prod_swap.norm.integral_prod_right' simp_rw [integrable_prod_iff' h_meas] - refine ⟨eventually_of_forall fun t => (L (f t)).integrable_comp (hg.comp_sub_right t), ?_⟩ + refine ⟨Eventually.of_forall fun t => (L (f t)).integrable_comp (hg.comp_sub_right t), ?_⟩ refine Integrable.mono' ?_ h2_meas - (eventually_of_forall fun t => (?_ : _ ≤ ‖L‖ * ‖f t‖ * ∫ x, ‖g (x - t)‖ ∂μ)) + (Eventually.of_forall fun t => (?_ : _ ≤ ‖L‖ * ‖f t‖ * ∫ x, ‖g (x - t)‖ ∂μ)) · simp only [integral_sub_right_eq_self (‖g ·‖)] exact (hf.norm.const_mul _).mul_const _ · simp_rw [← integral_mul_left] rw [Real.norm_of_nonneg (by positivity)] - exact integral_mono_of_nonneg (eventually_of_forall fun t => norm_nonneg _) - ((hg.comp_sub_right t).norm.const_mul _) (eventually_of_forall fun t => L.le_opNorm₂ _ _) + exact integral_mono_of_nonneg (Eventually.of_forall fun t => norm_nonneg _) + ((hg.comp_sub_right t).norm.const_mul _) (Eventually.of_forall fun t => L.le_opNorm₂ _ _) theorem Integrable.ae_convolution_exists (hf : Integrable f ν) (hg : Integrable g μ) : ∀ᵐ x ∂μ, ConvolutionExistsAt f g x L ν := @@ -545,7 +545,7 @@ theorem continuousOn_convolution_right_with_param {g : P → G → E'} {s : Set by_cases H : ∀ p ∈ s, ∀ x, g p x = 0 · apply (continuousOn_const (c := 0)).congr rintro ⟨p, x⟩ ⟨hp, -⟩ - apply integral_eq_zero_of_ae (eventually_of_forall (fun y ↦ ?_)) + apply integral_eq_zero_of_ae (Eventually.of_forall (fun y ↦ ?_)) simp [H p hp _] have : LocallyCompactSpace G := by push_neg at H @@ -618,10 +618,10 @@ theorem _root_.BddAbove.continuous_convolution_right_of_integrable filter_upwards with x; filter_upwards with t apply_rules [L.le_of_opNorm₂_le_of_le, le_rfl, le_ciSup hbg (x - t)] refine continuousAt_of_dominated ?_ this ?_ ?_ - · exact eventually_of_forall fun x => + · exact Eventually.of_forall fun x => hf.aestronglyMeasurable.convolution_integrand_snd' L hg.aestronglyMeasurable · exact (hf.norm.const_mul _).mul_const _ - · exact eventually_of_forall fun t => (L.continuous₂.comp₂ continuous_const <| + · exact Eventually.of_forall fun t => (L.continuous₂.comp₂ continuous_const <| hg.comp <| continuous_id.sub continuous_const).continuousAt end Group @@ -752,7 +752,7 @@ theorem dist_convolution_le' {x₀ : G} {R ε : ℝ} {z₀ : E'} (hε : 0 ≤ ε simp_rw [dist_eq_norm] at h2 ⊢ rw [← integral_sub hfg.integrable]; swap; · exact (L.flip z₀).integrable_comp hif refine (norm_integral_le_of_norm_le ((L.integrable_comp hif).norm.mul_const ε) - (eventually_of_forall h2)).trans ?_ + (Eventually.of_forall h2)).trans ?_ rw [integral_mul_right] refine mul_le_mul_of_nonneg_right ?_ hε have h3 : ∀ t, ‖L (f t)‖ ≤ ‖L‖ * ‖f t‖ := by @@ -899,7 +899,7 @@ theorem convolution_assoc (hL : ∀ (x : E) (y : E') (z : E''), L₂ (L x y) z = ((measurable_const.sub measurable_snd).sub measurable_fst) refine QuasiMeasurePreserving.absolutelyContinuous ?_ refine QuasiMeasurePreserving.prod_of_left - ((measurable_const.sub measurable_snd).sub measurable_fst) (eventually_of_forall fun y => ?_) + ((measurable_const.sub measurable_snd).sub measurable_fst) (Eventually.of_forall fun y => ?_) dsimp only exact quasiMeasurePreserving_sub_left_of_right_invariant μ _ have h2_meas : @@ -920,8 +920,8 @@ theorem convolution_assoc (hL : ∀ (x : E) (y : E') (z : E''), L₂ (L x y) z = (((quasiMeasurePreserving_sub_left_of_right_invariant ν x₀).ae hgk).mono fun t ht => ?_) simp_rw [convolution_def, mul_apply', mul_mul_mul_comm ‖L₃‖ ‖L₄‖, ← integral_mul_left] rw [Real.norm_of_nonneg (by positivity)] - refine integral_mono_of_nonneg (eventually_of_forall fun t => norm_nonneg _) - ((ht.const_mul _).const_mul _) (eventually_of_forall fun s => ?_) + refine integral_mono_of_nonneg (Eventually.of_forall fun t => norm_nonneg _) + ((ht.const_mul _).const_mul _) (Eventually.of_forall fun s => ?_) simp only [← mul_assoc ‖L₄‖] apply_rules [ContinuousLinearMap.le_of_opNorm₂_le_of_le, le_rfl] @@ -951,7 +951,7 @@ theorem _root_.HasCompactSupport.hasFDerivAt_convolution_right (hcg : HasCompact have : ProperSpace G := FiniteDimensional.proper_rclike 𝕜 G set L' := L.precompR G have h1 : ∀ᶠ x in 𝓝 x₀, AEStronglyMeasurable (fun t => L (f t) (g (x - t))) μ := - eventually_of_forall + Eventually.of_forall (hf.aestronglyMeasurable.convolution_integrand_snd L hg.continuous.aestronglyMeasurable) have h2 : ∀ x, AEStronglyMeasurable (fun t => L' (f t) (fderiv 𝕜 g (x - t))) μ := hf.aestronglyMeasurable.convolution_integrand_snd L' @@ -971,7 +971,7 @@ theorem _root_.HasCompactSupport.hasFDerivAt_convolution_right (hcg : HasCompact (ball_subset_closedBall hx) · rw [integrable_indicator_iff hK'.measurableSet] exact ((hf.integrableOn_isCompact hK').norm.const_mul _).mul_const _ - · exact eventually_of_forall fun t x _ => (L _).hasFDerivAt.comp x (h3 x t) + · exact Eventually.of_forall fun t x _ => (L _).hasFDerivAt.comp x (h3 x t) · exact hcg.convolutionExists_right L hf hg.continuous x₀ theorem _root_.HasCompactSupport.hasFDerivAt_convolution_left [IsNegInvariant μ] diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 2d2ee1ed85185..868adbc62f8a6 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -1070,7 +1070,7 @@ lemma integrable_pow_mul (f : 𝓢(D, V)) lemma integrable (f : 𝓢(D, V)) : Integrable f μ := (f.integrable_pow_mul μ 0).mono f.continuous.aestronglyMeasurable - (eventually_of_forall (fun _ ↦ by simp)) + (Eventually.of_forall (fun _ ↦ by simp)) variable (𝕜 μ) in /-- The integral as a continuous linear map from Schwartz space to the codomain. -/ diff --git a/Mathlib/Analysis/Fourier/AddCircle.lean b/Mathlib/Analysis/Fourier/AddCircle.lean index fe1f6f77fa8b1..17dac62db0a31 100644 --- a/Mathlib/Analysis/Fourier/AddCircle.lean +++ b/Mathlib/Analysis/Fourier/AddCircle.lean @@ -403,7 +403,7 @@ variable (f : C(AddCircle T, ℂ)) theorem fourierCoeff_toLp (n : ℤ) : fourierCoeff (toLp (E := ℂ) 2 haarAddCircle ℂ f) n = fourierCoeff f n := - integral_congr_ae (Filter.EventuallyEq.mul (Filter.eventually_of_forall (by tauto)) + integral_congr_ae (Filter.EventuallyEq.mul (Filter.Eventually.of_forall (by tauto)) (ContinuousMap.coeFn_toAEEqFun haarAddCircle f)) variable {f} diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index dd4ff0411f5f1..588af56dc65eb 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -218,7 +218,7 @@ theorem hasFDerivAt_fourierIntegral (fourierIntegral_convergent_iff continuous_fourierChar (by apply L.continuous₂ : Continuous (fun p : V × W ↦ L.toLinearMap₂ p.1 p.2)) w').2 hf have h1 : ∀ᶠ w' in 𝓝 w, AEStronglyMeasurable (F w') μ := - eventually_of_forall (fun w' ↦ (h0 w').aestronglyMeasurable) + Eventually.of_forall (fun w' ↦ (h0 w').aestronglyMeasurable) have h3 : AEStronglyMeasurable (F' w) μ := by refine .smul ?_ hf.1.fourierSMulRight refine (continuous_fourierChar.comp ?_).aestronglyMeasurable diff --git a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean index 59c58e215be50..9d6a346e16240 100644 --- a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean +++ b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean @@ -203,7 +203,7 @@ theorem tendsto_integral_exp_inner_smul_cocompact : (tendsto_integral_exp_inner_smul_cocompact_of_continuous_compact_support hg_cont hg_supp)) _ (div_pos hε two_pos)).mp - (eventually_of_forall fun w hI => ?_) + (Eventually.of_forall fun w hI => ?_) rw [dist_eq_norm] at hI ⊢ have : ‖(∫ v, 𝐞 (-⟪v, w⟫) • f v) - ∫ v, 𝐞 (-⟪v, w⟫) • g v‖ ≤ ε / 2 := by refine le_trans ?_ hfg diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 82122b6dce018..786527166c2b3 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -637,11 +637,11 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_eq [FiniteDimensional ℝ F] = eLpNorm (e.symm ∘ v) p' μ := by simp_rw [v, Function.comp, e.symm_apply_apply] _ ≤ C₁ • eLpNorm v p' μ := by apply eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul - exact eventually_of_forall (fun x ↦ (e.symm : F' →L[ℝ] F).le_opNNNorm _) + exact Eventually.of_forall (fun x ↦ (e.symm : F' →L[ℝ] F).le_opNNNorm _) _ = C₁ * eLpNorm v p' μ := rfl _ ≤ C₁ * C * eLpNorm (fderiv ℝ v) p μ := by rw [mul_assoc]; gcongr _ ≤ C₁ * C * (C₂ * eLpNorm (fderiv ℝ u) p μ) := by - gcongr; exact eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul (eventually_of_forall h4v) p + gcongr; exact eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul (Eventually.of_forall h4v) p _ = SNormLESNormFDerivOfEqConst F μ p * eLpNorm (fderiv ℝ u) p μ := by simp_rw [SNormLESNormFDerivOfEqConst] push_cast diff --git a/Mathlib/Analysis/InnerProductSpace/l2Space.lean b/Mathlib/Analysis/InnerProductSpace/l2Space.lean index 4552de407ddb4..4b30cfc401477 100644 --- a/Mathlib/Analysis/InnerProductSpace/l2Space.lean +++ b/Mathlib/Analysis/InnerProductSpace/l2Space.lean @@ -237,7 +237,7 @@ protected theorem range_linearIsometry [∀ i, CompleteSpace (G i)] : classical refine le_antisymm ?_ ?_ · rintro x ⟨f, rfl⟩ - refine mem_closure_of_tendsto (hV.hasSum_linearIsometry f) (eventually_of_forall ?_) + refine mem_closure_of_tendsto (hV.hasSum_linearIsometry f) (Eventually.of_forall ?_) intro s rw [SetLike.mem_coe] refine sum_mem ?_ @@ -441,7 +441,7 @@ protected theorem dense_span (b : HilbertBasis ι 𝕜 E) : classical rw [eq_top_iff] rintro x - - refine mem_closure_of_tendsto (b.hasSum_repr x) (eventually_of_forall ?_) + refine mem_closure_of_tendsto (b.hasSum_repr x) (Eventually.of_forall ?_) intro s simp only [SetLike.mem_coe] refine sum_mem ?_ diff --git a/Mathlib/Analysis/LocallyConvex/Bounded.lean b/Mathlib/Analysis/LocallyConvex/Bounded.lean index 680f38ea9ba4f..66eed51d7bf1e 100644 --- a/Mathlib/Analysis/LocallyConvex/Bounded.lean +++ b/Mathlib/Analysis/LocallyConvex/Bounded.lean @@ -99,7 +99,7 @@ theorem IsVonNBounded.of_boundedSpace [BoundedSpace 𝕜] {s : Set E} : IsVonNBo @[nontriviality] theorem IsVonNBounded.of_subsingleton [Subsingleton E] {s : Set E} : IsVonNBounded 𝕜 s := - fun U hU ↦ eventually_of_forall fun c ↦ calc + fun U hU ↦ .of_forall fun c ↦ calc s ⊆ univ := subset_univ s _ = c • U := .symm <| Subsingleton.eq_univ_of_nonempty <| (Filter.nonempty_of_mem hU).image _ @@ -239,7 +239,7 @@ theorem isVonNBounded_of_smul_tendsto_zero {ε : ι → 𝕜} {l : Filter ι} [l theorem isVonNBounded_iff_smul_tendsto_zero {ε : ι → 𝕜} {l : Filter ι} [l.NeBot] (hε : Tendsto ε l (𝓝[≠] 0)) {S : Set E} : IsVonNBounded 𝕜 S ↔ ∀ x : ι → E, (∀ n, x n ∈ S) → Tendsto (ε • x) l (𝓝 0) := - ⟨fun hS x hxS => hS.smul_tendsto_zero (eventually_of_forall hxS) (le_trans hε nhdsWithin_le_nhds), + ⟨fun hS x hxS => hS.smul_tendsto_zero (Eventually.of_forall hxS) (le_trans hε nhdsWithin_le_nhds), isVonNBounded_of_smul_tendsto_zero (by exact hε self_mem_nhdsWithin)⟩ end sequence @@ -255,7 +255,7 @@ theorem IsVonNBounded.extend_scalars [NontriviallyNormedField 𝕜] obtain ⟨ε, hε, hε₀⟩ : ∃ ε : ℕ → 𝕜, Tendsto ε atTop (𝓝 0) ∧ ∀ᶠ n in atTop, ε n ≠ 0 := by simpa only [tendsto_nhdsWithin_iff] using exists_seq_tendsto (𝓝[≠] (0 : 𝕜)) refine isVonNBounded_of_smul_tendsto_zero (ε := (ε · • 1)) (by simpa) fun x hx ↦ ?_ - have := h.smul_tendsto_zero (eventually_of_forall hx) hε + have := h.smul_tendsto_zero (.of_forall hx) hε simpa only [Pi.smul_def', smul_one_smul] section NormedField diff --git a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean index 2ca1c42f8b1e6..dc3db6a84bb37 100644 --- a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean +++ b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean @@ -647,7 +647,7 @@ protected theorem _root_.WithSeminorms.equicontinuous_TFAE {κ : Type*} tfae_have 4 → 1 -- This would work over any `NormedField` · intro ⟨p, hp, hfp⟩ exact Metric.equicontinuousAt_of_continuity_modulus p (map_zero p ▸ hp.tendsto 0) _ <| - eventually_of_forall fun x k ↦ by simpa using hfp k x + Eventually.of_forall fun x k ↦ by simpa using hfp k x tfae_finish theorem _root_.WithSeminorms.uniformEquicontinuous_iff_exists_continuous_seminorm {κ : Type*} diff --git a/Mathlib/Analysis/MellinTransform.lean b/Mathlib/Analysis/MellinTransform.lean index 81136cb0e8e3c..95ff3017bcc71 100644 --- a/Mathlib/Analysis/MellinTransform.lean +++ b/Mathlib/Analysis/MellinTransform.lean @@ -229,7 +229,7 @@ theorem mellin_convergent_zero_of_isBigO {b : ℝ} {f : ℝ → ℝ} rw [← IntegrableOn, ← integrableOn_Ioc_iff_integrableOn_Ioo, ← intervalIntegrable_iff_integrableOn_Ioc_of_le hε.le] exact intervalIntegral.intervalIntegrable_rpow' (by linarith) - · refine (ae_restrict_iff' measurableSet_Ioo).mpr (eventually_of_forall fun t ht => ?_) + · refine (ae_restrict_iff' measurableSet_Ioo).mpr (Eventually.of_forall fun t ht => ?_) rw [mul_comm, norm_mul] specialize hε' _ ht.1 · rw [dist_eq_norm, sub_zero, norm_of_nonneg (le_of_lt ht.1)] @@ -278,8 +278,8 @@ theorem isBigO_rpow_top_log_smul [NormedSpace ℝ E] {a b : ℝ} {f : ℝ → E} (fun t : ℝ => log t • f t) =O[atTop] (· ^ (-b)) := by refine ((isLittleO_log_rpow_atTop (sub_pos.mpr hab)).isBigO.smul hf).congr' - (eventually_of_forall fun t => by rfl) - ((eventually_gt_atTop 0).mp (eventually_of_forall fun t ht => ?_)) + (Eventually.of_forall fun t => by rfl) + ((eventually_gt_atTop 0).mp (Eventually.of_forall fun t ht => ?_)) simp only rw [smul_eq_mul, ← rpow_add ht, ← sub_eq_add_neg, sub_eq_add_neg a, add_sub_cancel_left] @@ -290,13 +290,13 @@ theorem isBigO_rpow_zero_log_smul [NormedSpace ℝ E] {a b : ℝ} {f : ℝ → E have : log =o[𝓝[>] 0] fun t : ℝ => t ^ (a - b) := by refine ((isLittleO_log_rpow_atTop (sub_pos.mpr hab)).neg_left.comp_tendsto tendsto_inv_zero_atTop).congr' - (eventually_nhdsWithin_iff.mpr <| eventually_of_forall fun t ht => ?_) - (eventually_nhdsWithin_iff.mpr <| eventually_of_forall fun t ht => ?_) + (eventually_nhdsWithin_iff.mpr <| Eventually.of_forall fun t ht => ?_) + (eventually_nhdsWithin_iff.mpr <| Eventually.of_forall fun t ht => ?_) · simp_rw [Function.comp_apply, ← one_div, log_div one_ne_zero (ne_of_gt ht), Real.log_one, zero_sub, neg_neg] · simp_rw [Function.comp_apply, inv_rpow (le_of_lt ht), ← rpow_neg (le_of_lt ht), neg_sub] - refine (this.isBigO.smul hf).congr' (eventually_of_forall fun t => by rfl) - (eventually_nhdsWithin_iff.mpr (eventually_of_forall fun t ht => ?_)) + refine (this.isBigO.smul hf).congr' (Eventually.of_forall fun t => by rfl) + (eventually_nhdsWithin_iff.mpr (Eventually.of_forall fun t ht => ?_)) simp_rw [smul_eq_mul, ← rpow_add ht] congr 1 abel @@ -319,7 +319,7 @@ theorem mellin_hasDerivAt_of_isBigO_rpow [NormedSpace ℂ E] {a b : ℝ} ⟨min w w', lt_min hw1 hw1', (min_le_right _ _).trans_lt hw2', (min_le_left _ _).trans_lt hw2⟩ let bound : ℝ → ℝ := fun t : ℝ => (t ^ (s.re + v - 1) + t ^ (s.re - v - 1)) * |log t| * ‖f t‖ have h1 : ∀ᶠ z : ℂ in 𝓝 s, AEStronglyMeasurable (F z) (volume.restrict <| Ioi 0) := by - refine eventually_of_forall fun z => AEStronglyMeasurable.smul ?_ hfc.aestronglyMeasurable + refine Eventually.of_forall fun z => AEStronglyMeasurable.smul ?_ hfc.aestronglyMeasurable refine ContinuousOn.aestronglyMeasurable ?_ measurableSet_Ioi refine ContinuousAt.continuousOn fun t ht => ?_ exact continuousAt_ofReal_cpow_const _ _ (Or.inr <| ne_of_gt ht) diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index cef2a8e2bc38e..6e710a4b5a224 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -596,7 +596,7 @@ lemma _root_.Subalgebra.isUnit_of_isUnit_val_of_eventually {l : Filter S} {a : S all_goals ext; simp apply hS.mem_of_tendsto hla₂ rw [Filter.eventually_map] - apply hl.mp <| eventually_of_forall fun x hx ↦ ?_ + apply hl.mono fun x hx ↦ ?_ suffices Ring.inverse (val S x) = (val S ↑hx.unit⁻¹) from this ▸ Subtype.property _ rw [← (hx.map (val S)).unit_spec, Ring.inverse_unit (hx.map (val S)).unit, val] apply Units.mul_eq_one_iff_inv_eq.mp diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index c7963a31c6d62..c401ebaa128e2 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -833,7 +833,7 @@ instance (priority := 100) NormedDivisionRing.to_hasContinuousInv₀ : HasContin -- Porting note: `ENNReal.{mul_sub, sub_mul}` should be `protected` _ = ‖r - e‖ / ‖r‖ / ‖e‖ := by field_simp [mul_comm] _ ≤ ‖r - e‖ / ‖r‖ / ε := by gcongr - refine squeeze_zero' (eventually_of_forall fun _ => norm_nonneg _) this ?_ + refine squeeze_zero' (Eventually.of_forall fun _ => norm_nonneg _) this ?_ refine (((continuous_const.sub continuous_id).norm.div_const _).div_const _).tendsto' _ _ ?_ simp diff --git a/Mathlib/Analysis/Normed/Group/AddTorsor.lean b/Mathlib/Analysis/Normed/Group/AddTorsor.lean index 8a26d2971bb7e..c4b97754790bb 100644 --- a/Mathlib/Analysis/Normed/Group/AddTorsor.lean +++ b/Mathlib/Analysis/Normed/Group/AddTorsor.lean @@ -283,7 +283,7 @@ theorem IsClosed.vadd_right_of_isCompact {s : Set V} {t : Set P} (hs : IsClosed choose! a ha v hv hav using husv rcases ht.isSeqCompact hv with ⟨q, hqt, φ, φ_mono, hφq⟩ refine ⟨p -ᵥ q, hs.mem_of_tendsto ((hup.comp φ_mono.tendsto_atTop).vsub hφq) - (eventually_of_forall fun n ↦ ?_), q, hqt, vsub_vadd _ _⟩ + (Eventually.of_forall fun n ↦ ?_), q, hqt, vsub_vadd _ _⟩ convert ha (φ n) using 1 exact (eq_vadd_iff_vsub_eq _ _ _).mp (hav (φ n)).symm diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index d5392d8c567b0..efa5924c33e43 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -739,7 +739,7 @@ real function `a` which tends to `0`, then `f` tends to `0`. In this pair of lem theorem squeeze_one_norm' {f : α → E} {a : α → ℝ} {t₀ : Filter α} (h : ∀ᶠ n in t₀, ‖f n‖ ≤ a n) (h' : Tendsto a t₀ (𝓝 0)) : Tendsto f t₀ (𝓝 1) := tendsto_one_iff_norm_tendsto_zero.2 <| - squeeze_zero' (eventually_of_forall fun _n => norm_nonneg' _) h h' + squeeze_zero' (Eventually.of_forall fun _n => norm_nonneg' _) h h' /-- Special case of the sandwich theorem: if the norm of `f` is bounded by a real function `a` which tends to `0`, then `f` tends to `1`. -/ @@ -747,7 +747,7 @@ tends to `0`, then `f` tends to `1`. -/ function `a` which tends to `0`, then `f` tends to `0`."] theorem squeeze_one_norm {f : α → E} {a : α → ℝ} {t₀ : Filter α} (h : ∀ n, ‖f n‖ ≤ a n) : Tendsto a t₀ (𝓝 0) → Tendsto f t₀ (𝓝 1) := - squeeze_one_norm' <| eventually_of_forall h + squeeze_one_norm' <| Eventually.of_forall h @[to_additive] theorem tendsto_norm_div_self (x : E) : Tendsto (fun a => ‖a / x‖) (𝓝 x) (𝓝 0) := by diff --git a/Mathlib/Analysis/Normed/Group/InfiniteSum.lean b/Mathlib/Analysis/Normed/Group/InfiniteSum.lean index fb02a2775c4d7..2eb196f059adf 100644 --- a/Mathlib/Analysis/Normed/Group/InfiniteSum.lean +++ b/Mathlib/Analysis/Normed/Group/InfiniteSum.lean @@ -65,7 +65,7 @@ theorem cauchySeq_finset_of_norm_bounded_eventually {f : ι → E} {g : ι → theorem cauchySeq_finset_of_norm_bounded {f : ι → E} (g : ι → ℝ) (hg : Summable g) (h : ∀ i, ‖f i‖ ≤ g i) : CauchySeq fun s : Finset ι => ∑ i ∈ s, f i := - cauchySeq_finset_of_norm_bounded_eventually hg <| eventually_of_forall h + cauchySeq_finset_of_norm_bounded_eventually hg <| Eventually.of_forall h /-- A version of the **direct comparison test** for conditionally convergent series. See `cauchySeq_finset_of_norm_bounded` for the same statement about absolutely convergent ones. -/ diff --git a/Mathlib/Analysis/Normed/Lp/lpSpace.lean b/Mathlib/Analysis/Normed/Lp/lpSpace.lean index 23bb44f6901d5..8bfe697508229 100644 --- a/Mathlib/Analysis/Normed/Lp/lpSpace.lean +++ b/Mathlib/Analysis/Normed/Lp/lpSpace.lean @@ -1066,9 +1066,9 @@ theorem memℓp_of_tendsto {F : ι → lp E p} (hF : Bornology.IsBounded (Set.ra · apply memℓp_infty use C rintro _ ⟨a, rfl⟩ - exact norm_apply_le_of_tendsto (eventually_of_forall hCF) hf a + exact norm_apply_le_of_tendsto (Eventually.of_forall hCF) hf a · apply memℓp_gen' - exact sum_rpow_le_of_tendsto hp.ne (eventually_of_forall hCF) hf + exact sum_rpow_le_of_tendsto hp.ne (Eventually.of_forall hCF) hf /-- If a sequence is Cauchy in the `lp E p` topology and pointwise convergent to an element `f` of `lp E p`, then it converges to `f` in the `lp E p` topology. -/ diff --git a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean index 6fbb936e9f60b..4c55f7d09dafa 100644 --- a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean +++ b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean @@ -508,7 +508,7 @@ lemma ProperSpace.of_locallyCompactSpace (𝕜 : Type*) [NontriviallyNormedField simpa [_root_.smul_closedBall' this] using hr.smul (c ^ n) have hTop : Tendsto (fun n ↦ ‖c‖^n * r) atTop atTop := Tendsto.atTop_mul_const rpos (tendsto_pow_atTop_atTop_of_one_lt hc) - exact .of_seq_closedBall hTop (eventually_of_forall hC) + exact .of_seq_closedBall hTop (Eventually.of_forall hC) @[deprecated (since := "2024-01-31")] alias properSpace_of_locallyCompactSpace := ProperSpace.of_locallyCompactSpace diff --git a/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean b/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean index b43fd4cefc393..439ff7a017041 100644 --- a/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean +++ b/Mathlib/Analysis/Normed/Operator/BoundedLinearMaps.lean @@ -317,7 +317,7 @@ protected theorem IsBoundedBilinearMap.isBigO (h : IsBoundedBilinearMap 𝕜 f) f =O[⊤] fun p : E × F => ‖p.1‖ * ‖p.2‖ := let ⟨C, Cpos, hC⟩ := h.bound Asymptotics.IsBigO.of_bound _ <| - Filter.eventually_of_forall fun ⟨x, y⟩ => by simpa [mul_assoc] using hC x y + Filter.Eventually.of_forall fun ⟨x, y⟩ => by simpa [mul_assoc] using hC x y theorem IsBoundedBilinearMap.isBigO_comp {α : Type*} (H : IsBoundedBilinearMap 𝕜 f) {g : α → E} {h : α → F} {l : Filter α} : (fun x => f (g x, h x)) =O[l] fun x => ‖g x‖ * ‖h x‖ := diff --git a/Mathlib/Analysis/NormedSpace/FunctionSeries.lean b/Mathlib/Analysis/NormedSpace/FunctionSeries.lean index b101059caa463..d4783a6874149 100644 --- a/Mathlib/Analysis/NormedSpace/FunctionSeries.lean +++ b/Mathlib/Analysis/NormedSpace/FunctionSeries.lean @@ -65,7 +65,7 @@ theorem continuousOn_tsum [TopologicalSpace β] {f : α → β → F} {s : Set (hf : ∀ i, ContinuousOn (f i) s) (hu : Summable u) (hfu : ∀ n x, x ∈ s → ‖f n x‖ ≤ u n) : ContinuousOn (fun x => ∑' n, f n x) s := by classical - refine (tendstoUniformlyOn_tsum hu hfu).continuousOn (eventually_of_forall ?_) + refine (tendstoUniformlyOn_tsum hu hfu).continuousOn (Eventually.of_forall ?_) intro t exact continuousOn_finset_sum _ fun i _ => hf i diff --git a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean index 62d8d4d30f47c..03c7738155ba1 100644 --- a/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean +++ b/Mathlib/Analysis/NormedSpace/OperatorNorm/Completeness.lean @@ -60,7 +60,7 @@ def ofTendstoOfBoundedRange {α : Type*} {l : Filter α} [l.NeBot] (f : E' → F (g : α → E' →SL[σ₁₂] F) (hf : Tendsto (fun a x => g a x) l (𝓝 f)) (hg : IsBounded (Set.range g)) : E' →SL[σ₁₂] F := ofMemClosureImageCoeBounded f hg <| mem_closure_of_tendsto hf <| - eventually_of_forall fun _ => mem_image_of_mem _ <| Set.mem_range_self _ + Eventually.of_forall fun _ => mem_image_of_mem _ <| Set.mem_range_self _ /-- If a Cauchy sequence of continuous linear map converges to a continuous linear map pointwise, then it converges to the same map in norm. This lemma is used to prove that the space of continuous diff --git a/Mathlib/Analysis/SpecialFunctions/Exp.lean b/Mathlib/Analysis/SpecialFunctions/Exp.lean index 33170348b5f1f..89a73b0dcb128 100644 --- a/Mathlib/Analysis/SpecialFunctions/Exp.lean +++ b/Mathlib/Analysis/SpecialFunctions/Exp.lean @@ -242,7 +242,7 @@ theorem tendsto_exp_atBot : Tendsto exp atBot (𝓝 0) := congr_arg exp <| neg_neg x theorem tendsto_exp_atBot_nhdsWithin : Tendsto exp atBot (𝓝[>] 0) := - tendsto_inf.2 ⟨tendsto_exp_atBot, tendsto_principal.2 <| eventually_of_forall exp_pos⟩ + tendsto_inf.2 ⟨tendsto_exp_atBot, tendsto_principal.2 <| Eventually.of_forall exp_pos⟩ @[simp] theorem isBoundedUnder_ge_exp_comp (l : Filter α) (f : α → ℝ) : @@ -387,7 +387,7 @@ theorem isLittleO_pow_exp_atTop {n : ℕ} : (fun x : ℝ => x ^ n) =o[atTop] Rea @[simp] theorem isBigO_exp_comp_exp_comp {f g : α → ℝ} : ((fun x => exp (f x)) =O[l] fun x => exp (g x)) ↔ IsBoundedUnder (· ≤ ·) l (f - g) := - Iff.trans (isBigO_iff_isBoundedUnder_le_div <| eventually_of_forall fun x => exp_ne_zero _) <| by + Iff.trans (isBigO_iff_isBoundedUnder_le_div <| Eventually.of_forall fun x => exp_ne_zero _) <| by simp only [norm_eq_abs, abs_exp, ← exp_sub, isBoundedUnder_le_exp_comp, Pi.sub_def] @[simp] diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean index 486e57156cc16..674fd91181743 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean @@ -350,7 +350,7 @@ theorem GammaSeq_tendsto_Gamma (s : ℂ) : Tendsto (GammaSeq s) atTop (𝓝 <| G rw [Nat.cast_zero, neg_zero] at hs rw [← Gamma_eq_GammaAux] · refine Tendsto.congr' ?_ (approx_Gamma_integral_tendsto_Gamma_integral hs) - refine (eventually_ne_atTop 0).mp (eventually_of_forall fun n hn => ?_) + refine (eventually_ne_atTop 0).mp (Eventually.of_forall fun n hn => ?_) exact (GammaSeq_eq_approx_Gamma_integral hs hn).symm · rwa [Nat.cast_zero, neg_lt_zero] · -- Induction step: use recurrence formulae in `s` for Gamma and GammaSeq @@ -358,7 +358,7 @@ theorem GammaSeq_tendsto_Gamma (s : ℂ) : Tendsto (GammaSeq s) atTop (𝓝 <| G rw [Nat.cast_succ, neg_add, ← sub_eq_add_neg, sub_lt_iff_lt_add, ← one_re, ← add_re] at hs rw [GammaAux] have := @Tendsto.congr' _ _ _ ?_ _ _ - ((eventually_ne_atTop 0).mp (eventually_of_forall fun n hn => ?_)) ((IH _ hs).div_const s) + ((eventually_ne_atTop 0).mp (Eventually.of_forall fun n hn => ?_)) ((IH _ hs).div_const s) pick_goal 3; · exact GammaSeq_add_one_left s hn -- doesn't work if inlined? conv at this => arg 1; intro n; rw [mul_comm] rwa [← mul_one (GammaAux m (s + 1) / s), tendsto_mul_iff_of_ne_zero _ (one_ne_zero' ℂ)] at this @@ -416,7 +416,7 @@ theorem Gamma_mul_Gamma_one_sub (z : ℂ) : Gamma z * Gamma (1 - z) = π / sin ( Complex.Gamma_neg_nat_eq_zero, mul_zero] refine tendsto_nhds_unique ((GammaSeq_tendsto_Gamma z).mul (GammaSeq_tendsto_Gamma <| 1 - z)) ?_ have : ↑π / sin (↑π * z) = 1 * (π / sin (π * z)) := by rw [one_mul] - convert Tendsto.congr' ((eventually_ne_atTop 0).mp (eventually_of_forall fun n hn => + convert Tendsto.congr' ((eventually_ne_atTop 0).mp (Eventually.of_forall fun n hn => (GammaSeq_mul z hn).symm)) (Tendsto.mul _ _) · convert tendsto_natCast_div_add_atTop (1 - z) using 1; ext1 n; rw [add_sub_assoc] · have : ↑π / sin (↑π * z) = 1 / (sin (π * z) / π) := by field_simp @@ -562,10 +562,10 @@ theorem Gamma_mul_Gamma_add_half (s : ℂ) : have h3 : Tendsto ((↑) : ℝ → ℂ) (𝓝[≠] 1) (𝓝[≠] 1) := by rw [tendsto_nhdsWithin_iff]; constructor · exact tendsto_nhdsWithin_of_tendsto_nhds continuous_ofReal.continuousAt - · exact eventually_nhdsWithin_iff.mpr (eventually_of_forall fun t ht => ofReal_ne_one.mpr ht) + · exact eventually_nhdsWithin_iff.mpr (Eventually.of_forall fun t ht => ofReal_ne_one.mpr ht) refine AnalyticOn.eq_of_frequently_eq h1 h2 (h3.frequently ?_) refine ((Eventually.filter_mono nhdsWithin_le_nhds) ?_).frequently - refine (eventually_gt_nhds zero_lt_one).mp (eventually_of_forall fun t ht => ?_) + refine (eventually_gt_nhds zero_lt_one).mp (Eventually.of_forall fun t ht => ?_) rw [← mul_inv, Gamma_ofReal, (by norm_num : (t : ℂ) + 1 / 2 = ↑(t + 1 / 2)), Gamma_ofReal, ← ofReal_mul, Gamma_mul_Gamma_add_half_of_pos ht, ofReal_mul, ofReal_mul, ← Gamma_ofReal, mul_inv, mul_inv, (by norm_num : 2 * (t : ℂ) = ↑(2 * t)), Gamma_ofReal, diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean index 62fae42a39804..7e05c6ddf0bc7 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/BohrMollerup.lean @@ -261,7 +261,7 @@ theorem tendsto_logGammaSeq (hf_conv : ConvexOn ℝ (Ioi 0) f) ∀ᶠ n : ℕ in atTop, logGammaSeq (x - 1) n = logGammaSeq x (n - 1) + x * (log (↑(n - 1) + 1) - log ↑(n - 1)) - log (x - 1) := by - refine Eventually.mp (eventually_ge_atTop 1) (eventually_of_forall fun n hn => ?_) + refine Eventually.mp (eventually_ge_atTop 1) (Eventually.of_forall fun n hn => ?_) have := logGammaSeq_add_one (x - 1) (n - 1) rw [sub_add_cancel, Nat.sub_add_cancel hn] at this rw [this] diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean index 4c44d7abdb5dd..a60fa19f4171c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian/FourierTransform.lean @@ -111,9 +111,9 @@ theorem tendsto_verticalIntegral (hb : 0 < b.re) (c : ℝ) : rw [tendsto_zero_iff_norm_tendsto_zero] refine tendsto_of_tendsto_of_tendsto_of_le_of_le' tendsto_const_nhds ?_ - (eventually_of_forall fun _ => norm_nonneg _) + (Eventually.of_forall fun _ => norm_nonneg _) ((eventually_ge_atTop (0 : ℝ)).mp - (eventually_of_forall fun T hT => verticalIntegral_norm_le hb c hT)) + (Eventually.of_forall fun T hT => verticalIntegral_norm_le hb c hT)) rw [(by ring : 0 = 2 * |c| * 0)] refine (tendsto_exp_atBot.comp (tendsto_neg_atTop_atBot.comp ?_)).const_mul _ apply tendsto_atTop_add_const_right diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean index e2a7aed0ba9c4..121367c2358ac 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean @@ -249,7 +249,7 @@ theorem continuousAt_gaussian_integral (b : ℂ) (hb : 0 < re b) : gcongr exact le_of_lt hc exact - continuousAt_of_dominated (eventually_of_forall f_meas) f_le_bd (integrable_exp_neg_mul_sq hd) + continuousAt_of_dominated (Eventually.of_forall f_meas) f_le_bd (integrable_exp_neg_mul_sq hd) (ae_of_all _ f_cts) theorem integral_gaussian_complex {b : ℂ} (hb : 0 < re b) : diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian/PoissonSummation.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian/PoissonSummation.lean index 9e3205f0be961..8347a8e9d6908 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian/PoissonSummation.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian/PoissonSummation.lean @@ -57,12 +57,12 @@ lemma cexp_neg_quadratic_isLittleO_abs_rpow_cocompact {a : ℂ} (ha : a.re < 0) rw [cocompact_eq_atBot_atTop, isLittleO_sup] constructor · refine ((cexp_neg_quadratic_isLittleO_rpow_atTop ha (-b) s).comp_tendsto - Filter.tendsto_neg_atBot_atTop).congr' (eventually_of_forall fun x ↦ ?_) ?_ + Filter.tendsto_neg_atBot_atTop).congr' (Eventually.of_forall fun x ↦ ?_) ?_ · simp only [neg_mul, Function.comp_apply, ofReal_neg, neg_sq, mul_neg, neg_neg] - · refine (eventually_lt_atBot 0).mp (eventually_of_forall fun x hx ↦ ?_) + · refine (eventually_lt_atBot 0).mp (Eventually.of_forall fun x hx ↦ ?_) simp only [Function.comp_apply, abs_of_neg hx] · refine (cexp_neg_quadratic_isLittleO_rpow_atTop ha b s).congr' EventuallyEq.rfl ?_ - refine (eventually_gt_atTop 0).mp (eventually_of_forall fun x hx ↦ ?_) + refine (eventually_gt_atTop 0).mp (Eventually.of_forall fun x hx ↦ ?_) simp_rw [abs_of_pos hx] theorem tendsto_rpow_abs_mul_exp_neg_mul_sq_cocompact {a : ℝ} (ha : 0 < a) (s : ℝ) : diff --git a/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean b/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean index 07c0aa8d4bd29..3ef729e02e9f1 100644 --- a/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/ImproperIntegrals.lean @@ -152,7 +152,7 @@ theorem integral_Ioi_cpow_of_lt {a : ℂ} (ha : a.re < -1) {c : ℝ} (hc : 0 < c suffices Tendsto (fun x : ℝ => ((x : ℂ) ^ (a + 1) - (c : ℂ) ^ (a + 1)) / (a + 1)) atTop (𝓝 <| -c ^ (a + 1) / (a + 1)) by - refine this.congr' ((eventually_gt_atTop 0).mp (eventually_of_forall fun x hx => ?_)) + refine this.congr' ((eventually_gt_atTop 0).mp (Eventually.of_forall fun x hx => ?_)) dsimp only rw [integral_cpow, id] refine Or.inr ⟨?_, not_mem_uIcc_of_lt hc hx⟩ @@ -164,7 +164,7 @@ theorem integral_Ioi_cpow_of_lt {a : ℂ} (ha : a.re < -1) {c : ℝ} (hc : 0 < c rw [tendsto_zero_iff_norm_tendsto_zero] refine (tendsto_rpow_neg_atTop (by linarith : 0 < -(a.re + 1))).congr' - ((eventually_gt_atTop 0).mp (eventually_of_forall fun x hx => ?_)) + ((eventually_gt_atTop 0).mp (Eventually.of_forall fun x hx => ?_)) simp_rw [neg_neg, Complex.norm_eq_abs, Complex.abs_cpow_eq_rpow_re_of_pos hx, Complex.add_re, Complex.one_re] diff --git a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean index 54269e1cd03fc..e596ded463560 100644 --- a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean +++ b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean @@ -95,14 +95,14 @@ theorem finite_integral_one_add_norm {r : ℝ} (hnr : (finrank ℝ E : ℝ) < r) -- We start by applying the layer cake formula have h_meas : Measurable fun ω : E => (1 + ‖ω‖) ^ (-r) := by fun_prop have h_pos : ∀ x : E, 0 ≤ (1 + ‖x‖) ^ (-r) := fun x ↦ by positivity - rw [lintegral_eq_lintegral_meas_le μ (eventually_of_forall h_pos) h_meas.aemeasurable] + rw [lintegral_eq_lintegral_meas_le μ (Eventually.of_forall h_pos) h_meas.aemeasurable] have h_int : ∀ t, 0 < t → μ {a : E | t ≤ (1 + ‖a‖) ^ (-r)} = μ (Metric.closedBall (0 : E) (t ^ (-r⁻¹) - 1)) := fun t ht ↦ by congr 1 ext x simp only [mem_setOf_eq, mem_closedBall_zero_iff] exact le_rpow_one_add_norm_iff_norm_le hr (mem_Ioi.mp ht) x - rw [setLIntegral_congr_fun measurableSet_Ioi (eventually_of_forall h_int)] + rw [setLIntegral_congr_fun measurableSet_Ioi (Eventually.of_forall h_int)] set f := fun t : ℝ ↦ μ (Metric.closedBall (0 : E) (t ^ (-r⁻¹) - 1)) set mB := μ (Metric.ball (0 : E) 1) -- the next two inequalities are in fact equalities but we don't need that @@ -142,7 +142,7 @@ theorem integrable_rpow_neg_one_add_norm_sq {r : ℝ} (hnr : (finrank ℝ E : Integrable (fun x ↦ ((1 : ℝ) + ‖x‖ ^ 2) ^ (-r / 2)) μ := by have hr : 0 < r := lt_of_le_of_lt (finrank ℝ E).cast_nonneg hnr refine ((integrable_one_add_norm hnr).const_mul <| (2 : ℝ) ^ (r / 2)).mono' - ?_ (eventually_of_forall fun x => ?_) + ?_ (Eventually.of_forall fun x => ?_) · apply Measurable.aestronglyMeasurable (by fun_prop) refine (abs_of_pos ?_).trans_le (rpow_neg_one_add_norm_sq_le x hr) positivity diff --git a/Mathlib/Analysis/SpecialFunctions/Polynomials.lean b/Mathlib/Analysis/SpecialFunctions/Polynomials.lean index 757dc664c8c4d..b69d21ca69f57 100644 --- a/Mathlib/Analysis/SpecialFunctions/Polynomials.lean +++ b/Mathlib/Analysis/SpecialFunctions/Polynomials.lean @@ -81,7 +81,7 @@ theorem abs_tendsto_atTop (hdeg : 0 < P.degree) : theorem abs_isBoundedUnder_iff : (IsBoundedUnder (· ≤ ·) atTop fun x => |eval x P|) ↔ P.degree ≤ 0 := by - refine ⟨fun h => ?_, fun h => ⟨|P.coeff 0|, eventually_map.mpr (eventually_of_forall + refine ⟨fun h => ?_, fun h => ⟨|P.coeff 0|, eventually_map.mpr (Eventually.of_forall (forall_imp (fun _ => le_of_eq) fun x => congr_arg abs <| _root_.trans (congr_arg (eval x) (eq_C_of_degree_le_zero h)) eval_C))⟩⟩ contrapose! h diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean index 0e75a747d5f90..a4b120ba2ef3c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean @@ -62,7 +62,7 @@ lemma tendsto_rpow_atTop_of_base_lt_one (b : ℝ) (hb₀ : -1 < b) (hb₁ : b < linarith case cos => rw [isBigO_iff] - exact ⟨1, eventually_of_forall fun x => by simp [Real.abs_cos_le_one]⟩ + exact ⟨1, Eventually.of_forall fun x => by simp [Real.abs_cos_le_one]⟩ case inr.inl => -- b = 0 refine Tendsto.mono_right ?_ (Iff.mpr pure_le_nhds_iff rfl) rw [tendsto_pure] diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean index 623700a7cbb97..d62254dfce817 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean @@ -441,10 +441,10 @@ theorem eventually_pow_one_div_le {x : ℝ≥0∞} (hx : x ≠ ∞) {y : ℝ≥0 ∀ᶠ n : ℕ in atTop, x ^ (1 / n : ℝ) ≤ y := by lift x to ℝ≥0 using hx by_cases h : y = ∞ - · exact eventually_of_forall fun n => h.symm ▸ le_top + · exact Eventually.of_forall fun n => h.symm ▸ le_top · lift y to ℝ≥0 using h have := NNReal.eventually_pow_one_div_le x (mod_cast hy : 1 < y) - refine this.congr (eventually_of_forall fun n => ?_) + refine this.congr (Eventually.of_forall fun n => ?_) rw [coe_rpow_of_nonneg x (by positivity : 0 ≤ (1 / n : ℝ)), coe_le_coe] private theorem continuousAt_rpow_const_of_pos {x : ℝ≥0∞} {y : ℝ} (h : 0 < y) : diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 903a92a213646..1292f0573acdb 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -99,7 +99,7 @@ statement simultaneously on `ℚ`, `ℝ` and `ℂ`. -/ theorem tendsto_natCast_div_add_atTop {𝕜 : Type*} [DivisionRing 𝕜] [TopologicalSpace 𝕜] [CharZero 𝕜] [Algebra ℝ 𝕜] [ContinuousSMul ℝ 𝕜] [TopologicalDivisionRing 𝕜] (x : 𝕜) : Tendsto (fun n : ℕ ↦ (n : 𝕜) / (n + x)) atTop (𝓝 1) := by - convert Tendsto.congr' ((eventually_ne_atTop 0).mp (eventually_of_forall fun n hn ↦ _)) _ + convert Tendsto.congr' ((eventually_ne_atTop 0).mp (Eventually.of_forall fun n hn ↦ _)) _ · exact fun n : ℕ ↦ 1 / (1 + x / n) · field_simp [Nat.cast_ne_zero.mpr hn] · have : 𝓝 (1 : 𝕜) = 𝓝 (1 / (1 + x * (0 : 𝕜))) := by @@ -167,7 +167,7 @@ theorem tendsto_pow_atTop_nhdsWithin_zero_of_lt_one {𝕜 : Type*} [LinearOrdere Tendsto (fun n : ℕ ↦ r ^ n) atTop (𝓝[>] 0) := tendsto_inf.2 ⟨tendsto_pow_atTop_nhds_zero_of_lt_one h₁.le h₂, - tendsto_principal.2 <| eventually_of_forall fun _ ↦ pow_pos h₁ _⟩ + tendsto_principal.2 <| Eventually.of_forall fun _ ↦ pow_pos h₁ _⟩ @[deprecated (since := "2024-01-31")] alias tendsto_pow_atTop_nhdsWithin_0_of_lt_1 := tendsto_pow_atTop_nhdsWithin_zero_of_lt_one @@ -605,7 +605,7 @@ theorem tendsto_factorial_div_pow_self_atTop : Tendsto (fun n ↦ n ! / (n : ℝ) ^ n : ℕ → ℝ) atTop (𝓝 0) := tendsto_of_tendsto_of_tendsto_of_le_of_le' tendsto_const_nhds (tendsto_const_div_atTop_nhds_zero_nat 1) - (eventually_of_forall fun n ↦ + (Eventually.of_forall fun n ↦ div_nonneg (mod_cast n.factorial_pos.le) (pow_nonneg (mod_cast n.zero_le) _)) (by diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index 21ee99951f714..549efc80b228a 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -624,7 +624,7 @@ theorem Monotone.cauchySeq_series_mul_of_tendsto_zero_of_bounded (hfa : Monotone simp_rw [Finset.sum_range_by_parts _ _ (Nat.succ _), sub_eq_add_neg, Nat.succ_sub_succ_eq_sub, tsub_zero] apply (NormedField.tendsto_zero_smul_of_tendsto_zero_of_bounded hf0 - ⟨b, eventually_map.mpr <| eventually_of_forall fun n ↦ hgb <| n + 1⟩).cauchySeq.add + ⟨b, eventually_map.mpr <| Eventually.of_forall fun n ↦ hgb <| n + 1⟩).cauchySeq.add refine CauchySeq.neg ?_ refine cauchySeq_range_of_norm_bounded _ ?_ (fun n ↦ ?_ : ∀ n, ‖(f (n + 1) + -f n) • (Finset.range (n + 1)).sum z‖ ≤ b * |f (n + 1) - f n|) diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index d25df9001edd4..afe58c1fc4266 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -399,7 +399,7 @@ lemma isLittleO_deriv_smoothingFn : deriv ε =o[atTop] fun x => x⁻¹ := calc rw [isLittleO_one_left_iff] exact Tendsto.comp tendsto_norm_atTop_atTop <| Tendsto.comp (tendsto_pow_atTop (by norm_num)) tendsto_log_atTop - · exact Filter.eventually_of_forall (fun x hx => by rw [mul_one] at hx; simp [hx]) + · exact Filter.Eventually.of_forall (fun x hx => by rw [mul_one] at hx; simp [hx]) _ = fun x => x⁻¹ := by simp lemma eventually_deriv_one_sub_smoothingFn : @@ -843,7 +843,7 @@ lemma isEquivalent_deriv_rpow_p_mul_one_sub_smoothingFn {p : ℝ} (hp : p ≠ 0) (fun z => z ^ (p-1) / (log z ^ 2)) =o[atTop] fun z => z ^ (p-1) / 1 := by simp_rw [div_eq_mul_inv] refine IsBigO.mul_isLittleO (isBigO_refl _ _) - (IsLittleO.inv_rev ?_ (by aesop (add safe eventually_of_forall))) + (IsLittleO.inv_rev ?_ (by aesop (add safe Eventually.of_forall))) rw [isLittleO_const_left] refine Or.inr <| Tendsto.comp tendsto_norm_atTop_atTop ?_ exact Tendsto.comp (g := fun z => z ^ 2) @@ -867,7 +867,7 @@ lemma isEquivalent_deriv_rpow_p_mul_one_add_smoothingFn {p : ℝ} (hp : p ≠ 0) (fun z => -(z ^ (p-1) / (log z ^ 2))) =o[atTop] fun z => z ^ (p-1) / 1 := by simp_rw [isLittleO_neg_left, div_eq_mul_inv] refine IsBigO.mul_isLittleO (isBigO_refl _ _) - (IsLittleO.inv_rev ?_ (by aesop (add safe eventually_of_forall))) + (IsLittleO.inv_rev ?_ (by aesop (add safe Eventually.of_forall))) rw [isLittleO_const_left] refine Or.inr <| Tendsto.comp tendsto_norm_atTop_atTop ?_ exact Tendsto.comp (g := fun z => z ^ 2) diff --git a/Mathlib/Dynamics/OmegaLimit.lean b/Mathlib/Dynamics/OmegaLimit.lean index 898b206987a20..61f82cc579837 100644 --- a/Mathlib/Dynamics/OmegaLimit.lean +++ b/Mathlib/Dynamics/OmegaLimit.lean @@ -93,7 +93,7 @@ theorem mapsTo_omegaLimit {α' β' : Type*} [TopologicalSpace β'] {f : Filter {ϕ' : τ → α' → β'} {ga : α → α'} {s' : Set α'} (hs : MapsTo ga s s') {gb : β → β'} (hg : ∀ t x, gb (ϕ t x) = ϕ' t (ga x)) (hgc : Continuous gb) : MapsTo gb (ω f ϕ s) (ω f ϕ' s') := - mapsTo_omegaLimit' _ hs (eventually_of_forall fun t x _hx ↦ hg t x) hgc + mapsTo_omegaLimit' _ hs (Eventually.of_forall fun t x _hx ↦ hg t x) hgc theorem omegaLimit_image_eq {α' : Type*} (ϕ : τ → α' → β) (f : Filter τ) (g : α → α') : ω f ϕ (g '' s) = ω f (fun t x ↦ ϕ t (g x)) s := by simp only [omegaLimit, image2_image_right] diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean index dd5807aa62123..99fb114ff5bf1 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean @@ -105,7 +105,7 @@ theorem measurable_limit_of_tendsto_metrizable_ae {ι} [Countable ι] [Nonempty classical inhabit ι rcases eq_or_neBot L with (rfl | hL) - · exact ⟨(hf default).mk _, (hf default).measurable_mk, eventually_of_forall fun x => tendsto_bot⟩ + · exact ⟨(hf default).mk _, (hf default).measurable_mk, Eventually.of_forall fun x => tendsto_bot⟩ let p : α → (ι → β) → Prop := fun x f' => ∃ l : β, Tendsto (fun n => f' n) L (𝓝 l) have hp_mem : ∀ x ∈ aeSeqSet hf p, p x fun n => f n x := fun x hx => aeSeq.fun_prop_of_mem_aeSeqSet hf hx diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean index e219a86ee54da..febc262889c89 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean @@ -449,9 +449,9 @@ theorem measure_ae_null_of_prod_null {s : Set (α × β)} (h : μ.prod ν s = 0) rw [measure_prod_null mt] at ht rw [eventuallyLE_antisymm_iff] exact - ⟨EventuallyLE.trans_eq (eventually_of_forall fun x => (measure_mono (preimage_mono hst) : _)) + ⟨EventuallyLE.trans_eq (Eventually.of_forall fun x => (measure_mono (preimage_mono hst) : _)) ht, - eventually_of_forall fun x => zero_le _⟩ + Eventually.of_forall fun x => zero_le _⟩ theorem AbsolutelyContinuous.prod [SFinite ν'] (h1 : μ ≪ μ') (h2 : ν ≪ ν') : μ.prod ν ≪ μ'.prod ν' := by diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean b/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean index 18a596110f89b..fd2189d76b79a 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Integral.lean @@ -104,12 +104,12 @@ theorem MeasureTheory.StronglyMeasurable.integral_prod_right [SFinite ν] ⦃f : refine tendsto_integral_of_dominated_convergence (fun y => ‖f x y‖ + ‖f x y‖) (fun n => (s' n x).aestronglyMeasurable) (hfx.norm.add hfx.norm) ?_ ?_ - · refine fun n => eventually_of_forall fun y => + · refine fun n => Eventually.of_forall fun y => SimpleFunc.norm_approxOn_zero_le ?_ ?_ (x, y) n -- Porting note: Lean 3 solved the following two subgoals on its own · exact hf.measurable · simp - · refine eventually_of_forall fun y => SimpleFunc.tendsto_approxOn ?_ ?_ ?_ + · refine Eventually.of_forall fun y => SimpleFunc.tendsto_approxOn ?_ ?_ ?_ -- Porting note: Lean 3 solved the following two subgoals on its own · exact hf.measurable.of_uncurry_left · simp @@ -287,7 +287,7 @@ theorem Integrable.prod_smul {𝕜 : Type*} [NontriviallyNormedField 𝕜] [Norm Integrable (fun z : α × β => f z.1 • g z.2) (μ.prod ν) := by refine (integrable_prod_iff ?_).2 ⟨?_, ?_⟩ · exact hf.1.fst.smul hg.1.snd - · exact eventually_of_forall fun x => hg.smul (f x) + · exact Eventually.of_forall fun x => hg.smul (f x) · simpa only [norm_smul, integral_mul_left] using hf.norm.mul_const _ theorem Integrable.prod_mul {L : Type*} [RCLike L] {f : α → L} {g : β → L} (hf : Integrable f μ) @@ -301,11 +301,11 @@ variable [NormedSpace ℝ E] theorem Integrable.integral_prod_left ⦃f : α × β → E⦄ (hf : Integrable f (μ.prod ν)) : Integrable (fun x => ∫ y, f (x, y) ∂ν) μ := Integrable.mono hf.integral_norm_prod_left hf.aestronglyMeasurable.integral_prod_right' <| - eventually_of_forall fun x => + Eventually.of_forall fun x => (norm_integral_le_integral_norm _).trans_eq <| (norm_of_nonneg <| integral_nonneg_of_ae <| - eventually_of_forall fun y => (norm_nonneg (f (x, y)) : _)).symm + Eventually.of_forall fun y => (norm_nonneg (f (x, y)) : _)).symm theorem Integrable.integral_prod_right [SFinite μ] ⦃f : α × β → E⦄ (hf : Integrable f (μ.prod ν)) : Integrable (fun y => ∫ x, f (x, y) ∂μ) ν := @@ -388,7 +388,7 @@ theorem continuous_integral_integral : rw [continuous_iff_continuousAt]; intro g refine tendsto_integral_of_L1 _ (L1.integrable_coeFn g).integral_prod_left - (eventually_of_forall fun h => (L1.integrable_coeFn h).integral_prod_left) ?_ + (Eventually.of_forall fun h => (L1.integrable_coeFn h).integral_prod_left) ?_ simp_rw [← lintegral_fn_integral_sub (fun x => (‖x‖₊ : ℝ≥0∞)) (L1.integrable_coeFn _) (L1.integrable_coeFn g)] @@ -535,7 +535,7 @@ lemma integral_integral_swap_of_hasCompactSupport apply (integrableOn_iff_integrable_of_support_subset (subset_tsupport f.uncurry)).mp refine ⟨(h'f.stronglyMeasurable_of_prod hf).aestronglyMeasurable, ?_⟩ obtain ⟨C, hC⟩ : ∃ C, ∀ p, ‖f.uncurry p‖ ≤ C := hf.bounded_above_of_compact_support h'f - exact hasFiniteIntegral_of_bounded (C := C) (eventually_of_forall hC) + exact hasFiniteIntegral_of_bounded (C := C) (Eventually.of_forall hC) _ = ∫ y, (∫ x in U, f x y ∂μ) ∂ν := by apply setIntegral_eq_integral_of_forall_compl_eq_zero (fun y hy ↦ ?_) have : ∀ x, f x y = 0 := by diff --git a/Mathlib/MeasureTheory/Covering/Differentiation.lean b/Mathlib/MeasureTheory/Covering/Differentiation.lean index cb24a16d2a1c8..abe8892caee43 100644 --- a/Mathlib/MeasureTheory/Covering/Differentiation.lean +++ b/Mathlib/MeasureTheory/Covering/Differentiation.lean @@ -539,7 +539,7 @@ theorem withDensity_le_mul {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht have M : MeasurableSet (s ∩ f ⁻¹' {0}) := hs.inter (f_meas (measurableSet_singleton _)) simp only [ν, nonpos_iff_eq_zero, M, withDensity_apply, lintegral_eq_zero_iff f_meas] apply (ae_restrict_iff' M).2 - exact eventually_of_forall fun x hx => hx.2 + exact Eventually.of_forall fun x hx => hx.2 have B : ν (s ∩ f ⁻¹' {∞}) ≤ ((t : ℝ≥0∞) ^ 2 • ρ :) (s ∩ f ⁻¹' {∞}) := by apply le_trans (le_of_eq _) (zero_le _) apply withDensity_absolutelyContinuous μ _ @@ -555,7 +555,7 @@ theorem withDensity_le_mul {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht simp only [ν, M, withDensity_apply, coe_nnreal_smul_apply] calc (∫⁻ x in s ∩ f ⁻¹' I, f x ∂μ) ≤ ∫⁻ _ in s ∩ f ⁻¹' I, (t : ℝ≥0∞) ^ (n + 1) ∂μ := - lintegral_mono_ae ((ae_restrict_iff' M).2 (eventually_of_forall fun x hx => hx.2.2.le)) + lintegral_mono_ae ((ae_restrict_iff' M).2 (Eventually.of_forall fun x hx => hx.2.2.le)) _ = (t : ℝ≥0∞) ^ (n + 1) * μ (s ∩ f ⁻¹' I) := by simp only [lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] _ = (t : ℝ≥0∞) ^ (2 : ℤ) * ((t : ℝ≥0∞) ^ (n - 1) * μ (s ∩ f ⁻¹' I)) := by @@ -626,7 +626,7 @@ theorem le_mul_withDensity {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht _ = ∫⁻ _ in s ∩ f ⁻¹' I, (t : ℝ≥0∞) ^ (n + 1) ∂μ := by simp only [lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] _ ≤ ∫⁻ x in s ∩ f ⁻¹' I, t * f x ∂μ := by - apply lintegral_mono_ae ((ae_restrict_iff' M).2 (eventually_of_forall fun x hx => ?_)) + apply lintegral_mono_ae ((ae_restrict_iff' M).2 (Eventually.of_forall fun x hx => ?_)) rw [add_comm, ENNReal.zpow_add t_ne_zero ENNReal.coe_ne_top, zpow_one] exact mul_le_mul_left' hx.2.1 _ _ = t * ∫⁻ x in s ∩ f ⁻¹' I, f x ∂μ := lintegral_const_mul _ f_meas @@ -875,7 +875,7 @@ theorem ae_tendsto_lintegral_nnnorm_sub_div {f : α → E} (hf : LocallyIntegrab filter_upwards [v.eventually_filterAt_subset_of_nhds ((u_open n).mem_nhds hn), v.eventually_filterAt_measurableSet x] with a ha h'a congr 1 - refine setLIntegral_congr_fun h'a (eventually_of_forall (fun y hy ↦ ?_)) + refine setLIntegral_congr_fun h'a (Eventually.of_forall (fun y hy ↦ ?_)) rw [indicator_of_mem (ha hy) f, indicator_of_mem hn f] /-- *Lebesgue differentiation theorem*: for almost every point `x`, the @@ -902,7 +902,7 @@ theorem ae_tendsto_average [NormedSpace ℝ E] [CompleteSpace E] {f : α → E} ∀ᵐ x ∂μ, Tendsto (fun a => ⨍ y in a, f y ∂μ) (v.filterAt x) (𝓝 (f x)) := by filter_upwards [v.ae_tendsto_average_norm_sub hf, v.ae_eventually_measure_pos] with x hx h'x rw [tendsto_iff_norm_sub_tendsto_zero] - refine squeeze_zero' (eventually_of_forall fun a => norm_nonneg _) ?_ hx + refine squeeze_zero' (Eventually.of_forall fun a => norm_nonneg _) ?_ hx filter_upwards [h'x, v.eventually_measure_lt_top x, v.eventually_filterAt_integrableOn x hf] with a ha h'a h''a nth_rw 1 [← setAverage_const ha.ne' h'a.ne (f x)] diff --git a/Mathlib/MeasureTheory/Covering/LiminfLimsup.lean b/Mathlib/MeasureTheory/Covering/LiminfLimsup.lean index c4bb3f2085e98..539f24411b7c9 100644 --- a/Mathlib/MeasureTheory/Covering/LiminfLimsup.lean +++ b/Mathlib/MeasureTheory/Covering/LiminfLimsup.lean @@ -105,7 +105,7 @@ theorem blimsup_cthickening_ae_le_of_eventually_mul_le_aux (p : ℕ → Prop) {s obtain ⟨η, hη, hη'⟩ := this replace hη' : 1 ≤ η := by simpa only [ENNReal.one_le_coe_iff] using - le_of_tendsto (hd' w (fun j => r₁ (f j)) hr <| eventually_of_forall hw') hη' + le_of_tendsto (hd' w (fun j => r₁ (f j)) hr <| Eventually.of_forall hw') hη' exact (lt_self_iff_false _).mp (lt_of_lt_of_le hη hη') refine ⟨1 - C⁻¹, tsub_lt_self zero_lt_one (inv_pos.mpr hC), ?_⟩ replace hC : C ≠ 0 := ne_of_gt hC @@ -198,13 +198,13 @@ theorem blimsup_cthickening_mul_ae_eq (p : ℕ → Prop) (s : ℕ → Set α) {M convert TendstoNhdsWithinIoi.const_mul hM hr <;> simp only [mul_zero] refine eventuallyLE_antisymm_iff.mpr ⟨?_, ?_⟩ · exact blimsup_cthickening_ae_le_of_eventually_mul_le μ p (inv_pos.mpr hM) hr' - (eventually_of_forall fun i => by rw [inv_mul_cancel_left₀ hM.ne' (r i)]) + (Eventually.of_forall fun i => by rw [inv_mul_cancel_left₀ hM.ne' (r i)]) · exact blimsup_cthickening_ae_le_of_eventually_mul_le μ p hM hr - (eventually_of_forall fun i => le_refl _) + (Eventually.of_forall fun i => le_refl _) let r' : ℕ → ℝ := fun i => if 0 < r i then r i else 1 / ((i : ℝ) + 1) have hr' : Tendsto r' atTop (𝓝[>] 0) := by refine tendsto_nhdsWithin_iff.mpr - ⟨Tendsto.if' hr tendsto_one_div_add_atTop_nhds_zero_nat, eventually_of_forall fun i => ?_⟩ + ⟨Tendsto.if' hr tendsto_one_div_add_atTop_nhds_zero_nat, Eventually.of_forall fun i => ?_⟩ by_cases hi : 0 < r i · simp [r', hi] · simp only [r', hi, one_div, mem_Ioi, if_false, inv_pos]; positivity @@ -220,8 +220,8 @@ theorem blimsup_cthickening_mul_ae_eq (p : ℕ → Prop) (s : ℕ → Set α) {M ext i; simp [← and_or_left, lt_or_le 0 (r i)] rw [hp, blimsup_or_eq_sup, blimsup_or_eq_sup] simp only [sup_eq_union] - rw [blimsup_congr (eventually_of_forall h₀), blimsup_congr (eventually_of_forall h₁), - blimsup_congr (eventually_of_forall h₂)] + rw [blimsup_congr (Eventually.of_forall h₀), blimsup_congr (Eventually.of_forall h₁), + blimsup_congr (Eventually.of_forall h₂)] exact ae_eq_set_union (this (fun i => p i ∧ 0 < r i) hr') (ae_eq_refl _) theorem blimsup_cthickening_ae_eq_blimsup_thickening {p : ℕ → Prop} {s : ℕ → Set α} {r : ℕ → ℝ} @@ -268,14 +268,14 @@ theorem blimsup_thickening_mul_ae_eq (p : ℕ → Prop) (s : ℕ → Set α) {M let q : ℕ → Prop := fun i => p i ∧ 0 < r i have h₁ : blimsup (fun i => thickening (r i) (s i)) atTop p = blimsup (fun i => thickening (r i) (s i)) atTop q := by - refine blimsup_congr' (eventually_of_forall fun i h => ?_) + refine blimsup_congr' (Eventually.of_forall fun i h => ?_) replace hi : 0 < r i := by contrapose! h; apply thickening_of_nonpos h simp only [q, hi, iff_self_and, imp_true_iff] have h₂ : blimsup (fun i => thickening (M * r i) (s i)) atTop p = blimsup (fun i => thickening (M * r i) (s i)) atTop q := by - refine blimsup_congr' (eventually_of_forall fun i h ↦ ?_) + refine blimsup_congr' (Eventually.of_forall fun i h ↦ ?_) replace h : 0 < r i := by rw [← mul_pos_iff_of_pos_left hM]; contrapose! h; apply thickening_of_nonpos h simp only [q, h, iff_self_and, imp_true_iff] rw [h₁, h₂] - exact blimsup_thickening_mul_ae_eq_aux μ q s hM r hr (eventually_of_forall fun i hi => hi.2) + exact blimsup_thickening_mul_ae_eq_aux μ q s hM r hr (Eventually.of_forall fun i hi => hi.2) diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index 5cc54ff5324a5..4842fd3b46f30 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -824,9 +824,9 @@ theorem haveLebesgueDecomposition_of_finiteMeasure [IsFiniteMeasure μ] [IsFinit refine Measurable.aemeasurable ?_ convert (iSup_mem_measurableLE _ hf₁ n).1 simp - · refine Filter.eventually_of_forall fun a ↦ ?_ + · refine Filter.Eventually.of_forall fun a ↦ ?_ simp [iSup_monotone' f _] - · refine Filter.eventually_of_forall fun a ↦ ?_ + · refine Filter.Eventually.of_forall fun a ↦ ?_ simp [tendsto_atTop_iSup (iSup_monotone' f a)] have hξm : Measurable ξ := by convert measurable_iSup fun n ↦ (iSup_mem_measurableLE _ hf₁ n).1 diff --git a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean index 13feeda01ee5a..054d7829fab2f 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean @@ -372,7 +372,7 @@ theorem exists_compl_positive_negative (s : SignedMeasure α) : have hA₂ : s ≤[A] 0 := restrict_le_restrict_iUnion _ _ hB₁ hB₂ have hA₃ : s A = sInf s.measureOfNegatives := by apply le_antisymm - · refine le_of_tendsto_of_tendsto tendsto_const_nhds hf₂ (eventually_of_forall fun n => ?_) + · refine le_of_tendsto_of_tendsto tendsto_const_nhds hf₂ (Eventually.of_forall fun n => ?_) rw [← (hB n).2, hA, ← Set.diff_union_of_subset (Set.subset_iUnion _ n), of_union Set.disjoint_sdiff_left _ (hB₁ n)] · refine add_le_of_nonpos_left ?_ diff --git a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean index 6c90b652e93b4..75e3cb37ba203 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean @@ -107,7 +107,7 @@ theorem ae_eq_zero_of_forall_dual [NormedAddCommGroup E] [NormedSpace 𝕜 E] [SecondCountableTopology E] {f : α → E} (hf : ∀ c : Dual 𝕜 E, (fun x => ⟪f x, c⟫) =ᵐ[μ] 0) : f =ᵐ[μ] 0 := ae_eq_zero_of_forall_dual_of_isSeparable 𝕜 (.of_separableSpace Set.univ) hf - (eventually_of_forall fun _ => Set.mem_univ _) + (Eventually.of_forall fun _ => Set.mem_univ _) variable {𝕜} @@ -260,7 +260,7 @@ theorem ae_nonneg_of_forall_setIntegral_nonneg (hf : Integrable f μ) have h_const_le : (∫ x in s, f x ∂μ) ≤ ∫ _ in s, b ∂μ := by refine setIntegral_mono_ae_restrict hf.integrableOn (integrableOn_const.mpr (Or.inr mus)) ?_ rw [EventuallyLE, ae_restrict_iff₀ (hs.mono μ.restrict_le_self)] - exact eventually_of_forall fun x hxs => hxs + exact Eventually.of_forall fun x hxs => hxs rwa [setIntegral_const, smul_eq_mul, mul_comm] at h_const_le contrapose! h_int_gt with H calc @@ -523,7 +523,7 @@ theorem ae_eq_zero_of_forall_setIntegral_eq_of_finStronglyMeasurable_trim (hm : haveI : SigmaFinite ((μ.restrict t).trim hm) := by rwa [restrict_trim hm μ ht_meas] at htμ have htf_zero : f =ᵐ[μ.restrict tᶜ] 0 := by rw [EventuallyEq, ae_restrict_iff' (MeasurableSet.compl (hm _ ht_meas))] - exact eventually_of_forall htf_zero + exact Eventually.of_forall htf_zero have hf_meas_m : StronglyMeasurable[m] f := hf.stronglyMeasurable suffices f =ᵐ[μ.restrict t] 0 from ae_of_ae_restrict_of_ae_restrict_compl _ this htf_zero diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean index 2f9b8035e0291..9aeca60eb12fa 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/AEMeasurable.lean @@ -379,7 +379,7 @@ theorem lpMeasSubgroupToLpTrim_neg (hm : m ≤ m0) (f : lpMeasSubgroup F m p μ) refine EventuallyEq.trans ?_ (EventuallyEq.neg (lpMeasSubgroupToLpTrim_ae_eq hm f).symm) refine (Lp.coeFn_neg _).trans ?_ simp_rw [lpMeasSubgroup_coe] - exact eventually_of_forall fun x => by rfl + exact Eventually.of_forall fun x => by rfl theorem lpMeasSubgroupToLpTrim_sub (hm : m ≤ m0) (f g : lpMeasSubgroup F m p μ) : lpMeasSubgroupToLpTrim F p μ hm (f - g) = diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean index 5feff3cfacb85..04cb417cd938d 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Basic.lean @@ -138,7 +138,7 @@ theorem condexp_ae_eq_condexpL1 (hm : m ≤ m0) [hμm : SigmaFinite (μ.trim hm) theorem condexp_ae_eq_condexpL1CLM (hm : m ≤ m0) [SigmaFinite (μ.trim hm)] (hf : Integrable f μ) : μ[f|m] =ᵐ[μ] condexpL1CLM F' hm μ (hf.toL1 f) := by - refine (condexp_ae_eq_condexpL1 hm f).trans (eventually_of_forall fun x => ?_) + refine (condexp_ae_eq_condexpL1 hm f).trans (Eventually.of_forall fun x => ?_) rw [condexpL1_eq hf] theorem condexp_undef (hf : ¬Integrable f μ) : μ[f|m] = 0 := by @@ -256,7 +256,7 @@ theorem condexp_bot_ae_eq (f : α → F') : μ[f|⊥] =ᵐ[μ] fun _ => (μ Set.univ).toReal⁻¹ • ∫ x, f x ∂μ := by rcases eq_zero_or_neZero μ with rfl | hμ · rw [ae_zero]; exact eventually_bot - · exact eventually_of_forall <| congr_fun (condexp_bot' f) + · exact Eventually.of_forall <| congr_fun (condexp_bot' f) theorem condexp_bot [IsProbabilityMeasure μ] (f : α → F') : μ[f|⊥] = fun _ => ∫ x, f x ∂μ := by refine (condexp_bot' f).trans ?_; rw [measure_univ, ENNReal.one_toReal, inv_one, one_smul] @@ -385,6 +385,6 @@ theorem tendsto_condexp_unique (fs gs : ℕ → α → F') (f g : α → F') have hcond_gs : Tendsto (fun n => condexpL1 hm μ (gs n)) atTop (𝓝 (condexpL1 hm μ g)) := tendsto_condexpL1_of_dominated_convergence hm _ (fun n => (hgs_int n).1) h_int_bound_gs hgs_bound hgs - exact tendsto_nhds_unique_of_eventuallyEq hcond_gs hcond_fs (eventually_of_forall hn_eq) + exact tendsto_nhds_unique_of_eventuallyEq hcond_gs hcond_fs (Eventually.of_forall hn_eq) end MeasureTheory diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean index 6980909f8a5ae..b900dbece2f95 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean @@ -93,7 +93,7 @@ theorem condexpIndL1Fin_add (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x y : refine EventuallyEq.trans ?_ (EventuallyEq.add (Memℒp.coeFn_toLp q).symm (Memℒp.coeFn_toLp q).symm) rw [condexpIndSMul_add] - refine (Lp.coeFn_add _ _).trans (eventually_of_forall fun a => ?_) + refine (Lp.coeFn_add _ _).trans (Eventually.of_forall fun a => ?_) rfl theorem condexpIndL1Fin_smul (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (c : ℝ) (x : G) : diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean index 3b2163fa2524d..44d7dd73837dd 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean @@ -127,12 +127,12 @@ theorem setIntegral_abs_condexp_le {s : Set α} (hs : MeasurableSet[m] s) (f : refine integral_congr_ae ?_ have : (fun x => |(μ[s.indicator f|m]) x|) =ᵐ[μ] fun x => |s.indicator (μ[f|m]) x| := (condexp_indicator hfint hs).fun_comp abs - refine EventuallyEq.trans (eventually_of_forall fun x => ?_) this.symm + refine EventuallyEq.trans (Eventually.of_forall fun x => ?_) this.symm rw [← Real.norm_eq_abs, norm_indicator_eq_indicator_norm] simp only [Real.norm_eq_abs] rw [this, ← integral_indicator (hnm _ hs)] refine (integral_abs_condexp_le _).trans - (le_of_eq <| integral_congr_ae <| eventually_of_forall fun x => ?_) + (le_of_eq <| integral_congr_ae <| Eventually.of_forall fun x => ?_) simp_rw [← Real.norm_eq_abs, norm_indicator_eq_indicator_norm] @[deprecated (since := "2024-04-17")] @@ -145,7 +145,7 @@ theorem ae_bdd_condexp_of_ae_bdd {R : ℝ≥0} {f : α → ℝ} (hbdd : ∀ᵐ x by_cases hnm : m ≤ m0 swap · simp_rw [condexp_of_not_le hnm, Pi.zero_apply, abs_zero] - exact eventually_of_forall fun _ => R.coe_nonneg + exact Eventually.of_forall fun _ => R.coe_nonneg by_cases hfint : Integrable f μ swap · simp_rw [condexp_undef hfint] @@ -268,19 +268,19 @@ theorem condexp_stronglyMeasurable_mul_of_bound (hm : m ≤ m0) [IsFiniteMeasure refine tendsto_condexp_unique (fun n x => fs n x * g x) (fun n x => fs n x * (μ[g|m]) x) (f * g) (f * μ[g|m]) ?_ ?_ ?_ ?_ (c * ‖g ·‖) ?_ (c * ‖(μ[g|m]) ·‖) ?_ ?_ ?_ ?_ · exact fun n => hg.bdd_mul' ((SimpleFunc.stronglyMeasurable (fs n)).mono hm).aestronglyMeasurable - (eventually_of_forall (hfs_bound n)) + (Eventually.of_forall (hfs_bound n)) · exact fun n => integrable_condexp.bdd_mul' ((SimpleFunc.stronglyMeasurable (fs n)).mono hm).aestronglyMeasurable - (eventually_of_forall (hfs_bound n)) + (Eventually.of_forall (hfs_bound n)) · filter_upwards [hfs_tendsto] with x hx exact hx.mul tendsto_const_nhds · filter_upwards [hfs_tendsto] with x hx exact hx.mul tendsto_const_nhds · exact hg.norm.const_mul c · exact integrable_condexp.norm.const_mul c - · refine fun n => eventually_of_forall fun x => ?_ + · refine fun n => Eventually.of_forall fun x => ?_ exact (norm_mul_le _ _).trans (mul_le_mul_of_nonneg_right (hfs_bound n x) (norm_nonneg _)) - · refine fun n => eventually_of_forall fun x => ?_ + · refine fun n => Eventually.of_forall fun x => ?_ exact (norm_mul_le _ _).trans (mul_le_mul_of_nonneg_right (hfs_bound n x) (norm_nonneg _)) · intro n simp_rw [← Pi.mul_apply] @@ -289,7 +289,7 @@ theorem condexp_stronglyMeasurable_mul_of_bound (hm : m ≤ m0) [IsFiniteMeasure ((SimpleFunc.stronglyMeasurable _).mul stronglyMeasurable_condexp) _] exact integrable_condexp.bdd_mul' ((SimpleFunc.stronglyMeasurable (fs n)).mono hm).aestronglyMeasurable - (eventually_of_forall (hfs_bound n)) + (Eventually.of_forall (hfs_bound n)) theorem condexp_stronglyMeasurable_mul_of_bound₀ (hm : m ≤ m0) [IsFiniteMeasure μ] {f g : α → ℝ} (hf : AEStronglyMeasurable' m f μ) (hg : Integrable g μ) (c : ℝ) diff --git a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean index 40edbfe36294c..1d97cdeb360fd 100644 --- a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean +++ b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean @@ -78,7 +78,7 @@ protected theorem congr' (h_left : ∀ᶠ i in l, f i =ᵐ[μ] f' i) (h_right : protected theorem congr (h_left : ∀ i, f i =ᵐ[μ] f' i) (h_right : g =ᵐ[μ] g') (h_tendsto : TendstoInMeasure μ f l g) : TendstoInMeasure μ f' l g' := - TendstoInMeasure.congr' (eventually_of_forall h_left) h_right h_tendsto + TendstoInMeasure.congr' (Eventually.of_forall h_left) h_right h_tendsto theorem congr_left (h : ∀ i, f i =ᵐ[μ] f' i) (h_tendsto : TendstoInMeasure μ f l g) : TendstoInMeasure μ f' l g := diff --git a/Mathlib/MeasureTheory/Function/Intersectivity.lean b/Mathlib/MeasureTheory/Function/Intersectivity.lean index a8aa6fa0f1d78..a9113bf78d52b 100644 --- a/Mathlib/MeasureTheory/Function/Intersectivity.lean +++ b/Mathlib/MeasureTheory/Function/Intersectivity.lean @@ -83,7 +83,7 @@ lemma bergelson' {s : ℕ → Set α} (hs : ∀ n, MeasurableSet (s n)) (hr₀ : have : ∫⁻ x, limsup (f · x) atTop ∂μ ≤ μ univ := by rw [← lintegral_one] exact lintegral_mono fun a ↦ limsup_le_of_le ⟨0, fun R _ ↦ bot_le⟩ $ - eventually_of_forall fun n ↦ hf₁ _ _ + Eventually.of_forall fun n ↦ hf₁ _ _ -- By the first moment method, there exists some `x ∉ N` such that `limsup f n x` is at least `r`. obtain ⟨x, hxN, hx⟩ := exists_not_mem_null_laverage_le hμ (ne_top_of_le_ne_top (measure_ne_top μ univ) this) hN₀ @@ -107,7 +107,7 @@ lemma bergelson' {s : ℕ → Set α} (hs : ∀ n, MeasurableSet (s n)) (hr₀ : using Finset.card_le_card fun m hm ↦ hxs.mem_toFinset.2 (Finset.mem_filter.1 hm).2 · simp_rw [← hu.mem_toFinset] exact hN₁ _ ⟨x, mem_iInter₂.2 fun n hn ↦ hux $ hu.mem_toFinset.1 hn, hxN⟩ - · refine eventually_of_forall fun n ↦ ?_ + · refine Eventually.of_forall fun n ↦ ?_ obtain rfl | _ := eq_zero_or_neZero μ · simp · rw [← laverage_const μ 1] diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index 2477fc7a63860..79ef7cb7efe73 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -183,7 +183,7 @@ theorem exists_closed_cover_approximatesLinearOn_of_hasFDerivWithinAt [SecondCou have L : Tendsto (fun k => f (a k)) atTop (𝓝 (f x)) := by apply (hf' x xs).continuousWithinAt.tendsto.comp apply tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within _ a_lim - exact eventually_of_forall fun k => (aM k).1 + exact Eventually.of_forall fun k => (aM k).1 apply Tendsto.sub (tendsto_const_nhds.sub L) exact ((f' z).continuous.tendsto _).comp (tendsto_const_nhds.sub a_lim) have L2 : Tendsto (fun k : ℕ => (r (f' z) : ℝ) * ‖y - a k‖) atTop (𝓝 (r (f' z) * ‖y - x‖)) := diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index d9a16a74df256..974593ec53552 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -222,7 +222,7 @@ theorem HasFiniteIntegral.norm {f : α → β} (hfi : HasFiniteIntegral f μ) : theorem hasFiniteIntegral_norm_iff (f : α → β) : HasFiniteIntegral (fun a => ‖f a‖) μ ↔ HasFiniteIntegral f μ := - hasFiniteIntegral_congr' <| eventually_of_forall fun x => norm_norm (f x) + hasFiniteIntegral_congr' <| Eventually.of_forall fun x => norm_norm (f x) theorem hasFiniteIntegral_toReal_of_lintegral_ne_top {f : α → ℝ≥0∞} (hf : (∫⁻ x, f x ∂μ) ≠ ∞) : HasFiniteIntegral (fun x => (f x).toReal) μ := by @@ -336,11 +336,11 @@ section PosPart theorem HasFiniteIntegral.max_zero {f : α → ℝ} (hf : HasFiniteIntegral f μ) : HasFiniteIntegral (fun a => max (f a) 0) μ := - hf.mono <| eventually_of_forall fun x => by simp [abs_le, le_abs_self] + hf.mono <| Eventually.of_forall fun x => by simp [abs_le, le_abs_self] theorem HasFiniteIntegral.min_zero {f : α → ℝ} (hf : HasFiniteIntegral f μ) : HasFiniteIntegral (fun a => min (f a) 0) μ := - hf.mono <| eventually_of_forall fun x => by simpa [abs_le] using neg_abs_le _ + hf.mono <| Eventually.of_forall fun x => by simpa [abs_le] using neg_abs_le _ end PosPart @@ -745,7 +745,7 @@ theorem Integrable.prod_mk {f : α → β} {g : α → γ} (hf : Integrable f μ Integrable (fun x => (f x, g x)) μ := ⟨hf.aestronglyMeasurable.prod_mk hg.aestronglyMeasurable, (hf.norm.add' hg.norm).mono <| - eventually_of_forall fun x => + Eventually.of_forall fun x => calc max ‖f x‖ ‖g x‖ ≤ ‖f x‖ + ‖g x‖ := max_le_add_of_nonneg (norm_nonneg _) (norm_nonneg _) _ ≤ ‖‖f x‖ + ‖g x‖‖ := le_abs_self _⟩ @@ -983,7 +983,7 @@ noncomputable def withDensitySMulLI {f : α → ℝ≥0} (f_meas : Measurable f) one_ne_zero, ENNReal.one_ne_top, ENNReal.one_toReal, if_false, eLpNorm', ENNReal.rpow_one, _root_.div_one, Lp.norm_def] rw [lintegral_withDensity_eq_lintegral_mul_non_measurable _ f_meas.coe_nnreal_ennreal - (Filter.eventually_of_forall fun x => ENNReal.coe_lt_top)] + (Filter.Eventually.of_forall fun x => ENNReal.coe_lt_top)] congr 1 apply lintegral_congr_ae filter_upwards [(memℒ1_smul_of_L1_withDensity f_meas u).coeFn_toLp] with x hx @@ -1419,7 +1419,7 @@ theorem ContinuousLinearMap.integrable_comp {φ : α → H} (L : H →L[𝕜] E) Integrable (fun a : α => L (φ a)) μ := ((Integrable.norm φ_int).const_mul ‖L‖).mono' (L.continuous.comp_aestronglyMeasurable φ_int.aestronglyMeasurable) - (eventually_of_forall fun a => L.le_opNorm (φ a)) + (Eventually.of_forall fun a => L.le_opNorm (φ a)) @[simp] theorem ContinuousLinearEquiv.integrable_comp_iff {φ : α → H} (L : H ≃L[𝕜] E) : diff --git a/Mathlib/MeasureTheory/Function/L2Space.lean b/Mathlib/MeasureTheory/Function/L2Space.lean index 1c7c073e397b5..7af0d94802b18 100644 --- a/Mathlib/MeasureTheory/Function/L2Space.lean +++ b/Mathlib/MeasureTheory/Function/L2Space.lean @@ -62,11 +62,11 @@ local notation "⟪" x ", " y "⟫" => @inner 𝕜 E _ x y theorem Memℒp.const_inner (c : E) {f : α → E} (hf : Memℒp f p μ) : Memℒp (fun a => ⟪c, f a⟫) p μ := hf.of_le_mul (AEStronglyMeasurable.inner aestronglyMeasurable_const hf.1) - (eventually_of_forall fun _ => norm_inner_le_norm _ _) + (Eventually.of_forall fun _ => norm_inner_le_norm _ _) theorem Memℒp.inner_const {f : α → E} (hf : Memℒp f p μ) (c : E) : Memℒp (fun a => ⟪f a, c⟫) p μ := hf.of_le_mul (AEStronglyMeasurable.inner hf.1 aestronglyMeasurable_const) - (eventually_of_forall fun x => by rw [mul_comm]; exact norm_inner_le_norm _ _) + (Eventually.of_forall fun x => by rw [mul_comm]; exact norm_inner_le_norm _ _) variable {f : α → E} @@ -148,7 +148,7 @@ theorem integral_inner_eq_sq_eLpNorm (f : α →₂[μ] E) : norm_cast rw [integral_eq_lintegral_of_nonneg_ae] rotate_left - · exact Filter.eventually_of_forall fun x => sq_nonneg _ + · exact Filter.Eventually.of_forall fun x => sq_nonneg _ · exact ((Lp.aestronglyMeasurable f).norm.aemeasurable.pow_const _).aestronglyMeasurable congr ext1 x diff --git a/Mathlib/MeasureTheory/Function/LpOrder.lean b/Mathlib/MeasureTheory/Function/LpOrder.lean index 51148571e059f..86290745da5ec 100644 --- a/Mathlib/MeasureTheory/Function/LpOrder.lean +++ b/Mathlib/MeasureTheory/Function/LpOrder.lean @@ -59,12 +59,12 @@ instance instOrderedAddCommGroup : OrderedAddCommGroup (Lp E p μ) := theorem _root_.MeasureTheory.Memℒp.sup {f g : α → E} (hf : Memℒp f p μ) (hg : Memℒp g p μ) : Memℒp (f ⊔ g) p μ := Memℒp.mono' (hf.norm.add hg.norm) (hf.1.sup hg.1) - (Filter.eventually_of_forall fun x => norm_sup_le_add (f x) (g x)) + (Filter.Eventually.of_forall fun x => norm_sup_le_add (f x) (g x)) theorem _root_.MeasureTheory.Memℒp.inf {f g : α → E} (hf : Memℒp f p μ) (hg : Memℒp g p μ) : Memℒp (f ⊓ g) p μ := Memℒp.mono' (hf.norm.add hg.norm) (hf.1.inf hg.1) - (Filter.eventually_of_forall fun x => norm_inf_le_add (f x) (g x)) + (Filter.Eventually.of_forall fun x => norm_inf_le_add (f x) (g x)) theorem _root_.MeasureTheory.Memℒp.abs {f : α → E} (hf : Memℒp f p μ) : Memℒp |f| p μ := hf.sup hf.neg diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 052668ca96ad7..c35df4f7ac93b 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -526,21 +526,21 @@ alias snorm_mono_ae_real := eLpNorm_mono_ae_real theorem eLpNorm_mono_nnnorm {f : α → F} {g : α → G} (h : ∀ x, ‖f x‖₊ ≤ ‖g x‖₊) : eLpNorm f p μ ≤ eLpNorm g p μ := - eLpNorm_mono_nnnorm_ae (eventually_of_forall fun x => h x) + eLpNorm_mono_nnnorm_ae (Eventually.of_forall fun x => h x) @[deprecated (since := "2024-07-27")] alias snorm_mono_nnnorm := eLpNorm_mono_nnnorm theorem eLpNorm_mono {f : α → F} {g : α → G} (h : ∀ x, ‖f x‖ ≤ ‖g x‖) : eLpNorm f p μ ≤ eLpNorm g p μ := - eLpNorm_mono_ae (eventually_of_forall fun x => h x) + eLpNorm_mono_ae (Eventually.of_forall fun x => h x) @[deprecated (since := "2024-07-27")] alias snorm_mono := eLpNorm_mono theorem eLpNorm_mono_real {f : α → F} {g : α → ℝ} (h : ∀ x, ‖f x‖ ≤ g x) : eLpNorm f p μ ≤ eLpNorm g p μ := - eLpNorm_mono_ae_real (eventually_of_forall fun x => h x) + eLpNorm_mono_ae_real (Eventually.of_forall fun x => h x) @[deprecated (since := "2024-07-27")] alias snorm_mono_real := eLpNorm_mono_real @@ -627,7 +627,7 @@ alias snorm'_norm := eLpNorm'_norm @[simp] theorem eLpNorm_norm (f : α → F) : eLpNorm (fun x => ‖f x‖) p μ = eLpNorm f p μ := - eLpNorm_congr_norm_ae <| eventually_of_forall fun _ => norm_norm _ + eLpNorm_congr_norm_ae <| Eventually.of_forall fun _ => norm_norm _ @[deprecated (since := "2024-07-27")] alias snorm_norm := eLpNorm_norm @@ -891,7 +891,7 @@ theorem Memℒp.right_of_add_measure {f : α → E} (h : Memℒp f p (μ + ν)) h.mono_measure <| Measure.le_add_left <| le_refl _ theorem Memℒp.norm {f : α → E} (h : Memℒp f p μ) : Memℒp (fun x => ‖f x‖) p μ := - h.of_le h.aestronglyMeasurable.norm (eventually_of_forall fun x => by simp) + h.of_le h.aestronglyMeasurable.norm (Eventually.of_forall fun x => by simp) theorem memℒp_norm_iff {f : α → E} (hf : AEStronglyMeasurable f μ) : Memℒp (fun x => ‖f x‖) p μ ↔ Memℒp f p μ := @@ -1182,7 +1182,7 @@ variable [BoundedSMul 𝕜 E] [BoundedSMul 𝕜 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 _ _) + eLpNorm'_le_nnreal_smul_eLpNorm'_of_ae_le_mul (Eventually.of_forall fun _ => nnnorm_smul_le _ _) hq_pos @[deprecated (since := "2024-07-27")] @@ -1191,14 +1191,14 @@ alias snorm'_const_smul_le := eLpNorm'_const_smul_le theorem eLpNormEssSup_const_smul_le (c : 𝕜) (f : α → F) : eLpNormEssSup (c • f) μ ≤ ‖c‖₊ • eLpNormEssSup f μ := eLpNormEssSup_le_nnreal_smul_eLpNormEssSup_of_ae_le_mul - (eventually_of_forall fun _ => by simp [nnnorm_smul_le]) + (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 μ := eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul - (eventually_of_forall fun _ => by simp [nnnorm_smul_le]) _ + (Eventually.of_forall fun _ => by simp [nnnorm_smul_le]) _ @[deprecated (since := "2024-07-27")] alias snorm_const_smul_le := eLpNorm_const_smul_le @@ -1286,7 +1286,7 @@ theorem Memℒp.re (hf : Memℒp f p μ) : Memℒp (fun x => RCLike.re (f x)) p intro x rw [one_mul] exact RCLike.norm_re_le_norm (f x) - refine hf.of_le_mul ?_ (eventually_of_forall this) + refine hf.of_le_mul ?_ (Eventually.of_forall this) exact RCLike.continuous_re.comp_aestronglyMeasurable hf.1 theorem Memℒp.im (hf : Memℒp f p μ) : Memℒp (fun x => RCLike.im (f x)) p μ := by @@ -1294,7 +1294,7 @@ theorem Memℒp.im (hf : Memℒp f p μ) : Memℒp (fun x => RCLike.im (f x)) p intro x rw [one_mul] exact RCLike.norm_im_le_norm (f x) - refine hf.of_le_mul ?_ (eventually_of_forall this) + refine hf.of_le_mul ?_ (Eventually.of_forall this) exact RCLike.continuous_im.comp_aestronglyMeasurable hf.1 end RCLike @@ -1340,7 +1340,7 @@ theorem ae_bdd_liminf_atTop_of_eLpNorm_bdd {p : ℝ≥0∞} (hp : p ≠ 0) {f : (lt_of_le_of_lt (hbdd n) <| ENNReal.lt_add_right ENNReal.coe_ne_top one_ne_zero) rw [← ae_all_iff] at this filter_upwards [this] with x hx using lt_of_le_of_lt - (liminf_le_of_frequently_le' <| frequently_of_forall fun n => (hx n).le) + (liminf_le_of_frequently_le' <| Frequently.of_forall fun n => (hx n).le) (ENNReal.add_lt_top.2 ⟨ENNReal.coe_lt_top, ENNReal.one_lt_top⟩) filter_upwards [ae_bdd_liminf_atTop_rpow_of_eLpNorm_bdd hfmeas hbdd] with x hx have hppos : 0 < p.toReal := ENNReal.toReal_pos hp hp' @@ -1369,7 +1369,7 @@ theorem _root_.Continuous.memℒp_top_of_hasCompactSupport {f : X → E} (hf : Continuous f) (h'f : HasCompactSupport f) (μ : Measure X) : Memℒp f ⊤ μ := by borelize E rcases hf.bounded_above_of_compact_support h'f with ⟨C, hC⟩ - apply memℒp_top_of_bound ?_ C (Filter.eventually_of_forall hC) + apply memℒp_top_of_bound ?_ C (Filter.Eventually.of_forall hC) exact (hf.stronglyMeasurable_of_hasCompactSupport h'f).aestronglyMeasurable section UnifTight diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean index 74680c8ebc928..04c5adbec044e 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean @@ -306,7 +306,7 @@ variable {𝕜 α E F : Type*} {m : MeasurableSpace α} {μ : Measure α} [Norme theorem eLpNorm_smul_le_eLpNorm_top_mul_eLpNorm (p : ℝ≥0∞) (hf : AEStronglyMeasurable f μ) (φ : α → 𝕜) : eLpNorm (φ • f) p μ ≤ eLpNorm φ ∞ μ * eLpNorm f p μ := (eLpNorm_le_eLpNorm_top_mul_eLpNorm p φ hf (· • ·) - (eventually_of_forall fun _ => nnnorm_smul_le _ _) : _) + (Eventually.of_forall fun _ => nnnorm_smul_le _ _) : _) @[deprecated (since := "2024-07-27")] alias snorm_smul_le_snorm_top_mul_snorm := eLpNorm_smul_le_eLpNorm_top_mul_eLpNorm @@ -314,7 +314,7 @@ alias snorm_smul_le_snorm_top_mul_snorm := eLpNorm_smul_le_eLpNorm_top_mul_eLpNo theorem eLpNorm_smul_le_eLpNorm_mul_eLpNorm_top (p : ℝ≥0∞) (f : α → E) {φ : α → 𝕜} (hφ : AEStronglyMeasurable φ μ) : eLpNorm (φ • f) p μ ≤ eLpNorm φ p μ * eLpNorm f ∞ μ := (eLpNorm_le_eLpNorm_mul_eLpNorm_top p hφ f (· • ·) - (eventually_of_forall fun _ => nnnorm_smul_le _ _) : _) + (Eventually.of_forall fun _ => nnnorm_smul_le _ _) : _) @[deprecated (since := "2024-07-27")] alias snorm_smul_le_snorm_mul_snorm_top := eLpNorm_smul_le_eLpNorm_mul_eLpNorm_top @@ -322,7 +322,7 @@ alias snorm_smul_le_snorm_mul_snorm_top := eLpNorm_smul_le_eLpNorm_mul_eLpNorm_t theorem eLpNorm'_smul_le_mul_eLpNorm' {p q r : ℝ} {f : α → E} (hf : AEStronglyMeasurable f μ) {φ : α → 𝕜} (hφ : AEStronglyMeasurable φ μ) (hp0_lt : 0 < p) (hpq : p < q) (hpqr : 1 / p = 1 / q + 1 / r) : eLpNorm' (φ • f) p μ ≤ eLpNorm' φ q μ * eLpNorm' f r μ := - eLpNorm'_le_eLpNorm'_mul_eLpNorm' hφ hf (· • ·) (eventually_of_forall fun _ => nnnorm_smul_le _ _) + eLpNorm'_le_eLpNorm'_mul_eLpNorm' hφ hf (· • ·) (Eventually.of_forall fun _ => nnnorm_smul_le _ _) hp0_lt hpq hpqr @[deprecated (since := "2024-07-27")] @@ -333,7 +333,7 @@ theorem eLpNorm_smul_le_mul_eLpNorm {p q r : ℝ≥0∞} {f : α → E} (hf : AE {φ : α → 𝕜} (hφ : AEStronglyMeasurable φ μ) (hpqr : 1 / p = 1 / q + 1 / r) : eLpNorm (φ • f) p μ ≤ eLpNorm φ q μ * eLpNorm f r μ := (eLpNorm_le_eLpNorm_mul_eLpNorm_of_nnnorm hφ hf (· • ·) - (eventually_of_forall fun _ => nnnorm_smul_le _ _) hpqr : + (Eventually.of_forall fun _ => nnnorm_smul_le _ _) hpqr : _) @[deprecated (since := "2024-07-27")] diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean index 36a7d7423dd74..76196abdb4abc 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/TriangleInequality.lean @@ -48,7 +48,7 @@ alias snorm'_add_le_of_le_one := eLpNorm'_add_le_of_le_one theorem eLpNormEssSup_add_le {f g : α → E} : eLpNormEssSup (f + g) μ ≤ eLpNormEssSup f μ + eLpNormEssSup g μ := by - refine le_trans (essSup_mono_ae (eventually_of_forall fun x => ?_)) (ENNReal.essSup_add_le _ _) + refine le_trans (essSup_mono_ae (Eventually.of_forall fun x => ?_)) (ENNReal.essSup_add_le _ _) simp_rw [Pi.add_apply, ← ENNReal.coe_add, ENNReal.coe_le_coe] exact nnnorm_add_le _ _ diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 77bda2bdf1659..cfd54d45970a3 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -502,7 +502,7 @@ variable {c : E} {f : α → E} {hf : AEStronglyMeasurable f μ} {s : Set α} theorem eLpNormEssSup_indicator_le (s : Set α) (f : α → G) : eLpNormEssSup (s.indicator f) μ ≤ eLpNormEssSup f μ := by - refine essSup_mono_ae (eventually_of_forall fun x => ?_) + refine essSup_mono_ae (Eventually.of_forall fun x => ?_) rw [ENNReal.coe_le_coe, nnnorm_indicator_eq_indicator_nnnorm] exact Set.indicator_le_self s _ x @@ -532,7 +532,7 @@ theorem eLpNormEssSup_indicator_const_eq (s : Set α) (c : G) (hμs : μ s ≠ 0 alias snormEssSup_indicator_const_eq := eLpNormEssSup_indicator_const_eq theorem eLpNorm_indicator_le (f : α → E) : eLpNorm (s.indicator f) p μ ≤ eLpNorm f p μ := by - refine eLpNorm_mono_ae (eventually_of_forall fun x => ?_) + refine eLpNorm_mono_ae (Eventually.of_forall fun x => ?_) suffices ‖s.indicator f x‖₊ ≤ ‖f x‖₊ by exact NNReal.coe_mono this rw [nnnorm_indicator_eq_indicator_nnnorm] exact s.indicator_le_self _ x @@ -1052,7 +1052,7 @@ theorem LipschitzWith.comp_memℒp {α E F} {K} [MeasurableSpace α] {μ : Measu have : ∀ x, ‖g (f x)‖ ≤ K * ‖f x‖ := fun x ↦ by -- TODO: add `LipschitzWith.nnnorm_sub_le` and `LipschitzWith.nnnorm_le` simpa [g0] using hg.norm_sub_le (f x) 0 - hL.of_le_mul (hg.continuous.comp_aestronglyMeasurable hL.1) (eventually_of_forall this) + hL.of_le_mul (hg.continuous.comp_aestronglyMeasurable hL.1) (Eventually.of_forall this) theorem MeasureTheory.Memℒp.of_comp_antilipschitzWith {α E F} {K'} [MeasurableSpace α] {μ : Measure α} [NormedAddCommGroup E] [NormedAddCommGroup F] {f : α → E} {g : E → F} @@ -1065,7 +1065,7 @@ theorem MeasureTheory.Memℒp.of_comp_antilipschitzWith {α E F} {K'} [Measurabl apply hg'.le_mul_dist have B : AEStronglyMeasurable f μ := (hg'.uniformEmbedding hg).embedding.aestronglyMeasurable_comp_iff.1 hL.1 - exact hL.of_le_mul B (Filter.eventually_of_forall A) + exact hL.of_le_mul B (Filter.Eventually.of_forall A) namespace LipschitzWith @@ -1547,7 +1547,7 @@ private theorem lintegral_rpow_tsum_coe_nnnorm_sub_le_tsum {f : ℕ → α → E (Finset.aemeasurable_sum (Finset.range (n + 1)) fun i _ => ((hf (i + 1)).sub (hf i)).ennnorm).pow_const _ - · exact liminf_le_of_frequently_le' (frequently_of_forall h) + · exact liminf_le_of_frequently_le' (Frequently.of_forall h) private theorem tsum_nnnorm_sub_ae_lt_top {f : ℕ → α → E} (hf : ∀ n, AEStronglyMeasurable (f n) μ) {p : ℝ} (hp1 : 1 ≤ p) {B : ℕ → ℝ≥0∞} (hB : ∑' i, B i ≠ ∞) diff --git a/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean b/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean index 74a783f2afe90..ecb76ae835229 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace/ContinuousCompMeasurePreserving.lean @@ -58,7 +58,7 @@ theorem compMeasurePreserving_continuous (hp : p ≠ ∞) : refine continuous_indicatorConstLp_set hp fun f ↦ ?_ apply tendsto_measure_symmDiff_preimage_nhds_zero continuousAt_subtype_val _ f.2 hs.nullMeasurableSet hνs.ne - exact eventually_of_forall Subtype.property + exact .of_forall Subtype.property end Lp diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index 26d57a24e44ae..163748ae34896 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -110,7 +110,7 @@ theorem tendsto_approxOn_Lp_eLpNorm [OpensMeasurableSpace E] {f : β → E} (hf ∀ n, (fun x => (‖approxOn f hf s y₀ h₀ n x - f x‖₊ : ℝ≥0∞) ^ p.toReal) ≤ᵐ[μ] fun x => (‖f x - y₀‖₊ : ℝ≥0∞) ^ p.toReal := fun n => - eventually_of_forall fun x => + Eventually.of_forall fun x => rpow_le_rpow (coe_mono (nnnorm_approxOn_le hf h₀ x n)) toReal_nonneg -- (3) The bounding function `fun x => ‖f x - y₀‖ ^ p.toReal` has finite integral have h_fin : (∫⁻ a : β, (‖f a - y₀‖₊ : ℝ≥0∞) ^ p.toReal ∂μ) ≠ ⊤ := @@ -275,7 +275,7 @@ theorem memℒp_zero (f : α →ₛ E) (μ : Measure α) : Memℒp f 0 μ := theorem memℒp_top (f : α →ₛ E) (μ : Measure α) : Memℒp f ∞ μ := let ⟨C, hfC⟩ := f.exists_forall_norm_le - memℒp_top_of_bound f.aestronglyMeasurable C <| eventually_of_forall hfC + memℒp_top_of_bound f.aestronglyMeasurable C <| Eventually.of_forall hfC protected theorem eLpNorm'_eq {p : ℝ} (f : α →ₛ F) (μ : Measure α) : eLpNorm' f p μ = (∑ y ∈ f.range, (‖y‖₊ : ℝ≥0∞) ^ p * μ (f ⁻¹' {y})) ^ (1 / p) := by @@ -357,7 +357,7 @@ theorem integrable_pair {f : α →ₛ E} {g : α →ₛ F} : theorem memℒp_of_isFiniteMeasure (f : α →ₛ E) (p : ℝ≥0∞) (μ : Measure α) [IsFiniteMeasure μ] : Memℒp f p μ := let ⟨C, hfC⟩ := f.exists_forall_norm_le - Memℒp.of_bound f.aestronglyMeasurable C <| eventually_of_forall hfC + Memℒp.of_bound f.aestronglyMeasurable C <| Eventually.of_forall hfC theorem integrable_of_isFiniteMeasure [IsFiniteMeasure μ] (f : α →ₛ E) : Integrable f μ := memℒp_one_iff_integrable.mp (f.memℒp_of_isFiniteMeasure 1 μ) diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 4428be6393921..0c844b249cd24 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -1755,7 +1755,7 @@ theorem exists_set_sigmaFinite (hf : AEFinStronglyMeasurable f μ) : refine ⟨t, ht, ?_, htμ⟩ refine EventuallyEq.trans (ae_restrict_of_ae hfg) ?_ rw [EventuallyEq, ae_restrict_iff' ht.compl] - exact eventually_of_forall hgt_zero + exact Eventually.of_forall hgt_zero /-- A measurable set `t` such that `f =ᵐ[μ.restrict tᶜ] 0` and `sigma_finite (μ.restrict t)`. -/ def sigmaFiniteSet (hf : AEFinStronglyMeasurable f μ) : Set α := diff --git a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean index f337d65513ca3..383456a232000 100644 --- a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean @@ -872,7 +872,7 @@ theorem UniformIntegrable.spec' (hp : p ≠ 0) (hp' : p ≠ ∞) (hf : ∀ i, St _ ≤ eLpNorm ({ x | C ≤ ‖f (ℐ C) x‖₊ }.indicator (f (ℐ C))) p μ := by refine le_eLpNorm_of_bddBelow hp hp' _ (measurableSet_le measurable_const (hf _).nnnorm.measurable) - (eventually_of_forall fun x hx => ?_) + (Eventually.of_forall fun x hx => ?_) rwa [nnnorm_indicator_eq_indicator_nnnorm, Set.indicator_of_mem hx] _ ≤ eLpNorm (f (ℐ C)) p μ := eLpNorm_indicator_le _ specialize this (2 * max M 1 * δ⁻¹ ^ (1 / p.toReal)) diff --git a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean index 37cbb63632c50..561111dd29111 100644 --- a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean +++ b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean @@ -84,7 +84,7 @@ is a fundamental domain for the action of `G` on `α`. -/ theorem mk' (h_meas : NullMeasurableSet s μ) (h_exists : ∀ x : α, ∃! g : G, g • x ∈ s) : IsFundamentalDomain G s μ where nullMeasurableSet := h_meas - ae_covers := eventually_of_forall fun x => (h_exists x).exists + ae_covers := Eventually.of_forall fun x => (h_exists x).exists aedisjoint a b hab := Disjoint.aedisjoint <| disjoint_left.2 fun x hxa hxb => by rw [mem_smul_set_iff_inv_smul_mem] at hxa hxb exact hab (inv_injective <| (h_exists x).unique hxa hxb) diff --git a/Mathlib/MeasureTheory/Group/Prod.lean b/Mathlib/MeasureTheory/Group/Prod.lean index f2e659ca8ab0e..7e695cc33d3e1 100644 --- a/Mathlib/MeasureTheory/Group/Prod.lean +++ b/Mathlib/MeasureTheory/Group/Prod.lean @@ -83,7 +83,7 @@ There, the map in this lemma is called `S`. -/ theorem measurePreserving_prod_mul [IsMulLeftInvariant ν] : MeasurePreserving (fun z : G × G => (z.1, z.1 * z.2)) (μ.prod ν) (μ.prod ν) := (MeasurePreserving.id μ).skew_product measurable_mul <| - Filter.eventually_of_forall <| map_mul_left_eq_self ν + Filter.Eventually.of_forall <| map_mul_left_eq_self ν /-- The map `(x, y) ↦ (y, yx)` sends the measure `μ × ν` to `ν × μ`. This is the map `SR` in [Halmos, §59]. @@ -342,7 +342,7 @@ section RightInvariant theorem measurePreserving_prod_mul_right [IsMulRightInvariant ν] : MeasurePreserving (fun z : G × G => (z.1, z.2 * z.1)) (μ.prod ν) (μ.prod ν) := MeasurePreserving.skew_product (g := fun x y => y * x) (MeasurePreserving.id μ) - (measurable_snd.mul measurable_fst) <| Filter.eventually_of_forall <| map_mul_right_eq_self ν + (measurable_snd.mul measurable_fst) <| Filter.Eventually.of_forall <| map_mul_right_eq_self ν /-- The map `(x, y) ↦ (y, xy)` sends the measure `μ × ν` to `ν × μ`. -/ @[to_additive measurePreserving_prod_add_swap_right @@ -422,7 +422,7 @@ theorem quasiMeasurePreserving_div_left_of_right_invariant [IsMulRightInvariant @[to_additive] theorem quasiMeasurePreserving_div_of_right_invariant [IsMulRightInvariant μ] : QuasiMeasurePreserving (fun p : G × G => p.1 / p.2) (μ.prod ν) μ := by - refine QuasiMeasurePreserving.prod_of_left measurable_div (eventually_of_forall fun y => ?_) + refine QuasiMeasurePreserving.prod_of_left measurable_div (Eventually.of_forall fun y => ?_) exact (measurePreserving_div_right μ y).quasiMeasurePreserving @[to_additive] diff --git a/Mathlib/MeasureTheory/Integral/Average.lean b/Mathlib/MeasureTheory/Integral/Average.lean index 8ec1ba5cb9b64..2b93b827a0339 100644 --- a/Mathlib/MeasureTheory/Integral/Average.lean +++ b/Mathlib/MeasureTheory/Integral/Average.lean @@ -752,7 +752,7 @@ theorem tendsto_integral_smul_of_tendsto_average_norm_sub rw [← integrableOn_iff_integrable_of_support_subset A] apply Integrable.smul_of_top_right hif exact memℒp_top_of_bound hig.aestronglyMeasurable.restrict - (K / (μ (a i)).toReal) (eventually_of_forall hibound) + (K / (μ (a i)).toReal) (Eventually.of_forall hibound) · exact hig.smul_const _ have L0 : Tendsto (fun i ↦ ∫ y, g i y • (f y - c) ∂μ) l (𝓝 0) := by have := hf.const_mul K @@ -773,8 +773,8 @@ theorem tendsto_integral_smul_of_tendsto_average_norm_sub have : g i x = 0 := by rw [← Function.nmem_support]; exact fun h ↦ hx (hi h) simp [this] rw [← setIntegral_eq_integral_of_forall_compl_eq_zero this (μ := μ)] - refine integral_mono_of_nonneg (eventually_of_forall (fun x ↦ by positivity)) ?_ - (eventually_of_forall (fun x ↦ ?_)) + refine integral_mono_of_nonneg (Eventually.of_forall (fun x ↦ by positivity)) ?_ + (Eventually.of_forall (fun x ↦ ?_)) · apply (Integrable.sub h''i _).norm.const_mul change IntegrableOn (fun _ ↦ c) (a i) μ simp [integrableOn_const, mu_ai] diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 74bbe759404b4..ab1f1a86e1f93 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1077,7 +1077,7 @@ theorem integral_nonneg_of_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ] f) : 0 ≤ ∫ theorem lintegral_coe_eq_integral (f : α → ℝ≥0) (hfi : Integrable (fun x => (f x : ℝ)) μ) : ∫⁻ a, f a ∂μ = ENNReal.ofReal (∫ a, f a ∂μ) := by - simp_rw [integral_eq_lintegral_of_nonneg_ae (eventually_of_forall fun x => (f x).coe_nonneg) + simp_rw [integral_eq_lintegral_of_nonneg_ae (Eventually.of_forall fun x => (f x).coe_nonneg) hfi.aestronglyMeasurable, ← ENNReal.coe_nnreal_eq] rw [ENNReal.ofReal_toReal] rw [← lt_top_iff_ne_top] @@ -1096,7 +1096,7 @@ theorem integral_toReal {f : α → ℝ≥0∞} (hfm : AEMeasurable f μ) (hf : ∫ a, (f a).toReal ∂μ = (∫⁻ a, f a ∂μ).toReal := by rw [integral_eq_lintegral_of_nonneg_ae _ hfm.ennreal_toReal.aestronglyMeasurable, lintegral_congr_ae (ofReal_toReal_ae_eq hf)] - exact eventually_of_forall fun x => ENNReal.toReal_nonneg + exact Eventually.of_forall fun x => ENNReal.toReal_nonneg theorem lintegral_coe_le_coe_iff_integral_le {f : α → ℝ≥0} (hfi : Integrable (fun x => (f x : ℝ)) μ) {b : ℝ≥0} : ∫⁻ a, f a ∂μ ≤ b ↔ ∫ a, (f a : ℝ) ∂μ ≤ b := by @@ -1110,7 +1110,7 @@ theorem integral_coe_le_of_lintegral_coe_le {f : α → ℝ≥0} {b : ℝ≥0} ( · rw [integral_undef hf]; exact b.2 theorem integral_nonneg {f : α → ℝ} (hf : 0 ≤ f) : 0 ≤ ∫ a, f a ∂μ := - integral_nonneg_of_ae <| eventually_of_forall hf + integral_nonneg_of_ae <| Eventually.of_forall hf theorem integral_nonpos_of_ae {f : α → ℝ} (hf : f ≤ᵐ[μ] 0) : ∫ a, f a ∂μ ≤ 0 := by have hf : 0 ≤ᵐ[μ] -f := hf.mono fun a h => by rwa [Pi.neg_apply, Pi.zero_apply, neg_nonneg] @@ -1118,7 +1118,7 @@ theorem integral_nonpos_of_ae {f : α → ℝ} (hf : f ≤ᵐ[μ] 0) : ∫ a, f rwa [integral_neg, neg_nonneg] at this theorem integral_nonpos {f : α → ℝ} (hf : f ≤ 0) : ∫ a, f a ∂μ ≤ 0 := - integral_nonpos_of_ae <| eventually_of_forall hf + integral_nonpos_of_ae <| Eventually.of_forall hf theorem integral_eq_zero_iff_of_nonneg_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ] f) (hfi : Integrable f μ) : ∫ x, f x ∂μ = 0 ↔ f =ᵐ[μ] 0 := by @@ -1132,7 +1132,7 @@ theorem integral_eq_zero_iff_of_nonneg_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ] f) theorem integral_eq_zero_iff_of_nonneg {f : α → ℝ} (hf : 0 ≤ f) (hfi : Integrable f μ) : ∫ x, f x ∂μ = 0 ↔ f =ᵐ[μ] 0 := - integral_eq_zero_iff_of_nonneg_ae (eventually_of_forall hf) hfi + integral_eq_zero_iff_of_nonneg_ae (Eventually.of_forall hf) hfi lemma integral_eq_iff_of_ae_le {f g : α → ℝ} (hf : Integrable f μ) (hg : Integrable g μ) (hfg : f ≤ᵐ[μ] g) : @@ -1150,7 +1150,7 @@ theorem integral_pos_iff_support_of_nonneg_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ theorem integral_pos_iff_support_of_nonneg {f : α → ℝ} (hf : 0 ≤ f) (hfi : Integrable f μ) : (0 < ∫ x, f x ∂μ) ↔ 0 < μ (Function.support f) := - integral_pos_iff_support_of_nonneg_ae (eventually_of_forall hf) hfi + integral_pos_iff_support_of_nonneg_ae (Eventually.of_forall hf) hfi lemma integral_exp_pos {μ : Measure α} {f : α → ℝ} [hμ : NeZero μ] (hf : Integrable (fun x ↦ Real.exp (f x)) μ) : @@ -1301,7 +1301,7 @@ variable {H : Type*} [NormedAddCommGroup H] theorem L1.norm_eq_integral_norm (f : α →₁[μ] H) : ‖f‖ = ∫ a, ‖f a‖ ∂μ := by simp only [eLpNorm, eLpNorm', ENNReal.one_toReal, ENNReal.rpow_one, Lp.norm_def, if_false, ENNReal.one_ne_top, one_ne_zero, _root_.div_one] - rw [integral_eq_lintegral_of_nonneg_ae (eventually_of_forall (by simp [norm_nonneg])) + rw [integral_eq_lintegral_of_nonneg_ae (Eventually.of_forall (by simp [norm_nonneg])) (Lp.aestronglyMeasurable f).norm] simp [ofReal_norm_eq_coe_nnnorm] @@ -1341,7 +1341,7 @@ theorem integral_mono_ae {f g : α → ℝ} (hf : Integrable f μ) (hg : Integra @[mono] theorem integral_mono {f g : α → ℝ} (hf : Integrable f μ) (hg : Integrable g μ) (h : f ≤ g) : ∫ a, f a ∂μ ≤ ∫ a, g a ∂μ := - integral_mono_ae hf hg <| eventually_of_forall h + integral_mono_ae hf hg <| Eventually.of_forall h theorem integral_mono_of_nonneg {f g : α → ℝ} (hf : 0 ≤ᵐ[μ] f) (hgi : Integrable g μ) (h : f ≤ᵐ[μ] g) : ∫ a, f a ∂μ ≤ ∫ a, g a ∂μ := by @@ -1362,7 +1362,7 @@ theorem integral_mono_measure {f : α → ℝ} {ν} (hle : μ ≤ ν) (hf : 0 ((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 _ + have le_ae : ∀ᵐ a ∂μ, 0 ≤ ‖f a‖ := Eventually.of_forall fun a => norm_nonneg _ by_cases h : AEStronglyMeasurable f μ · calc ‖∫ a, f a ∂μ‖ ≤ ENNReal.toReal (∫⁻ a, ENNReal.ofReal ‖f a‖ ∂μ) := @@ -1375,7 +1375,7 @@ theorem norm_integral_le_of_norm_le {f : α → G} {g : α → ℝ} (hg : Integr (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ g x) : ‖∫ x, f x ∂μ‖ ≤ ∫ x, g x ∂μ := calc ‖∫ x, f x ∂μ‖ ≤ ∫ x, ‖f x‖ ∂μ := norm_integral_le_integral_norm f - _ ≤ ∫ x, g x ∂μ := integral_mono_of_nonneg (eventually_of_forall fun _ => norm_nonneg _) hg h + _ ≤ ∫ x, g x ∂μ := integral_mono_of_nonneg (Eventually.of_forall fun _ => norm_nonneg _) hg h theorem SimpleFunc.integral_eq_integral (f : α →ₛ E) (hfi : Integrable f μ) : f.integral μ = ∫ x, f x ∂μ := by @@ -1422,7 +1422,7 @@ theorem tendsto_integral_approxOn_of_measurable_of_range_subset [MeasurableSpace Tendsto (fun n => (SimpleFunc.approxOn f fmeas s 0 (hs <| by simp) n).integral μ) atTop (𝓝 <| ∫ x, f x ∂μ) := by apply tendsto_integral_approxOn_of_measurable hf fmeas _ _ (integrable_zero _ _ _) - exact eventually_of_forall fun x => subset_closure (hs (Set.mem_union_left _ (mem_range_self _))) + exact Eventually.of_forall fun x => subset_closure (hs (Set.mem_union_left _ (mem_range_self _))) -- We redeclare `E` here to temporarily avoid -- the `[CompleteSpace E]` and `[NormedSpace ℝ E]` instances. @@ -1661,11 +1661,11 @@ theorem integral_mul_norm_le_Lp_mul_Lq {E} [NormedAddCommGroup E] {f g : α → rw [integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae] rotate_left - · exact eventually_of_forall fun x => Real.rpow_nonneg (norm_nonneg _) _ + · exact Eventually.of_forall fun x => Real.rpow_nonneg (norm_nonneg _) _ · exact (hg.1.norm.aemeasurable.pow aemeasurable_const).aestronglyMeasurable - · exact eventually_of_forall fun x => Real.rpow_nonneg (norm_nonneg _) _ + · exact Eventually.of_forall fun x => Real.rpow_nonneg (norm_nonneg _) _ · exact (hf.1.norm.aemeasurable.pow aemeasurable_const).aestronglyMeasurable - · exact eventually_of_forall fun x => mul_nonneg (norm_nonneg _) (norm_nonneg _) + · exact Eventually.of_forall fun x => mul_nonneg (norm_nonneg _) (norm_nonneg _) · exact hf.1.norm.mul hg.1.norm rw [ENNReal.toReal_rpow, ENNReal.toReal_rpow, ← ENNReal.toReal_mul] -- replace norms by nnnorm @@ -1830,12 +1830,12 @@ theorem integral_trim (hm : m ≤ m0) {f : β → G} (hf : StronglyMeasurable[m] have hf_seq_eq : ∀ n, ∫ x, f_seq n x ∂μ = ∫ x, f_seq n x ∂μ.trim hm := fun n => integral_trim_simpleFunc hm (f_seq n) (hf_seq_int n) have h_lim_1 : atTop.Tendsto (fun n => ∫ x, f_seq n x ∂μ) (𝓝 (∫ x, f x ∂μ)) := by - refine tendsto_integral_of_L1 f hf_int (eventually_of_forall hf_seq_int) ?_ + refine tendsto_integral_of_L1 f hf_int (Eventually.of_forall hf_seq_int) ?_ exact SimpleFunc.tendsto_approxOn_range_L1_nnnorm (hf.mono hm).measurable hf_int have h_lim_2 : atTop.Tendsto (fun n => ∫ x, f_seq n x ∂μ) (𝓝 (∫ x, f x ∂μ.trim hm)) := by simp_rw [hf_seq_eq] refine @tendsto_integral_of_L1 β G _ _ m (μ.trim hm) _ f (hf_int.trim hm hf) _ _ - (eventually_of_forall hf_seq_int_m) ?_ + (Eventually.of_forall hf_seq_int_m) ?_ exact @SimpleFunc.tendsto_approxOn_range_L1_nnnorm β G m _ _ _ f _ _ hf.measurable (hf_int.trim hm hf) exact tendsto_nhds_unique h_lim_1 h_lim_2 diff --git a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean index 1a10dec86151f..509fc54a3bb5b 100644 --- a/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean +++ b/Mathlib/MeasureTheory/Integral/BoundedContinuousFunction.lean @@ -64,7 +64,7 @@ theorem toReal_lintegral_coe_eq_integral [OpensMeasurableSpace X] (f : X →ᵇ rw [integral_eq_lintegral_of_nonneg_ae _ (by simpa [Function.comp_apply] using (NNReal.continuous_coe.comp f.continuous).measurable.aestronglyMeasurable)] · simp only [ENNReal.ofReal_coe_nnreal] - · exact eventually_of_forall (by simp only [Pi.zero_apply, NNReal.zero_le_coe, imp_true_iff]) + · exact Eventually.of_forall (by simp only [Pi.zero_apply, NNReal.zero_le_coe, imp_true_iff]) end NNRealValued diff --git a/Mathlib/MeasureTheory/Integral/CircleIntegral.lean b/Mathlib/MeasureTheory/Integral/CircleIntegral.lean index a2409b3431d6b..761aa3eb77f15 100644 --- a/Mathlib/MeasureTheory/Integral/CircleIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/CircleIntegral.lean @@ -532,10 +532,10 @@ theorem hasSum_two_pi_I_cauchyPowerSeries_integral {f : ℂ → E} {c : ℂ} {R · fun_prop · fun_prop · simp [norm_smul, abs_of_pos hR, mul_left_comm R, inv_mul_cancel_left₀ hR.ne', mul_comm ‖_‖] - · exact eventually_of_forall fun _ _ => (summable_geometric_of_lt_one hwR.1 hwR.2).mul_left _ + · exact Eventually.of_forall fun _ _ => (summable_geometric_of_lt_one hwR.1 hwR.2).mul_left _ · simpa only [tsum_mul_left, tsum_geometric_of_lt_one hwR.1 hwR.2] using hf.norm.mul_continuousOn continuousOn_const - · refine eventually_of_forall fun θ _ => HasSum.const_smul _ ?_ + · refine Eventually.of_forall fun θ _ => HasSum.const_smul _ ?_ simp only [smul_smul] refine HasSum.smul_const ?_ _ have : ‖w / (circleMap c R θ - c)‖ < 1 := by simpa [abs_of_pos hR] using hwR.2 diff --git a/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean b/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean index bdd921a649991..dfb6dddd4aaef 100644 --- a/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean +++ b/Mathlib/MeasureTheory/Integral/DivergenceTheorem.lean @@ -170,7 +170,7 @@ theorem integral_divergence_of_hasFDerivWithinAt_off_countable_aux₂ (I : Box ( exact tendsto_setIntegral_of_monotone (fun k => (J k).measurableSet_Ioo) (Box.Ioo.comp J).monotone Hi -- Thus it suffices to prove the same about the RHS. - refine tendsto_nhds_unique_of_eventuallyEq hI_tendsto ?_ (eventually_of_forall HJ_eq) + refine tendsto_nhds_unique_of_eventuallyEq hI_tendsto ?_ (Eventually.of_forall HJ_eq) clear hI_tendsto rw [tendsto_pi_nhds] at hJl hJu /- We'll need to prove a similar statement about the integrals over the front sides and the @@ -187,7 +187,7 @@ theorem integral_divergence_of_hasFDerivWithinAt_off_countable_aux₂ (I : Box ( /- First we prove that the integrals of the restriction of `f` to `{x | x i = d}` over increasing boxes `((J k).face i).Icc` tend to the desired limit. The proof mostly repeats the one above. -/ have hd : d ∈ Icc (I.lower i) (I.upper i) := - isClosed_Icc.mem_of_tendsto hcd (eventually_of_forall hc) + isClosed_Icc.mem_of_tendsto hcd (Eventually.of_forall hc) have Hic : ∀ k, IntegrableOn (fun x => f (i.insertNth (c k) x) i) (Box.Icc (I.face i)) := fun k => (Box.continuousOn_face_Icc ((continuous_apply i).comp_continuousOn Hc) (hc k)).integrableOn_Icc have Hid : IntegrableOn (fun x => f (i.insertNth d x) i) (Box.Icc (I.face i)) := diff --git a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean index 50035d3ae4173..0aa5977e8f9e6 100644 --- a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean +++ b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean @@ -93,7 +93,7 @@ theorem hasSum_integral_of_dominated_convergence {ι} [Countable ι] {F : ι → simp only [HasSum, ← integral_finset_sum _ fun n _ => hF_integrable n] refine tendsto_integral_filter_of_dominated_convergence (fun a => ∑' n, bound n a) ?_ ?_ bound_integrable h_lim - · exact eventually_of_forall fun s => s.aestronglyMeasurable_sum fun n _ => hF_meas n + · exact Eventually.of_forall fun s => s.aestronglyMeasurable_sum fun n _ => hF_meas n · filter_upwards with s filter_upwards [eventually_countable_forall.2 h_bound, hb_nonneg, bound_summable] with a hFa ha0 has @@ -177,7 +177,7 @@ theorem _root_.Antitone.tendsto_setIntegral (hsm : ∀ i, MeasurableSet (s i)) ( · rw [integrable_indicator_iff (hsm 0)] exact hfi.norm · simp_rw [norm_indicator_eq_indicator_norm] - refine fun n => eventually_of_forall fun x => ?_ + refine fun n => Eventually.of_forall fun x => ?_ exact indicator_le_indicator_of_subset (h_anti (zero_le n)) (fun a => norm_nonneg _) _ · filter_upwards [] with a using le_trans (h_anti.tendsto_indicator _ _ _) (pure_le_nhds _) @@ -293,7 +293,7 @@ theorem continuous_of_dominated_interval {F : X → ℝ → E} {bound : ℝ → (h_cont : ∀ᵐ t ∂μ, t ∈ Ι a b → Continuous fun x => F x t) : Continuous fun x => ∫ t in a..b, F x t ∂μ := continuous_iff_continuousAt.mpr fun _ => - continuousAt_of_dominated_interval (eventually_of_forall hF_meas) (eventually_of_forall h_bound) + continuousAt_of_dominated_interval (Eventually.of_forall hF_meas) (Eventually.of_forall h_bound) bound_integrable <| h_cont.mono fun _ himp hx => (himp hx).continuousAt @@ -395,7 +395,7 @@ theorem continuousAt_parametric_primitive_of_dominated [FirstCountableTopology X have hiF : ∀ {x a₀ b₀}, (∀ᵐ t : ℝ ∂μ.restrict (Ι a b), ‖F x t‖ ≤ bound t) → a₀ ∈ Ioo a b → b₀ ∈ Ioo a b → IntervalIntegrable (F x) μ a₀ b₀ := fun {x a₀ b₀} hx ha₀ hb₀ ↦ - (bound_integrable.mono_set_ae <| eventually_of_forall <| hsub ha₀ hb₀).mono_fun' + (bound_integrable.mono_set_ae <| Eventually.of_forall <| hsub ha₀ hb₀).mono_fun' ((hF_meas x).mono_set <| hsub ha₀ hb₀) (ae_restrict_of_ae_restrict_of_subset (hsub ha₀ hb₀) hx) rw [intervalIntegral.integral_sub, add_assoc, add_sub_cancel, @@ -407,10 +407,10 @@ theorem continuousAt_parametric_primitive_of_dominated [FirstCountableTopology X rw [continuousAt_congr this]; clear this refine (ContinuousAt.add ?_ ?_).add ?_ · exact (intervalIntegral.continuousAt_of_dominated_interval - (eventually_of_forall fun x ↦ (hF_meas x).mono_set <| hsub ha₀ hb₀) + (Eventually.of_forall fun x ↦ (hF_meas x).mono_set <| hsub ha₀ hb₀) (h_bound.mono fun x hx ↦ ae_imp_of_ae_restrict <| ae_restrict_of_ae_restrict_of_subset (hsub ha₀ hb₀) hx) - (bound_integrable.mono_set_ae <| eventually_of_forall <| hsub ha₀ hb₀) <| + (bound_integrable.mono_set_ae <| Eventually.of_forall <| hsub ha₀ hb₀) <| ae_imp_of_ae_restrict <| ae_restrict_of_ae_restrict_of_subset (hsub ha₀ hb₀) h_cont).fst' · refine (?_ : ContinuousAt (fun t ↦ ∫ s in b₀..t, F x₀ s ∂μ) b₀).snd' apply ContinuousWithinAt.continuousAt _ (Icc_mem_nhds hb₀.1 hb₀.2) @@ -577,17 +577,17 @@ theorem continuous_parametric_primitive_of_continuous gcongr · apply setIntegral_mono_set · exact (hf.uncurry_left _).norm.integrableOn_Icc - · exact eventually_of_forall (fun x ↦ norm_nonneg _) + · exact Eventually.of_forall (fun x ↦ norm_nonneg _) · have : Ι b₀ s ⊆ Icc (b₀ - δ) (b₀ + δ) := by apply uIoc_subset_uIcc.trans (uIcc_subset_Icc ?_ ⟨hs.1.le, hs.2.le⟩ ) simp [δpos.le] - exact eventually_of_forall this + exact Eventually.of_forall this · apply setIntegral_mono_set · exact ((hf.uncurry_left _).sub (hf.uncurry_left _)).norm.integrableOn_Icc - · exact eventually_of_forall (fun x ↦ norm_nonneg _) + · exact Eventually.of_forall (fun x ↦ norm_nonneg _) · have : Ι a₀ b₀ ⊆ Icc a b := uIoc_subset_uIcc.trans (uIcc_subset_Icc ⟨a_lt.1.le, lt_b.1.le⟩ ⟨a_lt.2.le, lt_b.2.le⟩) - exact eventually_of_forall this + exact Eventually.of_forall this _ ≤ ∫ t in Icc (b₀ - δ) (b₀ + δ), M + 1 ∂μ + ∫ _t in Icc a b, δ ∂μ := by gcongr · apply setIntegral_mono_on diff --git a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean index 958b5b34a0ce4..20ef2efcbbe01 100644 --- a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean +++ b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean @@ -122,7 +122,7 @@ theorem integrableOn_congr_fun_ae (hst : f =ᵐ[μ.restrict s] g) : theorem IntegrableOn.congr_fun (h : IntegrableOn f s μ) (hst : EqOn f g s) (hs : MeasurableSet s) : IntegrableOn g s μ := - h.congr_fun_ae ((ae_restrict_iff' hs).2 (eventually_of_forall hst)) + h.congr_fun_ae ((ae_restrict_iff' hs).2 (Eventually.of_forall hst)) theorem integrableOn_congr_fun (hst : EqOn f g s) (hs : MeasurableSet s) : IntegrableOn f s μ ↔ IntegrableOn g s μ := @@ -294,7 +294,7 @@ theorem IntegrableOn.of_ae_diff_eq_zero (hf : IntegrableOn f s μ) (ht : NullMea if `t` is measurable. -/ theorem IntegrableOn.of_forall_diff_eq_zero (hf : IntegrableOn f s μ) (ht : MeasurableSet t) (h't : ∀ x ∈ t \ s, f x = 0) : IntegrableOn f t μ := - hf.of_ae_diff_eq_zero ht.nullMeasurableSet (eventually_of_forall h't) + hf.of_ae_diff_eq_zero ht.nullMeasurableSet (Eventually.of_forall h't) /-- If a function is integrable on a set `s` and vanishes almost everywhere on its complement, then it is integrable. -/ @@ -308,7 +308,7 @@ theorem IntegrableOn.integrable_of_ae_not_mem_eq_zero (hf : IntegrableOn f s μ) then it is integrable. -/ theorem IntegrableOn.integrable_of_forall_not_mem_eq_zero (hf : IntegrableOn f s μ) (h't : ∀ x, x ∉ s → f x = 0) : Integrable f μ := - hf.integrable_of_ae_not_mem_eq_zero (eventually_of_forall fun x hx => h't x hx) + hf.integrable_of_ae_not_mem_eq_zero (Eventually.of_forall fun x hx => h't x hx) theorem integrableOn_iff_integrable_of_support_subset (h1s : support f ⊆ s) : IntegrableOn f s μ ↔ Integrable f μ := by @@ -444,7 +444,7 @@ theorem Measure.FiniteAtFilter.integrableAtFilter {l : Filter α} [IsMeasurablyG ⟨s, hsl, hsm, hfm, hμ, hC⟩ refine ⟨s, hsl, ⟨hfm, hasFiniteIntegral_restrict_of_bounded hμ (C := C) ?_⟩⟩ rw [ae_restrict_eq hsm, eventually_inf_principal] - exact eventually_of_forall hC + exact Eventually.of_forall hC theorem Measure.FiniteAtFilter.integrableAtFilter_of_tendsto_ae {l : Filter α} [IsMeasurablyGenerated l] (hfm : StronglyMeasurableAtFilter f l μ) (hμ : μ.FiniteAtFilter l) {b} diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index cbb80865f8332..4b9143054ef0c 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -474,8 +474,8 @@ theorem AECover.integral_tendsto_of_countably_generated [l.IsCountablyGenerated] suffices h : Tendsto (fun i => ∫ x : α, (φ i).indicator f x ∂μ) l (𝓝 (∫ x : α, f x ∂μ)) from by convert h using 2; rw [integral_indicator (hφ.measurableSet _)] tendsto_integral_filter_of_dominated_convergence (fun x => ‖f x‖) - (eventually_of_forall fun i => hfi.aestronglyMeasurable.indicator <| hφ.measurableSet i) - (eventually_of_forall fun i => ae_of_all _ fun x => norm_indicator_le_norm_self _ _) hfi.norm + (Eventually.of_forall fun i => hfi.aestronglyMeasurable.indicator <| hφ.measurableSet i) + (Eventually.of_forall fun i => ae_of_all _ fun x => norm_indicator_le_norm_self _ _) hfi.norm (hφ.ae_tendsto_indicator f) /-- Slight reformulation of diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 53d163b674b03..3e19fb4e62ced 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -199,7 +199,7 @@ theorem mono_set_ae (hf : IntervalIntegrable f μ a b) (h : Ι c d ≤ᵐ[μ] Ι theorem mono_set' (hf : IntervalIntegrable f μ a b) (hsub : Ι c d ⊆ Ι a b) : IntervalIntegrable f μ c d := - hf.mono_set_ae <| eventually_of_forall hsub + hf.mono_set_ae <| Eventually.of_forall hsub theorem mono_fun [NormedAddCommGroup F] {g : ℝ → F} (hf : IntervalIntegrable f μ a b) (hgm : AEStronglyMeasurable g (μ.restrict (Ι a b))) @@ -516,7 +516,7 @@ theorem norm_integral_le_of_norm_le_const_ae {a b C : ℝ} {f : ℝ → E} theorem norm_integral_le_of_norm_le_const {a b C : ℝ} {f : ℝ → E} (h : ∀ x ∈ Ι a b, ‖f x‖ ≤ C) : ‖∫ x in a..b, f x‖ ≤ C * |b - a| := - norm_integral_le_of_norm_le_const_ae <| eventually_of_forall h + norm_integral_le_of_norm_le_const_ae <| Eventually.of_forall h @[simp] nonrec theorem integral_add (hf : IntervalIntegrable f μ a b) (hg : IntervalIntegrable g μ a b) : @@ -1004,7 +1004,7 @@ theorem integral_lt_integral_of_continuousOn_of_le_of_exists_lt {f g : ℝ → have h_eq : f =ᵐ[volume.restrict (Ioc a b)] g := by simp only [← not_le, ← ae_iff] at hlt exact EventuallyLE.antisymm ((ae_restrict_iff' measurableSet_Ioc).2 <| - eventually_of_forall hle) hlt + Eventually.of_forall hle) hlt rw [Measure.restrict_congr_set Ioc_ae_eq_Icc] at h_eq exact fun c hc ↦ (Measure.eqOn_Icc_of_ae_eq volume hab.ne h_eq hfc hgc hc).ge @@ -1017,7 +1017,7 @@ theorem integral_nonneg_of_ae (hab : a ≤ b) (hf : 0 ≤ᵐ[μ] f) : 0 ≤ ∫ integral_nonneg_of_ae_restrict hab <| ae_restrict_of_ae hf theorem integral_nonneg_of_forall (hab : a ≤ b) (hf : ∀ u, 0 ≤ f u) : 0 ≤ ∫ u in a..b, f u ∂μ := - integral_nonneg_of_ae hab <| eventually_of_forall hf + integral_nonneg_of_ae hab <| Eventually.of_forall hf theorem integral_nonneg (hab : a ≤ b) (hf : ∀ u, u ∈ Icc a b → 0 ≤ f u) : 0 ≤ ∫ u in a..b, f u ∂μ := integral_nonneg_of_ae_restrict hab <| (ae_restrict_iff' measurableSet_Icc).mpr <| ae_of_all μ hf diff --git a/Mathlib/MeasureTheory/Integral/Layercake.lean b/Mathlib/MeasureTheory/Integral/Layercake.lean index 9c055c3e3bd82..220f6a09f2e21 100644 --- a/Mathlib/MeasureTheory/Integral/Layercake.lean +++ b/Mathlib/MeasureTheory/Integral/Layercake.lean @@ -343,7 +343,7 @@ theorem lintegral_comp_eq_lintegral_meas_le_mul_of_measurable (μ : Measure α) have meas : MeasurableSet {a | M < f a} := measurableSet_lt measurable_const f_mble have I : ∫⁻ ω in {a | M < f a}ᶜ, ENNReal.ofReal (∫ t in (0).. f ω, g t) ∂μ = ∫⁻ _ in {a | M < f a}ᶜ, 0 ∂μ := by - apply setLIntegral_congr_fun meas.compl (eventually_of_forall (fun s hs ↦ ?_)) + apply setLIntegral_congr_fun meas.compl (Eventually.of_forall (fun s hs ↦ ?_)) have : ∫ (t : ℝ) in (0)..f s, g t = ∫ (t : ℝ) in (0)..f s, 0 := by simp_rw [intervalIntegral.integral_of_le (f_nonneg s)] apply integral_congr_ae @@ -364,7 +364,7 @@ theorem lintegral_comp_eq_lintegral_meas_le_mul_of_measurable (μ : Measure α) simp [ht] have B2 : ∫⁻ t in Ioi M, μ {a : α | t ≤ f a} * ENNReal.ofReal (g t) = ∫⁻ t in Ioi M, ν {a : α | t ≤ f a} * ENNReal.ofReal (g t) := by - apply setLIntegral_congr_fun measurableSet_Ioi (eventually_of_forall (fun t ht ↦ ?_)) + apply setLIntegral_congr_fun measurableSet_Ioi (Eventually.of_forall (fun t ht ↦ ?_)) rw [Measure.restrict_apply (measurableSet_le measurable_const f_mble)] congr 3 exact (inter_eq_left.2 (fun a ha ↦ (mem_Ioi.1 ht).trans_le ha)).symm @@ -448,7 +448,7 @@ theorem lintegral_eq_lintegral_meas_le (μ : Measure α) (f_nn : 0 ≤ᵐ[μ] f) intervalIntegrable_const have key := lintegral_comp_eq_lintegral_meas_le_mul μ f_nn f_mble cst_intble - (eventually_of_forall fun _ => zero_le_one) + (Eventually.of_forall fun _ => zero_le_one) simp_rw [cst, ENNReal.ofReal_one, mul_one] at key rw [← key] congr with ω @@ -528,9 +528,9 @@ theorem Integrable.integral_eq_integral_meas_lt (fun t ↦ ENNReal.toReal (μ {a : α | t < f a})) ?_ ?_ · rw [aux] congr 1 - apply setLIntegral_congr_fun measurableSet_Ioi (eventually_of_forall _) + apply setLIntegral_congr_fun measurableSet_Ioi (Eventually.of_forall _) exact fun t t_pos ↦ ENNReal.ofReal_toReal (rhs_integrand_finite t t_pos).ne - · exact eventually_of_forall (fun x ↦ by simp only [Pi.zero_apply, ENNReal.toReal_nonneg]) + · exact Eventually.of_forall (fun x ↦ by simp only [Pi.zero_apply, ENNReal.toReal_nonneg]) · apply Measurable.aestronglyMeasurable refine Measurable.ennreal_toReal ?_ exact Antitone.measurable (fun _ _ hst ↦ measure_mono (fun _ h ↦ lt_of_le_of_lt hst h)) @@ -549,7 +549,7 @@ lemma Integrable.integral_eq_integral_Ioc_meas_le {f : α → ℝ} {M : ℝ} rw [f_intble.integral_eq_integral_meas_le f_nn] rw [setIntegral_eq_of_subset_of_ae_diff_eq_zero measurableSet_Ioi.nullMeasurableSet Ioc_subset_Ioi_self ?_] - apply eventually_of_forall (fun t ht ↦ ?_) + apply Eventually.of_forall (fun t ht ↦ ?_) have htM : M < t := by simp_all only [mem_diff, mem_Ioi, mem_Ioc, not_and, not_le] have obs : μ {a | M < f a} = 0 := by rw [measure_zero_iff_ae_nmem] diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index 78b990a4b3679..f85ddf5e5f9b9 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -339,7 +339,7 @@ theorem lintegral_nnnorm_eq_of_ae_nonneg {f : α → ℝ} (h_nonneg : 0 ≤ᵐ[ theorem lintegral_nnnorm_eq_of_nonneg {f : α → ℝ} (h_nonneg : 0 ≤ f) : ∫⁻ x, ‖f x‖₊ ∂μ = ∫⁻ x, ENNReal.ofReal (f x) ∂μ := - lintegral_nnnorm_eq_of_ae_nonneg (Filter.eventually_of_forall h_nonneg) + lintegral_nnnorm_eq_of_ae_nonneg (Filter.Eventually.of_forall h_nonneg) /-- **Monotone convergence theorem** -- sometimes called **Beppo-Levi convergence**. See `lintegral_iSup_directed` for a more general form. -/ @@ -1410,7 +1410,7 @@ theorem _root_.MeasurableEmbedding.lintegral_map [MeasurableSpace β] {g : α exact le_iSup_of_le (comp f₀ g hg.measurable) (by exact le_iSup (α := ℝ≥0∞) _ this) · rw [← f₀.extend_comp_eq hg (const _ 0), ← SimpleFunc.lintegral_map, ← SimpleFunc.lintegral_eq_lintegral, ← lintegral] - refine lintegral_mono_ae (hg.ae_map_iff.2 <| eventually_of_forall fun x => ?_) + refine lintegral_mono_ae (hg.ae_map_iff.2 <| Eventually.of_forall fun x => ?_) exact (extend_apply _ _ _ _).trans_le (hf₀ _) /-- The `lintegral` transforms appropriately under a measurable equivalence `g : α ≃ᵐ β`. @@ -1753,7 +1753,7 @@ lemma tendsto_of_lintegral_tendsto_of_monotone {α : Type*} {mα : MeasurableSpa · exact fun n ↦ (M n).aemeasurable · simp_rw [Int_eq] exact hf_tendsto - · exact eventually_of_forall (fun x ↦ monotone_nat_of_le_succ (fun n ↦ le_max_right _ _)) + · exact Eventually.of_forall (fun x ↦ monotone_nat_of_le_succ (fun n ↦ le_max_right _ _)) · filter_upwards [h_bound, I'] with x h'x hx n using (hx n).trans (h'x n) filter_upwards [this, I', h_bound] with x hx h'x h''x exact tendsto_of_tendsto_of_tendsto_of_le_of_le hx tendsto_const_nhds h'x h''x @@ -2029,10 +2029,10 @@ lemma tendsto_measure_of_ae_tendsto_indicator {μ : Measure α} (A_mble : Measur simp_rw [← MeasureTheory.lintegral_indicator_one A_mble, ← MeasureTheory.lintegral_indicator_one (As_mble _)] refine tendsto_lintegral_filter_of_dominated_convergence (B.indicator (1 : α → ℝ≥0∞)) - (eventually_of_forall ?_) ?_ ?_ ?_ + (Eventually.of_forall ?_) ?_ ?_ ?_ · exact fun i ↦ Measurable.indicator measurable_const (As_mble i) · filter_upwards [As_le_B] with i hi - exact eventually_of_forall (fun x ↦ indicator_le_indicator_of_subset hi (by simp) x) + exact Eventually.of_forall (fun x ↦ indicator_le_indicator_of_subset hi (by simp) x) · rwa [← lintegral_indicator_one B_mble] at B_finmeas · simpa only [Pi.one_def, tendsto_indicator_const_apply_iff_eventually] using h_lim @@ -2044,7 +2044,7 @@ lemma tendsto_measure_of_ae_tendsto_indicator_of_isFiniteMeasure (As_mble : ∀ i, MeasurableSet (As i)) (h_lim : ∀ᵐ x ∂μ, ∀ᶠ i in L, x ∈ As i ↔ x ∈ A) : Tendsto (fun i ↦ μ (As i)) L (𝓝 (μ A)) := tendsto_measure_of_ae_tendsto_indicator L A_mble As_mble MeasurableSet.univ - (measure_ne_top μ univ) (eventually_of_forall (fun i ↦ subset_univ (As i))) h_lim + (measure_ne_top μ univ) (Eventually.of_forall (fun i ↦ subset_univ (As i))) h_lim end TendstoIndicator -- section diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index a3e735eec79cc..8db2aa64138ec 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -144,8 +144,8 @@ theorem tendsto_setIntegral_peak_smul_of_integrableOn_of_tendsto_aux _ ≤ ∫ x in t, ‖φ i x‖ * δ ∂μ := by apply setIntegral_mono_set · exact I.norm.mul_const _ - · exact eventually_of_forall fun x => mul_nonneg (norm_nonneg _) δpos.le - · exact eventually_of_forall ut + · exact Eventually.of_forall fun x => mul_nonneg (norm_nonneg _) δpos.le + · exact Eventually.of_forall ut _ = ∫ x in t, φ i x * δ ∂μ := by apply setIntegral_congr ht fun x hx => ?_ rw [Real.norm_of_nonneg (hφpos _ (hts hx))] @@ -319,7 +319,7 @@ theorem tendsto_setIntegral_pow_smul_of_unique_maximum_of_isCompact_of_measure_n · intro x hx exact pow_le_pow_left t'_pos.le (le_of_lt (hv hx)) _ _ ≤ ∫ y in s, c y ^ n ∂μ := - setIntegral_mono_set (I n) (J n) (eventually_of_forall inter_subset_right) + setIntegral_mono_set (I n) (J n) (Eventually.of_forall inter_subset_right) simp_rw [φ, ← div_eq_inv_mul, div_pow, div_div] apply div_le_div (pow_nonneg t_pos n) _ _ B · exact pow_le_pow_left (hnc _ hx.1) (ht x hx) _ @@ -341,10 +341,10 @@ theorem tendsto_setIntegral_pow_smul_of_unique_maximum_of_isCompact_of_measure_n have B : Tendsto (fun i ↦ ∫ (x : α) in s, φ i x ∂μ) atTop (𝓝 1) := tendsto_const_nhds.congr (fun n ↦ (hiφ n).symm) have C : ∀ᶠ (i : ℕ) in atTop, AEStronglyMeasurable (fun x ↦ φ i x) (μ.restrict s) := by - apply eventually_of_forall (fun n ↦ ((I n).const_mul _).aestronglyMeasurable) + apply Eventually.of_forall (fun n ↦ ((I n).const_mul _).aestronglyMeasurable) exact tendsto_setIntegral_peak_smul_of_integrableOn_of_tendsto hs.measurableSet hs.measurableSet (Subset.rfl) (self_mem_nhdsWithin) - hs.measure_lt_top.ne (eventually_of_forall hnφ) A B C hmg hcg + hs.measure_lt_top.ne (Eventually.of_forall hnφ) A B C hmg hcg convert this simp_rw [φ, ← smul_smul, integral_smul] diff --git a/Mathlib/MeasureTheory/Integral/Periodic.lean b/Mathlib/MeasureTheory/Integral/Periodic.lean index 00da4d383fa2e..892fa83114568 100644 --- a/Mathlib/MeasureTheory/Integral/Periodic.lean +++ b/Mathlib/MeasureTheory/Integral/Periodic.lean @@ -112,7 +112,7 @@ theorem volume_closedBall {x : AddCircle T} (ε : ℝ) : · simp [I, hε, min_eq_left (by linarith : T ≤ 2 * ε)] instance : IsUnifLocDoublingMeasure (volume : Measure (AddCircle T)) := by - refine ⟨⟨Real.toNNReal 2, Filter.eventually_of_forall fun ε x => ?_⟩⟩ + refine ⟨⟨Real.toNNReal 2, Filter.Eventually.of_forall fun ε x => ?_⟩⟩ simp only [volume_closedBall] erw [← ENNReal.ofReal_mul zero_le_two] apply ENNReal.ofReal_le_ofReal diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 44f6063be9b79..9982a8ab04974 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -84,14 +84,14 @@ alias set_integral_congr_ae := setIntegral_congr_ae theorem setIntegral_congr₀ (hs : NullMeasurableSet s μ) (h : EqOn f g s) : ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := - setIntegral_congr_ae₀ hs <| eventually_of_forall h + setIntegral_congr_ae₀ hs <| Eventually.of_forall h @[deprecated (since := "2024-04-17")] alias set_integral_congr₀ := setIntegral_congr₀ theorem setIntegral_congr (hs : MeasurableSet s) (h : EqOn f g s) : ∫ x in s, f x ∂μ = ∫ x in s, g x ∂μ := - setIntegral_congr_ae hs <| eventually_of_forall h + setIntegral_congr_ae hs <| Eventually.of_forall h @[deprecated (since := "2024-04-17")] alias set_integral_congr := setIntegral_congr @@ -312,7 +312,7 @@ alias set_integral_eq_zero_of_ae_eq_zero := setIntegral_eq_zero_of_ae_eq_zero theorem setIntegral_eq_zero_of_forall_eq_zero (ht_eq : ∀ x ∈ t, f x = 0) : ∫ x in t, f x ∂μ = 0 := - setIntegral_eq_zero_of_ae_eq_zero (eventually_of_forall ht_eq) + setIntegral_eq_zero_of_ae_eq_zero (Eventually.of_forall ht_eq) @[deprecated (since := "2024-04-17")] alias set_integral_eq_zero_of_forall_eq_zero := setIntegral_eq_zero_of_forall_eq_zero @@ -353,7 +353,7 @@ theorem integral_union_eq_left_of_ae (ht_eq : ∀ᵐ x ∂μ.restrict t, f x = 0 theorem integral_union_eq_left_of_forall₀ {f : X → E} (ht : NullMeasurableSet t μ) (ht_eq : ∀ x ∈ t, f x = 0) : ∫ x in s ∪ t, f x ∂μ = ∫ x in s, f x ∂μ := - integral_union_eq_left_of_ae ((ae_restrict_iff'₀ ht).2 (eventually_of_forall ht_eq)) + integral_union_eq_left_of_ae ((ae_restrict_iff'₀ ht).2 (Eventually.of_forall ht_eq)) theorem integral_union_eq_left_of_forall {f : X → E} (ht : MeasurableSet t) (ht_eq : ∀ x ∈ t, f x = 0) : ∫ x in s ∪ t, f x ∂μ = ∫ x in s, f x ∂μ := @@ -418,7 +418,7 @@ and `t` coincide if `t` is measurable. -/ theorem setIntegral_eq_of_subset_of_forall_diff_eq_zero (ht : MeasurableSet t) (hts : s ⊆ t) (h't : ∀ x ∈ t \ s, f x = 0) : ∫ x in t, f x ∂μ = ∫ x in s, f x ∂μ := setIntegral_eq_of_subset_of_ae_diff_eq_zero ht.nullMeasurableSet hts - (eventually_of_forall fun x hx => h't x hx) + (Eventually.of_forall fun x hx => h't x hx) @[deprecated (since := "2024-04-17")] alias set_integral_eq_of_subset_of_forall_diff_eq_zero := @@ -440,7 +440,7 @@ alias set_integral_eq_integral_of_ae_compl_eq_zero := setIntegral_eq_integral_of whole space. -/ theorem setIntegral_eq_integral_of_forall_compl_eq_zero (h : ∀ x, x ∉ s → f x = 0) : ∫ x in s, f x ∂μ = ∫ x, f x ∂μ := - setIntegral_eq_integral_of_ae_compl_eq_zero (eventually_of_forall h) + setIntegral_eq_integral_of_ae_compl_eq_zero (Eventually.of_forall h) @[deprecated (since := "2024-04-17")] alias set_integral_eq_integral_of_forall_compl_eq_zero := @@ -605,14 +605,14 @@ alias norm_set_integral_le_of_norm_le_const_ae'' := norm_setIntegral_le_of_norm_ theorem norm_setIntegral_le_of_norm_le_const {C : ℝ} (hs : μ s < ∞) (hC : ∀ x ∈ s, ‖f x‖ ≤ C) (hfm : AEStronglyMeasurable f (μ.restrict s)) : ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).toReal := - norm_setIntegral_le_of_norm_le_const_ae' hs (eventually_of_forall hC) hfm + norm_setIntegral_le_of_norm_le_const_ae' hs (Eventually.of_forall hC) hfm @[deprecated (since := "2024-04-17")] alias norm_set_integral_le_of_norm_le_const := norm_setIntegral_le_of_norm_le_const theorem norm_setIntegral_le_of_norm_le_const' {C : ℝ} (hs : μ s < ∞) (hsm : MeasurableSet s) (hC : ∀ x ∈ s, ‖f x‖ ≤ C) : ‖∫ x in s, f x ∂μ‖ ≤ C * (μ s).toReal := - norm_setIntegral_le_of_norm_le_const_ae'' hs hsm <| eventually_of_forall hC + norm_setIntegral_le_of_norm_le_const_ae'' hs hsm <| Eventually.of_forall hC @[deprecated (since := "2024-04-17")] alias norm_set_integral_le_of_norm_le_const' := norm_setIntegral_le_of_norm_le_const' @@ -648,7 +648,7 @@ theorem setIntegral_gt_gt {R : ℝ} {f : X → ℝ} (hR : 0 ≤ R) rwa [Set.inter_eq_self_of_subset_right] exact fun x hx => Ne.symm (ne_of_lt <| sub_pos.2 hx) · rw [Pi.zero_def, EventuallyLE, ae_restrict_iff₀] - · exact eventually_of_forall fun x hx => sub_nonneg.2 <| le_of_lt hx + · exact Eventually.of_forall fun x hx => sub_nonneg.2 <| le_of_lt hx · exact nullMeasurableSet_le aemeasurable_zero (hfint.1.aemeasurable.sub aemeasurable_const) · exact Integrable.sub hfint this @@ -896,7 +896,7 @@ lemma integral_le_measure {f : X → ℝ} {s : Set X} apply ENNReal.ofReal_le_ofReal exact integral_mono H g_int (fun x ↦ le_max_left _ _) apply this.trans - rw [ofReal_integral_eq_lintegral_ofReal g_int (eventually_of_forall (fun x ↦ le_max_right _ _))] + rw [ofReal_integral_eq_lintegral_ofReal g_int (Eventually.of_forall (fun x ↦ le_max_right _ _))] apply lintegral_le_meas · intro x apply ENNReal.ofReal_le_of_le_toReal @@ -1563,7 +1563,7 @@ lemma continuousOn_integral_bilinear_of_locally_integrable_of_compact_support apply StronglyMeasurable.aestronglyMeasurable apply Continuous.stronglyMeasurable_of_support_subset_isCompact (A p hp) hk apply support_subset_iff'.2 (fun y hy ↦ hfs p y hp hy) - · apply eventually_of_forall (fun y ↦ (le_opNorm₂ L (g y) (f p y)).trans ?_) + · apply Eventually.of_forall (fun y ↦ (le_opNorm₂ L (g y) (f p y)).trans ?_) gcongr apply hC filter_upwards [v_mem, self_mem_nhdsWithin] with p hp h'p @@ -1578,7 +1578,7 @@ lemma continuousOn_integral_bilinear_of_locally_integrable_of_compact_support _ = ‖∫ x in k, L (g x) (f p x) - L (g x) (f q x) ∂μ‖ := by rw [integral_sub (I p h'p) (I q hq)] _ ≤ ∫ x in k, ‖L (g x) (f p x) - L (g x) (f q x)‖ ∂μ := norm_integral_le_integral_norm _ _ ≤ ∫ x in k, ‖L‖ * ‖g x‖ * δ ∂μ := by - apply integral_mono_of_nonneg (eventually_of_forall (fun y ↦ by positivity)) + apply integral_mono_of_nonneg (Eventually.of_forall (fun y ↦ by positivity)) · exact (hg.norm.const_mul _).mul_const _ · filter_upwards with y by_cases hy : y ∈ k diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index c177aeb5d4943..55416b4d8f253 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -1384,7 +1384,7 @@ theorem tendsto_setToFun_approxOn_of_measurable (hT : DominatedFinMeasAdditive Tendsto (fun n => setToFun μ T hT (SimpleFunc.approxOn f hfm s y₀ h₀ n)) atTop (𝓝 <| setToFun μ T hT f) := tendsto_setToFun_of_L1 hT _ hfi - (eventually_of_forall (SimpleFunc.integrable_approxOn hfm hfi h₀ h₀i)) + (Eventually.of_forall (SimpleFunc.integrable_approxOn hfm hfi h₀ h₀i)) (SimpleFunc.tendsto_approxOn_L1_nnnorm hfm _ hs (hfi.sub h₀i).2) theorem tendsto_setToFun_approxOn_of_measurable_of_range_subset @@ -1394,7 +1394,7 @@ theorem tendsto_setToFun_approxOn_of_measurable_of_range_subset Tendsto (fun n => setToFun μ T hT (SimpleFunc.approxOn f fmeas s 0 (hs <| by simp) n)) atTop (𝓝 <| setToFun μ T hT f) := by refine tendsto_setToFun_approxOn_of_measurable hT hf fmeas ?_ _ (integrable_zero _ _ _) - exact eventually_of_forall fun x => subset_closure (hs (Set.mem_union_left _ (mem_range_self _))) + exact Eventually.of_forall fun x => subset_closure (hs (Set.mem_union_left _ (mem_range_self _))) /-- Auxiliary lemma for `setToFun_congr_measure`: the function sending `f : α →₁[μ] G` to `f : α →₁[μ'] G` is continuous when `μ' ≤ c' • μ` for `c' ≠ ∞`. -/ @@ -1639,8 +1639,8 @@ theorem continuous_setToFun_of_dominated (hT : DominatedFinMeasAdditive μ T C) (h_bound : ∀ x, ∀ᵐ a ∂μ, ‖fs x a‖ ≤ bound a) (bound_integrable : Integrable bound μ) (h_cont : ∀ᵐ a ∂μ, Continuous fun x => fs x a) : Continuous fun x => setToFun μ T hT (fs x) := continuous_iff_continuousAt.mpr fun x₀ => - continuousAt_setToFun_of_dominated hT (eventually_of_forall hfs_meas) - (eventually_of_forall h_bound) ‹_› <| + continuousAt_setToFun_of_dominated hT (Eventually.of_forall hfs_meas) + (Eventually.of_forall h_bound) ‹_› <| h_cont.mono fun _ => Continuous.continuousAt end Function diff --git a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean index 7894d2996f639..15a4b3437ddd3 100644 --- a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean +++ b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean @@ -297,9 +297,9 @@ theorem exists_lt_lowerSemicontinuous_integral_gt_nnreal [SigmaFinite μ] (f : _ < ENNReal.toReal (∫⁻ a : α, f a ∂μ) + ε := add_lt_add_left hδε _ _ = (∫⁻ a : α, ENNReal.ofReal ↑(f a) ∂μ).toReal + ε := by simp - · apply Filter.eventually_of_forall fun x => _; simp + · apply Filter.Eventually.of_forall fun x => _; simp · exact fmeas.coe_nnreal_real.aestronglyMeasurable - · apply Filter.eventually_of_forall fun x => _; simp + · apply Filter.Eventually.of_forall fun x => _; simp · apply gcont.measurable.ennreal_toReal.aemeasurable.aestronglyMeasurable /-! ### Upper semicontinuous lower bound for nonnegative functions -/ @@ -421,16 +421,16 @@ theorem exists_upperSemicontinuous_le_integral_le (f : α → ℝ≥0) refine ⟨g, gf, gcont, ?_, ?_⟩ · refine Integrable.mono fint gcont.measurable.coe_nnreal_real.aemeasurable.aestronglyMeasurable ?_ - exact Filter.eventually_of_forall fun x => by simp [gf x] + exact Filter.Eventually.of_forall fun x => by simp [gf x] · rw [integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae] · rw [sub_le_iff_le_add] convert ENNReal.toReal_mono _ gint · simp · rw [ENNReal.toReal_add Ig.ne ENNReal.coe_ne_top]; simp · simpa using Ig.ne - · apply Filter.eventually_of_forall; simp + · apply Filter.Eventually.of_forall; simp · exact gcont.measurable.coe_nnreal_real.aemeasurable.aestronglyMeasurable - · apply Filter.eventually_of_forall; simp + · apply Filter.Eventually.of_forall; simp · exact fint.aestronglyMeasurable /-! ### Vitali-Carathéodory theorem -/ diff --git a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean index c1459922a78e9..f9ad11c6fbc66 100644 --- a/Mathlib/MeasureTheory/Measure/AEMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/AEMeasurable.lean @@ -208,7 +208,7 @@ theorem subtype_mk (h : AEMeasurable f μ) {s : Set β} {hfs : ∀ x, f x ∈ s} AEMeasurable (codRestrict f s hfs) μ := by nontriviality α; inhabit α obtain ⟨g, g_meas, hg, fg⟩ : ∃ g : α → β, Measurable g ∧ range g ⊆ s ∧ f =ᵐ[μ] g := - h.exists_ae_eq_range_subset (eventually_of_forall hfs) ⟨_, hfs default⟩ + h.exists_ae_eq_range_subset (Eventually.of_forall hfs) ⟨_, hfs default⟩ refine ⟨codRestrict g s fun x => hg (mem_range_self _), Measurable.subtype_mk g_meas, ?_⟩ filter_upwards [fg] with x hx simpa [Subtype.ext_iff] diff --git a/Mathlib/MeasureTheory/Measure/ContinuousPreimage.lean b/Mathlib/MeasureTheory/Measure/ContinuousPreimage.lean index 2de20eb3e8ac1..e6124f9f8875b 100644 --- a/Mathlib/MeasureTheory/Measure/ContinuousPreimage.lean +++ b/Mathlib/MeasureTheory/Measure/ContinuousPreimage.lean @@ -113,7 +113,7 @@ theorem isClosed_setOf_preimage_ae_eq {f : Z → C(X, Y)} (hf : Continuous f) apply gt_mem_nhds rwa [pos_iff_ne_zero, ne_eq, measure_symmDiff_eq_zero_iff] filter_upwards [(tendsto_measure_symmDiff_preimage_nhds_zero (hf.tendsto z) - (eventually_of_forall hfm) (hfm z) htm ht).eventually hz] with w hw + (.of_forall hfm) (hfm z) htm ht).eventually hz] with w hw intro (hw' : f w ⁻¹' t =ᵐ[μ] s) rw [measure_congr (hw'.symmDiff (ae_eq_refl _)), symmDiff_comm] at hw exact hw.false diff --git a/Mathlib/MeasureTheory/Measure/DiracProba.lean b/Mathlib/MeasureTheory/Measure/DiracProba.lean index 4dd9303f1e13b..18bc9a7f40a5e 100644 --- a/Mathlib/MeasureTheory/Measure/DiracProba.lean +++ b/Mathlib/MeasureTheory/Measure/DiracProba.lean @@ -101,7 +101,7 @@ lemma not_tendsto_diracProba_of_not_tendsto [CompletelyRegularSpace X] {x : X} ( lintegral_dirac' _ (measurable_coe_nnreal_ennreal_iff.mpr f.continuous.measurable)] apply not_tendsto_iff_exists_frequently_nmem.mpr refine ⟨Ioi 0, Ioi_mem_nhds (by simp only [ENNReal.coe_one, zero_lt_one]), - hU.mp (eventually_of_forall ?_)⟩ + hU.mp (Eventually.of_forall ?_)⟩ intro x x_notin_U rw [f_vanishes_outside x (compl_subset_compl.mpr (show interior U ⊆ U from interior_subset) x_notin_U)] diff --git a/Mathlib/MeasureTheory/Measure/EverywherePos.lean b/Mathlib/MeasureTheory/Measure/EverywherePos.lean index 753d1e7f9e58c..c3975a2f9c5e2 100644 --- a/Mathlib/MeasureTheory/Measure/EverywherePos.lean +++ b/Mathlib/MeasureTheory/Measure/EverywherePos.lean @@ -249,7 +249,7 @@ lemma IsEverywherePos.IsGdelta_of_isMulLeftInvariant have J : x * (y i) ⁻¹ ∈ V i := by simpa [← hvy i] using hv i exact I J have B : μ (((x * z ⁻¹) • k) \ k) = 0 := - le_antisymm (ge_of_tendsto u_lim (eventually_of_forall A)) bot_le + le_antisymm (ge_of_tendsto u_lim (Eventually.of_forall A)) bot_le have C : μ (k \ (z * x⁻¹) • k) = 0 := by have : μ ((z * x⁻¹) • (((x * z ⁻¹) • k) \ k)) = 0 := by rwa [measure_smul] rw [← this, smul_set_sdiff, smul_smul] diff --git a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean index ac8e1902e55e3..049c09311b973 100644 --- a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean @@ -540,8 +540,8 @@ theorem tendsto_lintegral_nn_of_le_const (μ : FiniteMeasure Ω) {fs : ℕ → (fs_lim : ∀ ω, Tendsto (fun n ↦ fs n ω) atTop (𝓝 (f ω))) : Tendsto (fun n ↦ ∫⁻ ω, fs n ω ∂(μ : Measure Ω)) atTop (𝓝 (∫⁻ ω, f ω ∂(μ : Measure Ω))) := tendsto_lintegral_nn_filter_of_le_const μ - (eventually_of_forall fun n ↦ eventually_of_forall (fs_le_const n)) - (eventually_of_forall fs_lim) + (.of_forall fun n ↦ .of_forall (fs_le_const n)) + (.of_forall fs_lim) /-- A bounded convergence theorem for a finite measure: If bounded continuous non-negative functions are uniformly bounded by a constant and tend to a @@ -580,8 +580,8 @@ theorem tendsto_testAgainstNN_of_le_const {μ : FiniteMeasure Ω} {fs : ℕ → (fs_lim : ∀ ω, Tendsto (fun n ↦ fs n ω) atTop (𝓝 (f ω))) : Tendsto (fun n ↦ μ.testAgainstNN (fs n)) atTop (𝓝 (μ.testAgainstNN f)) := tendsto_testAgainstNN_filter_of_le_const - (eventually_of_forall fun n ↦ eventually_of_forall (fs_le_const n)) - (eventually_of_forall fs_lim) + (.of_forall fun n ↦ .of_forall (fs_le_const n)) + (.of_forall fs_lim) end FiniteMeasureBoundedConvergence @@ -611,8 +611,8 @@ theorem tendsto_of_forall_integral_tendsto {γ : Type*} {F : Filter γ} {μs : have f₀_eq : ⇑f₀ = ((↑) : ℝ≥0 → ℝ) ∘ ⇑f := rfl have f₀_nn : 0 ≤ ⇑f₀ := fun _ ↦ by simp only [f₀_eq, Pi.zero_apply, Function.comp_apply, NNReal.zero_le_coe] - have f₀_ae_nn : 0 ≤ᵐ[(μ : Measure Ω)] ⇑f₀ := eventually_of_forall f₀_nn - have f₀_ae_nns : ∀ i, 0 ≤ᵐ[(μs i : Measure Ω)] ⇑f₀ := fun i ↦ eventually_of_forall f₀_nn + have f₀_ae_nn : 0 ≤ᵐ[(μ : Measure Ω)] ⇑f₀ := .of_forall f₀_nn + have f₀_ae_nns : ∀ i, 0 ≤ᵐ[(μs i : Measure Ω)] ⇑f₀ := fun i ↦ .of_forall f₀_nn have aux := integral_eq_lintegral_of_nonneg_ae f₀_ae_nn f₀.continuous.measurable.aestronglyMeasurable have auxs := fun i ↦ diff --git a/Mathlib/MeasureTheory/Measure/Haar/Disintegration.lean b/Mathlib/MeasureTheory/Measure/Haar/Disintegration.lean index df47302132fae..1a085463ecac8 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Disintegration.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Disintegration.lean @@ -146,6 +146,6 @@ suffices to check it almost everywhere along all translates of a given vector su instance of a disintegration argument for additive Haar measures. -/ lemma ae_mem_of_ae_add_linearMap_mem [LocallyCompactSpace F] {s : Set F} (hs : MeasurableSet s) (h : ∀ y, ∀ᵐ x ∂μ, y + L x ∈ s) : ∀ᵐ y ∂ν, y ∈ s := - (ae_ae_add_linearMap_mem_iff L μ ν hs).1 (Filter.eventually_of_forall h) + (ae_ae_add_linearMap_mem_iff L μ ν hs).1 (Filter.Eventually.of_forall h) end MeasureTheory diff --git a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean index cd5d79900442d..51a8903fe9a11 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Unique.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Unique.lean @@ -81,7 +81,7 @@ lemma IsCompact.measure_eq_biInf_integral_hasCompactSupport · simp only [le_iInf_iff] intro f f_cont f_comp fk f_nonneg apply (f_cont.integrable_of_hasCompactSupport f_comp).measure_le_integral - · exact eventually_of_forall f_nonneg + · exact Eventually.of_forall f_nonneg · exact fun x hx ↦ by simp [fk hx] · apply le_of_forall_lt' (fun r hr ↦ ?_) simp only [iInf_lt_iff, exists_prop, exists_and_left] @@ -468,7 +468,7 @@ lemma measure_preimage_isMulLeftInvariant_eq_smul_of_hasCompactSupport · exact fun n ↦ (vf_cont n).aestronglyMeasurable · apply IntegrableOn.integrable_indicator _ (isClosed_tsupport f).measurableSet simpa using IsCompact.measure_lt_top h'f - · refine fun n ↦ eventually_of_forall (fun x ↦ ?_) + · refine fun n ↦ Eventually.of_forall (fun x ↦ ?_) by_cases hx : x ∈ tsupport f · simp only [v, Real.norm_eq_abs, NNReal.abs_eq, hx, indicator_of_mem] norm_cast diff --git a/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean b/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean index b3d825fc3b697..ed741f37cbf85 100644 --- a/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean +++ b/Mathlib/MeasureTheory/Measure/HasOuterApproxClosed.lean @@ -59,7 +59,7 @@ theorem tendsto_lintegral_nn_filter_of_le_const {ι : Type*} {L : Filter ι} [L. (fs_lim : ∀ᵐ ω : Ω ∂μ, Tendsto (fun i ↦ fs i ω) L (𝓝 (f ω))) : Tendsto (fun i ↦ ∫⁻ ω, fs i ω ∂μ) L (𝓝 (∫⁻ ω, f ω ∂μ)) := by refine tendsto_lintegral_filter_of_dominated_convergence (fun _ ↦ c) - (eventually_of_forall fun i ↦ (ENNReal.continuous_coe.comp (fs i).continuous).measurable) ?_ + (Eventually.of_forall fun i ↦ (ENNReal.continuous_coe.comp (fs i).continuous).measurable) ?_ (@lintegral_const_lt_top _ _ μ _ _ (@ENNReal.coe_ne_top c)).ne ?_ · simpa only [Function.comp_apply, ENNReal.coe_le_coe] using fs_le_const · simpa only [Function.comp_apply, ENNReal.tendsto_coe] using fs_lim @@ -100,7 +100,7 @@ theorem measure_of_cont_bdd_of_tendsto_indicator rw [tendsto_pi_nhds] at fs_lim exact fun ω ↦ fs_lim ω apply measure_of_cont_bdd_of_tendsto_filter_indicator μ E_mble fs - (eventually_of_forall fun n ↦ eventually_of_forall (fs_bdd n)) (eventually_of_forall fs_lim') + (Eventually.of_forall fun n ↦ Eventually.of_forall (fs_bdd n)) (Eventually.of_forall fs_lim') /-- The integrals of thickened indicators of a closed set against a finite measure tend to the measure of the closed set if the thickening radii tend to zero. -/ diff --git a/Mathlib/MeasureTheory/Measure/Hausdorff.lean b/Mathlib/MeasureTheory/Measure/Hausdorff.lean index 28921d6562aa6..2bbc642378587 100644 --- a/Mathlib/MeasureTheory/Measure/Hausdorff.lean +++ b/Mathlib/MeasureTheory/Measure/Hausdorff.lean @@ -279,7 +279,7 @@ theorem tendsto_pre (m : Set X → ℝ≥0∞) (s : Set X) : theorem tendsto_pre_nat (m : Set X → ℝ≥0∞) (s : Set X) : Tendsto (fun n : ℕ => pre m n⁻¹ s) atTop (𝓝 <| mkMetric' m s) := by refine (tendsto_pre m s).comp (tendsto_inf.2 ⟨ENNReal.tendsto_inv_nat_nhds_zero, ?_⟩) - refine tendsto_principal.2 (eventually_of_forall fun n => ?_) + refine tendsto_principal.2 (Eventually.of_forall fun n => ?_) simp theorem eq_iSup_nat (m : Set X → ℝ≥0∞) : mkMetric' m = ⨆ n : ℕ, mkMetric'.pre m n⁻¹ := by diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 666c2b7b7b87f..7411012c4e1c6 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -601,7 +601,7 @@ theorem tendsto_addHaar_inter_smul_zero_of_density_zero_aux1 (s : Set E) (x : E) have A : Tendsto (fun r : ℝ => μ (s ∩ ({x} + r • t)) / μ (closedBall x r)) (𝓝[>] 0) (𝓝 0) := by apply tendsto_of_tendsto_of_tendsto_of_le_of_le' tendsto_const_nhds h - (eventually_of_forall fun b => zero_le _) + (Eventually.of_forall fun b => zero_le _) filter_upwards [self_mem_nhdsWithin] rintro r (rpos : 0 < r) rw [← affinity_unitClosedBall rpos.le, singleton_add, ← image_vadd] @@ -798,7 +798,7 @@ theorem tendsto_addHaar_inter_smul_one_of_density_one (s : Set E) (x : E) tendsto_addHaar_inter_smul_one_of_density_one_aux μ _ (measurableSet_toMeasurable _ _) _ _ t ht h't h''t apply tendsto_of_tendsto_of_tendsto_of_le_of_le' h tendsto_const_nhds - · refine eventually_of_forall fun r ↦ ?_ + · refine Eventually.of_forall fun r ↦ ?_ gcongr apply subset_toMeasurable · filter_upwards [self_mem_nhdsWithin] diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean index 50a9a66664f98..f4d39f02991d6 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean @@ -35,7 +35,7 @@ theorem volume_regionBetween_eq_integral [SigmaFinite μ] (f_int : IntegrableOn (g_int : IntegrableOn g s μ) (hs : MeasurableSet s) (hfg : ∀ x ∈ s, f x ≤ g x) : μ.prod volume (regionBetween f g s) = ENNReal.ofReal (∫ y in s, (g - f) y ∂μ) := volume_regionBetween_eq_integral' f_int g_int hs - ((ae_restrict_iff' hs).mpr (eventually_of_forall hfg)) + ((ae_restrict_iff' hs).mpr (Eventually.of_forall hfg)) end regionBetween diff --git a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean index 87d03ba19bd52..9fe65ea8e8dea 100644 --- a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean +++ b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean @@ -303,7 +303,7 @@ lemma BoundedContinuousFunction.integral_eq_integral_meas_le_of_hasFiniteIntegra ∫ ω, f ω ∂μ = ∫ t in Ioc 0 ‖f‖, ENNReal.toReal (μ {a : α | t ≤ f a}) := by rw [Integrable.integral_eq_integral_Ioc_meas_le (M := ‖f‖) ?_ f_nn ?_] · refine ⟨f.continuous.measurable.aestronglyMeasurable, hf⟩ - · exact eventually_of_forall (fun x ↦ BoundedContinuousFunction.apply_le_norm f x) + · exact Eventually.of_forall (fun x ↦ BoundedContinuousFunction.apply_le_norm f x) /-- A version of the layer cake formula for bounded continuous functions and finite measures: ∫ f dμ = ∫ t in (0, ‖f‖], μ {x | f(x) ≥ t} dt. -/ @@ -333,7 +333,7 @@ lemma BoundedContinuousFunction.integral_le_of_levyProkhorovEDist_lt (μ ν : Me apply Measure.integrableOn_of_bounded (M := ENNReal.toReal (μ univ)) measure_Ioc_lt_top.ne · apply (Measurable.ennreal_toReal (Antitone.measurable ?_)).aestronglyMeasurable exact fun _ _ hst ↦ measure_mono (fun _ h ↦ hst.trans h) - · apply eventually_of_forall <| fun t ↦ ?_ + · 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 _) @@ -342,7 +342,7 @@ lemma BoundedContinuousFunction.integral_le_of_levyProkhorovEDist_lt (μ ν : Me apply Measure.integrableOn_of_bounded (M := ENNReal.toReal (ν univ)) measure_Ioc_lt_top.ne · apply (Measurable.ennreal_toReal (Antitone.measurable ?_)).aestronglyMeasurable exact fun _ _ hst ↦ measure_mono <| thickening_subset_of_subset ε (fun _ h ↦ hst.trans h) - · apply eventually_of_forall <| fun t ↦ ?_ + · 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 _) @@ -366,17 +366,17 @@ lemma tendsto_integral_meas_thickening_le (f : Ω →ᵇ ℝ) (𝓝 (∫ t in A, ENNReal.toReal (μ {a | t ≤ f a}))) := by apply tendsto_integral_filter_of_dominated_convergence (G := ℝ) (μ := volume.restrict A) (F := fun ε t ↦ (μ (thickening ε {a | t ≤ f a}))) (f := fun t ↦ (μ {a | t ≤ f a})) 1 - · apply eventually_of_forall fun n ↦ Measurable.aestronglyMeasurable ?_ + · apply Eventually.of_forall fun n ↦ Measurable.aestronglyMeasurable ?_ simp only [measurable_coe_nnreal_real_iff] apply measurable_toNNReal.comp <| Antitone.measurable (fun s t hst ↦ ?_) exact measure_mono <| thickening_subset_of_subset _ <| fun ω h ↦ hst.trans h - · apply eventually_of_forall (fun i ↦ ?_) - apply eventually_of_forall (fun t ↦ ?_) + · 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 · have aux : IsFiniteMeasure (volume.restrict A) := ⟨by simp [lt_top_iff_ne_top, A_finmeas]⟩ apply integrable_const - · apply eventually_of_forall (fun t ↦ ?_) + · apply Eventually.of_forall (fun t ↦ ?_) simp only [NNReal.tendsto_coe] apply (ENNReal.tendsto_toNNReal _).comp · apply tendsto_measure_thickening_of_isClosed ?_ ?_ @@ -405,7 +405,7 @@ lemma LevyProkhorov.continuous_toProbabilityMeasure : have ε_of_room := Tendsto.add (tendsto_iff_dist_tendsto_zero.mp hμs) εs_lim have ε_of_room' : Tendsto (fun n ↦ dist (μs n) ν + εs n) atTop (𝓝[>] 0) := by rw [tendsto_nhdsWithin_iff] - refine ⟨by simpa using ε_of_room, eventually_of_forall fun n ↦ ?_⟩ + refine ⟨by simpa using ε_of_room, Eventually.of_forall fun n ↦ ?_⟩ · rw [mem_Ioi] linarith [εs_pos n, dist_nonneg (x := μs n) (y := ν)] rw [add_zero] at ε_of_room @@ -428,7 +428,7 @@ lemma LevyProkhorov.continuous_toProbabilityMeasure : δ / 2 + ‖f‖ * (dist (μs n) ν + εs n) _ ≤ δ / 2 + ‖f‖ * (‖f‖⁻¹ * δ / 2) := by gcongr _ = δ := by field_simp; ring - · exact eventually_of_forall f_nn + · exact Eventually.of_forall f_nn · positivity · rw [ENNReal.ofReal_add (by positivity) (by positivity), ← add_zero (levyProkhorovEDist _ _)] apply ENNReal.add_lt_add_of_le_of_lt (levyProkhorovEDist_ne_top _ _) @@ -436,7 +436,7 @@ lemma LevyProkhorov.continuous_toProbabilityMeasure : rw [LevyProkhorov.dist_def, levyProkhorovDist, ofReal_toReal (levyProkhorovEDist_ne_top _ _)] simp only [Ps, P, LevyProkhorov.toProbabilityMeasure] - · exact eventually_of_forall f_nn + · exact Eventually.of_forall f_nn · simp only [IsCoboundedUnder, IsCobounded, eventually_map, eventually_atTop, forall_exists_index] refine ⟨0, fun a i hia ↦ le_trans (integral_nonneg f_nn) (hia i le_rfl)⟩ @@ -498,7 +498,7 @@ lemma ProbabilityMeasure.toMeasure_add_pos_gt_mem_nhds (P : ProbabilityMeasure {G : Set Ω} (G_open : IsOpen G) {ε : ℝ≥0∞} (ε_pos : 0 < ε) : {Q | P.toMeasure G < Q.toMeasure G + ε} ∈ 𝓝 P := by by_cases easy : P.toMeasure G < ε - · exact eventually_of_forall (fun _ ↦ lt_of_lt_of_le easy le_add_self) + · exact Eventually.of_forall (fun _ ↦ lt_of_lt_of_le easy le_add_self) by_cases ε_top : ε = ∞ · simp [ε_top, measure_lt_top] simp only [not_lt] at easy diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean index 9ada0295ca07b..18b574d4208a0 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpace.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpace.lean @@ -603,7 +603,7 @@ theorem measure_limsup_eq_zero {s : ℕ → Set α} (hs : (∑' i, μ (s i)) ≠ suffices μ (limsup t atTop) = 0 by have A : s ≤ t := fun n => subset_toMeasurable μ (s n) -- TODO default args fail - exact measure_mono_null (limsup_le_limsup (eventually_of_forall (Pi.le_def.mp A))) this + exact measure_mono_null (limsup_le_limsup (Eventually.of_forall (Pi.le_def.mp A))) this -- Next we unfold `limsup` for sets and replace equality with an inequality simp only [limsup_eq_iInf_iSup_of_nat', Set.iInf_eq_iInter, Set.iSup_eq_iUnion, ← nonpos_iff_eq_zero] diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 19386844f64da..37d3edc5f8472 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -205,13 +205,13 @@ theorem tendsto_measure_of_le_liminf_measure_of_limsup_measure_le {ι : Type*} { μ E = μ E₀ := measure_congr E₀_ae_eq_E.symm _ ≤ L.liminf fun i ↦ μs i E₀ := h_E₀ _ ≤ L.liminf fun i ↦ μs i E := - liminf_le_liminf (eventually_of_forall fun _ ↦ measure_mono E₀_subset) + liminf_le_liminf (.of_forall fun _ ↦ measure_mono E₀_subset) · have E_ae_eq_E₁ : E =ᵐ[μ] E₁ := EventuallyLE.antisymm subset_E₁.eventuallyLE ((ae_le_set.mpr nulldiff).trans E₀_subset.eventuallyLE) calc (L.limsup fun i ↦ μs i E) ≤ L.limsup fun i ↦ μs i E₁ := - limsup_le_limsup (eventually_of_forall fun _ ↦ measure_mono subset_E₁) + limsup_le_limsup (.of_forall fun _ ↦ measure_mono subset_E₁) _ ≤ μ E₁ := h_E₁ _ = μ E := measure_congr E_ae_eq_E₁.symm · infer_param @@ -425,7 +425,7 @@ lemma limsup_measure_closed_le_of_forall_tendsto_measure simp only [mem_map, mem_atTop_sets, ge_iff_le, mem_preimage, mem_Iio] at rs_lim obtain ⟨m, hm⟩ := rs_lim have aux : (fun i ↦ (μs i F)) ≤ᶠ[L] (fun i ↦ μs i (Metric.thickening (rs m) F)) := - eventually_of_forall <| fun i ↦ measure_mono (Metric.self_subset_thickening (rs_pos m) F) + .of_forall <| fun i ↦ measure_mono (Metric.self_subset_thickening (rs_pos m) F) refine (limsup_le_limsup aux).trans ?_ rw [Tendsto.limsup_eq (key m)] apply (measure_mono (Metric.thickening_subset_cthickening (rs m) F)).trans (hm m rfl.le).le @@ -468,7 +468,7 @@ lemma lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure {μ : Measure Ω} {μs : ℕ → Measure Ω} {f : Ω → ℝ} (f_cont : Continuous f) (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : ∫⁻ x, ENNReal.ofReal (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫⁻ x, ENNReal.ofReal (f x) ∂ (μs i)) := by - simp_rw [lintegral_eq_lintegral_meas_lt _ (eventually_of_forall f_nn) f_cont.aemeasurable] + simp_rw [lintegral_eq_lintegral_meas_lt _ (Eventually.of_forall f_nn) f_cont.aemeasurable] calc ∫⁻ (t : ℝ) in Set.Ioi 0, μ {a | t < f a} ≤ ∫⁻ (t : ℝ) in Set.Ioi 0, atTop.liminf (fun i ↦ (μs i) {a | t < f a}) := ?_ -- (i) _ ≤ atTop.liminf (fun i ↦ ∫⁻ (t : ℝ) in Set.Ioi 0, (μs i) {a | t < f a}) := ?_ -- (ii) @@ -486,23 +486,23 @@ lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by have same := lintegral_le_liminf_lintegral_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) + rw [@integral_eq_lintegral_of_nonneg_ae Ω _ μ f (Eventually.of_forall f_nn) f.continuous.measurable.aestronglyMeasurable] convert (ENNReal.toReal_le_toReal ?_ ?_).mpr same - · simp only [fun i ↦ @integral_eq_lintegral_of_nonneg_ae Ω _ (μs i) f (eventually_of_forall f_nn) + · 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 have bound : ∀ i, ∫⁻ x, ENNReal.ofReal (f x) ∂(μs i) ≤ nndist 0 g := fun i ↦ by 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) + 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 apply lt_of_le_of_lt _ (show (‖f‖₊ : ℝ≥0∞) < ∞ from ENNReal.coe_lt_top) apply liminf_le_of_le - · refine ⟨0, eventually_of_forall (by simp only [ge_iff_le, zero_le, forall_const])⟩ + · refine ⟨0, .of_forall (by simp only [ge_iff_le, zero_le, forall_const])⟩ · intro x hx obtain ⟨i, hi⟩ := hx.exists apply le_trans hi diff --git a/Mathlib/MeasureTheory/Measure/Restrict.lean b/Mathlib/MeasureTheory/Measure/Restrict.lean index 515994e89c22a..c38f9cc1c8cef 100644 --- a/Mathlib/MeasureTheory/Measure/Restrict.lean +++ b/Mathlib/MeasureTheory/Measure/Restrict.lean @@ -570,7 +570,7 @@ theorem _root_.Filter.EventuallyEq.restrict {f g : α → δ} {s : Set α} (hfg exact Measure.absolutelyContinuous_of_le Measure.restrict_le_self theorem ae_restrict_mem₀ (hs : NullMeasurableSet s μ) : ∀ᵐ x ∂μ.restrict s, x ∈ s := - (ae_restrict_iff'₀ hs).2 (Filter.eventually_of_forall fun _ => id) + (ae_restrict_iff'₀ hs).2 (Filter.Eventually.of_forall fun _ => id) theorem ae_restrict_mem (hs : MeasurableSet s) : ∀ᵐ x ∂μ.restrict s, x ∈ s := ae_restrict_mem₀ hs.nullMeasurableSet diff --git a/Mathlib/MeasureTheory/Measure/Stieltjes.lean b/Mathlib/MeasureTheory/Measure/Stieltjes.lean index 98968697286c0..7399a3bc6bab9 100644 --- a/Mathlib/MeasureTheory/Measure/Stieltjes.lean +++ b/Mathlib/MeasureTheory/Measure/Stieltjes.lean @@ -383,7 +383,7 @@ theorem measure_singleton (a : ℝ) : f.measure {a} = ofReal (f a - leftLim f a) apply (f.mono.tendsto_leftLim a).comp exact tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within _ u_lim - (eventually_of_forall fun n => u_lt_a n) + (Eventually.of_forall fun n => u_lt_a n) exact ENNReal.continuous_ofReal.continuousAt.tendsto.comp (tendsto_const_nhds.sub this) exact tendsto_nhds_unique L1 L2 diff --git a/Mathlib/MeasureTheory/Order/UpperLower.lean b/Mathlib/MeasureTheory/Order/UpperLower.lean index 65f4ad6fa888a..ef8fa8580acfc 100644 --- a/Mathlib/MeasureTheory/Order/UpperLower.lean +++ b/Mathlib/MeasureTheory/Order/UpperLower.lean @@ -64,7 +64,7 @@ private lemma aux₀ intro H obtain ⟨ε, -, hε', hε₀⟩ := exists_seq_strictAnti_tendsto_nhdsWithin (0 : ℝ) refine not_eventually.2 - (frequently_of_forall fun _ ↦ lt_irrefl $ ENNReal.ofReal $ 4⁻¹ ^ Fintype.card ι) + (Frequently.of_forall fun _ ↦ lt_irrefl $ ENNReal.ofReal $ 4⁻¹ ^ Fintype.card ι) ((Filter.Tendsto.eventually_lt (H.comp hε₀) tendsto_const_nhds ?_).mono fun n ↦ lt_of_le_of_lt ?_) on_goal 2 => @@ -92,7 +92,7 @@ private lemma aux₁ intro H obtain ⟨ε, -, hε', hε₀⟩ := exists_seq_strictAnti_tendsto_nhdsWithin (0 : ℝ) refine not_eventually.2 - (frequently_of_forall fun _ ↦ lt_irrefl $ 1 - ENNReal.ofReal (4⁻¹ ^ Fintype.card ι)) + (Frequently.of_forall fun _ ↦ lt_irrefl $ 1 - ENNReal.ofReal (4⁻¹ ^ Fintype.card ι)) ((Filter.Tendsto.eventually_lt tendsto_const_nhds (H.comp hε₀) $ ENNReal.sub_lt_self ENNReal.one_ne_top one_ne_zero ?_).mono fun n ↦ lt_of_le_of_lt' ?_) diff --git a/Mathlib/MeasureTheory/OuterMeasure/AE.lean b/Mathlib/MeasureTheory/OuterMeasure/AE.lean index 9291c540a5bcf..2b8af44469a76 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/AE.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/AE.lean @@ -85,7 +85,7 @@ theorem measure_zero_iff_ae_nmem {s : Set α} : μ s = 0 ↔ ∀ᵐ a ∂μ, a compl_mem_ae_iff.symm theorem ae_of_all {p : α → Prop} (μ : F) : (∀ a, p a) → ∀ᵐ a ∂μ, p a := - eventually_of_forall + Eventually.of_forall instance instCountableInterFilter : CountableInterFilter (ae μ) := by unfold ae; infer_instance diff --git a/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean b/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean index e18b71b7b04b6..c5c6b9529953e 100644 --- a/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean +++ b/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean @@ -127,7 +127,7 @@ lemma term_tsum_one : HasSum (fun n ↦ term (n + 1) 1) (1 - γ) := by simp_rw [term_sum_one, sub_eq_neg_add] refine Tendsto.add ?_ tendsto_const_nhds have := (tendsto_eulerMascheroniSeq'.comp (tendsto_add_atTop_nat 1)).neg - refine this.congr' (eventually_of_forall (fun n ↦ ?_)) + refine this.congr' (Eventually.of_forall (fun n ↦ ?_)) simp_rw [Function.comp_apply, eulerMascheroniSeq', if_false] push_cast abel diff --git a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/MDifferentiable.lean b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/MDifferentiable.lean index 209318f22a221..8b310863a3576 100644 --- a/Mathlib/NumberTheory/ModularForms/EisensteinSeries/MDifferentiable.lean +++ b/Mathlib/NumberTheory/ModularForms/EisensteinSeries/MDifferentiable.lean @@ -55,7 +55,7 @@ theorem eisensteinSeries_SIF_MDifferentiable {k : ℤ} {N : ℕ} (hk : 3 ≤ k) refine DifferentiableOn.differentiableAt ?_ ((isOpen_lt continuous_const Complex.continuous_im).mem_nhds τ.2) exact (eisensteinSeries_tendstoLocallyUniformlyOn hk a).differentiableOn - (eventually_of_forall fun s ↦ DifferentiableOn.sum + (Eventually.of_forall fun s ↦ DifferentiableOn.sum fun _ _ ↦ eisSummand_extension_differentiableOn _ _) (isOpen_lt continuous_const continuous_im) diff --git a/Mathlib/NumberTheory/WellApproximable.lean b/Mathlib/NumberTheory/WellApproximable.lean index 69b68a0ebfab8..fa77ff16edf76 100644 --- a/Mathlib/NumberTheory/WellApproximable.lean +++ b/Mathlib/NumberTheory/WellApproximable.lean @@ -289,7 +289,7 @@ theorem addWellApproximable_ae_empty_or_univ (δ : ℕ → ℝ) (hδ : Tendsto let e := (AddAction.toPerm (u p) : Equiv.Perm 𝕊).toOrderIsoSet change e (C p) = C p rw [OrderIso.apply_blimsup e, ← hu₀ p] - exact blimsup_congr (eventually_of_forall fun n hn => + exact blimsup_congr (Eventually.of_forall fun n hn => approxAddOrderOf.vadd_eq_of_mul_dvd (δ n) hn.1 hn.2) by_cases h : ∀ p : Nat.Primes, A p =ᵐ[μ] (∅ : Set 𝕊) ∧ B p =ᵐ[μ] (∅ : Set 𝕊) · replace h : ∀ p : Nat.Primes, (u p +ᵥ E : Set _) =ᵐ[μ] E := by diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index 692f4d4c39e72..aaa43927f9f6c 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -343,7 +343,7 @@ theorem tendsto_atBot_mono' [Preorder β] (l : Filter α) ⦃f₁ f₂ : α → theorem tendsto_atTop_mono [Preorder β] {l : Filter α} {f g : α → β} (h : ∀ n, f n ≤ g n) : Tendsto f l atTop → Tendsto g l atTop := - tendsto_atTop_mono' l <| eventually_of_forall h + tendsto_atTop_mono' l <| Eventually.of_forall h theorem tendsto_atBot_mono [Preorder β] {l : Filter α} {f g : α → β} (h : ∀ n, f n ≤ g n) : Tendsto g l atBot → Tendsto f l atBot := @@ -577,7 +577,7 @@ theorem tendsto_atBot_add_nonpos_left' (hf : ∀ᶠ x in l, f x ≤ 0) (hg : Ten theorem tendsto_atTop_add_nonneg_left (hf : ∀ x, 0 ≤ f x) (hg : Tendsto g l atTop) : Tendsto (fun x => f x + g x) l atTop := - tendsto_atTop_add_nonneg_left' (eventually_of_forall hf) hg + tendsto_atTop_add_nonneg_left' (Eventually.of_forall hf) hg theorem tendsto_atBot_add_nonpos_left (hf : ∀ x, f x ≤ 0) (hg : Tendsto g l atBot) : Tendsto (fun x => f x + g x) l atBot := @@ -593,7 +593,7 @@ theorem tendsto_atBot_add_nonpos_right' (hf : Tendsto f l atBot) (hg : ∀ᶠ x theorem tendsto_atTop_add_nonneg_right (hf : Tendsto f l atTop) (hg : ∀ x, 0 ≤ g x) : Tendsto (fun x => f x + g x) l atTop := - tendsto_atTop_add_nonneg_right' hf (eventually_of_forall hg) + tendsto_atTop_add_nonneg_right' hf (Eventually.of_forall hg) theorem tendsto_atBot_add_nonpos_right (hf : Tendsto f l atBot) (hg : ∀ x, g x ≤ 0) : Tendsto (fun x => f x + g x) l atBot := @@ -1737,7 +1737,7 @@ lemma frequently_iff_seq_forall {ι : Type*} {l : Filter ι} {p : ι → Prop} [l.IsCountablyGenerated] : (∃ᶠ n in l, p n) ↔ ∃ ns : ℕ → ι, Tendsto ns atTop l ∧ ∀ n, p (ns n) := ⟨exists_seq_forall_of_frequently, fun ⟨_ns, hnsl, hpns⟩ ↦ - hnsl.frequently <| frequently_of_forall hpns⟩ + hnsl.frequently <| Frequently.of_forall hpns⟩ /-- A sequence converges if every subsequence has a convergent subsequence. -/ theorem tendsto_of_subseq_tendsto {ι : Type*} {x : ι → α} {f : Filter α} {l : Filter ι} diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 9111ee66adb85..37961c2404ebb 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -978,9 +978,11 @@ protected theorem Eventually.and {p q : α → Prop} {f : Filter α} : @[simp] theorem eventually_true (f : Filter α) : ∀ᶠ _ in f, True := univ_mem -theorem eventually_of_forall {p : α → Prop} {f : Filter α} (hp : ∀ x, p x) : ∀ᶠ x in f, p x := +theorem Eventually.of_forall {p : α → Prop} {f : Filter α} (hp : ∀ x, p x) : ∀ᶠ x in f, p x := univ_mem' hp +@[deprecated (since := "2024-08-02")] alias eventually_of_forall := Eventually.of_forall + @[simp] theorem eventually_false_iff_eq_bot {f : Filter α} : (∀ᶠ _ in f, False) ↔ f = ⊥ := empty_mem_iff_bot @@ -1003,7 +1005,7 @@ theorem Eventually.mp {p q : α → Prop} {f : Filter α} (hp : ∀ᶠ x in f, p theorem Eventually.mono {p q : α → Prop} {f : Filter α} (hp : ∀ᶠ x in f, p x) (hq : ∀ x, p x → q x) : ∀ᶠ x in f, q x := - hp.mp (eventually_of_forall hq) + hp.mp (Eventually.of_forall hq) theorem forall_eventually_of_eventually_forall {f : Filter α} {p : α → β → Prop} (h : ∀ᶠ x in f, ∀ y, p x y) : ∀ y, ∀ᶠ x in f, p x y := @@ -1111,9 +1113,11 @@ theorem Eventually.frequently {f : Filter α} [NeBot f] {p : α → Prop} (h : ∃ᶠ x in f, p x := compl_not_mem h -theorem frequently_of_forall {f : Filter α} [NeBot f] {p : α → Prop} (h : ∀ x, p x) : +theorem Frequently.of_forall {f : Filter α} [NeBot f] {p : α → Prop} (h : ∀ x, p x) : ∃ᶠ x in f, p x := - Eventually.frequently (eventually_of_forall h) + Eventually.frequently (Eventually.of_forall h) + +@[deprecated (since := "2024-08-02")] alias frequently_of_forall := Frequently.of_forall theorem Frequently.mp {p q : α → Prop} {f : Filter α} (h : ∃ᶠ x in f, p x) (hpq : ∀ᶠ x in f, p x → q x) : ∃ᶠ x in f, q x := @@ -1125,7 +1129,7 @@ theorem Frequently.filter_mono {p : α → Prop} {f g : Filter α} (h : ∃ᶠ x theorem Frequently.mono {p q : α → Prop} {f : Filter α} (h : ∃ᶠ x in f, p x) (hpq : ∀ x, p x → q x) : ∃ᶠ x in f, q x := - h.mp (eventually_of_forall hpq) + h.mp (Eventually.of_forall hpq) theorem Frequently.and_eventually {p q : α → Prop} {f : Filter α} (hp : ∃ᶠ x in f, p x) (hq : ∀ᶠ x in f, q x) : ∃ᶠ x in f, p x ∧ q x := by @@ -1138,7 +1142,7 @@ theorem Eventually.and_frequently {p q : α → Prop} {f : Filter α} (hp : ∀ theorem Frequently.exists {p : α → Prop} {f : Filter α} (hp : ∃ᶠ x in f, p x) : ∃ x, p x := by by_contra H - replace H : ∀ᶠ x in f, ¬p x := eventually_of_forall (not_exists.1 H) + replace H : ∀ᶠ x in f, ¬p x := Eventually.of_forall (not_exists.1 H) exact hp H theorem Eventually.exists {p : α → Prop} {f : Filter α} [NeBot f] (hp : ∀ᶠ x in f, p x) : @@ -1274,7 +1278,7 @@ theorem EventuallyEq.rw {l : Filter α} {f g : α → β} (h : f =ᶠ[l] g) (p : hf.congr <| h.mono fun _ hx => hx ▸ Iff.rfl theorem eventuallyEq_set {s t : Set α} {l : Filter α} : s =ᶠ[l] t ↔ ∀ᶠ x in l, x ∈ s ↔ x ∈ t := - eventually_congr <| eventually_of_forall fun _ ↦ eq_iff_iff + eventually_congr <| Eventually.of_forall fun _ ↦ eq_iff_iff alias ⟨EventuallyEq.mem_iff, Eventually.set_eq⟩ := eventuallyEq_set @@ -1300,7 +1304,7 @@ theorem EventuallyEq.filter_mono {l l' : Filter α} {f g : α → β} (h₁ : f @[refl, simp] theorem EventuallyEq.refl (l : Filter α) (f : α → β) : f =ᶠ[l] f := - eventually_of_forall fun _ => rfl + Eventually.of_forall fun _ => rfl protected theorem EventuallyEq.rfl {l : Filter α} {f : α → β} : f =ᶠ[l] f := EventuallyEq.refl l f @@ -2832,7 +2836,7 @@ theorem Set.EqOn.eventuallyEq_of_mem {α β} {s : Set α} {l : Filter α} {f g : h.eventuallyEq.filter_mono <| Filter.le_principal_iff.2 hl theorem HasSubset.Subset.eventuallyLE {α} {l : Filter α} {s t : Set α} (h : s ⊆ t) : s ≤ᶠ[l] t := - Filter.eventually_of_forall h + Filter.Eventually.of_forall h theorem Set.MapsTo.tendsto {α β} {s : Set α} {t : Set β} {f : α → β} (h : MapsTo f s t) : Filter.Tendsto f (𝓟 s) (𝓟 t) := diff --git a/Mathlib/Order/Filter/ENNReal.lean b/Mathlib/Order/Filter/ENNReal.lean index 3c8f70423bc95..8acfdf4fae082 100644 --- a/Mathlib/Order/Filter/ENNReal.lean +++ b/Mathlib/Order/Filter/ENNReal.lean @@ -82,6 +82,6 @@ theorem limsup_liminf_le_liminf_limsup {β} [Countable β] {f : Filter α} [Coun have h1 : ∀ᶠ a in f, ∀ b, u a b ≤ f.limsup fun a' => u a' b := by rw [eventually_countable_forall] exact fun b => ENNReal.eventually_le_limsup fun a => u a b - sInf_le <| h1.mono fun x hx => Filter.liminf_le_liminf (Filter.eventually_of_forall hx) + sInf_le <| h1.mono fun x hx => Filter.liminf_le_liminf (Filter.Eventually.of_forall hx) end ENNReal diff --git a/Mathlib/Order/Filter/FilterProduct.lean b/Mathlib/Order/Filter/FilterProduct.lean index 827c876602408..99c01de709197 100644 --- a/Mathlib/Order/Filter/FilterProduct.lean +++ b/Mathlib/Order/Filter/FilterProduct.lean @@ -79,7 +79,7 @@ theorem lt_def [Preorder β] : ((· < ·) : β* → β* → Prop) = LiftRel (· instance isTotal [LE β] [IsTotal β (· ≤ ·)] : IsTotal β* (· ≤ ·) := ⟨fun f g => - inductionOn₂ f g fun _f _g => eventually_or.1 <| eventually_of_forall fun _x => total_of _ _ _⟩ + inductionOn₂ f g fun _f _g => eventually_or.1 <| Eventually.of_forall fun _x => total_of _ _ _⟩ open Classical in /-- If `φ` is an ultrafilter then the ultraproduct is a linear order. -/ diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index 52f1af5d388e6..dd4039e19d250 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -74,7 +74,7 @@ def Germ (l : Filter α) (β : Type*) : Type _ := def productSetoid (l : Filter α) (ε : α → Type*) : Setoid ((a : _) → ε a) where r f g := ∀ᶠ a in l, f a = g a iseqv := - ⟨fun _ => eventually_of_forall fun _ => rfl, fun h => h.mono fun _ => Eq.symm, + ⟨fun _ => Eventually.of_forall fun _ => rfl, fun h => h.mono fun _ => Eq.symm, fun h1 h2 => h1.congr (h2.mono fun _ hx => hx ▸ Iff.rfl)⟩ /-- The filter product `(a : α) → ε a` at a filter `l`. This is a dependent version of @@ -118,7 +118,7 @@ def IsConstant {l : Filter α} (P : Germ l β) : Prop := exact fun f g b hfg hf ↦ (hfg.symm).trans hf theorem isConstant_coe {l : Filter α} {b} (h : ∀ x', f x' = b) : (↑f : Germ l β).IsConstant := - ⟨b, eventually_of_forall (fun x ↦ h x)⟩ + ⟨b, Eventually.of_forall (fun x ↦ h x)⟩ @[simp] theorem isConstant_coe_const {l : Filter α} {b : β} : (fun _ : α ↦ b : Germ l β).IsConstant := by @@ -291,7 +291,7 @@ theorem liftPred_coe {p : β → Prop} {f : α → β} : LiftPred p (f : Germ l Iff.rfl theorem liftPred_const {p : β → Prop} {x : β} (hx : p x) : LiftPred p (↑x : Germ l β) := - eventually_of_forall fun _y => hx + Eventually.of_forall fun _y => hx @[simp] theorem liftPred_const_iff [NeBot l] {p : β → Prop} {x : β} : LiftPred p (↑x : Germ l β) ↔ p x := @@ -309,7 +309,7 @@ theorem liftRel_coe {r : β → γ → Prop} {f : α → β} {g : α → γ} : theorem liftRel_const {r : β → γ → Prop} {x : β} {y : γ} (h : r x y) : LiftRel r (↑x : Germ l β) ↑y := - eventually_of_forall fun _ => h + Eventually.of_forall fun _ => h @[simp] theorem liftRel_const_iff [NeBot l] {r : β → γ → Prop} {x : β} {y : γ} : @@ -707,10 +707,10 @@ theorem const_top [Top β] : (↑(⊤ : β) : Germ l β) = ⊤ := rfl instance instOrderBot [LE β] [OrderBot β] : OrderBot (Germ l β) where - bot_le f := inductionOn f fun _ => eventually_of_forall fun _ => bot_le + bot_le f := inductionOn f fun _ => Eventually.of_forall fun _ => bot_le instance instOrderTop [LE β] [OrderTop β] : OrderTop (Germ l β) where - le_top f := inductionOn f fun _ => eventually_of_forall fun _ => le_top + le_top f := inductionOn f fun _ => Eventually.of_forall fun _ => le_top instance instBoundedOrder [LE β] [BoundedOrder β] : BoundedOrder (Germ l β) where __ := instOrderBot @@ -728,13 +728,13 @@ theorem const_inf [Inf β] (a b : β) : ↑(a ⊓ b) = (↑a ⊓ ↑b : Germ l rfl instance instSemilatticeSup [SemilatticeSup β] : SemilatticeSup (Germ l β) where - le_sup_left f g := inductionOn₂ f g fun _f _g => eventually_of_forall fun _x ↦ le_sup_left - le_sup_right f g := inductionOn₂ f g fun _f _g ↦ eventually_of_forall fun _x ↦ le_sup_right + le_sup_left f g := inductionOn₂ f g fun _f _g => Eventually.of_forall fun _x ↦ le_sup_left + le_sup_right f g := inductionOn₂ f g fun _f _g ↦ Eventually.of_forall fun _x ↦ le_sup_right sup_le f₁ f₂ g := inductionOn₃ f₁ f₂ g fun _f₁ _f₂ _g h₁ h₂ ↦ h₂.mp <| h₁.mono fun _x ↦ sup_le instance instSemilatticeInf [SemilatticeInf β] : SemilatticeInf (Germ l β) where - inf_le_left f g := inductionOn₂ f g fun _f _g ↦ eventually_of_forall fun _x ↦ inf_le_left - inf_le_right f g := inductionOn₂ f g fun _f _g ↦ eventually_of_forall fun _x ↦ inf_le_right + inf_le_left f g := inductionOn₂ f g fun _f _g ↦ Eventually.of_forall fun _x ↦ inf_le_left + inf_le_right f g := inductionOn₂ f g fun _f _g ↦ Eventually.of_forall fun _x ↦ inf_le_right le_inf f₁ f₂ g := inductionOn₃ f₁ f₂ g fun _f₁ _f₂ _g h₁ h₂ ↦ h₂.mp <| h₁.mono fun _x ↦ le_inf instance instLattice [Lattice β] : Lattice (Germ l β) where @@ -742,7 +742,7 @@ instance instLattice [Lattice β] : Lattice (Germ l β) where __ := instSemilatticeInf instance instDistribLattice [DistribLattice β] : DistribLattice (Germ l β) where - le_sup_inf f g h := inductionOn₃ f g h fun _f _g _h ↦ eventually_of_forall fun _ ↦ le_sup_inf + le_sup_inf f g h := inductionOn₃ f g h fun _f _g _h ↦ Eventually.of_forall fun _ ↦ le_sup_inf @[to_additive] instance instExistsMulOfLE [Mul β] [LE β] [ExistsMulOfLE β] : ExistsMulOfLE (Germ l β) where diff --git a/Mathlib/Order/Filter/Germ/OrderedMonoid.lean b/Mathlib/Order/Filter/Germ/OrderedMonoid.lean index f2d7d62077b7e..a3446d9e484a3 100644 --- a/Mathlib/Order/Filter/Germ/OrderedMonoid.lean +++ b/Mathlib/Order/Filter/Germ/OrderedMonoid.lean @@ -39,6 +39,6 @@ instance instOrderedCancelCommMonoid [OrderedCancelCommMonoid β] : instance instCanonicallyOrderedCommMonoid [CanonicallyOrderedCommMonoid β] : CanonicallyOrderedCommMonoid (Germ l β) where __ := instExistsMulOfLE - le_self_mul x y := inductionOn₂ x y fun _ _ ↦ eventually_of_forall fun _ ↦ le_self_mul + le_self_mul x y := inductionOn₂ x y fun _ _ ↦ Eventually.of_forall fun _ ↦ le_self_mul end Filter.Germ diff --git a/Mathlib/Order/Filter/IndicatorFunction.lean b/Mathlib/Order/Filter/IndicatorFunction.lean index 9bc67a4c136d9..d7338ff33d231 100644 --- a/Mathlib/Order/Filter/IndicatorFunction.lean +++ b/Mathlib/Order/Filter/IndicatorFunction.lean @@ -120,7 +120,7 @@ theorem Filter.EventuallyEq.of_mulIndicator [One β] {l : Filter α} {f : α → @[to_additive] theorem Filter.EventuallyEq.of_mulIndicator_const [One β] {l : Filter α} {c : β} (hc : c ≠ 1) {s t : Set α} (h : s.mulIndicator (fun _ ↦ c) =ᶠ[l] t.mulIndicator fun _ ↦ c) : s =ᶠ[l] t := - .of_mulIndicator (eventually_of_forall fun _ ↦ hc) h + .of_mulIndicator (Eventually.of_forall fun _ ↦ hc) h @[to_additive] theorem Filter.mulIndicator_const_eventuallyEq [One β] {l : Filter α} {c : β} (hc : c ≠ 1) diff --git a/Mathlib/Order/Filter/Interval.lean b/Mathlib/Order/Filter/Interval.lean index 0bb96d5502cad..e5b7087c92c4f 100644 --- a/Mathlib/Order/Filter/Interval.lean +++ b/Mathlib/Order/Filter/Interval.lean @@ -110,7 +110,7 @@ theorem tendstoIxxClass_inf {l₁ l₁' l₂ l₂' : Filter α} {Ixx} [h : Tends theorem tendstoIxxClass_of_subset {l₁ l₂ : Filter α} {Ixx Ixx' : α → α → Set α} (h : ∀ a b, Ixx a b ⊆ Ixx' a b) [h' : TendstoIxxClass Ixx' l₁ l₂] : TendstoIxxClass Ixx l₁ l₂ := - ⟨h'.1.smallSets_mono <| eventually_of_forall <| Prod.forall.2 h⟩ + ⟨h'.1.smallSets_mono <| Eventually.of_forall <| Prod.forall.2 h⟩ theorem HasBasis.tendstoIxxClass {ι : Type*} {p : ι → Prop} {s} {l : Filter α} (hl : l.HasBasis p s) {Ixx : α → α → Set α} diff --git a/Mathlib/Order/Filter/ModEq.lean b/Mathlib/Order/Filter/ModEq.lean index 5014ac919f8fd..0b9c0b89ea5f7 100644 --- a/Mathlib/Order/Filter/ModEq.lean +++ b/Mathlib/Order/Filter/ModEq.lean @@ -21,7 +21,7 @@ namespace Nat /-- Infinitely many natural numbers are equal to `d` mod `n`. -/ theorem frequently_modEq {n : ℕ} (h : n ≠ 0) (d : ℕ) : ∃ᶠ m in atTop, m ≡ d [MOD n] := ((tendsto_add_atTop_nat d).comp (tendsto_id.nsmul_atTop h.bot_lt)).frequently <| - frequently_of_forall fun m => by simp [Nat.modEq_iff_dvd, ← sub_sub] + Frequently.of_forall fun m => by simp [Nat.modEq_iff_dvd, ← sub_sub] theorem frequently_mod_eq {d n : ℕ} (h : d < n) : ∃ᶠ m in atTop, m % n = d := by simpa only [Nat.ModEq, mod_eq_of_lt h] using frequently_modEq h.ne_bot d diff --git a/Mathlib/Order/Filter/Ring.lean b/Mathlib/Order/Filter/Ring.lean index d4431d4758418..3a1e434b51dd2 100644 --- a/Mathlib/Order/Filter/Ring.lean +++ b/Mathlib/Order/Filter/Ring.lean @@ -33,7 +33,7 @@ theorem EventuallyLE.mul_nonneg [OrderedSemiring β] {l : Filter α} {f g : α theorem eventually_sub_nonneg [OrderedRing β] {l : Filter α} {f g : α → β} : 0 ≤ᶠ[l] g - f ↔ f ≤ᶠ[l] g := - eventually_congr <| eventually_of_forall fun _ => sub_nonneg + eventually_congr <| Eventually.of_forall fun _ => sub_nonneg namespace Germ diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 11b9e050b8436..b407ed6f446a0 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -63,7 +63,7 @@ theorem isBounded_iff : f.IsBounded r ↔ ∃ s ∈ f.sets, ∃ b, s ⊆ { x | r /-- A bounded function `u` is in particular eventually bounded. -/ theorem isBoundedUnder_of {f : Filter β} {u : β → α} : (∃ b, ∀ x, r (u x) b) → f.IsBoundedUnder r u - | ⟨b, hb⟩ => ⟨b, show ∀ᶠ x in f, r (u x) b from eventually_of_forall hb⟩ + | ⟨b, hb⟩ => ⟨b, show ∀ᶠ x in f, r (u x) b from Eventually.of_forall hb⟩ theorem isBounded_bot : IsBounded r ⊥ ↔ Nonempty α := by simp [IsBounded, exists_true_iff_nonempty] @@ -95,7 +95,7 @@ theorem IsBoundedUnder.mono_ge [Preorder β] {l : Filter α} {u v : α → β} IsBoundedUnder.mono_le (β := βᵒᵈ) hu hv theorem isBoundedUnder_const [IsRefl α r] {l : Filter β} {a : α} : IsBoundedUnder r l fun _ => a := - ⟨a, eventually_map.2 <| eventually_of_forall fun _ => refl _⟩ + ⟨a, eventually_map.2 <| Eventually.of_forall fun _ => refl _⟩ theorem IsBounded.isBoundedUnder {q : β → β → Prop} {u : α → β} (hu : ∀ a₀ a₁, r a₀ a₁ → q (u a₀) (u a₁)) : f.IsBounded r → f.IsBoundedUnder q u @@ -238,12 +238,12 @@ lemma isCoboundedUnder_ge_of_eventually_le [Preorder α] (l : Filter ι) [NeBot lemma isCoboundedUnder_le_of_le [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} {x : α} (hf : ∀ i, x ≤ f i) : IsCoboundedUnder (· ≤ ·) l f := - isCoboundedUnder_le_of_eventually_le l (eventually_of_forall hf) + isCoboundedUnder_le_of_eventually_le l (Eventually.of_forall hf) lemma isCoboundedUnder_ge_of_le [Preorder α] (l : Filter ι) [NeBot l] {f : ι → α} {x : α} (hf : ∀ i, f i ≤ x) : IsCoboundedUnder (· ≥ ·) l f := - isCoboundedUnder_ge_of_eventually_le l (eventually_of_forall hf) + isCoboundedUnder_ge_of_eventually_le l (Eventually.of_forall hf) theorem isCobounded_bot : IsCobounded r ⊥ ↔ ∃ b, ∀ x, r b x := by simp [IsCobounded] @@ -350,10 +350,10 @@ theorem isCobounded_ge_of_top [Preorder α] [OrderTop α] {f : Filter α} : f.Is ⟨⊤, fun _ _ => le_top⟩ theorem isBounded_le_of_top [Preorder α] [OrderTop α] {f : Filter α} : f.IsBounded (· ≤ ·) := - ⟨⊤, eventually_of_forall fun _ => le_top⟩ + ⟨⊤, Eventually.of_forall fun _ => le_top⟩ theorem isBounded_ge_of_bot [Preorder α] [OrderBot α] {f : Filter α} : f.IsBounded (· ≥ ·) := - ⟨⊥, eventually_of_forall fun _ => bot_le⟩ + ⟨⊥, Eventually.of_forall fun _ => bot_le⟩ @[simp] theorem _root_.OrderIso.isBoundedUnder_le_comp [Preorder α] [Preorder β] (e : α ≃o β) {l : Filter γ} @@ -388,8 +388,8 @@ theorem isBoundedUnder_le_sup [SemilatticeSup α] {f : Filter β} {u v : β → (f.IsBoundedUnder (· ≤ ·) fun a => u a ⊔ v a) ↔ f.IsBoundedUnder (· ≤ ·) u ∧ f.IsBoundedUnder (· ≤ ·) v := ⟨fun h => - ⟨h.mono_le <| eventually_of_forall fun _ => le_sup_left, - h.mono_le <| eventually_of_forall fun _ => le_sup_right⟩, + ⟨h.mono_le <| Eventually.of_forall fun _ => le_sup_left, + h.mono_le <| Eventually.of_forall fun _ => le_sup_right⟩, fun h => h.1.sup h.2⟩ theorem IsBoundedUnder.inf [SemilatticeInf α] {f : Filter β} {u v : β → α} : @@ -728,7 +728,7 @@ theorem bliminf_false {f : Filter β} {u : β → α} : (bliminf u f fun _ => Fa @[simp] theorem limsup_const_bot {f : Filter β} : limsup (fun _ : β => (⊥ : α)) f = (⊥ : α) := by rw [limsup_eq, eq_bot_iff] - exact sInf_le (eventually_of_forall fun _ => le_rfl) + exact sInf_le (Eventually.of_forall fun _ => le_rfl) /-- Same as limsup_const applied to `⊤` but without the `NeBot f` assumption -/ @[simp] @@ -753,10 +753,10 @@ theorem limsInf_eq_iSup_sInf {f : Filter α} : limsInf f = ⨆ s ∈ f, sInf s : limsSup_eq_iInf_sSup (α := αᵒᵈ) theorem limsup_le_iSup {f : Filter β} {u : β → α} : limsup u f ≤ ⨆ n, u n := - limsup_le_of_le (by isBoundedDefault) (eventually_of_forall (le_iSup u)) + limsup_le_of_le (by isBoundedDefault) (Eventually.of_forall (le_iSup u)) theorem iInf_le_liminf {f : Filter β} {u : β → α} : ⨅ n, u n ≤ liminf u f := - le_liminf_of_le (by isBoundedDefault) (eventually_of_forall (iInf_le u)) + le_liminf_of_le (by isBoundedDefault) (Eventually.of_forall (iInf_le u)) /-- In a complete lattice, the limsup of a function is the infimum over sets `s` in the filter of the supremum of the function over `s` -/ @@ -894,13 +894,13 @@ theorem mono_blimsup' (h : ∀ᶠ x in f, p x → u x ≤ v x) : blimsup u f p sInf_le_sInf fun _ ha => (ha.and h).mono fun _ hx hx' => (hx.2 hx').trans (hx.1 hx') theorem mono_blimsup (h : ∀ x, p x → u x ≤ v x) : blimsup u f p ≤ blimsup v f p := - mono_blimsup' <| eventually_of_forall h + mono_blimsup' <| Eventually.of_forall h theorem mono_bliminf' (h : ∀ᶠ x in f, p x → u x ≤ v x) : bliminf u f p ≤ bliminf v f p := sSup_le_sSup fun _ ha => (ha.and h).mono fun _ hx hx' => (hx.1 hx').trans (hx.2 hx') theorem mono_bliminf (h : ∀ x, p x → u x ≤ v x) : bliminf u f p ≤ bliminf v f p := - mono_bliminf' <| eventually_of_forall h + mono_bliminf' <| Eventually.of_forall h theorem bliminf_antitone_filter (h : f ≤ g) : bliminf u g p ≤ bliminf u f p := sSup_le_sSup fun _ ha => ha.filter_mono h @@ -1480,8 +1480,8 @@ theorem limsup_max [ConditionallyCompleteLinearOrder β] {f : Filter α} {u v : have hv := eventually_lt_of_limsup_lt (lt_of_le_of_lt (le_max_right _ _) hb) h₄ refine mem_of_superset (inter_mem hu hv) (fun _ ↦ by simp) · exact max_le (c := limsup (fun a ↦ max (u a) (v a)) f) - (limsup_le_limsup (eventually_of_forall (fun a : α ↦ le_max_left (u a) (v a))) h₁ bddmax) - (limsup_le_limsup (eventually_of_forall (fun a : α ↦ le_max_right (u a) (v a))) h₂ bddmax) + (limsup_le_limsup (Eventually.of_forall (fun a : α ↦ le_max_left (u a) (v a))) h₁ bddmax) + (limsup_le_limsup (Eventually.of_forall (fun a : α ↦ le_max_right (u a) (v a))) h₂ bddmax) theorem liminf_min [ConditionallyCompleteLinearOrder β] {f : Filter α} {u v : α → β} (h₁ : f.IsCoboundedUnder (· ≥ ·) u := by isBoundedDefault) @@ -1562,7 +1562,7 @@ theorem limsup_finset_sup' [ConditionallyCompleteLinearOrder β] {f : Filter α} exact lt_of_le_of_lt (Finset.le_sup' (f := fun i ↦ limsup (F i) f) i_s) hb · simp only [mem_iInter, mem_setOf_eq, Finset.sup'_apply, sup'_lt_iff, imp_self, implies_true] · apply Finset.sup'_le hs (fun i ↦ limsup (F i) f) - refine fun i i_s ↦ limsup_le_limsup (eventually_of_forall (fun a ↦ ?_)) (h₁ i i_s) bddsup + refine fun i i_s ↦ limsup_le_limsup (Eventually.of_forall (fun a ↦ ?_)) (h₁ i i_s) bddsup simp only [Finset.sup'_apply, le_sup'_iff] use i, i_s @@ -1608,7 +1608,7 @@ lemma IsCobounded.frequently_ge [NeBot F] (cobdd : IsCobounded (· ≤ ·) F) : ∃ l, ∃ᶠ x in F, l ≤ x := by obtain ⟨t, ht⟩ := cobdd by_cases tbot : IsBot t - · refine ⟨t, frequently_of_forall fun r ↦ tbot r⟩ + · refine ⟨t, Frequently.of_forall fun r ↦ tbot r⟩ obtain ⟨t', ht'⟩ : ∃ t', t' < t := by by_contra! exact tbot this diff --git a/Mathlib/Probability/BorelCantelli.lean b/Mathlib/Probability/BorelCantelli.lean index fead097056f50..a1f404dfaedf6 100644 --- a/Mathlib/Probability/BorelCantelli.lean +++ b/Mathlib/Probability/BorelCantelli.lean @@ -84,7 +84,7 @@ theorem measure_limsup_eq_one {s : ℕ → Set Ω} (hsm : ∀ n, MeasurableSet ( rw [← sub_nonneg, Finset.sum_range_succ_sub_sum] exact ENNReal.toReal_nonneg · rintro ⟨B, hB⟩ - refine not_eventually.2 (frequently_of_forall fun n => ?_) (htends B.toNNReal) + refine not_eventually.2 (Frequently.of_forall fun n => ?_) (htends B.toNNReal) rw [mem_upperBounds] at hB specialize hB (∑ k ∈ Finset.range n, μ (s (k + 1))).toReal _ · refine ⟨n, ?_⟩ diff --git a/Mathlib/Probability/ConditionalExpectation.lean b/Mathlib/Probability/ConditionalExpectation.lean index e9c83057f3fe8..66719c1935974 100644 --- a/Mathlib/Probability/ConditionalExpectation.lean +++ b/Mathlib/Probability/ConditionalExpectation.lean @@ -63,7 +63,7 @@ theorem condexp_indep_eq (hle₁ : m₁ ≤ m) (hle₂ : m₂ ≤ m) [SigmaFinit (fun f : Lp E 1 μ => ∫ x in s, f x ∂μ) ∘ Submodule.subtypeL _ := by refine funext fun f => integral_congr_ae (ae_restrict_of_ae ?_) simp_rw [Submodule.coe_subtypeL', Submodule.coeSubtype] - exact eventually_of_forall fun _ => (by trivial) + exact Eventually.of_forall fun _ => (by trivial) refine isClosed_eq (Continuous.const_smul ?_ _) ?_ · rw [heq₁] exact continuous_integral.comp (ContinuousLinearMap.continuous _) diff --git a/Mathlib/Probability/Distributions/Exponential.lean b/Mathlib/Probability/Distributions/Exponential.lean index ad5ce86b273a2..fb141c97472bc 100644 --- a/Mathlib/Probability/Distributions/Exponential.lean +++ b/Mathlib/Probability/Distributions/Exponential.lean @@ -139,7 +139,7 @@ lemma lintegral_exponentialPDF_eq_antiDeriv {r : ℝ} (hr : 0 < r) (x : ℝ) : rw [setLIntegral_congr_fun measurableSet_Icc (ae_of_all _ (by intro a ⟨(hle : _ ≤ a), _⟩; rw [if_pos hle]))] rw [← ENNReal.toReal_eq_toReal _ ENNReal.ofReal_ne_top, ← integral_eq_lintegral_of_nonneg_ae - (eventually_of_forall fun _ ↦ le_of_lt (mul_pos hr (exp_pos _)))] + (Eventually.of_forall fun _ ↦ le_of_lt (mul_pos hr (exp_pos _)))] · have : ∫ a in uIoc 0 x, r * rexp (-(r * a)) = ∫ a in (0)..x, r * rexp (-(r * a)) := by rw [intervalIntegral.intervalIntegral_eq_integral_uIoc, smul_eq_mul, if_pos h, one_mul] rw [integral_Icc_eq_integral_Ioc, ← uIoc_of_le h, this] diff --git a/Mathlib/Probability/Independence/Kernel.lean b/Mathlib/Probability/Independence/Kernel.lean index 8e6e7abb69925..65ebb710ed088 100644 --- a/Mathlib/Probability/Independence/Kernel.lean +++ b/Mathlib/Probability/Independence/Kernel.lean @@ -181,7 +181,7 @@ theorem indep_bot_right (m' : MeasurableSpace Ω) {_mΩ : MeasurableSpace Ω} Indep m' ⊥ κ μ := by intros s t _ ht rw [Set.mem_setOf_eq, MeasurableSpace.measurableSet_bot_iff] at ht - refine Filter.eventually_of_forall (fun a ↦ ?_) + refine Filter.Eventually.of_forall (fun a ↦ ?_) cases' ht with ht ht · rw [ht, Set.inter_empty, measure_empty, mul_zero] · rw [ht, Set.inter_univ, measure_univ, mul_one] diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 412687e4436d6..7c4d7a7363b7e 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -294,8 +294,8 @@ theorem ae_null_of_compProd_null (h : (κ ⊗ₖ η) a s = 0) : rw [Filter.eventuallyLE_antisymm_iff] exact ⟨Filter.EventuallyLE.trans_eq - (Filter.eventually_of_forall fun x => (measure_mono (Set.preimage_mono hst) : _)) ht, - Filter.eventually_of_forall fun x => zero_le _⟩ + (Filter.Eventually.of_forall fun x => (measure_mono (Set.preimage_mono hst) : _)) ht, + Filter.Eventually.of_forall fun x => zero_le _⟩ theorem ae_ae_of_ae_compProd {p : β × γ → Prop} (h : ∀ᵐ bc ∂(κ ⊗ₖ η) a, p bc) : ∀ᵐ b ∂κ a, ∀ᵐ c ∂η (a, b), p (b, c) := diff --git a/Mathlib/Probability/Kernel/CondDistrib.lean b/Mathlib/Probability/Kernel/CondDistrib.lean index b91377a2e288d..35ea2d1cf3e43 100644 --- a/Mathlib/Probability/Kernel/CondDistrib.lean +++ b/Mathlib/Probability/Kernel/CondDistrib.lean @@ -209,7 +209,7 @@ theorem condDistrib_ae_eq_condexp (hX : Measurable X) (hY : Measurable Y) (hs : · exact fun t _ _ => (integrable_toReal_condDistrib hX.aemeasurable hs).integrableOn · intro t ht _ rw [integral_toReal ((measurable_condDistrib hs).mono hX.comap_le le_rfl).aemeasurable - (eventually_of_forall fun ω => measure_lt_top (condDistrib Y X μ (X ω)) _), + (Eventually.of_forall fun ω => measure_lt_top (condDistrib Y X μ (X ω)) _), integral_indicator_const _ (hY hs), Measure.restrict_apply (hY hs), smul_eq_mul, mul_one, inter_comm, setLIntegral_condDistrib_of_measurableSet hX hY.aemeasurable hs ht] · refine (Measurable.stronglyMeasurable ?_).aeStronglyMeasurable' diff --git a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean index 2f7e71e4e8064..a27ba1c121a56 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CDFToKernel.lean @@ -585,7 +585,7 @@ lemma setLIntegral_toKernel_prod [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ · exact (Kernel.measurable_coe (hf.toKernel f) ht).comp measurable_prod_mk_left · rw [ht_lintegral] exact measure_ne_top _ _ - · exact eventually_of_forall fun a ↦ measure_mono (subset_univ _) + · exact Eventually.of_forall fun a ↦ measure_mono (subset_univ _) _ = κ a (s ×ˢ univ) - κ a (s ×ˢ t) := by rw [setLIntegral_toKernel_univ hf a hs, ht_lintegral] _ = κ a (s ×ˢ tᶜ) := by @@ -625,13 +625,13 @@ lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) rw [← lintegral_add_compl _ ht₁] have h_eq1 : ∫⁻ x in t₁, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = ∫⁻ x in t₁, hf.toKernel f (a, x) t₂ ∂(ν a) := by - refine setLIntegral_congr_fun ht₁ (eventually_of_forall fun a ha ↦ ?_) + refine setLIntegral_congr_fun ht₁ (Eventually.of_forall fun a ha ↦ ?_) rw [h_prod_eq_snd a ha] have h_eq2 : ∫⁻ x in t₁ᶜ, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} ∂(ν a) = 0 := by suffices h_eq_zero : ∀ x ∈ t₁ᶜ, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t₁ ×ˢ t₂} = 0 by - rw [setLIntegral_congr_fun ht₁.compl (eventually_of_forall h_eq_zero)] + rw [setLIntegral_congr_fun ht₁.compl (Eventually.of_forall h_eq_zero)] simp only [lintegral_const, zero_mul] intro a hat₁ rw [mem_compl_iff] at hat₁ @@ -650,7 +650,7 @@ lemma lintegral_toKernel_mem [IsFiniteKernel κ] (hf : IsCondKernelCDF f κ ν) ∫⁻ x, hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t} ∂(ν a) := by have h_le : (fun x ↦ hf.toKernel f (a, x) {y : ℝ | (x, y) ∈ t}) ≤ᵐ[ν a] fun x ↦ hf.toKernel f (a, x) univ := - eventually_of_forall fun _ ↦ measure_mono (subset_univ _) + Eventually.of_forall fun _ ↦ measure_mono (subset_univ _) rw [lintegral_sub _ _ h_le] · exact Kernel.measurable_kernel_prod_mk_left' ht a refine ((lintegral_mono_ae h_le).trans_lt ?_).ne diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index ddd69fe353d44..8b2f314be1286 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -501,21 +501,21 @@ lemma measurable_density_right (κ : Kernel α (γ × β)) (ν : Kernel α γ) lemma density_mono_set (hκν : fst κ ≤ ν) (a : α) (x : γ) {s s' : Set β} (h : s ⊆ s') : density κ ν a x s ≤ density κ ν a x s' := by refine limsup_le_limsup ?_ ?_ ?_ - · exact eventually_of_forall (fun n ↦ densityProcess_mono_set hκν n a x h) + · exact Eventually.of_forall (fun n ↦ densityProcess_mono_set hκν n a x h) · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ lemma density_nonneg (hκν : fst κ ≤ ν) (a : α) (x : γ) (s : Set β) : 0 ≤ density κ ν a x s := by refine le_limsup_of_frequently_le ?_ ?_ - · exact frequently_of_forall (fun n ↦ densityProcess_nonneg _ _ _ _ _ _) + · exact Frequently.of_forall (fun n ↦ densityProcess_nonneg _ _ _ _ _ _) · exact isBoundedUnder_of ⟨1, fun n ↦ densityProcess_le_one hκν _ _ _ _⟩ lemma density_le_one (hκν : fst κ ≤ ν) (a : α) (x : γ) (s : Set β) : density κ ν a x s ≤ 1 := by refine limsup_le_of_le ?_ ?_ · exact isCoboundedUnder_le_of_le atTop (fun i ↦ densityProcess_nonneg _ _ _ _ _ _) - · exact eventually_of_forall (fun n ↦ densityProcess_le_one hκν _ _ _ _) + · exact Eventually.of_forall (fun n ↦ densityProcess_le_one hκν _ _ _ _) section Integral @@ -543,7 +543,7 @@ lemma tendsto_setIntegral_densityProcess (hκν : fst κ ≤ ν) (𝓝 (∫ x in A, density κ ν a x s ∂(ν a))) := by refine tendsto_setIntegral_of_L1' (μ := ν a) (fun x ↦ density κ ν a x s) (integrable_density hκν a hs) (F := fun i x ↦ densityProcess κ ν i a x s) (l := atTop) - (eventually_of_forall (fun n ↦ integrable_densityProcess hκν _ _ hs)) ?_ A + (Eventually.of_forall (fun n ↦ integrable_densityProcess hκν _ _ hs)) ?_ A refine (tendsto_congr fun n ↦ ?_).mp (tendsto_eLpNorm_one_densityProcess_limitProcess hκν a hs) refine eLpNorm_congr_ae ?_ exact EventuallyEq.rfl.sub (density_ae_eq_limitProcess hκν a hs).symm diff --git a/Mathlib/Probability/Kernel/Disintegration/Integral.lean b/Mathlib/Probability/Kernel/Disintegration/Integral.lean index 5f5de66ae8cd4..20389e93f27ce 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Integral.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Integral.lean @@ -294,10 +294,10 @@ theorem Integrable.integral_norm_condKernel {f : α × Ω → F} (hf_int : Integ theorem Integrable.norm_integral_condKernel {f : α × Ω → E} (hf_int : Integrable f ρ) : Integrable (fun x ↦ ‖∫ y, f (x, y) ∂ρ.condKernel x‖) ρ.fst := by refine hf_int.integral_norm_condKernel.mono hf_int.1.integral_condKernel.norm ?_ - refine Filter.eventually_of_forall fun x ↦ ?_ + refine Filter.Eventually.of_forall fun x ↦ ?_ rw [norm_norm] refine (norm_integral_le_integral_norm _).trans_eq (Real.norm_of_nonneg ?_).symm - exact integral_nonneg_of_ae (Filter.eventually_of_forall fun y ↦ norm_nonneg _) + exact integral_nonneg_of_ae (Filter.Eventually.of_forall fun y ↦ norm_nonneg _) theorem Integrable.integral_condKernel {f : α × Ω → E} (hf_int : Integrable f ρ) : Integrable (fun x ↦ ∫ y, f (x, y) ∂ρ.condKernel x) ρ.fst := diff --git a/Mathlib/Probability/Kernel/IntegralCompProd.lean b/Mathlib/Probability/Kernel/IntegralCompProd.lean index 25fa15ae04f0a..6e080b423d194 100644 --- a/Mathlib/Probability/Kernel/IntegralCompProd.lean +++ b/Mathlib/Probability/Kernel/IntegralCompProd.lean @@ -84,7 +84,7 @@ theorem hasFiniteIntegral_compProd_iff ⦃f : β × γ → E⦄ (h1f : StronglyM HasFiniteIntegral (fun x => ∫ y, ‖f (x, y)‖ ∂η (a, x)) (κ a) := by simp only [HasFiniteIntegral] rw [Kernel.lintegral_compProd _ _ _ h1f.ennnorm] - have : ∀ x, ∀ᵐ y ∂η (a, x), 0 ≤ ‖f (x, y)‖ := fun x => eventually_of_forall fun y => norm_nonneg _ + have : ∀ x, ∀ᵐ y ∂η (a, x), 0 ≤ ‖f (x, y)‖ := fun x => Eventually.of_forall fun y => norm_nonneg _ simp_rw [integral_eq_lintegral_of_nonneg_ae (this _) (h1f.norm.comp_measurable measurable_prod_mk_left).aestronglyMeasurable, ennnorm_eq_ofReal toReal_nonneg, ofReal_norm_eq_coe_nnnorm] @@ -130,11 +130,11 @@ theorem _root_.MeasureTheory.Integrable.integral_compProd [NormedSpace ℝ E] ⦃f : β × γ → E⦄ (hf : Integrable f ((κ ⊗ₖ η) a)) : Integrable (fun x => ∫ y, f (x, y) ∂η (a, x)) (κ a) := Integrable.mono hf.integral_norm_compProd hf.aestronglyMeasurable.integral_kernel_compProd <| - eventually_of_forall fun x => + Eventually.of_forall fun x => (norm_integral_le_integral_norm _).trans_eq <| (norm_of_nonneg <| integral_nonneg_of_ae <| - eventually_of_forall fun y => (norm_nonneg (f (x, y)) : _)).symm + Eventually.of_forall fun y => (norm_nonneg (f (x, y)) : _)).symm /-! ### Bochner integral -/ @@ -199,7 +199,7 @@ theorem Kernel.continuous_integral_integral : rw [continuous_iff_continuousAt]; intro g refine tendsto_integral_of_L1 _ (L1.integrable_coeFn g).integral_compProd - (eventually_of_forall fun h => (L1.integrable_coeFn h).integral_compProd) ?_ + (Eventually.of_forall fun h => (L1.integrable_coeFn h).integral_compProd) ?_ simp_rw [← Kernel.lintegral_fn_integral_sub (fun x => (‖x‖₊ : ℝ≥0∞)) (L1.integrable_coeFn _) (L1.integrable_coeFn g)] diff --git a/Mathlib/Probability/Kernel/MeasurableIntegral.lean b/Mathlib/Probability/Kernel/MeasurableIntegral.lean index 98543badf7199..6790cc25384d2 100644 --- a/Mathlib/Probability/Kernel/MeasurableIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableIntegral.lean @@ -285,11 +285,11 @@ theorem StronglyMeasurable.integral_kernel_prod_right ⦃f : α → β → E⦄ tendsto_integral_of_dominated_convergence (fun y => ‖f x y‖ + ‖f x y‖) (fun n => (s' n x).aestronglyMeasurable) (hfx.norm.add hfx.norm) ?_ ?_ · -- Porting note: was - -- exact fun n => eventually_of_forall fun y => + -- exact fun n => Eventually.of_forall fun y => -- SimpleFunc.norm_approxOn_zero_le _ _ (x, y) n - exact fun n => eventually_of_forall fun y => + exact fun n => Eventually.of_forall fun y => SimpleFunc.norm_approxOn_zero_le hf.measurable (by simp) (x, y) n - · refine eventually_of_forall fun y => SimpleFunc.tendsto_approxOn hf.measurable (by simp) ?_ + · refine Eventually.of_forall fun y => SimpleFunc.tendsto_approxOn hf.measurable (by simp) ?_ apply subset_closure simp [-uncurry_apply_pair] · simp [f', hfx, integral_undef] diff --git a/Mathlib/Probability/Martingale/Basic.lean b/Mathlib/Probability/Martingale/Basic.lean index f4d4655f4bb85..08032df92e947 100644 --- a/Mathlib/Probability/Martingale/Basic.lean +++ b/Mathlib/Probability/Martingale/Basic.lean @@ -241,10 +241,10 @@ protected theorem sup {f g : ι → Ω → ℝ} (hf : Submartingale f ℱ μ) (h refine EventuallyLE.sup_le ?_ ?_ · exact EventuallyLE.trans (hf.2.1 i j hij) (condexp_mono (hf.integrable _) (Integrable.sup (hf.integrable j) (hg.integrable j)) - (eventually_of_forall fun x => le_max_left _ _)) + (Eventually.of_forall fun x => le_max_left _ _)) · exact EventuallyLE.trans (hg.2.1 i j hij) (condexp_mono (hg.integrable _) (Integrable.sup (hf.integrable j) (hg.integrable j)) - (eventually_of_forall fun x => le_max_right _ _)) + (Eventually.of_forall fun x => le_max_right _ _)) protected theorem pos {f : ι → Ω → ℝ} (hf : Submartingale f ℱ μ) : Submartingale (f⁺) ℱ μ := hf.sup (martingale_zero _ _ _).submartingale @@ -489,7 +489,7 @@ theorem Submartingale.sum_mul_sub [IsFiniteMeasure μ] {R : ℝ} {ξ f : ℕ → refine submartingale_of_condexp_sub_nonneg_nat hadp hint fun i => ?_ simp only [← Finset.sum_Ico_eq_sub _ (Nat.le_succ _), Finset.sum_apply, Pi.mul_apply, Pi.sub_apply, Nat.Ico_succ_singleton, Finset.sum_singleton] - exact EventuallyLE.trans (EventuallyLE.mul_nonneg (eventually_of_forall (hnonneg _)) + exact EventuallyLE.trans (EventuallyLE.mul_nonneg (Eventually.of_forall (hnonneg _)) (hf.condexp_sub_nonneg (Nat.le_succ _))) (condexp_stronglyMeasurable_mul (hξ _) (((hf.integrable _).sub (hf.integrable _)).bdd_mul hξ.stronglyMeasurable.aestronglyMeasurable (hξbdd _)) diff --git a/Mathlib/Probability/Martingale/BorelCantelli.lean b/Mathlib/Probability/Martingale/BorelCantelli.lean index a210e0652fd50..bf606fa145ff1 100644 --- a/Mathlib/Probability/Martingale/BorelCantelli.lean +++ b/Mathlib/Probability/Martingale/BorelCantelli.lean @@ -347,8 +347,8 @@ theorem tendsto_sum_indicator_atTop_iff' [IsFiniteMeasure μ] {s : ℕ → Set (s (k + 1)).indicator (1 : Ω → ℝ) ω) atTop atTop ↔ Tendsto (fun n => ∑ k ∈ Finset.range n, (μ[(s (k + 1)).indicator (1 : Ω → ℝ)|ℱ k]) ω) atTop atTop := by - have := tendsto_sum_indicator_atTop_iff (eventually_of_forall fun ω n => ?_) (adapted_process hs) - (integrable_process μ hs) (eventually_of_forall <| process_difference_le s) + have := tendsto_sum_indicator_atTop_iff (Eventually.of_forall fun ω n => ?_) (adapted_process hs) + (integrable_process μ hs) (Eventually.of_forall <| process_difference_le s) swap · rw [process, process, ← sub_nonneg, Finset.sum_apply, Finset.sum_apply, Finset.sum_range_succ_sub_sum] diff --git a/Mathlib/Probability/Martingale/OptionalSampling.lean b/Mathlib/Probability/Martingale/OptionalSampling.lean index 9f77e4e6fc6c7..3132d43cb03f9 100644 --- a/Mathlib/Probability/Martingale/OptionalSampling.lean +++ b/Mathlib/Probability/Martingale/OptionalSampling.lean @@ -76,7 +76,7 @@ theorem stoppedValue_ae_eq_restrict_eq (h : Martingale f ℱ μ) (hτ : IsStoppi refine Filter.EventuallyEq.trans ?_ (condexp_stopping_time_ae_eq_restrict_eq_const_of_le_const h hτ hτ_le i).symm rw [Filter.EventuallyEq, ae_restrict_iff' (ℱ.le _ _ (hτ.measurableSet_eq i))] - refine Filter.eventually_of_forall fun x hx => ?_ + refine Filter.Eventually.of_forall fun x hx => ?_ rw [Set.mem_setOf_eq] at hx simp_rw [stoppedValue, hx] diff --git a/Mathlib/Probability/Martingale/Upcrossing.lean b/Mathlib/Probability/Martingale/Upcrossing.lean index 70a525bec0762..148f91df06297 100644 --- a/Mathlib/Probability/Martingale/Upcrossing.lean +++ b/Mathlib/Probability/Martingale/Upcrossing.lean @@ -605,7 +605,7 @@ theorem integral_mul_upcrossingsBefore_le_integral [IsFiniteMeasure μ] (hf : Su μ[∑ k ∈ Finset.range N, upcrossingStrat a b f N k * (f (k + 1) - f k)] := by rw [← integral_mul_left] refine integral_mono_of_nonneg ?_ ((hf.sum_upcrossingStrat_mul a b N).integrable N) ?_ - · exact eventually_of_forall fun ω => mul_nonneg (sub_nonneg.2 hab.le) (Nat.cast_nonneg _) + · exact Eventually.of_forall fun ω => mul_nonneg (sub_nonneg.2 hab.le) (Nat.cast_nonneg _) · filter_upwards with ω simpa using mul_upcrossingsBefore_le (hfN ω) hab _ ≤ μ[f N] - μ[f 0] := hf.sum_mul_upcrossingStrat_le @@ -794,7 +794,7 @@ theorem Submartingale.mul_lintegral_upcrossings_le_lintegral_pos_part [IsFiniteM intro N rw [ofReal_integral_eq_lintegral_ofReal] · exact (hf.sub_martingale (martingale_const _ _ _)).pos.integrable _ - · exact eventually_of_forall fun ω => posPart_nonneg _ + · exact Eventually.of_forall fun ω => posPart_nonneg _ rw [lintegral_iSup'] · simp_rw [this, ENNReal.mul_iSup, iSup_le_iff] intro N diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 2f8ff3eb320c1..12a5969e58728 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -119,7 +119,7 @@ theorem truncation_nonneg {f : α → ℝ} (A : ℝ) {x : α} (h : 0 ≤ f x) : theorem _root_.MeasureTheory.AEStronglyMeasurable.memℒp_truncation [IsFiniteMeasure μ] (hf : AEStronglyMeasurable f μ) {A : ℝ} {p : ℝ≥0∞} : Memℒp (truncation f A) p μ := - Memℒp.of_bound hf.truncation |A| (eventually_of_forall fun _ => abs_truncation_le_bound _ _ _) + Memℒp.of_bound hf.truncation |A| (Eventually.of_forall fun _ => abs_truncation_le_bound _ _ _) theorem _root_.MeasureTheory.AEStronglyMeasurable.integrable_truncation [IsFiniteMeasure μ] (hf : AEStronglyMeasurable f μ) {A : ℝ} : Integrable (truncation f A) μ := by @@ -153,7 +153,7 @@ theorem moment_truncation_eq_intervalIntegral_of_nonneg (hf : AEStronglyMeasurab zero_eq_neg] apply integral_eq_zero_of_ae have : ∀ᵐ x ∂Measure.map f μ, (0 : ℝ) ≤ x := - (ae_map_iff hf.aemeasurable measurableSet_Ici).2 (eventually_of_forall h'f) + (ae_map_iff hf.aemeasurable measurableSet_Ici).2 (Eventually.of_forall h'f) filter_upwards [this] with x hx simp only [indicator, Set.mem_Ioc, Pi.zero_apply, ite_eq_right_iff, and_imp] intro _ h''x @@ -172,7 +172,7 @@ theorem integral_truncation_eq_intervalIntegral_of_nonneg (hf : AEStronglyMeasur theorem integral_truncation_le_integral_of_nonneg (hf : Integrable f μ) (h'f : 0 ≤ f) {A : ℝ} : ∫ x, truncation f A x ∂μ ≤ ∫ x, f x ∂μ := by apply integral_mono_of_nonneg - (eventually_of_forall fun x => ?_) hf (eventually_of_forall fun x => ?_) + (Eventually.of_forall fun x => ?_) hf (Eventually.of_forall fun x => ?_) · exact truncation_nonneg _ (h'f x) · calc truncation f A x ≤ |truncation f A x| := le_abs_self _ @@ -184,7 +184,7 @@ integral of the whole function. -/ theorem tendsto_integral_truncation {f : α → ℝ} (hf : Integrable f μ) : Tendsto (fun A => ∫ x, truncation f A x ∂μ) atTop (𝓝 (∫ x, f x ∂μ)) := by refine tendsto_integral_filter_of_dominated_convergence (fun x => abs (f x)) ?_ ?_ ?_ ?_ - · exact eventually_of_forall fun A ↦ hf.aestronglyMeasurable.truncation + · exact Eventually.of_forall fun A ↦ hf.aestronglyMeasurable.truncation · filter_upwards with A filter_upwards with x rw [Real.norm_eq_abs] @@ -285,7 +285,7 @@ theorem tsum_prob_mem_Ioi_lt_top {X : Ω → ℝ} (hint : Integrable X) (hnonneg (∑' j : ℕ, ℙ {ω | X ω ∈ Set.Ioi (j : ℝ)}) < ∞ := by suffices ∀ K : ℕ, ∑ j ∈ range K, ℙ {ω | X ω ∈ Set.Ioi (j : ℝ)} ≤ ENNReal.ofReal (𝔼[X] + 1) from (le_of_tendsto_of_tendsto (ENNReal.tendsto_nat_tsum _) tendsto_const_nhds - (eventually_of_forall this)).trans_lt ENNReal.ofReal_lt_top + (Eventually.of_forall this)).trans_lt ENNReal.ofReal_lt_top intro K have A : Tendsto (fun N : ℕ => ∑ j ∈ range K, ℙ {ω | X ω ∈ Set.Ioc (j : ℝ) N}) atTop (𝓝 (∑ j ∈ range K, ℙ {ω | X ω ∈ Set.Ioi (j : ℝ)})) := by @@ -725,7 +725,7 @@ lemma strong_law_ae_of_measurable -- check that, when both convergences above hold, then the strong law is satisfied filter_upwards [A, B] with ω hω h'ω rw [tendsto_iff_norm_sub_tendsto_zero, tendsto_order] - refine ⟨fun c hc ↦ eventually_of_forall (fun n ↦ hc.trans_le (norm_nonneg _)), ?_⟩ + refine ⟨fun c hc ↦ Eventually.of_forall (fun n ↦ hc.trans_le (norm_nonneg _)), ?_⟩ -- start with some positive `ε` (the desired precision), and fix `δ` with `3 δ < ε`. intro ε (εpos : 0 < ε) obtain ⟨δ, δpos, hδ⟩ : ∃ δ, 0 < δ ∧ δ + δ + δ < ε := ⟨ε/4, by positivity, by linarith⟩ diff --git a/Mathlib/Probability/Variance.lean b/Mathlib/Probability/Variance.lean index b737727f95f56..d2036c459f53a 100644 --- a/Mathlib/Probability/Variance.lean +++ b/Mathlib/Probability/Variance.lean @@ -210,7 +210,7 @@ theorem variance_le_expectation_sq [@IsProbabilityMeasure Ω _ ℙ] {X : Ω → apply hX convert A.add B simp - · exact eventually_of_forall fun x => sq_nonneg _ + · exact Eventually.of_forall fun x => sq_nonneg _ · exact (AEMeasurable.pow_const (hm.aemeasurable.sub_const _) _).aestronglyMeasurable theorem evariance_def' [@IsProbabilityMeasure Ω _ ℙ] {X : Ω → ℝ} (hX : AEStronglyMeasurable X ℙ) : diff --git a/Mathlib/Tactic/Peel.lean b/Mathlib/Tactic/Peel.lean index 89c1c50bfa47a..b3ab2443710a5 100644 --- a/Mathlib/Tactic/Peel.lean +++ b/Mathlib/Tactic/Peel.lean @@ -10,7 +10,7 @@ import Mathlib.Tactic.Basic # The `peel` tactic `peel h with h' idents*` tries to apply `forall_imp` (or `Exists.imp`, or `Filter.Eventually.mp`, -`Filter.Frequently.mp` and `Filter.eventually_of_forall`) with the argument `h` and uses `idents*` +`Filter.Frequently.mp` and `Filter.Eventually.of_forall`) with the argument `h` and uses `idents*` to introduce variables with the supplied names, giving the "peeled" argument the name `h'`. One can provide a numeric argument as in `peel 4 h` which will peel 4 quantifiers off @@ -78,7 +78,7 @@ immediately. In particular, `peel h using e` is equivalent to `peel h; exact e`. may be paired with any of the other features of `peel`. This tactic works by repeatedly applying lemmas such as `forall_imp`, `Exists.imp`, -`Filter.Eventually.mp`, `Filter.Frequently.mp`, and `Filter.eventually_of_forall`. +`Filter.Eventually.mp`, `Filter.Frequently.mp`, and `Filter.Eventually.of_forall`. -/ syntax (name := peel) "peel" (num)? (ppSpace colGt term)? @@ -88,11 +88,11 @@ private lemma and_imp_left_of_imp_imp {p q r : Prop} (h : r → p → q) : r ∧ private theorem eventually_imp {α : Type*} {p q : α → Prop} {f : Filter α} (hq : ∀ (x : α), p x → q x) (hp : ∀ᶠ (x : α) in f, p x) : ∀ᶠ (x : α) in f, q x := - Filter.Eventually.mp hp (Filter.eventually_of_forall hq) + Filter.Eventually.mp hp (Filter.Eventually.of_forall hq) private theorem frequently_imp {α : Type*} {p q : α → Prop} {f : Filter α} (hq : ∀ (x : α), p x → q x) (hp : ∃ᶠ (x : α) in f, p x) : ∃ᶠ (x : α) in f, q x := - Filter.Frequently.mp hp (Filter.eventually_of_forall hq) + Filter.Frequently.mp hp (Filter.Eventually.of_forall hq) private theorem eventually_congr {α : Type*} {p q : α → Prop} {f : Filter α} (hq : ∀ (x : α), p x ↔ q x) : (∀ᶠ (x : α) in f, p x) ↔ ∀ᶠ (x : α) in f, q x := by diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean index c73c289acf3f5..b21eaf9c595e2 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean @@ -209,7 +209,7 @@ theorem cauchySeq_finset_iff_tprod_vanishing : refine ⟨s, fun t hts ↦ oe ?_⟩ by_cases ht : Multipliable fun a : t ↦ f a · classical - refine o_closed.mem_of_tendsto ht.hasProd (eventually_of_forall fun t' ↦ ?_) + refine o_closed.mem_of_tendsto ht.hasProd (Eventually.of_forall fun t' ↦ ?_) rw [← prod_subtype_map_embedding fun _ _ ↦ by rfl] apply hs simp_rw [Finset.mem_map] diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean index e0e206d0522bb..a22caf1d21fd1 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Order.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Order.lean @@ -287,4 +287,4 @@ theorem Summable.tendsto_atTop_of_pos [LinearOrderedField α] [TopologicalSpace {f : ℕ → α} (hf : Summable f⁻¹) (hf' : ∀ n, 0 < f n) : Tendsto f atTop atTop := inv_inv f ▸ Filter.Tendsto.inv_tendsto_zero <| tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within _ hf.tendsto_atTop_zero <| - eventually_of_forall fun _ ↦ inv_pos.2 (hf' _) + Eventually.of_forall fun _ ↦ inv_pos.2 (hf' _) diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index d075322b1f185..03ce9dac62eb3 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -338,7 +338,7 @@ def linearMapOfMemClosureRangeCoe (f : M₁ → M₂) def linearMapOfTendsto (f : M₁ → M₂) (g : α → M₁ →ₛₗ[σ] M₂) [l.NeBot] (h : Tendsto (fun a x => g a x) l (𝓝 f)) : M₁ →ₛₗ[σ] M₂ := linearMapOfMemClosureRangeCoe f <| - mem_closure_of_tendsto h <| eventually_of_forall fun _ => Set.mem_range_self _ + mem_closure_of_tendsto h <| Eventually.of_forall fun _ => Set.mem_range_self _ variable (M₁ M₂ σ) diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index dd2ee7af35ad6..7d976664104f1 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -336,7 +336,7 @@ monoid homomorphisms"] def monoidHomOfTendsto (f : M₁ → M₂) (g : α → F) [l.NeBot] (h : Tendsto (fun a x => g a x) l (𝓝 f)) : M₁ →* M₂ := monoidHomOfMemClosureRangeCoe f <| - mem_closure_of_tendsto h <| eventually_of_forall fun _ => mem_range_self _ + mem_closure_of_tendsto h <| Eventually.of_forall fun _ => mem_range_self _ variable (M₁ M₂) diff --git a/Mathlib/Topology/Algebra/Order/Floor.lean b/Mathlib/Topology/Algebra/Order/Floor.lean index 1eb2f5f1f2ad0..3877c5331e7bf 100644 --- a/Mathlib/Topology/Algebra/Order/Floor.lean +++ b/Mathlib/Topology/Algebra/Order/Floor.lean @@ -146,7 +146,7 @@ theorem tendsto_fract_left' [OrderClosedTopology α] [TopologicalAddGroup α] (n theorem tendsto_fract_left [OrderClosedTopology α] [TopologicalAddGroup α] (n : ℤ) : Tendsto (fract : α → α) (𝓝[<] n) (𝓝[<] 1) := tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within _ (tendsto_fract_left' _) - (eventually_of_forall fract_lt_one) + (Eventually.of_forall fract_lt_one) theorem tendsto_fract_right' [OrderClosedTopology α] [TopologicalAddGroup α] (n : ℤ) : Tendsto (fract : α → α) (𝓝[≥] n) (𝓝 0) := @@ -155,7 +155,7 @@ theorem tendsto_fract_right' [OrderClosedTopology α] [TopologicalAddGroup α] ( theorem tendsto_fract_right [OrderClosedTopology α] [TopologicalAddGroup α] (n : ℤ) : Tendsto (fract : α → α) (𝓝[≥] n) (𝓝[≥] 0) := tendsto_nhdsWithin_of_tendsto_nhds_of_eventually_within _ (tendsto_fract_right' _) - (eventually_of_forall fract_nonneg) + (Eventually.of_forall fract_nonneg) local notation "I" => (Icc 0 1 : Set α) diff --git a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean index fa6d36430e5cf..dc3ec2436b49f 100644 --- a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean +++ b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean @@ -136,13 +136,13 @@ instance (priority := 100) OrderBot.to_BoundedGENhdsClass [OrderBot α] : Bounde instance (priority := 100) OrderTopology.to_BoundedLENhdsClass [IsDirected α (· ≤ ·)] [OrderTopology α] : BoundedLENhdsClass α := ⟨fun a ↦ - ((isTop_or_exists_gt a).elim fun h ↦ ⟨a, eventually_of_forall h⟩) <| + ((isTop_or_exists_gt a).elim fun h ↦ ⟨a, Eventually.of_forall h⟩) <| Exists.imp fun _b ↦ ge_mem_nhds⟩ -- 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⟩) <| + ⟨fun a ↦ ((isBot_or_exists_lt a).elim fun h ↦ ⟨a, Eventually.of_forall h⟩) <| Exists.imp fun _b ↦ le_mem_nhds⟩ end Preorder @@ -245,13 +245,13 @@ variable [FirstCountableTopology α] {f : Filter β} [CountableInterFilter f] {u theorem eventually_le_limsup (hf : IsBoundedUnder (· ≤ ·) f u := by isBoundedDefault) : ∀ᶠ b in f, u b ≤ f.limsup u := by obtain ha | ha := isTop_or_exists_gt (f.limsup u) - · exact eventually_of_forall fun _ => ha _ + · exact Eventually.of_forall fun _ => ha _ by_cases H : IsGLB (Set.Ioi (f.limsup u)) (f.limsup u) · obtain ⟨u, -, -, hua, hu⟩ := H.exists_seq_antitone_tendsto ha have := fun n => eventually_lt_of_limsup_lt (hu n) hf exact (eventually_countable_forall.2 this).mono fun b hb => - ge_of_tendsto hua <| eventually_of_forall fun n => (hb _).le + ge_of_tendsto hua <| Eventually.of_forall fun n => (hb _).le · obtain ⟨x, hx, xa⟩ : ∃ x, (∀ ⦃b⦄, f.limsup u < b → x ≤ b) ∧ f.limsup u < x := by simp only [IsGLB, IsGreatest, lowerBounds, upperBounds, Set.mem_Ioi, Set.mem_setOf_eq, not_and, not_forall, not_le, exists_prop] at H @@ -274,7 +274,7 @@ variable [CompleteLinearOrder α] [TopologicalSpace α] [FirstCountableTopology @[simp] theorem limsup_eq_bot : f.limsup u = ⊥ ↔ u =ᶠ[f] ⊥ := ⟨fun h => - (EventuallyLE.trans eventually_le_limsup <| eventually_of_forall fun _ => h.le).mono fun x hx => + (EventuallyLE.trans eventually_le_limsup <| Eventually.of_forall fun _ => h.le).mono fun x hx => le_antisymm hx bot_le, fun h => by rw [limsup_congr h] @@ -328,7 +328,7 @@ theorem Antitone.map_limsSup_of_continuousAt {F : Filter R} [NeBot F] {f : R → by_contra! H have not_bot : ¬ IsBot F.limsSup := fun maybe_bot ↦ lt_irrefl (F.liminf f) <| lt_of_le_of_lt - (liminf_le_of_frequently_le (frequently_of_forall (fun r ↦ f_decr (maybe_bot r))) + (liminf_le_of_frequently_le (Frequently.of_forall (fun r ↦ f_decr (maybe_bot r))) (bdd_above.isBoundedUnder f_decr)) H obtain ⟨l, l_lt, h'l⟩ : ∃ l < F.limsSup, Set.Ioc l F.limsSup ⊆ { x : R | f x < F.liminf f } := by diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index 4b628fa42662d..d7d6a27b5b8c0 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -1451,7 +1451,7 @@ theorem Filter.EventuallyEq.continuousAt (h : f =ᶠ[𝓝 x] fun _ => y) : theorem continuous_of_const (h : ∀ x y, f x = f y) : Continuous f := continuous_iff_continuousAt.mpr fun x => - Filter.EventuallyEq.continuousAt <| eventually_of_forall fun y => h y x + Filter.EventuallyEq.continuousAt <| Eventually.of_forall fun y => h y x theorem continuousAt_id : ContinuousAt id x := continuous_id.continuousAt diff --git a/Mathlib/Topology/Bornology/Absorbs.lean b/Mathlib/Topology/Bornology/Absorbs.lean index ea21c3717d826..186c664199d67 100644 --- a/Mathlib/Topology/Bornology/Absorbs.lean +++ b/Mathlib/Topology/Bornology/Absorbs.lean @@ -129,7 +129,7 @@ protected lemma add [AddZeroClass E] [DistribSMul M E] h₂.mp <| h₁.eventually.mono fun x hx₁ hx₂ ↦ by rw [smul_add]; exact add_subset_add hx₁ hx₂ protected lemma zero [Zero E] [SMulZeroClass M E] {s : Set E} (hs : 0 ∈ s) : Absorbs M s 0 := - eventually_of_forall fun _ ↦ zero_subset.2 <| zero_mem_smul_set hs + Eventually.of_forall fun _ ↦ zero_subset.2 <| zero_mem_smul_set hs end AddZero diff --git a/Mathlib/Topology/ContinuousFunction/Bounded.lean b/Mathlib/Topology/ContinuousFunction/Bounded.lean index fe002c06a38ed..80e9caca4ffa6 100644 --- a/Mathlib/Topology/ContinuousFunction/Bounded.lean +++ b/Mathlib/Topology/ContinuousFunction/Bounded.lean @@ -228,12 +228,12 @@ theorem tendsto_iff_tendstoUniformly {ι : Type*} {F : ι → α →ᵇ β} {f : (fun h => tendstoUniformly_iff.2 fun ε ε0 => (Metric.tendsto_nhds.mp h ε ε0).mp - (eventually_of_forall fun n hn x => + (Eventually.of_forall fun n hn x => lt_of_le_of_lt (dist_coe_le_dist x) (dist_comm (F n) f ▸ hn))) fun h => Metric.tendsto_nhds.mpr fun _ ε_pos => (h _ (dist_mem_uniformity <| half_pos ε_pos)).mp - (eventually_of_forall fun n hn => + (Eventually.of_forall fun n hn => lt_of_le_of_lt ((dist_le (half_pos ε_pos).le).mpr fun x => dist_comm (f x) (F n x) ▸ le_of_lt (hn x)) (half_lt_self ε_pos)) @@ -307,7 +307,7 @@ instance instCompleteSpace [CompleteSpace β] : CompleteSpace (α →ᵇ β) := refine ((tendsto_order.1 b_lim).2 ε ε0).mono fun n hn x => ?_ rw [dist_comm] exact lt_of_le_of_lt (fF_bdd x n) hn - exact this.continuous (eventually_of_forall fun N => (f N).continuous) + exact this.continuous (Eventually.of_forall fun N => (f N).continuous) · -- Check that `F` is bounded rcases (f 0).bounded with ⟨C, hC⟩ refine ⟨C + (b 0 + b 0), fun x y => ?_⟩ diff --git a/Mathlib/Topology/ContinuousFunction/Weierstrass.lean b/Mathlib/Topology/ContinuousFunction/Weierstrass.lean index dfc2335f23237..90730fc0e1d83 100644 --- a/Mathlib/Topology/ContinuousFunction/Weierstrass.lean +++ b/Mathlib/Topology/ContinuousFunction/Weierstrass.lean @@ -32,7 +32,7 @@ theorem polynomialFunctions_closure_eq_top' : (polynomialFunctions I).topologica rintro f - refine Filter.Frequently.mem_closure ?_ refine Filter.Tendsto.frequently (bernsteinApproximation_uniform f) ?_ - apply frequently_of_forall + apply Frequently.of_forall intro n simp only [SetLike.mem_coe] apply Subalgebra.sum_mem diff --git a/Mathlib/Topology/ContinuousFunction/ZeroAtInfty.lean b/Mathlib/Topology/ContinuousFunction/ZeroAtInfty.lean index 586b078cd9837..579bbebe2285c 100644 --- a/Mathlib/Topology/ContinuousFunction/ZeroAtInfty.lean +++ b/Mathlib/Topology/ContinuousFunction/ZeroAtInfty.lean @@ -416,7 +416,7 @@ theorem isClosed_range_toBCF : IsClosed (range (toBCF : C₀(α, β) → α → refine Metric.tendsto_nhds.mpr fun ε hε => ?_ obtain ⟨_, hg, g, rfl⟩ := hf (ball f (ε / 2)) (ball_mem_nhds f <| half_pos hε) refine (Metric.tendsto_nhds.mp (zero_at_infty g) (ε / 2) (half_pos hε)).mp - (eventually_of_forall fun x hx => ?_) + (Eventually.of_forall fun x hx => ?_) calc dist (f x) 0 ≤ dist (g.toBCF x) (f x) + dist (g x) 0 := dist_triangle_left _ _ _ _ < dist g.toBCF f + ε / 2 := add_lt_add_of_le_of_lt (dist_coe_le_dist x) hx diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index 44d171d882c3f..cc5cc3ae5b508 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -384,7 +384,7 @@ theorem tendsto_nhdsWithin_iff {a : α} {l : Filter β} {s : Set α} {f : β → theorem tendsto_nhdsWithin_range {a : α} {l : Filter β} {f : β → α} : Tendsto f l (𝓝[range f] a) ↔ Tendsto f l (𝓝 a) := ⟨fun h => h.mono_right inf_le_left, fun h => - tendsto_inf.2 ⟨h, tendsto_principal.2 <| eventually_of_forall mem_range_self⟩⟩ + tendsto_inf.2 ⟨h, tendsto_principal.2 <| Eventually.of_forall mem_range_self⟩⟩ theorem Filter.EventuallyEq.eq_of_nhdsWithin {s : Set α} {f g : α → β} {a : α} (h : f =ᶠ[𝓝[s] a] g) (hmem : a ∈ s) : f a = g a := @@ -711,7 +711,7 @@ theorem continuousWithinAt_update_same [DecidableEq α] {f : α → β} {s : Set ContinuousWithinAt (update f x y) s x ↔ Tendsto (update f x y) (𝓝[s \ {x}] x) (𝓝 y) := by { rw [← continuousWithinAt_diff_self, ContinuousWithinAt, update_same] } _ ↔ Tendsto f (𝓝[s \ {x}] x) (𝓝 y) := - tendsto_congr' <| eventually_nhdsWithin_iff.2 <| eventually_of_forall + tendsto_congr' <| eventually_nhdsWithin_iff.2 <| Eventually.of_forall fun z hz => update_noteq hz.2 _ _ @[simp] diff --git a/Mathlib/Topology/DenseEmbedding.lean b/Mathlib/Topology/DenseEmbedding.lean index b6754581c6062..7957b2b180fe4 100644 --- a/Mathlib/Topology/DenseEmbedding.lean +++ b/Mathlib/Topology/DenseEmbedding.lean @@ -165,7 +165,7 @@ theorem extend_unique_at [T2Space γ] {b : β} {f : α → γ} {g : β → γ} ( theorem extend_unique [T2Space γ] {f : α → γ} {g : β → γ} (di : DenseInducing i) (hf : ∀ x, g (i x) = f x) (hg : Continuous g) : di.extend f = g := - funext fun _ => extend_unique_at di (eventually_of_forall hf) hg.continuousAt + funext fun _ => extend_unique_at di (Eventually.of_forall hf) hg.continuousAt theorem continuousAt_extend [T3Space γ] {b : β} {f : α → γ} (di : DenseInducing i) (hf : ∀ᶠ x in 𝓝 b, ∃ c, Tendsto f (comap i <| 𝓝 x) (𝓝 c)) : ContinuousAt (di.extend f) b := by diff --git a/Mathlib/Topology/Germ.lean b/Mathlib/Topology/Germ.lean index ea5f439b093ec..5b1e7856f5d62 100644 --- a/Mathlib/Topology/Germ.lean +++ b/Mathlib/Topology/Germ.lean @@ -117,7 +117,7 @@ theorem forall_restrictGermPredicate_iff {P : ∀ x : X, Germ (𝓝 x) Y → Pro theorem forall_restrictGermPredicate_of_forall {P : ∀ x : X, Germ (𝓝 x) Y → Prop} (h : ∀ x, P x f) : ∀ x, RestrictGermPredicate P A x f := - forall_restrictGermPredicate_iff.mpr (eventually_of_forall h) + forall_restrictGermPredicate_iff.mpr (Eventually.of_forall h) end RestrictGermPredicate namespace Filter.Germ diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index bfe23ec8d795c..0b4d7a5befa98 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -432,7 +432,7 @@ theorem le_of_forall_lt_one_mul_le {x y : ℝ≥0∞} (h : ∀ a < 1, a * x ≤ have : Tendsto (· * x) (𝓝[<] 1) (𝓝 (1 * x)) := (ENNReal.continuousAt_mul_const (Or.inr one_ne_zero)).mono_left inf_le_left rw [one_mul] at this - exact le_of_tendsto this (eventually_nhdsWithin_iff.2 <| eventually_of_forall h) + exact le_of_tendsto this (eventually_nhdsWithin_iff.2 <| Eventually.of_forall h) theorem iInf_mul_left' {ι} {f : ι → ℝ≥0∞} {a : ℝ≥0∞} (h : a = ∞ → ⨅ i, f i = 0 → ∃ i, f i = 0) (h0 : a = 0 → Nonempty ι) : ⨅ i, a * f i = a * ⨅ i, f i := by @@ -1499,7 +1499,7 @@ lemma liminf_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} ( have key := Monotone.map_liminf_of_continuousAt (F := F) (monotone_truncateToReal b_ne_top) xs (continuous_truncateToReal b_ne_top).continuousAt (IsBoundedUnder.isCoboundedUnder_ge ⟨b, by simpa only [eventually_map] using le_b⟩) - ⟨0, eventually_of_forall (by simp)⟩ + ⟨0, Eventually.of_forall (by simp)⟩ rw [key] rfl @@ -1516,7 +1516,7 @@ lemma limsup_toReal_eq {ι : Type*} {F : Filter ι} [NeBot F] {b : ℝ≥0∞} ( have key := Monotone.map_limsup_of_continuousAt (F := F) (monotone_truncateToReal b_ne_top) xs (continuous_truncateToReal b_ne_top).continuousAt ⟨b, by simpa only [eventually_map] using le_b⟩ - (IsBoundedUnder.isCoboundedUnder_le ⟨0, eventually_of_forall (by simp)⟩) + (IsBoundedUnder.isCoboundedUnder_le ⟨0, Eventually.of_forall (by simp)⟩) rw [key] rfl diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index 9da4bbc6a3b9b..a14e8778f0703 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -165,7 +165,7 @@ private lemma limsup_add_le_of_lt (ha : limsup u f < a) (hb : limsup v f < b) : · simp only [limsup_bot, bot_le] rw [← @limsup_const EReal α _ f _ (a + b)] apply limsup_le_limsup (Eventually.mp (Eventually.and (eventually_lt_of_limsup_lt ha) - (eventually_lt_of_limsup_lt hb)) (eventually_of_forall _)) + (eventually_lt_of_limsup_lt hb)) (Eventually.of_forall _)) simp only [Pi.add_apply, and_imp] intro x exact fun ux_lt_a vx_lt_b ↦ add_le_add (le_of_lt ux_lt_a) (le_of_lt vx_lt_b) diff --git a/Mathlib/Topology/LocallyConstant/Basic.lean b/Mathlib/Topology/LocallyConstant/Basic.lean index 4e3ac23a0989d..48f981df57cff 100644 --- a/Mathlib/Topology/LocallyConstant/Basic.lean +++ b/Mathlib/Topology/LocallyConstant/Basic.lean @@ -99,7 +99,7 @@ theorem iff_continuous {_ : TopologicalSpace Y} [DiscreteTopology Y] (f : X → ⟨IsLocallyConstant.continuous, fun h s => h.isOpen_preimage s (isOpen_discrete _)⟩ theorem of_constant (f : X → Y) (h : ∀ x y, f x = f y) : IsLocallyConstant f := - (iff_eventually_eq f).2 fun _ => eventually_of_forall fun _ => h _ _ + (iff_eventually_eq f).2 fun _ => Eventually.of_forall fun _ => h _ _ protected theorem const (y : Y) : IsLocallyConstant (Function.const X y) := of_constant _ fun _ _ => rfl diff --git a/Mathlib/Topology/Maps/Proper/Basic.lean b/Mathlib/Topology/Maps/Proper/Basic.lean index d2534f6fc6eac..4fa716956cb8b 100644 --- a/Mathlib/Topology/Maps/Proper/Basic.lean +++ b/Mathlib/Topology/Maps/Proper/Basic.lean @@ -441,7 +441,7 @@ theorem isProperMap_iff_isClosedMap_filter {X : Type u} {Y : Type v} [Topologica -- `𝒰`. Furthermore, each `(f, pure)(x) = (f × id)(x, pure x)` is clearly an element of -- the closed set `(f × id) '' F`, thus the limit `(y, 𝒰)` also belongs to that set. this.mem_of_tendsto (hy.prod_mk_nhds (Filter.tendsto_pure_self (𝒰 : Filter X))) - (eventually_of_forall fun x ↦ ⟨⟨x, pure x⟩, subset_closure rfl, rfl⟩) + (Eventually.of_forall fun x ↦ ⟨⟨x, pure x⟩, subset_closure rfl, rfl⟩) -- The above shows that `(y, 𝒰) = (f x, 𝒰)`, for some `x : X` such that `(x, 𝒰) ∈ F`. rcases this with ⟨⟨x, _⟩, hx, ⟨_, _⟩⟩ -- We already know that `f x = y`, so to finish the proof we just have to check that `𝒰` tends diff --git a/Mathlib/Topology/Maps/Proper/UniversallyClosed.lean b/Mathlib/Topology/Maps/Proper/UniversallyClosed.lean index aee1d663c4d09..1ad05fe6e8fab 100644 --- a/Mathlib/Topology/Maps/Proper/UniversallyClosed.lean +++ b/Mathlib/Topology/Maps/Proper/UniversallyClosed.lean @@ -31,7 +31,7 @@ theorem isProperMap_iff_isClosedMap_ultrafilter {X : Type u} {Y : Type v} [Topol have := H.2 F isClosed_closure have : (y, 𝒰) ∈ Prod.map f id '' F := this.mem_of_tendsto (hy.prod_mk_nhds (Ultrafilter.tendsto_pure_self 𝒰)) - (eventually_of_forall fun x ↦ ⟨⟨x, pure x⟩, subset_closure rfl, rfl⟩) + (Eventually.of_forall fun x ↦ ⟨⟨x, pure x⟩, subset_closure rfl, rfl⟩) rcases this with ⟨⟨x, _⟩, hx, ⟨_, _⟩⟩ refine ⟨x, rfl, fun U hU ↦ Ultrafilter.compl_not_mem_iff.mp fun hUc ↦ ?_⟩ rw [mem_closure_iff_nhds] at hx diff --git a/Mathlib/Topology/MetricSpace/Lipschitz.lean b/Mathlib/Topology/MetricSpace/Lipschitz.lean index 68c14362debf9..75a374b12a72f 100644 --- a/Mathlib/Topology/MetricSpace/Lipschitz.lean +++ b/Mathlib/Topology/MetricSpace/Lipschitz.lean @@ -320,7 +320,7 @@ variable [PseudoMetricSpace α] [PseudoMetricSpace β] {f : α → β} theorem continuousAt_of_locally_lipschitz {x : α} {r : ℝ} (hr : 0 < r) (K : ℝ) (h : ∀ y, dist y x < r → dist (f y) (f x) ≤ K * dist y x) : ContinuousAt f x := by -- We use `h` to squeeze `dist (f y) (f x)` between `0` and `K * dist y x` - refine tendsto_iff_dist_tendsto_zero.2 (squeeze_zero' (eventually_of_forall fun _ => dist_nonneg) + refine tendsto_iff_dist_tendsto_zero.2 (squeeze_zero' (Eventually.of_forall fun _ => dist_nonneg) (mem_of_superset (ball_mem_nhds _ hr) h) ?_) -- Then show that `K * dist y x` tends to zero as `y → x` refine (continuous_const.mul (continuous_id.dist continuous_const)).tendsto' _ _ ?_ diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean b/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean index 7bde82bf0007b..2a58189df2473 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Lemmas.lean @@ -32,7 +32,7 @@ lemma squeeze_zero' {α} {f g : α → ℝ} {t₀ : Filter α} (hf : ∀ᶠ t in and `tendsto_of_tendsto_of_tendsto_of_le_of_le'` for the general case. -/ lemma squeeze_zero {α} {f g : α → ℝ} {t₀ : Filter α} (hf : ∀ t, 0 ≤ f t) (hft : ∀ t, f t ≤ g t) (g0 : Tendsto g t₀ (𝓝 0)) : Tendsto f t₀ (𝓝 0) := - squeeze_zero' (eventually_of_forall hf) (eventually_of_forall hft) g0 + squeeze_zero' (Eventually.of_forall hf) (Eventually.of_forall hft) g0 /-- If `u` is a neighborhood of `x`, then for small enough `r`, the closed ball `Metric.closedBall x r` is contained in `u`. -/ diff --git a/Mathlib/Topology/MetricSpace/Sequences.lean b/Mathlib/Topology/MetricSpace/Sequences.lean index 2f1d5f1bb5afb..e407168843967 100644 --- a/Mathlib/Topology/MetricSpace/Sequences.lean +++ b/Mathlib/Topology/MetricSpace/Sequences.lean @@ -39,4 +39,4 @@ theorem tendsto_subseq_of_frequently_bounded (hs : IsBounded s) {x : ℕ → X} every bounded sequence has a converging subsequence. -/ theorem tendsto_subseq_of_bounded (hs : IsBounded s) {x : ℕ → X} (hx : ∀ n, x n ∈ s) : ∃ a ∈ closure s, ∃ φ : ℕ → ℕ, StrictMono φ ∧ Tendsto (x ∘ φ) atTop (𝓝 a) := - tendsto_subseq_of_frequently_bounded hs <| frequently_of_forall hx + tendsto_subseq_of_frequently_bounded hs <| Frequently.of_forall hx diff --git a/Mathlib/Topology/Order/Basic.lean b/Mathlib/Topology/Order/Basic.lean index e7570fcfdd2cf..3e2cc3f13ebe9 100644 --- a/Mathlib/Topology/Order/Basic.lean +++ b/Mathlib/Topology/Order/Basic.lean @@ -144,8 +144,8 @@ hold everywhere. -/ theorem tendsto_of_tendsto_of_tendsto_of_le_of_le [OrderTopology α] {f g h : β → α} {b : Filter β} {a : α} (hg : Tendsto g b (𝓝 a)) (hh : Tendsto h b (𝓝 a)) (hgf : g ≤ f) (hfh : f ≤ h) : Tendsto f b (𝓝 a) := - tendsto_of_tendsto_of_tendsto_of_le_of_le' hg hh (eventually_of_forall hgf) - (eventually_of_forall hfh) + tendsto_of_tendsto_of_tendsto_of_le_of_le' hg hh (Eventually.of_forall hgf) + (Eventually.of_forall hfh) theorem nhds_order_unbounded [OrderTopology α] {a : α} (hu : ∃ u, a < u) (hl : ∃ l, l < a) : 𝓝 a = ⨅ (l) (_ : l < a) (u) (_ : a < u), 𝓟 (Ioo l u) := by @@ -320,11 +320,11 @@ theorem tendsto_nhds_bot_mono [TopologicalSpace β] [Preorder β] [OrderBot β] theorem tendsto_nhds_top_mono' [TopologicalSpace β] [Preorder β] [OrderTop β] [OrderTopology β] {l : Filter α} {f g : α → β} (hf : Tendsto f l (𝓝 ⊤)) (hg : f ≤ g) : Tendsto g l (𝓝 ⊤) := - tendsto_nhds_top_mono hf (eventually_of_forall hg) + tendsto_nhds_top_mono hf (Eventually.of_forall hg) theorem tendsto_nhds_bot_mono' [TopologicalSpace β] [Preorder β] [OrderBot β] [OrderTopology β] {l : Filter α} {f g : α → β} (hf : Tendsto f l (𝓝 ⊥)) (hg : g ≤ f) : Tendsto g l (𝓝 ⊥) := - tendsto_nhds_bot_mono hf (eventually_of_forall hg) + tendsto_nhds_bot_mono hf (Eventually.of_forall hg) section LinearOrder diff --git a/Mathlib/Topology/Order/MonotoneConvergence.lean b/Mathlib/Topology/Order/MonotoneConvergence.lean index 320b7751381fa..bad7ee58b291e 100644 --- a/Mathlib/Topology/Order/MonotoneConvergence.lean +++ b/Mathlib/Topology/Order/MonotoneConvergence.lean @@ -67,7 +67,7 @@ instance (priority := 100) LinearOrder.supConvergenceClass [TopologicalSpace α] · rcases ha.exists_between hb with ⟨c, hcs, bc, bca⟩ lift c to s using hcs exact (eventually_ge_atTop c).mono fun x hx => bc.trans_le hx - · exact eventually_of_forall fun x => (ha.1 x.2).trans_lt hb + · exact Eventually.of_forall fun x => (ha.1 x.2).trans_lt hb -- see Note [lower instance priority] instance (priority := 100) LinearOrder.infConvergenceClass [TopologicalSpace α] [LinearOrder α] diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index f4a952a37cc45..7a476daeaf513 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -129,7 +129,7 @@ theorem le_of_tendsto {x : Filter β} [NeBot x] (lim : Tendsto f x (𝓝 a)) theorem le_of_tendsto' {x : Filter β} [NeBot x] (lim : Tendsto f x (𝓝 a)) (h : ∀ c, f c ≤ b) : a ≤ b := - le_of_tendsto lim (eventually_of_forall h) + le_of_tendsto lim (Eventually.of_forall h) @[simp] lemma upperBounds_closure (s : Set α) : upperBounds (closure s : Set α) = upperBounds s := ext fun a ↦ by simp_rw [mem_upperBounds_iff_subset_Iic, isClosed_Iic.closure_subset_iff] @@ -352,7 +352,7 @@ theorem ge_of_tendsto {x : Filter β} [NeBot x] (lim : Tendsto f x (𝓝 a)) theorem ge_of_tendsto' {x : Filter β} [NeBot x] (lim : Tendsto f x (𝓝 a)) (h : ∀ c, b ≤ f c) : b ≤ a := - ge_of_tendsto lim (eventually_of_forall h) + ge_of_tendsto lim (Eventually.of_forall h) @[simp] lemma lowerBounds_closure (s : Set α) : lowerBounds (closure s : Set α) = lowerBounds s := ext fun a ↦ by simp_rw [mem_lowerBounds_iff_subset_Ici, isClosed_Ici.closure_subset_iff] @@ -580,7 +580,7 @@ alias tendsto_le_of_eventuallyLE := le_of_tendsto_of_tendsto theorem le_of_tendsto_of_tendsto' {f g : β → α} {b : Filter β} {a₁ a₂ : α} [NeBot b] (hf : Tendsto f b (𝓝 a₁)) (hg : Tendsto g b (𝓝 a₂)) (h : ∀ x, f x ≤ g x) : a₁ ≤ a₂ := - le_of_tendsto_of_tendsto hf hg (eventually_of_forall h) + le_of_tendsto_of_tendsto hf hg (Eventually.of_forall h) @[simp] theorem closure_le_eq [TopologicalSpace β] {f g : β → α} (hf : Continuous f) (hg : Continuous g) : diff --git a/Mathlib/Topology/RestrictGenTopology.lean b/Mathlib/Topology/RestrictGenTopology.lean index faa7f586f6d16..9d244532a38d5 100644 --- a/Mathlib/Topology/RestrictGenTopology.lean +++ b/Mathlib/Topology/RestrictGenTopology.lean @@ -83,7 +83,7 @@ lemma of_seq [SequentialSpace X] rcases isClosed_induced_iff.1 (ht _ (h hux)) with ⟨s, hsc, hst⟩ rw [Subtype.preimage_val_eq_preimage_val_iff, Set.ext_iff] at hst suffices x ∈ s by specialize hst x; simp_all - refine hsc.mem_of_tendsto hux <| eventually_of_forall fun k ↦ ?_ + refine hsc.mem_of_tendsto hux <| Eventually.of_forall fun k ↦ ?_ specialize hst (u k) simp_all diff --git a/Mathlib/Topology/Semicontinuous.lean b/Mathlib/Topology/Semicontinuous.lean index b48b959356b6c..28c129606e374 100644 --- a/Mathlib/Topology/Semicontinuous.lean +++ b/Mathlib/Topology/Semicontinuous.lean @@ -169,10 +169,10 @@ theorem LowerSemicontinuous.lowerSemicontinuousOn (h : LowerSemicontinuous f) (s theorem lowerSemicontinuousWithinAt_const : LowerSemicontinuousWithinAt (fun _x => z) s x := - fun _y hy => Filter.eventually_of_forall fun _x => hy + fun _y hy => Filter.Eventually.of_forall fun _x => hy theorem lowerSemicontinuousAt_const : LowerSemicontinuousAt (fun _x => z) x := fun _y hy => - Filter.eventually_of_forall fun _x => hy + Filter.Eventually.of_forall fun _x => hy theorem lowerSemicontinuousOn_const : LowerSemicontinuousOn (fun _x => z) s := fun _x _hx => lowerSemicontinuousWithinAt_const @@ -193,7 +193,7 @@ theorem IsOpen.lowerSemicontinuous_indicator (hs : IsOpen s) (hy : 0 ≤ y) : by_cases h : x ∈ s <;> simp [h] at hz · filter_upwards [hs.mem_nhds h] simp (config := { contextual := true }) [hz] - · refine Filter.eventually_of_forall fun x' => ?_ + · refine Filter.Eventually.of_forall fun x' => ?_ by_cases h' : x' ∈ s <;> simp [h', hz.trans_le hy, hz] theorem IsOpen.lowerSemicontinuousOn_indicator (hs : IsOpen s) (hy : 0 ≤ y) : @@ -212,7 +212,7 @@ theorem IsClosed.lowerSemicontinuous_indicator (hs : IsClosed s) (hy : y ≤ 0) LowerSemicontinuous (indicator s fun _x => y) := by intro x z hz by_cases h : x ∈ s <;> simp [h] at hz - · refine Filter.eventually_of_forall fun x' => ?_ + · refine Filter.Eventually.of_forall fun x' => ?_ by_cases h' : x' ∈ s <;> simp [h', hz, hz.trans_le hy] · filter_upwards [hs.isOpen_compl.mem_nhds h] simp (config := { contextual := true }) [hz] @@ -358,7 +358,7 @@ theorem ContinuousAt.comp_lowerSemicontinuousWithinAt {g : γ → δ} {f : α _ ≤ g (f a) := gmon (min_le_right _ _) · simp only [not_exists, not_lt] at h - exact Filter.eventually_of_forall fun a => hy.trans_le (gmon (h (f a))) + exact Filter.Eventually.of_forall fun a => hy.trans_le (gmon (h (f a))) theorem ContinuousAt.comp_lowerSemicontinuousAt {g : γ → δ} {f : α → γ} (hg : ContinuousAt g (f x)) (hf : LowerSemicontinuousAt f x) (gmon : Monotone g) : LowerSemicontinuousAt (g ∘ f) x := by @@ -473,7 +473,7 @@ theorem LowerSemicontinuousWithinAt.add' {f g : α → γ} (hf : LowerSemicontin y < f x + min (g z) (g x) := h this _ ≤ f z + g z := add_le_add (hx₁ (f z)) (min_le_left _ _) · simp only [not_exists, not_lt] at hx₁ hx₂ - apply Filter.eventually_of_forall + apply Filter.Eventually.of_forall intro z have : (f x, g x) ∈ u ×ˢ v := ⟨xu, xv⟩ calc @@ -625,7 +625,7 @@ theorem lowerSemicontinuousOn_biSup {p : ι → Prop} {f : ∀ i, p i → α → theorem lowerSemicontinuous_ciSup {f : ι → α → δ'} (bdd : ∀ x, BddAbove (range fun i => f i x)) (h : ∀ i, LowerSemicontinuous (f i)) : LowerSemicontinuous fun x' => ⨆ i, f i x' := fun x => - lowerSemicontinuousAt_ciSup (eventually_of_forall bdd) fun i => h i x + lowerSemicontinuousAt_ciSup (Eventually.of_forall bdd) fun i => h i x theorem lowerSemicontinuous_iSup {f : ι → α → δ} (h : ∀ i, LowerSemicontinuous (f i)) : LowerSemicontinuous fun x' => ⨆ i, f i x' := @@ -711,10 +711,10 @@ theorem UpperSemicontinuous.upperSemicontinuousOn (h : UpperSemicontinuous f) (s theorem upperSemicontinuousWithinAt_const : UpperSemicontinuousWithinAt (fun _x => z) s x := - fun _y hy => Filter.eventually_of_forall fun _x => hy + fun _y hy => Filter.Eventually.of_forall fun _x => hy theorem upperSemicontinuousAt_const : UpperSemicontinuousAt (fun _x => z) x := fun _y hy => - Filter.eventually_of_forall fun _x => hy + Filter.Eventually.of_forall fun _x => hy theorem upperSemicontinuousOn_const : UpperSemicontinuousOn (fun _x => z) s := fun _x _hx => upperSemicontinuousWithinAt_const @@ -1054,7 +1054,7 @@ theorem upperSemicontinuousOn_biInf {p : ι → Prop} {f : ∀ i, p i → α → theorem upperSemicontinuous_ciInf {f : ι → α → δ'} (bdd : ∀ x, BddBelow (range fun i => f i x)) (h : ∀ i, UpperSemicontinuous (f i)) : UpperSemicontinuous fun x' => ⨅ i, f i x' := fun x => - upperSemicontinuousAt_ciInf (eventually_of_forall bdd) fun i => h i x + upperSemicontinuousAt_ciInf (Eventually.of_forall bdd) fun i => h i x theorem upperSemicontinuous_iInf {f : ι → α → δ} (h : ∀ i, UpperSemicontinuous (f i)) : UpperSemicontinuous fun x' => ⨅ i, f i x' := fun x => upperSemicontinuousAt_iInf fun i => h i x @@ -1094,7 +1094,7 @@ theorem continuousWithinAt_iff_lower_upperSemicontinuousWithinAt {f : α → γ} apply hu exact ⟨Hl (f a), lfa⟩ · simp only [not_exists, not_lt] at Hu - apply Filter.eventually_of_forall + apply Filter.Eventually.of_forall intro a have : f a = f x := le_antisymm (Hu _) (Hl _) rw [this] diff --git a/Mathlib/Topology/Sequences.lean b/Mathlib/Topology/Sequences.lean index 4729f9ccd86bf..c3f13a33c72e3 100644 --- a/Mathlib/Topology/Sequences.lean +++ b/Mathlib/Topology/Sequences.lean @@ -94,7 +94,7 @@ theorem isSeqClosed_iff {s : Set X} : IsSeqClosed s ↔ seqClosure s = s := /-- A set is sequentially closed if it is closed. -/ protected theorem IsClosed.isSeqClosed {s : Set X} (hc : IsClosed s) : IsSeqClosed s := - fun _u _x hu hx => hc.mem_of_tendsto hx (eventually_of_forall hu) + fun _u _x hu hx => hc.mem_of_tendsto hx (Eventually.of_forall hu) theorem seqClosure_eq_closure [FrechetUrysohnSpace X] (s : Set X) : seqClosure s = closure s := seqClosure_subset_closure.antisymm <| FrechetUrysohnSpace.closure_subset_seqClosure s @@ -121,7 +121,7 @@ theorem tendsto_nhds_iff_seq_tendsto [FrechetUrysohnSpace X] {f : X → Y} {a : refine ⟨closure (f ⁻¹' s), ⟨mt ?_ hbs, isClosed_closure⟩, fun x => mt fun hx => subset_closure hx⟩ rw [← seqClosure_eq_closure] rintro ⟨u, hus, hu⟩ - exact hsc.mem_of_tendsto (h u hu) (eventually_of_forall hus) + exact hsc.mem_of_tendsto (h u hu) (Eventually.of_forall hus) /-- An alternative construction for `FrechetUrysohnSpace`: if sequential convergence implies convergence, then the space is a Fréchet-Urysohn space. -/ @@ -254,7 +254,7 @@ open FirstCountableTopology protected theorem IsCompact.isSeqCompact {s : Set X} (hs : IsCompact s) : IsSeqCompact s := fun _x x_in => - let ⟨a, a_in, ha⟩ := hs (tendsto_principal.mpr (eventually_of_forall x_in)) + let ⟨a, a_in, ha⟩ := hs (tendsto_principal.mpr (Eventually.of_forall x_in)) ⟨a, a_in, tendsto_subseq ha⟩ theorem IsCompact.tendsto_subseq' {s : Set X} {x : ℕ → X} (hs : IsCompact s) @@ -315,7 +315,7 @@ theorem IsSeqCompact.exists_tendsto_of_frequently_mem (hs : IsSeqCompact s) {u : theorem IsSeqCompact.exists_tendsto (hs : IsSeqCompact s) {u : ℕ → X} (hu : ∀ n, u n ∈ s) (huc : CauchySeq u) : ∃ x ∈ s, Tendsto u atTop (𝓝 x) := - hs.exists_tendsto_of_frequently_mem (frequently_of_forall hu) huc + hs.exists_tendsto_of_frequently_mem (Frequently.of_forall hu) huc /-- A sequentially compact set in a uniform space is totally bounded. -/ protected theorem IsSeqCompact.totallyBounded (h : IsSeqCompact s) : TotallyBounded s := by diff --git a/Mathlib/Topology/UniformSpace/Ascoli.lean b/Mathlib/Topology/UniformSpace/Ascoli.lean index 2be2900a38f6e..6ceb7a0371b6f 100644 --- a/Mathlib/Topology/UniformSpace/Ascoli.lean +++ b/Mathlib/Topology/UniformSpace/Ascoli.lean @@ -379,7 +379,7 @@ theorem EquicontinuousOn.isClosed_range_pi_of_uniformOnFun' mapClusterPt_iff_ultrafilter, range_comp, Subtype.coe_injective.surjective_comp_right.forall, ← restrict_eq, ← EquicontinuousOn.tendsto_uniformOnFun_iff_pi' 𝔖_compact F_eqcont] exact fun f ⟨u, _, hu⟩ ↦ mem_image_of_mem _ <| H.mem_of_tendsto hu <| - eventually_of_forall mem_range_self + Eventually.of_forall mem_range_self /-- Let `X` be a topological space, `𝔖` a covering of `X` by compact subsets, and `α` a uniform space. An equicontinuous subset of `X → α` is closed in the topology of uniform diff --git a/Mathlib/Topology/UniformSpace/Equicontinuity.lean b/Mathlib/Topology/UniformSpace/Equicontinuity.lean index d047acb2e03e0..98b21924662ae 100644 --- a/Mathlib/Topology/UniformSpace/Equicontinuity.lean +++ b/Mathlib/Topology/UniformSpace/Equicontinuity.lean @@ -209,12 +209,12 @@ lemma uniformEquicontinuous_restrict_iff (F : ι → β → α) {S : Set β} : @[simp] lemma equicontinuousAt_empty [h : IsEmpty ι] (F : ι → X → α) (x₀ : X) : EquicontinuousAt F x₀ := - fun _ _ ↦ eventually_of_forall (fun _ ↦ h.elim) + fun _ _ ↦ Eventually.of_forall (fun _ ↦ h.elim) @[simp] lemma equicontinuousWithinAt_empty [h : IsEmpty ι] (F : ι → X → α) (S : Set X) (x₀ : X) : EquicontinuousWithinAt F S x₀ := - fun _ _ ↦ eventually_of_forall (fun _ ↦ h.elim) + fun _ _ ↦ Eventually.of_forall (fun _ ↦ h.elim) @[simp] lemma equicontinuous_empty [IsEmpty ι] (F : ι → X → α) : @@ -229,12 +229,12 @@ lemma equicontinuousOn_empty [IsEmpty ι] (F : ι → X → α) (S : Set X) : @[simp] lemma uniformEquicontinuous_empty [h : IsEmpty ι] (F : ι → β → α) : UniformEquicontinuous F := - fun _ _ ↦ eventually_of_forall (fun _ ↦ h.elim) + fun _ _ ↦ Eventually.of_forall (fun _ ↦ h.elim) @[simp] lemma uniformEquicontinuousOn_empty [h : IsEmpty ι] (F : ι → β → α) (S : Set β) : UniformEquicontinuousOn F S := - fun _ _ ↦ eventually_of_forall (fun _ ↦ h.elim) + fun _ _ ↦ Eventually.of_forall (fun _ ↦ h.elim) /-! ### Finite index type @@ -880,13 +880,13 @@ theorem Filter.Tendsto.continuousAt_of_equicontinuousAt {l : Filter ι} [l.NeBot {f : X → α} {x₀ : X} (h₁ : Tendsto F l (𝓝 f)) (h₂ : EquicontinuousAt F x₀) : ContinuousAt f x₀ := (equicontinuousAt_iff_range.mp h₂).closure.continuousAt - ⟨f, mem_closure_of_tendsto h₁ <| eventually_of_forall mem_range_self⟩ + ⟨f, mem_closure_of_tendsto h₁ <| Eventually.of_forall mem_range_self⟩ theorem Filter.Tendsto.uniformContinuous_of_uniformEquicontinuous {l : Filter ι} [l.NeBot] {F : ι → β → α} {f : β → α} (h₁ : Tendsto F l (𝓝 f)) (h₂ : UniformEquicontinuous F) : UniformContinuous f := (uniformEquicontinuous_iff_range.mp h₂).closure.uniformContinuous - ⟨f, mem_closure_of_tendsto h₁ <| eventually_of_forall mem_range_self⟩ + ⟨f, mem_closure_of_tendsto h₁ <| Eventually.of_forall mem_range_self⟩ ``` Unfortunately, the proofs get painful when dealing with the relative case as one needs to change @@ -905,7 +905,7 @@ theorem Filter.Tendsto.continuousWithinAt_of_equicontinuousWithinAt {l : Filter rcases mem_uniformity_isClosed hV with ⟨W, hW, hWclosed, hWV⟩ filter_upwards [h₃ W hW, eventually_mem_nhdsWithin] with x hx hxS using hVU <| ball_mono hWV (f x₀) <| hWclosed.mem_of_tendsto (h₂.prod_mk_nhds (h₁ x hxS)) <| - eventually_of_forall hx + Eventually.of_forall hx /-- If `𝓕 : ι → X → α` tends to `f : X → α` *pointwise* along some nontrivial filter, and if the family `𝓕` is equicontinuous at some `x₀ : X`, then the limit is continuous at `x₀`. -/ @@ -940,7 +940,7 @@ theorem Filter.Tendsto.uniformContinuousOn_of_uniformEquicontinuousOn {l : Filte filter_upwards [h₂ V hV, mem_inf_of_right (mem_principal_self _)] rintro ⟨x, y⟩ hxy ⟨hxS, hyS⟩ exact hVU <| hVclosed.mem_of_tendsto ((h₁ x hxS).prod_mk_nhds (h₁ y hyS)) <| - eventually_of_forall hxy + Eventually.of_forall hxy /-- If `𝓕 : ι → β → α` tends to `f : β → α` *pointwise* along some nontrivial filter, and if the family `𝓕` is uniformly equicontinuous, then the limit is uniformly continuous. -/ From 9ab6a770bd46219a989b6ee12ec047137ddb1ef0 Mon Sep 17 00:00:00 2001 From: grunweg Date: Fri, 16 Aug 2024 16:22:00 +0000 Subject: [PATCH 339/359] chore(Logic): remove bare open Classical (#15836) --- Mathlib/Logic/Denumerable.lean | 21 +++++++++------------ Mathlib/Logic/Nontrivial/Basic.lean | 5 ++--- Mathlib/Logic/Nontrivial/Defs.lean | 4 +--- Mathlib/Logic/Small/Basic.lean | 2 -- Mathlib/Logic/Small/Defs.lean | 2 -- 5 files changed, 12 insertions(+), 22 deletions(-) diff --git a/Mathlib/Logic/Denumerable.lean b/Mathlib/Logic/Denumerable.lean index 4869749bcd209..94274fd7d2531 100644 --- a/Mathlib/Logic/Denumerable.lean +++ b/Mathlib/Logic/Denumerable.lean @@ -184,22 +184,19 @@ open Function Encodable /-! ### Subsets of `ℕ` -/ - variable {s : Set ℕ} [Infinite s] section Classical -open scoped Classical - -theorem exists_succ (x : s) : ∃ n, (x : ℕ) + n + 1 ∈ s := - _root_.by_contradiction fun h => - have : ∀ (a : ℕ) (_ : a ∈ s), a < x + 1 := fun a ha => - lt_of_not_ge fun hax => h ⟨a - (x + 1), by rwa [add_right_comm, Nat.add_sub_cancel' hax]⟩ - Fintype.false - ⟨(((Multiset.range (succ x)).filter (· ∈ s)).pmap - (fun (y : ℕ) (hy : y ∈ s) => Subtype.mk y hy) - (by simp [-Multiset.range_succ])).toFinset, - by simpa [Subtype.ext_iff_val, Multiset.mem_filter, -Multiset.range_succ] ⟩ +theorem exists_succ (x : s) : ∃ n, (x : ℕ) + n + 1 ∈ s := by + by_contra h + have : ∀ (a : ℕ) (_ : a ∈ s), a < x + 1 := fun a ha => + lt_of_not_ge fun hax => h ⟨a - (x + 1), by rwa [add_right_comm, Nat.add_sub_cancel' hax]⟩ + classical + exact Fintype.false + ⟨(((Multiset.range (succ x)).filter (· ∈ s)).pmap + (fun (y : ℕ) (hy : y ∈ s) => Subtype.mk y hy) (by simp [-Multiset.range_succ])).toFinset, + by simpa [Subtype.ext_iff_val, Multiset.mem_filter, -Multiset.range_succ] ⟩ end Classical diff --git a/Mathlib/Logic/Nontrivial/Basic.lean b/Mathlib/Logic/Nontrivial/Basic.lean index fb49821633797..ee2fcbb4f7759 100644 --- a/Mathlib/Logic/Nontrivial/Basic.lean +++ b/Mathlib/Logic/Nontrivial/Basic.lean @@ -16,11 +16,8 @@ import Mathlib.Tactic.Attr.Register Results about `Nontrivial`. -/ - variable {α : Type*} {β : Type*} -open scoped Classical - -- `x` and `y` are explicit here, as they are often needed to guide typechecking of `h`. theorem nontrivial_of_lt [Preorder α] (x y : α) (h : x < y) : Nontrivial α := ⟨⟨x, y, ne_of_lt h⟩⟩ @@ -36,6 +33,7 @@ theorem Subtype.nontrivial_iff_exists_ne (p : α → Prop) (x : Subtype p) : Nontrivial (Subtype p) ↔ ∃ (y : α) (_ : p y), y ≠ x := by simp only [_root_.nontrivial_iff_exists_ne x, Subtype.exists, Ne, Subtype.ext_iff] +open Classical in /-- An inhabited type is either nontrivial, or has a unique element. -/ noncomputable def nontrivialPSumUnique (α : Type*) [Inhabited α] : Nontrivial α ⊕' Unique α := @@ -76,6 +74,7 @@ namespace Pi variable {I : Type*} {f : I → Type*} +open Classical in /-- A pi type is nontrivial if it's nonempty everywhere and nontrivial somewhere. -/ theorem nontrivial_at (i' : I) [inst : ∀ i, Nonempty (f i)] [Nontrivial (f i')] : Nontrivial (∀ i : I, f i) := by diff --git a/Mathlib/Logic/Nontrivial/Defs.lean b/Mathlib/Logic/Nontrivial/Defs.lean index a0173e706b52d..51b62d4be17b3 100644 --- a/Mathlib/Logic/Nontrivial/Defs.lean +++ b/Mathlib/Logic/Nontrivial/Defs.lean @@ -19,11 +19,8 @@ We introduce a typeclass `Nontrivial` formalizing this property. Basic results about nontrivial types are in `Mathlib.Logic.Nontrivial.Basic`. -/ - variable {α : Type*} {β : Type*} -open scoped Classical - /-- Predicate typeclass for expressing that a type is not reduced to a single element. In rings, this is equivalent to `0 ≠ 1`. In vector spaces, this is equivalent to positive dimension. -/ class Nontrivial (α : Type*) : Prop where @@ -44,6 +41,7 @@ protected theorem Decidable.exists_ne [Nontrivial α] [DecidableEq α] (x : α) exact ⟨y', h.symm⟩ · exact ⟨y, Ne.symm hx⟩ +open Classical in theorem exists_ne [Nontrivial α] (x : α) : ∃ y, y ≠ x := Decidable.exists_ne x -- `x` and `y` are explicit here, as they are often needed to guide typechecking of `h`. diff --git a/Mathlib/Logic/Small/Basic.lean b/Mathlib/Logic/Small/Basic.lean index db322f081308c..902d79ca3fdf2 100644 --- a/Mathlib/Logic/Small/Basic.lean +++ b/Mathlib/Logic/Small/Basic.lean @@ -16,8 +16,6 @@ universe u w v v' section -open scoped Classical - instance small_subtype (α : Type v) [Small.{w} α] (P : α → Prop) : Small.{w} { x // P x } := small_map (equivShrink α).subtypeEquivOfSubtype' diff --git a/Mathlib/Logic/Small/Defs.lean b/Mathlib/Logic/Small/Defs.lean index 6a6f1cf0785a9..c93c2ba60d7d1 100644 --- a/Mathlib/Logic/Small/Defs.lean +++ b/Mathlib/Logic/Small/Defs.lean @@ -88,8 +88,6 @@ theorem small_type : Small.{max (u + 1) v} (Type u) := section -open scoped Classical - theorem small_congr {α : Type*} {β : Type*} (e : α ≃ β) : Small.{w} α ↔ Small.{w} β := ⟨fun h => @small_map _ _ h e.symm, fun h => @small_map _ _ h e⟩ From ab340af9ad48cbdb7c8c329c867a920c11c88862 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Fri, 16 Aug 2024 16:22:01 +0000 Subject: [PATCH 340/359] chore: turn `induction'` into `induction` (#15889) This PR turns `induction'` (soon to be deprecated) into `induction` and golfs a few proofs. --- Mathlib/Algebra/Ring/Hom/Defs.lean | 2 +- .../DoldKan/NReflectsIso.lean | 8 ++-- Mathlib/Analysis/Analytic/Meromorphic.lean | 12 +++--- Mathlib/Analysis/Calculus/Taylor.lean | 39 +++++++++-------- .../Convex/SpecificFunctions/Deriv.lean | 23 +++++----- Mathlib/Analysis/Convolution.lean | 14 ++++--- .../Analysis/Distribution/SchwartzSpace.lean | 26 ++++++------ Mathlib/Analysis/Hofer.lean | 42 +++++++++---------- Mathlib/Analysis/PSeries.lean | 25 +++++------ .../SpecialFunctions/Gamma/Basic.lean | 15 ++++--- 10 files changed, 108 insertions(+), 98 deletions(-) diff --git a/Mathlib/Algebra/Ring/Hom/Defs.lean b/Mathlib/Algebra/Ring/Hom/Defs.lean index 8b6d0d07f6d18..e24b81bf70f81 100644 --- a/Mathlib/Algebra/Ring/Hom/Defs.lean +++ b/Mathlib/Algebra/Ring/Hom/Defs.lean @@ -598,7 +598,7 @@ instance instMonoid : Monoid (α →+* α) where mul_one := comp_id one_mul := id_comp mul_assoc f g h := comp_assoc _ _ _ - npow n f := (npowRec n f).copy f^[n] $ by induction' n <;> simp [npowRec, *] + npow n f := (npowRec n f).copy f^[n] $ by induction n <;> simp [npowRec, *] npow_succ n f := DFunLike.coe_injective $ Function.iterate_succ _ _ @[simp, norm_cast] lemma coe_pow (f : α →+* α) (n : ℕ) : ⇑(f ^ n) = f^[n] := rfl diff --git a/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean b/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean index 45f68a713e024..ff1f2cf7b3dbd 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean @@ -48,15 +48,17 @@ instance : (N₁ : SimplicialObject C ⥤ Karoubi (ChainComplex C ℕ)).Reflects AlternatingFaceMapComplex.map_f, N₁_obj_p, Karoubi.id_f, assoc] at h₁ h₂ h₃ -- we have to construct an inverse to f in degree n, by induction on n intro n - induction' n with n hn + induction n with -- degree 0 - · use (inv (N₁.map f)).f.f 0 + | zero => + use (inv (N₁.map f)).f.f 0 have h₁₀ := h₁ 0 have h₂₀ := h₂ 0 dsimp at h₁₀ h₂₀ simp only [id_comp, comp_id] at h₁₀ h₂₀ tauto - · haveI := hn + | succ n hn => + haveI := hn use φ { a := PInfty.f (n + 1) ≫ (inv (N₁.map f)).f.f (n + 1) b := fun i => inv (f.app (op [n])) ≫ X.σ i } simp only [MorphComponents.id, ← id_φ, ← preComp_φ, preComp, ← postComp_φ, postComp, diff --git a/Mathlib/Analysis/Analytic/Meromorphic.lean b/Mathlib/Analysis/Analytic/Meromorphic.lean index a447d8a51c6d6..12ec5f6379e2c 100644 --- a/Mathlib/Analysis/Analytic/Meromorphic.lean +++ b/Mathlib/Analysis/Analytic/Meromorphic.lean @@ -133,14 +133,14 @@ lemma div {f g : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (hg : Meromo (div_eq_mul_inv f g).symm ▸ (hf.mul hg.inv) lemma pow {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (n : ℕ) : MeromorphicAt (f ^ n) x := by - induction' n with m hm - · simpa only [pow_zero] using MeromorphicAt.const 1 x - · simpa only [pow_succ] using hm.mul hf + induction n with + | zero => simpa only [pow_zero] using MeromorphicAt.const 1 x + | succ m hm => simpa only [pow_succ] using hm.mul hf lemma zpow {f : 𝕜 → 𝕜} {x : 𝕜} (hf : MeromorphicAt f x) (n : ℤ) : MeromorphicAt (f ^ n) x := by - induction' n with m m - · simpa only [Int.ofNat_eq_coe, zpow_natCast] using hf.pow m - · simpa only [zpow_negSucc, inv_iff] using hf.pow (m + 1) + induction n with + | ofNat m => simpa only [Int.ofNat_eq_coe, zpow_natCast] using hf.pow m + | negSucc m => simpa only [zpow_negSucc, inv_iff] using hf.pow (m + 1) /-- The order of vanishing of a meromorphic function, as an element of `ℤ ∪ ∞` (to include the case of functions identically 0 near `x`). -/ diff --git a/Mathlib/Analysis/Calculus/Taylor.lean b/Mathlib/Analysis/Calculus/Taylor.lean index d48e71fcee876..9f73b88cf066d 100644 --- a/Mathlib/Analysis/Calculus/Taylor.lean +++ b/Mathlib/Analysis/Calculus/Taylor.lean @@ -98,17 +98,18 @@ theorem taylor_within_zero_eval (f : ℝ → E) (s : Set ℝ) (x₀ x : ℝ) : @[simp] theorem taylorWithinEval_self (f : ℝ → E) (n : ℕ) (s : Set ℝ) (x₀ : ℝ) : taylorWithinEval f n s x₀ x₀ = f x₀ := by - induction' n with k hk - · exact taylor_within_zero_eval _ _ _ _ - simp [hk] + induction n with + | zero => exact taylor_within_zero_eval _ _ _ _ + | succ k hk => simp [hk] theorem taylor_within_apply (f : ℝ → E) (n : ℕ) (s : Set ℝ) (x₀ x : ℝ) : taylorWithinEval f n s x₀ x = ∑ k ∈ Finset.range (n + 1), ((k ! : ℝ)⁻¹ * (x - x₀) ^ k) • iteratedDerivWithin k f s x₀ := by - induction' n with k hk - · simp - rw [taylorWithinEval_succ, Finset.sum_range_succ, hk] - simp [Nat.factorial] + induction n with + | zero => simp + | succ k hk => + rw [taylorWithinEval_succ, Finset.sum_range_succ, hk] + simp [Nat.factorial] /-- If `f` is `n` times continuous differentiable on a set `s`, then the Taylor polynomial `taylorWithinEval f n s x₀ x` is continuous in `x₀`. -/ @@ -166,21 +167,23 @@ theorem hasDerivWithinAt_taylorWithinEval {f : ℝ → E} {x y : ℝ} {n : ℕ} (hf' : DifferentiableWithinAt ℝ (iteratedDerivWithin n f s) s y) : HasDerivWithinAt (fun t => taylorWithinEval f n s t x) (((n ! : ℝ)⁻¹ * (x - y) ^ n) • iteratedDerivWithin (n + 1) f s y) s' y := by - induction' n with k hk - · simp only [taylor_within_zero_eval, Nat.factorial_zero, Nat.cast_one, inv_one, pow_zero, + induction n with + | zero => + simp only [taylor_within_zero_eval, Nat.factorial_zero, Nat.cast_one, inv_one, pow_zero, mul_one, zero_add, one_smul] simp only [iteratedDerivWithin_zero] at hf' rw [iteratedDerivWithin_one (hs_unique _ (h hy))] exact hf'.hasDerivWithinAt.mono h - simp_rw [Nat.add_succ, taylorWithinEval_succ] - 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 - specialize hk hf.of_succ ((hdiff y hy).mono_of_mem hs') - convert hk.add (hasDerivWithinAt_taylor_coeff_within hs'_unique - (nhdsWithin_mono _ h self_mem_nhdsWithin) hf') using 1 - exact (add_sub_cancel _ _).symm + | succ k hk => + simp_rw [Nat.add_succ, taylorWithinEval_succ] + 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 + specialize hk hf.of_succ ((hdiff y hy).mono_of_mem hs') + convert hk.add (hasDerivWithinAt_taylor_coeff_within hs'_unique + (nhdsWithin_mono _ h self_mem_nhdsWithin) hf') using 1 + exact (add_sub_cancel _ _).symm /-- Calculate the derivative of the Taylor polynomial with respect to `x₀`. diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean index 0d78d9d0e1048..9d5442089d921 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean @@ -67,17 +67,18 @@ theorem Finset.prod_nonneg_of_card_nonpos_even {α β : Type*} [LinearOrderedCom theorem int_prod_range_nonneg (m : ℤ) (n : ℕ) (hn : Even n) : 0 ≤ ∏ k ∈ Finset.range n, (m - k) := by rcases hn with ⟨n, rfl⟩ - induction' n with n ihn - · simp - rw [← two_mul] at ihn - rw [← two_mul, mul_add, mul_one, ← one_add_one_eq_two, ← add_assoc, - Finset.prod_range_succ, Finset.prod_range_succ, mul_assoc] - refine mul_nonneg ihn ?_; generalize (1 + 1) * n = k - rcases le_or_lt m k with hmk | hmk - · have : m ≤ k + 1 := hmk.trans (lt_add_one (k : ℤ)).le - convert mul_nonneg_of_nonpos_of_nonpos (sub_nonpos_of_le hmk) _ - convert sub_nonpos_of_le this - · exact mul_nonneg (sub_nonneg_of_le hmk.le) (sub_nonneg_of_le hmk) + induction n with + | zero => simp + | succ n ihn => + rw [← two_mul] at ihn + rw [← two_mul, mul_add, mul_one, ← one_add_one_eq_two, ← add_assoc, + Finset.prod_range_succ, Finset.prod_range_succ, mul_assoc] + refine mul_nonneg ihn ?_; generalize (1 + 1) * n = k + rcases le_or_lt m k with hmk | hmk + · have : m ≤ k + 1 := hmk.trans (lt_add_one (k : ℤ)).le + convert mul_nonneg_of_nonpos_of_nonpos (sub_nonpos_of_le hmk) _ + convert sub_nonpos_of_le this + · exact mul_nonneg (sub_nonneg_of_le hmk.le) (sub_nonneg_of_le hmk) theorem int_prod_range_pos {m : ℤ} {n : ℕ} (hn : Even n) (hm : m ∉ Ico (0 : ℤ) n) : 0 < ∏ k ∈ Finset.range n, (m - k) := by diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 7beb76c9e195a..4b2168a084049 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1178,10 +1178,12 @@ theorem contDiffOn_convolution_right_with_param_aux {G : Type uP} {E' : Type uP} `n` (but for this we need the spaces at the different steps of the induction to live in the same universe, which is why we make the assumption in the lemma that all the relevant spaces come from the same universe). -/ - induction' n using ENat.nat_induction with n ih ih generalizing g E' F - · rw [contDiffOn_zero] at hg ⊢ + induction n using ENat.nat_induction generalizing g E' F with + | h0 => + rw [contDiffOn_zero] at hg ⊢ exact continuousOn_convolution_right_with_param L hk hgs hf hg - · let f' : P → G → P × G →L[𝕜] F := fun p a => + | hsuc n ih => + 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₀ := @@ -1205,9 +1207,9 @@ theorem contDiffOn_convolution_right_with_param_aux {G : Type uP} {E' : Type uP} exact hgs p y hp hy apply ih (L.precompR (P × G) : _) B convert hg.2 - · rw [contDiffOn_top] at hg ⊢ - intro n - exact ih n L hgs (hg n) + | htop ih => + rw [contDiffOn_top] at hg ⊢ + exact fun n ↦ ih n L hgs (hg n) /-- The convolution `f * g` is `C^n` when `f` is locally integrable and `g` is `C^n` and compactly supported. Version where `g` depends on an additional parameter in an open subset `s` of a diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 868adbc62f8a6..c8404661ec6dc 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -1002,20 +1002,22 @@ theorem iteratedPDeriv_succ_left {n : ℕ} (m : Fin (n + 1) → E) (f : 𝓢(E, theorem iteratedPDeriv_succ_right {n : ℕ} (m : Fin (n + 1) → E) (f : 𝓢(E, F)) : iteratedPDeriv 𝕜 m f = iteratedPDeriv 𝕜 (Fin.init m) (pderivCLM 𝕜 (m (Fin.last n)) f) := by - induction' n with n IH - · rw [iteratedPDeriv_zero, iteratedPDeriv_one] + induction n with + | zero => + rw [iteratedPDeriv_zero, iteratedPDeriv_one] rfl -- The proof is `∂^{n + 2} = ∂ ∂^{n + 1} = ∂ ∂^n ∂ = ∂^{n+1} ∂` - have hmzero : Fin.init m 0 = m 0 := by simp only [Fin.init_def, Fin.castSucc_zero] - have hmtail : Fin.tail m (Fin.last n) = m (Fin.last n.succ) := by - simp only [Fin.tail_def, Fin.succ_last] - calc - _ = pderivCLM 𝕜 (m 0) (iteratedPDeriv 𝕜 _ f) := iteratedPDeriv_succ_left _ _ _ - _ = pderivCLM 𝕜 (m 0) ((iteratedPDeriv 𝕜 _) ((pderivCLM 𝕜 _) f)) := by - congr 1 - exact IH _ - _ = _ := by - simp only [hmtail, iteratedPDeriv_succ_left, hmzero, Fin.tail_init_eq_init_tail] + | succ n IH => + have hmzero : Fin.init m 0 = m 0 := by simp only [Fin.init_def, Fin.castSucc_zero] + have hmtail : Fin.tail m (Fin.last n) = m (Fin.last n.succ) := by + simp only [Fin.tail_def, Fin.succ_last] + calc + _ = pderivCLM 𝕜 (m 0) (iteratedPDeriv 𝕜 _ f) := iteratedPDeriv_succ_left _ _ _ + _ = pderivCLM 𝕜 (m 0) ((iteratedPDeriv 𝕜 _) ((pderivCLM 𝕜 _) f)) := by + congr 1 + exact IH _ + _ = _ := by + simp only [hmtail, iteratedPDeriv_succ_left, hmzero, Fin.tail_init_eq_init_tail] theorem iteratedPDeriv_eq_iteratedFDeriv {n : ℕ} {m : Fin n → E} {f : 𝓢(E, F)} {x : E} : iteratedPDeriv 𝕜 m f x = iteratedFDeriv ℝ n f x m := by diff --git a/Mathlib/Analysis/Hofer.lean b/Mathlib/Analysis/Hofer.lean index 6ed5d0e71bb32..b008d74901b49 100644 --- a/Mathlib/Analysis/Hofer.lean +++ b/Mathlib/Analysis/Hofer.lean @@ -37,7 +37,6 @@ theorem hofer {X : Type*} [MetricSpace X] [CompleteSpace X] (x : X) (ε : ℝ) ( push_neg at H have := H (ε / 2 ^ k) (by positivity) x' (by simp [ε_pos.le, one_le_two]) simpa [reformulation] using this - clear reformulation haveI : Nonempty X := ⟨x⟩ choose! F hF using H -- Use the axiom of choice @@ -48,32 +47,29 @@ theorem hofer {X : Type*} [MetricSpace X] [CompleteSpace X] (x : X) (ε : ℝ) ( ∀ n, d (u n) x ≤ 2 * ε ∧ 2 ^ n * ϕ x ≤ ϕ (u n) → d (u n) (u <| n + 1) ≤ ε / 2 ^ n ∧ 2 * ϕ (u n) < ϕ (u <| n + 1) := by - intro n - exact hF n (u n) - clear hF + exact fun n ↦ hF n (u n) -- Key properties of u, to be proven by induction have key : ∀ n, d (u n) (u (n + 1)) ≤ ε / 2 ^ n ∧ 2 * ϕ (u n) < ϕ (u (n + 1)) := by intro n - induction' n using Nat.case_strong_induction_on with n IH - · simpa [u, ε_pos.le] using hu 0 - have A : d (u (n + 1)) x ≤ 2 * ε := by - rw [dist_comm] - let r := range (n + 1) -- range (n+1) = {0, ..., n} - calc - d (u 0) (u (n + 1)) ≤ ∑ i ∈ r, d (u i) (u <| i + 1) := dist_le_range_sum_dist u (n + 1) - _ ≤ ∑ i ∈ r, ε / 2 ^ i := - (sum_le_sum fun i i_in => (IH i <| Nat.lt_succ_iff.mp <| Finset.mem_range.mp i_in).1) - _ = (∑ i ∈ r, (1 / 2 : ℝ) ^ i) * ε := by - rw [Finset.sum_mul] - congr with i - field_simp - _ ≤ 2 * ε := by gcongr; apply sum_geometric_two_le - have B : 2 ^ (n + 1) * ϕ x ≤ ϕ (u (n + 1)) := by - refine @geom_le (ϕ ∘ u) _ zero_le_two (n + 1) fun m hm => ?_ - exact (IH _ <| Nat.lt_add_one_iff.1 hm).2.le - exact hu (n + 1) ⟨A, B⟩ + induction n using Nat.case_strong_induction_on with + | hz => simpa [u, ε_pos.le] using hu 0 + | hi n IH => + have A : d (u (n + 1)) x ≤ 2 * ε := by + rw [dist_comm] + let r := range (n + 1) -- range (n+1) = {0, ..., n} + calc + d (u 0) (u (n + 1)) ≤ ∑ i ∈ r, d (u i) (u <| i + 1) := dist_le_range_sum_dist u (n + 1) + _ ≤ ∑ i ∈ r, ε / 2 ^ i := + (sum_le_sum fun i i_in => (IH i <| Nat.lt_succ_iff.mp <| Finset.mem_range.mp i_in).1) + _ = (∑ i ∈ r, (1 / 2 : ℝ) ^ i) * ε := by + rw [Finset.sum_mul] + field_simp + _ ≤ 2 * ε := by gcongr; apply sum_geometric_two_le + have B : 2 ^ (n + 1) * ϕ x ≤ ϕ (u (n + 1)) := by + refine @geom_le (ϕ ∘ u) _ zero_le_two (n + 1) fun m hm => ?_ + exact (IH _ <| Nat.lt_add_one_iff.1 hm).2.le + exact hu (n + 1) ⟨A, B⟩ cases' forall_and.mp key with key₁ key₂ - clear hu key -- Hence u is Cauchy have cauchy_u : CauchySeq u := by refine cauchySeq_of_le_geometric _ ε one_half_lt_one fun n => ?_ diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 12be1fc3643b6..eca94e33eea6e 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -81,18 +81,19 @@ theorem le_sum_condensed (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f m) theorem sum_schlomilch_le' (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f m) (h_pos : ∀ n, 0 < u n) (hu : Monotone u) (n : ℕ) : (∑ k ∈ range n, (u (k + 1) - u k) • f (u (k + 1))) ≤ ∑ k ∈ Ico (u 0 + 1) (u n + 1), f k := by - induction' n with n ihn - · simp - suffices (u (n + 1) - u n) • f (u (n + 1)) ≤ ∑ k ∈ Ico (u n + 1) (u (n + 1) + 1), f k by - rw [sum_range_succ, ← sum_Ico_consecutive] - exacts [add_le_add ihn this, - (add_le_add_right (hu n.zero_le) _ : u 0 + 1 ≤ u n + 1), - add_le_add_right (hu n.le_succ) _] - have : ∀ k ∈ Ico (u n + 1) (u (n + 1) + 1), f (u (n + 1)) ≤ f k := fun k hk => - hf (Nat.lt_of_le_of_lt (Nat.succ_le_of_lt (h_pos n)) <| (Nat.lt_succ_of_le le_rfl).trans_le - (mem_Ico.mp hk).1) (Nat.le_of_lt_succ <| (mem_Ico.mp hk).2) - convert sum_le_sum this - simp [pow_succ, mul_two] + induction n with + | zero => simp + | succ n ihn => + suffices (u (n + 1) - u n) • f (u (n + 1)) ≤ ∑ k ∈ Ico (u n + 1) (u (n + 1) + 1), f k by + rw [sum_range_succ, ← sum_Ico_consecutive] + exacts [add_le_add ihn this, + (add_le_add_right (hu n.zero_le) _ : u 0 + 1 ≤ u n + 1), + add_le_add_right (hu n.le_succ) _] + have : ∀ k ∈ Ico (u n + 1) (u (n + 1) + 1), f (u (n + 1)) ≤ f k := fun k hk => + hf (Nat.lt_of_le_of_lt (Nat.succ_le_of_lt (h_pos n)) <| (Nat.lt_succ_of_le le_rfl).trans_le + (mem_Ico.mp hk).1) (Nat.le_of_lt_succ <| (mem_Ico.mp hk).2) + convert sum_le_sum this + simp [pow_succ, mul_two] theorem sum_condensed_le' (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f m) (n : ℕ) : (∑ k ∈ range n, 2 ^ k • f (2 ^ (k + 1))) ≤ ∑ k ∈ Ico 2 (2 ^ n + 1), f k := by diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index 3c575bbb3d974..d07803c56bc15 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -343,9 +343,10 @@ theorem Gamma_zero : Gamma 0 = 0 := by /-- At `-n` for `n ∈ ℕ`, the Gamma function is undefined; by convention we assign it the value 0. -/ theorem Gamma_neg_nat_eq_zero (n : ℕ) : Gamma (-n) = 0 := by - induction' n with n IH - · rw [Nat.cast_zero, neg_zero, Gamma_zero] - · have A : -(n.succ : ℂ) ≠ 0 := by + induction n with + | zero => rw [Nat.cast_zero, neg_zero, Gamma_zero] + | succ n IH => + have A : -(n.succ : ℂ) ≠ 0 := by rw [neg_ne_zero, Nat.cast_ne_zero] apply Nat.succ_ne_zero have : -(n : ℂ) = -↑n.succ + 1 := by simp @@ -581,11 +582,13 @@ theorem Gamma_ne_zero {s : ℝ} (hs : ∀ m : ℕ, s ≠ -m) : Gamma s ≠ 0 := rw [neg_lt, Nat.cast_add, Nat.cast_one] exact Nat.lt_floor_add_one _ intro n - induction' n with _ n_ih generalizing s - · intro hs + induction n generalizing s with + | zero => + intro hs refine (Gamma_pos_of_pos ?_).ne' rwa [Nat.cast_zero, neg_zero] at hs - · intro hs' + | succ _ n_ih => + intro hs' have : Gamma (s + 1) ≠ 0 := by apply n_ih · intro m From d1f6634fa2a5edd9fc4651dd96b2fedb8f7df864 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 16 Aug 2024 17:25:41 +0000 Subject: [PATCH 341/359] chore(.github/workflows/nightly_detect_failure.yml): fix the sha hash (#15888) The sha hash was only printed to the github env if the git tag "nightly-testing-${version}" did not yet exist. We move this out of the if clause, so that it is always printed. --- .github/workflows/nightly_detect_failure.yml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/.github/workflows/nightly_detect_failure.yml b/.github/workflows/nightly_detect_failure.yml index 2b430940e7fa4..e1cfd9147086b 100644 --- a/.github/workflows/nightly_detect_failure.yml +++ b/.github/workflows/nightly_detect_failure.yml @@ -50,9 +50,10 @@ jobs: git tag "nightly-testing-${version}" git push origin "nightly-testing-${version}" hash="$(git rev-parse "nightly-testing-${version}")" - printf 'SHA=%s\n' "${hash}" >> "${GITHUB_ENV}" curl -X POST "http://speed.lean-fro.org/mathlib4/api/queue/commit/e7b27246-a3e6-496a-b552-ff4b45c7236e/$hash" -u "admin:${{ secrets.SPEED }}" fi + hash="$(git rev-parse "nightly-testing-${version}")" + printf 'SHA=%s\n' "${hash}" >> "${GITHUB_ENV}" else echo "Error: The file lean-toolchain does not contain the expected pattern." exit 1 From 249c5361016c6265d2c82d9e24fa4702b713f11d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Aug 2024 17:37:55 +0000 Subject: [PATCH 342/359] feat: Doubling and difference constants of two finsets (#15882) Define the doubling and difference constants of two finsets in a group. The doubling and difference constants being small are one way to measure how close a set is from being a subgroup. This makes it a central concept in additive combinatorics and related fields. From LeanAPAP --- Mathlib.lean | 1 + .../Combinatorics/Additive/DoublingConst.lean | 215 ++++++++++++++++++ 2 files changed, 216 insertions(+) create mode 100644 Mathlib/Combinatorics/Additive/DoublingConst.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2e127b5d5f086..d1da137ddf3a0 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1864,6 +1864,7 @@ import Mathlib.Combinatorics.Additive.AP.Three.Defs import Mathlib.Combinatorics.Additive.Corner.Defs import Mathlib.Combinatorics.Additive.Corner.Roth import Mathlib.Combinatorics.Additive.Dissociation +import Mathlib.Combinatorics.Additive.DoublingConst import Mathlib.Combinatorics.Additive.ETransform import Mathlib.Combinatorics.Additive.Energy import Mathlib.Combinatorics.Additive.ErdosGinzburgZiv diff --git a/Mathlib/Combinatorics/Additive/DoublingConst.lean b/Mathlib/Combinatorics/Additive/DoublingConst.lean new file mode 100644 index 0000000000000..f9f65e418d246 --- /dev/null +++ b/Mathlib/Combinatorics/Additive/DoublingConst.lean @@ -0,0 +1,215 @@ +/- +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.Combinatorics.Additive.PluenneckeRuzsa +import Mathlib.Data.Finset.Density +import Mathlib.Tactic.Positivity.Basic +import Mathlib.Tactic.Positivity.Finset +import Mathlib.Tactic.Ring + +/-! +# Doubling and difference constants + +This file defines the doubling and difference constants of two finsets in a group. +-/ + +open Finset +open scoped Pointwise + +namespace Finset +section Group +variable {G G' : Type*} [Group G] [AddGroup G'] [DecidableEq G] [DecidableEq G'] {A B : Finset G} + {a : G} + +/-- The doubling constant `σₘ[A, B]` of two finsets `A` and `B` in a group is `|A * B| / |A|`. + +The notation `σₘ[A, B]` is available in scope `Combinatorics.Additive`. -/ +@[to_additive +"The doubling constant `σ[A, B]` of two finsets `A` and `B` in a group is `|A + B| / |A|`. + +The notation `σ[A, B]` is available in scope `Combinatorics.Additive`."] +def mulConst (A B : Finset G) : ℚ≥0 := (A * B).card / A.card + +/-- The difference constant `δₘ[A, B]` of two finsets `A` and `B` in a group is `|A / B| / |A|`. + +The notation `δₘ[A, B]` is available in scope `Combinatorics.Additive`. -/ +@[to_additive +"The difference constant `σ[A, B]` of two finsets `A` and `B` in a group is `|A - B| / |A|`. + +The notation `δ[A, B]` is available in scope `Combinatorics.Additive`."] +def divConst (A B : Finset G) : ℚ≥0 := (A / B).card / A.card + +/-- The doubling constant `σₘ[A, B]` of two finsets `A` and `B` in a group is `|A * B| / |A|`. -/ +scoped[Combinatorics.Additive] notation3:max "σₘ[" A ", " B "]" => Finset.mulConst A B + +/-- The doubling constant `σₘ[A]` of a finset `A` in a group is `|A * A| / |A|`. -/ +scoped[Combinatorics.Additive] notation3:max "σₘ[" A "]" => Finset.mulConst A A + +/-- The doubling constant `σ[A, B]` of two finsets `A` and `B` in a group is `|A + B| / |A|`. -/ +scoped[Combinatorics.Additive] notation3:max "σ[" A ", " B "]" => Finset.addConst A B + +/-- The doubling constant `σ[A]` of a finset `A` in a group is `|A + A| / |A|`. -/ +scoped[Combinatorics.Additive] notation3:max "σ[" A "]" => Finset.addConst A A + +/-- The difference constant `σₘ[A, B]` of two finsets `A` and `B` in a group is `|A / B| / |A|`. -/ +scoped[Combinatorics.Additive] notation3:max "δₘ[" A ", " B "]" => Finset.divConst A B + +/-- The difference constant `σₘ[A]` of a finset `A` in a group is `|A / A| / |A|`. -/ +scoped[Combinatorics.Additive] notation3:max "δₘ[" A "]" => Finset.divConst A A + +/-- The difference constant `σ[A, B]` of two finsets `A` and `B` in a group is `|A - B| / |A|`. -/ +scoped[Combinatorics.Additive] notation3:max "δ[" A ", " B "]" => Finset.subConst A B + +/-- The difference constant `σ[A]` of a finset `A` in a group is `|A - A| / |A|`. -/ +scoped[Combinatorics.Additive] notation3:max "δ[" A "]" => Finset.subConst A A + +open scoped Combinatorics.Additive + +@[to_additive (attr := simp) addConst_mul_card] +lemma mulConst_mul_card (A B : Finset G) : σₘ[A, B] * A.card = (A * B).card := by + obtain rfl | hA := A.eq_empty_or_nonempty + · simp + · exact div_mul_cancel₀ _ (by positivity) + +@[to_additive (attr := simp) subConst_mul_card] +lemma divConst_mul_card (A B : Finset G) : δₘ[A, B] * A.card = (A / B).card := by + obtain rfl | hA := A.eq_empty_or_nonempty + · simp + · exact div_mul_cancel₀ _ (by positivity) + +@[to_additive (attr := simp) card_mul_addConst] +lemma card_mul_mulConst (A B : Finset G) : A.card * σₘ[A, B] = (A * B).card := by + rw [mul_comm, mulConst_mul_card] + +@[to_additive (attr := simp) card_mul_subConst] +lemma card_mul_divConst (A B : Finset G) : A.card * δₘ[A, B] = (A / B).card := by + rw [mul_comm, divConst_mul_card] + +@[to_additive (attr := simp)] +lemma mulConst_empty_left (B : Finset G) : σₘ[∅, B] = 0 := by simp [mulConst] + +@[to_additive (attr := simp)] +lemma divConst_empty_left (B : Finset G) : δₘ[∅, B] = 0 := by simp [divConst] + +@[to_additive (attr := simp)] +lemma mulConst_empty_right (A : Finset G) : σₘ[A, ∅] = 0 := by simp [mulConst] + +@[to_additive (attr := simp)] +lemma divConst_empty_right (A : Finset G) : δₘ[A, ∅] = 0 := by simp [divConst] + +@[to_additive (attr := simp)] +lemma mulConst_inv_right (A B : Finset G) : σₘ[A, B⁻¹] = δₘ[A, B] := by + rw [mulConst, divConst, ← div_eq_mul_inv] + +@[to_additive (attr := simp)] +lemma divConst_inv_right (A B : Finset G) : δₘ[A, B⁻¹] = σₘ[A, B] := by + rw [mulConst, divConst, div_inv_eq_mul] + +section Fintype +variable [Fintype G] + +/-- Dense sets have small doubling. -/ +@[to_additive addConst_le_inv_dens "Dense sets have small doubling."] +lemma mulConst_le_inv_dens : σₘ[A, B] ≤ A.dens⁻¹ := by + rw [dens, inv_div, mulConst]; gcongr; exact card_le_univ _ + +/-- Dense sets have small difference constant. -/ +@[to_additive subConst_le_inv_dens "Dense sets have small difference constant."] +lemma divConst_le_inv_dens : δₘ[A, B] ≤ A.dens⁻¹ := by + rw [dens, inv_div, divConst]; gcongr; exact card_le_univ _ + +end Fintype + +variable {𝕜 : Type*} [Semifield 𝕜] [CharZero 𝕜] + +lemma cast_addConst (A B : Finset G') : (σ[A, B] : 𝕜) = (A + B).card / A.card := by + simp [addConst] + +lemma cast_subConst (A B : Finset G') : (δ[A, B] : 𝕜) = (A - B).card / A.card := by + simp [subConst] + +@[to_additive existing] +lemma cast_mulConst (A B : Finset G) : (σₘ[A, B] : 𝕜) = (A * B).card / A.card := by simp [mulConst] + +@[to_additive existing] +lemma cast_divConst (A B : Finset G) : (δₘ[A, B] : 𝕜) = (A / B).card / A.card := by simp [divConst] + +lemma cast_addConst_mul_card (A B : Finset G') : (σ[A, B] * A.card : 𝕜) = (A + B).card := by + norm_cast; exact addConst_mul_card _ _ + +lemma cast_subConst_mul_card (A B : Finset G') : (δ[A, B] * A.card : 𝕜) = (A - B).card := by + norm_cast; exact subConst_mul_card _ _ + +lemma card_mul_cast_addConst (A B : Finset G') : (A.card * σ[A, B] : 𝕜) = (A + B).card := by + norm_cast; exact card_mul_addConst _ _ + +lemma card_mul_cast_subConst (A B : Finset G') : (A.card * δ[A, B] : 𝕜) = (A - B).card := by + norm_cast; exact card_mul_subConst _ _ + +@[to_additive (attr := simp) existing cast_addConst_mul_card] +lemma cast_mulConst_mul_card (A B : Finset G) : (σₘ[A, B] * A.card : 𝕜) = (A * B).card := by + norm_cast; exact mulConst_mul_card _ _ + +@[to_additive (attr := simp) existing cast_subConst_mul_card] +lemma cast_divConst_mul_card (A B : Finset G) : (δₘ[A, B] * A.card : 𝕜) = (A / B).card := by + norm_cast; exact divConst_mul_card _ _ + +@[to_additive (attr := simp) existing card_mul_cast_addConst] +lemma card_mul_cast_mulConst (A B : Finset G) : (A.card * σₘ[A, B] : 𝕜) = (A * B).card := by + norm_cast; exact card_mul_mulConst _ _ + +@[to_additive (attr := simp) existing card_mul_cast_subConst] +lemma card_mul_cast_divConst (A B : Finset G) : (A.card * δₘ[A, B] : 𝕜) = (A / B).card := by + norm_cast; exact card_mul_divConst _ _ + +end Group + +open scoped Combinatorics.Additive + +section CommGroup +variable {G : Type*} [CommGroup G] [DecidableEq G] {A B : Finset G} {a : G} + +@[to_additive (attr := simp)] +lemma mulConst_inv_left (A B : Finset G) : σₘ[A⁻¹, B] = δₘ[A, B] := by + rw [mulConst, divConst, card_inv, ← card_inv, mul_inv_rev, inv_inv, inv_mul_eq_div] + +@[to_additive (attr := simp)] +lemma divConst_inv_left (A B : Finset G) : δₘ[A⁻¹, B] = σₘ[A, B] := by + rw [mulConst, divConst, card_inv, ← card_inv, inv_div, div_inv_eq_mul, mul_comm] + +/-- If `A` has small difference, then it has small doubling, with the constant squared. + +This is a consequence of the Ruzsa triangle inequality. -/ +@[to_additive +"If `A` has small difference, then it has small doubling, with the constant squared. + +This is a consequence of the Ruzsa triangle inequality."] +lemma mulConst_le_divConst_sq : σₘ[A] ≤ δₘ[A] ^ 2 := by + obtain rfl | hA' := A.eq_empty_or_nonempty + · simp + refine le_of_mul_le_mul_right ?_ (by positivity : (0 : ℚ≥0) < A.card * A.card) + calc + _ = (A * A).card * (A.card : ℚ≥0) := by rw [← mul_assoc, mulConst_mul_card] + _ ≤ (A / A).card * (A / A).card := by norm_cast; exact ruzsa_triangle_inequality_mul_div_div .. + _ = _ := by rw [← divConst_mul_card]; ring + +/-- If `A` has small doubling, then it has small difference, with the constant squared. + +This is a consequence of the Ruzsa triangle inequality. -/ +@[to_additive +"If `A` has small doubling, then it has small difference, with the constant squared. + +This is a consequence of the Ruzsa triangle inequality."] +lemma divConst_le_mulConst_sq : δₘ[A] ≤ σₘ[A] ^ 2 := by + obtain rfl | hA' := A.eq_empty_or_nonempty + · simp + refine le_of_mul_le_mul_right ?_ (by positivity : (0 : ℚ≥0) < A.card * A.card) + calc + _ = (A / A).card * (A.card : ℚ≥0) := by rw [← mul_assoc, divConst_mul_card] + _ ≤ (A * A).card * (A * A).card := by norm_cast; exact ruzsa_triangle_inequality_div_mul_mul .. + _ = _ := by rw [← mulConst_mul_card]; ring + +end CommGroup +end Finset From 224578708b4f82f7a3cbdd4f506a7ff98a337364 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 16 Aug 2024 19:53:03 +0000 Subject: [PATCH 343/359] fix(Filter/Pointwise): fix additive name (#15900) --- Mathlib/Order/Filter/Pointwise.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/Mathlib/Order/Filter/Pointwise.lean b/Mathlib/Order/Filter/Pointwise.lean index 4e0a809056265..3bfff41d94cf1 100644 --- a/Mathlib/Order/Filter/Pointwise.lean +++ b/Mathlib/Order/Filter/Pointwise.lean @@ -116,10 +116,12 @@ theorem eventually_one {p : α → Prop} : (∀ᶠ x in 1, p x) ↔ p 1 := theorem tendsto_one {a : Filter β} {f : β → α} : Tendsto f a 1 ↔ ∀ᶠ x in a, f x = 1 := tendsto_pure -@[to_additive (attr := simp)] +@[to_additive (attr := simp) zero_prod_zero] theorem one_prod_one [One β] : (1 : Filter α) ×ˢ (1 : Filter β) = 1 := prod_pure_pure +@[deprecated (since := "2024-08-16")] alias zero_sum_zero := zero_prod_zero + /-- `pure` as a `OneHom`. -/ @[to_additive "`pure` as a `ZeroHom`."] def pureOneHom : OneHom α (Filter α) where From 537997c9ab39c53e49e52947399dbf0c9239d353 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 16 Aug 2024 22:06:34 +0000 Subject: [PATCH 344/359] chore(scripts): script for adaptation PRs should not bail on merge conflicts (#15886) - **chore(scripts/create-adaptation-pr): don't bail on merge conflicts** - **enumerate steps in conflict resolution** --- scripts/create-adaptation-pr.sh | 27 +++++++++++++++------------ 1 file changed, 15 insertions(+), 12 deletions(-) diff --git a/scripts/create-adaptation-pr.sh b/scripts/create-adaptation-pr.sh index f9611d8e63e3b..416fbdd944766 100755 --- a/scripts/create-adaptation-pr.sh +++ b/scripts/create-adaptation-pr.sh @@ -61,7 +61,7 @@ echo "### [auto] checkout 'bump/$BUMPVERSION' and merge the latest changes from git checkout "bump/$BUMPVERSION" git pull -git merge origin/master +git merge origin/master || true # ignore error if there are conflicts # Check if there are merge conflicts if git diff --name-only --diff-filter=U | grep -q .; then @@ -78,9 +78,10 @@ if git diff --name-only --diff-filter=U | grep -q .; then echo "### [user] Conflict resolution" echo "We are merging the latest changes from 'origin/master' into 'bump/$BUMPVERSION'" echo "There seem to be conflicts: please resolve them" - echo "Open `pwd` in a new terminal and run 'git status'" - echo "Make sure to commit the resolved conflicts, but do not push them" - read -p "Press enter to continue, when you are done" + echo "" + echo " 1) Open `pwd` in a new terminal and run 'git status'" + echo " 2) Make sure to commit the resolved conflicts, but do not push them" + read -p " 3) Press enter to continue, when you are done" fi git push @@ -89,7 +90,7 @@ echo echo "### [auto] create a new branch 'bump/nightly-$NIGHTLYDATE' and merge the latest changes from 'origin/nightly-testing'" git checkout -b "bump/nightly-$NIGHTLYDATE" -git merge origin/nightly-testing +git merge origin/nightly-testing || true # ignore error if there are conflicts # Check if there are merge conflicts if git diff --name-only --diff-filter=U | grep -q .; then @@ -106,9 +107,10 @@ if git diff --name-only --diff-filter=U | grep -q .; then echo "### [user] Conflict resolution" echo "We are merging the latest changes from 'origin/nightly-testing' into 'bump/nightly-$NIGHTLYDATE'" echo "There seem to be conflicts: please resolve them" - echo "Open `pwd` in a new terminal and run 'git status'" - echo "Run 'git add' on the resolved files, but do not commit" - read -p "Press enter to continue, when you are done" + echo "" + echo " 1) Open `pwd` in a new terminal and run 'git status'" + echo " 2) Run 'git add' on the resolved files, but do not commit" + read -p " 3) Press enter to continue, when you are done" fi echo @@ -161,7 +163,7 @@ echo "### [auto] checkout the 'nightly-testing' branch and merge the new branch git checkout nightly-testing git pull -git merge "bump/nightly-$NIGHTLYDATE" +git merge "bump/nightly-$NIGHTLYDATE" || true # ignore error if there are conflicts # Check if there are merge conflicts if git diff --name-only --diff-filter=U | grep -q .; then @@ -178,9 +180,10 @@ if git diff --name-only --diff-filter=U | grep -q .; then echo "### [user] Conflict resolution" echo "We are merging the new PR "bump/nightly-$NIGHTLYDATE" into 'nightly-testing'" echo "There seem to be conflicts: please resolve them" - echo "Open `pwd` in a new terminal and run 'git status'" - echo "Make sure to commit the resolved conflicts, but do not push them" - read -p "Press enter to continue, when you are done" + echo "" + echo " 1) Open `pwd` in a new terminal and run 'git status'" + echo " 2) Make sure to commit the resolved conflicts, but do not push them" + read -p " 3) Press enter to continue, when you are done" fi git push From abe811a85f9a2bbc6eca9ba9166351af839bd1e1 Mon Sep 17 00:00:00 2001 From: Jon Eugster Date: Fri, 16 Aug 2024 22:15:42 +0000 Subject: [PATCH 345/359] doc: update docstring of MonoidHomClass (#15804) Fix typo Coauthored by: @alissa-tung --- Mathlib/Algebra/Group/Hom/Defs.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index a64a2118ac2da..6c96b0483e734 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -312,7 +312,7 @@ variable [MulOneClass M] [MulOneClass N] /-- `M →* N` is the type of functions `M → N` that preserve the `Monoid` structure. `MonoidHom` is also used for group homomorphisms. -When possible, instead of parametrizing results over `(f : M →+ N)`, +When possible, instead of parametrizing results over `(f : M →* N)`, you should parametrize over `(F : Type*) [MonoidHomClass F M N] (f : F)`. When you extend this structure, make sure to extend `MonoidHomClass`. From d3a9d9c231e653a2351abaf1d53ece94aabc0912 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 16 Aug 2024 22:15:43 +0000 Subject: [PATCH 346/359] chore: Let `gcongr` know about `Nat.cast_lt` (#15879) From LeanAPAP --- Mathlib/Computability/AkraBazzi/AkraBazzi.lean | 4 ++-- Mathlib/Data/Nat/Cast/Order/Basic.lean | 3 +++ 2 files changed, 5 insertions(+), 2 deletions(-) diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index afe58c1fc4266..73a3749b858d0 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -1053,7 +1053,7 @@ lemma rpow_p_mul_one_sub_smoothingFn_le : case bn_gt_one => calc 1 = b i * (b i)⁻¹ := by rw [mul_inv_cancel₀ (by positivity)] _ ≤ b i * ⌈(b i)⁻¹⌉₊ := by gcongr; exact Nat.le_ceil _ - _ < b i * n := by gcongr; rw [Nat.cast_lt]; exact hn + _ < b i * n := by gcongr case le => calc b i * n ≤ 1 * n := by have := R.b_lt_one i; gcongr _ = n := by rw [one_mul] positivity @@ -1149,7 +1149,7 @@ lemma rpow_p_mul_one_add_smoothingFn_ge : case bn_gt_one => calc 1 = b i * (b i)⁻¹ := by rw [mul_inv_cancel₀ (by positivity)] _ ≤ b i * ⌈(b i)⁻¹⌉₊ := by gcongr; exact Nat.le_ceil _ - _ < b i * n := by gcongr; rw [Nat.cast_lt]; exact hn + _ < b i * n := by gcongr case le => calc b i * n ≤ 1 * n := by have := R.b_lt_one i; gcongr _ = n := by rw [one_mul] positivity diff --git a/Mathlib/Data/Nat/Cast/Order/Basic.lean b/Mathlib/Data/Nat/Cast/Order/Basic.lean index 489bdfacffcc2..0bd91bdef4a52 100644 --- a/Mathlib/Data/Nat/Cast/Order/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Order/Basic.lean @@ -69,6 +69,9 @@ variable [CharZero α] {m n : ℕ} theorem strictMono_cast : StrictMono (Nat.cast : ℕ → α) := mono_cast.strictMono_of_injective cast_injective +@[gcongr] +lemma _root_.GCongr.natCast_lt_natCast {a b : ℕ} (h : a < b) : (a : α) < b := strictMono_cast h + /-- `Nat.cast : ℕ → α` as an `OrderEmbedding` -/ @[simps! (config := .asFn)] def castOrderEmbedding : ℕ ↪o α := From ff84cbdbda144e6f8597c701b740c0cdbd216732 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 16 Aug 2024 23:11:50 +0000 Subject: [PATCH 347/359] feat: better proof of spectral permanence (#15603) This uses a much cleaner argument than the previous one in order to prove the *spectral permanence* property of C*-algebras. The argument now goes through some generic results for Banach algebras, a corollary of which applies to selfadjoint elements. The old results which led to spectral permanence have been removed, although the spectral permanence results themselves remain unchanged (except that they have been moved into `Analysis.CStarAlgebra.Spectrum` from `Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Basic`, which seems a better home). Some module documentation has also been updated. --- .../ContinuousFunctionalCalculus/Basic.lean | 163 +----------------- .../Instances.lean | 2 +- Mathlib/Analysis/CStarAlgebra/Spectrum.lean | 110 +++++++++++- .../Complex/UpperHalfPlane/Topology.lean | 14 ++ Mathlib/Analysis/Normed/Algebra/Spectrum.lean | 34 +++- 5 files changed, 156 insertions(+), 167 deletions(-) diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Basic.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Basic.lean index 2dc9d1ad67c79..a475c96675c2d 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Basic.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Basic.lean @@ -25,8 +25,10 @@ description is perfectly accurate in finite dimension, but only heuristic in inf there might be no genuine eigenvector. In particular, when `f` is a polynomial `∑ cᵢ Xⁱ`, then `f a` is `∑ cᵢ aⁱ`. Also, `id a = a`. -This file also includes a proof of the **spectral permanence** theorem for (unital) C⋆-algebras -(see `StarSubalgebra.spectrum_eq`) +The result we have established here is the strongest possible, but it is not the version which is +most useful in practice. The generic API for the continuous functional calculus can be found in +`Analysis.CStarAlgebra.ContinuousFunctionalCalculus` in the `Unital` and `NonUnital` files. The +relevant instances on C⋆-algebra can be found in the `Instances` file. ## Main definitions @@ -39,19 +41,7 @@ This file also includes a proof of the **spectral permanence** theorem for (unit algebra homomorphism. Moreover, this map is continuous and bijective and since the spaces involved are compact Hausdorff, it is a homeomorphism. -## Main statements - -* `StarSubalgebra.coe_isUnit`: for `x : S` in a C⋆-Subalgebra `S` of `A`, then `↑x : A` is a Unit - if and only if `x` is a unit. -* `StarSubalgebra.spectrum_eq`: **spectral_permanence** for `x : S`, where `S` is a C⋆-Subalgebra - of `A`, `spectrum ℂ x = spectrum ℂ (x : A)`. - -## Notes - -The result we have established here is the strongest possible, but it is likely not the version -which will be most useful in practice. Future work will include developing an appropriate API for -the continuous functional calculus (including one for real-valued functions with real argument that -applies to self-adjoint elements of the algebra). -/ + -/ open scoped Pointwise ENNReal NNReal ComplexOrder @@ -67,142 +57,8 @@ instance {R A : Type*} [CommRing R] [StarRing R] [NormedRing A] [Algebra R A] [S { SubringClass.toNormedRing (elementalStarAlgebra R a) with mul_comm := mul_comm } --- Porting note: these hack instances no longer seem to be necessary - variable [CompleteSpace A] (a : A) [IsStarNormal a] (S : StarSubalgebra ℂ A) -/-- This lemma is used in the proof of `elementalStarAlgebra.isUnit_of_isUnit_of_isStarNormal`, -which in turn is the key to spectral permanence `StarSubalgebra.spectrum_eq`, which is itself -necessary for the continuous functional calculus. Using the continuous functional calculus, this -lemma can be superseded by one that omits the `IsStarNormal` hypothesis. -/ -theorem spectrum_star_mul_self_of_isStarNormal : - spectrum ℂ (star a * a) ⊆ Set.Icc (0 : ℂ) ‖star a * a‖ := by - -- this instance should be found automatically, but without providing it Lean goes on a wild - -- goose chase when trying to apply `spectrum.gelfandTransform_eq`. - --letI := elementalStarAlgebra.Complex.normedAlgebra a - rcases subsingleton_or_nontrivial A with ⟨⟩ - · simp only [spectrum.of_subsingleton, Set.empty_subset] - · set a' : elementalStarAlgebra ℂ a := ⟨a, self_mem ℂ a⟩ - refine (spectrum.subset_subalgebra (star a' * a')).trans ?_ - rw [← spectrum.gelfandTransform_eq (star a' * a'), ContinuousMap.spectrum_eq_range] - rintro - ⟨φ, rfl⟩ - rw [gelfandTransform_apply_apply ℂ _ (star a' * a') φ, map_mul φ, map_star φ] - rw [Complex.eq_coe_norm_of_nonneg (star_mul_self_nonneg _), ← map_star, ← map_mul] - exact ⟨by positivity, Complex.real_le_real.2 (AlgHom.norm_apply_le_self φ (star a' * a'))⟩ - -variable {a} - -/-- This is the key lemma on the way to establishing spectral permanence for C⋆-algebras, which is -established in `StarSubalgebra.spectrum_eq`. This lemma is superseded by -`StarSubalgebra.coe_isUnit`, which does not require an `IsStarNormal` hypothesis and holds for -any closed star subalgebra. -/ -theorem elementalStarAlgebra.isUnit_of_isUnit_of_isStarNormal (h : IsUnit a) : - IsUnit (⟨a, self_mem ℂ a⟩ : elementalStarAlgebra ℂ a) := by - /- Sketch of proof: Because `a` is normal, it suffices to prove that `star a * a` is invertible - in `elementalStarAlgebra ℂ a`. For this it suffices to prove that it is sufficiently close to a - unit, namely `algebraMap ℂ _ ‖star a * a‖`, and in this case the required distance is - `‖star a * a‖`. So one must show `‖star a * a - algebraMap ℂ _ ‖star a * a‖‖ < ‖star a * a‖`. - Since `star a * a - algebraMap ℂ _ ‖star a * a‖` is selfadjoint, by a corollary of Gelfand's - formula for the spectral radius (`IsSelfAdjoint.spectralRadius_eq_nnnorm`) its norm is the - supremum of the norms of elements in its spectrum (we may use the spectrum in `A` here because - the norm in `A` and the norm in the subalgebra coincide). - - By `spectrum_star_mul_self_of_isStarNormal`, the spectrum (in the algebra `A`) of `star a * a` - is contained in the interval `[0, ‖star a * a‖]`, and since `a` (and hence `star a * a`) is - invertible in `A`, we may omit `0` from this interval. Therefore, by basic spectral mapping - properties, the spectrum (in the algebra `A`) of `star a * a - algebraMap ℂ _ ‖star a * a‖` is - contained in `[0, ‖star a * a‖)`. The supremum of the (norms of) elements of the spectrum must - be *strictly* less that `‖star a * a‖` because the spectrum is compact, which completes the - proof. -/ - /- We may assume `A` is nontrivial. It suffices to show that `star a * a` is invertible in the - commutative (because `a` is normal) ring `elementalStarAlgebra ℂ a`. Indeed, by commutativity, - if `star a * a` is invertible, then so is `a`. -/ - nontriviality A - set a' : elementalStarAlgebra ℂ a := ⟨a, self_mem ℂ a⟩ - suffices IsUnit (star a' * a') from (IsUnit.mul_iff.1 this).2 - replace h := (show Commute (star a) a from star_comm_self' a).isUnit_mul_iff.2 ⟨h.star, h⟩ - /- Since `a` is invertible, `‖star a * a‖ ≠ 0`, so `‖star a * a‖ • 1` is invertible in - `elementalStarAlgebra ℂ a`, and so it suffices to show that the distance between this unit and - `star a * a` is less than `‖star a * a‖`. -/ - have h₁ : (‖star a * a‖ : ℂ) ≠ 0 := Complex.ofReal_ne_zero.mpr (norm_ne_zero_iff.mpr h.ne_zero) - set u : Units (elementalStarAlgebra ℂ a) := - Units.map (algebraMap ℂ (elementalStarAlgebra ℂ a)).toMonoidHom (Units.mk0 _ h₁) - refine ⟨u.ofNearby _ ?_, rfl⟩ - simp only [u, Units.coe_map, Units.val_inv_eq_inv_val, RingHom.toMonoidHom_eq_coe, Units.val_mk0, - Units.coe_map_inv, MonoidHom.coe_coe, norm_algebraMap', norm_inv, Complex.norm_eq_abs, - Complex.abs_ofReal, abs_norm, inv_inv] - /- Since `a` is invertible, by `spectrum_star_mul_self_of_isStarNormal`, the spectrum (in `A`) - of `star a * a` is contained in the half-open interval `(0, ‖star a * a‖]`. Therefore, by basic - spectral mapping properties, the spectrum of `‖star a * a‖ • 1 - star a * a` is contained in - `[0, ‖star a * a‖)`. -/ - have h₂ : ∀ z ∈ spectrum ℂ (algebraMap ℂ A ‖star a * a‖ - star a * a), ‖z‖₊ < ‖star a * a‖₊ := by - intro z hz - rw [← spectrum.singleton_sub_eq, Set.singleton_sub] at hz - have h₃ : z ∈ Set.Icc (0 : ℂ) ‖star a * a‖ := by - replace hz := Set.image_subset _ (spectrum_star_mul_self_of_isStarNormal a) hz - rwa [Set.image_const_sub_Icc, sub_self, sub_zero] at hz - refine lt_of_le_of_ne (Complex.real_le_real.1 <| Complex.eq_coe_norm_of_nonneg h₃.1 ▸ h₃.2) ?_ - · intro hz' - replace hz' := congr_arg (fun x : ℝ≥0 => ((x : ℝ) : ℂ)) hz' - simp only [coe_nnnorm] at hz' - rw [← Complex.eq_coe_norm_of_nonneg h₃.1] at hz' - obtain ⟨w, hw₁, hw₂⟩ := hz - refine (spectrum.zero_not_mem_iff ℂ).mpr h ?_ - rw [hz', sub_eq_self] at hw₂ - rwa [hw₂] at hw₁ - /- The norm of `‖star a * a‖ • 1 - star a * a` in the subalgebra and in `A` coincide. In `A`, - because this element is selfadjoint, by `IsSelfAdjoint.spectralRadius_eq_nnnorm`, its norm is - the supremum of the norms of the elements of the spectrum, which is strictly less than - `‖star a * a‖` by `h₂` and because the spectrum is compact. -/ - exact ENNReal.coe_lt_coe.1 - (calc - (‖star a' * a' - algebraMap ℂ _ ‖star a * a‖‖₊ : ℝ≥0∞) = - ‖algebraMap ℂ A ‖star a * a‖ - star a * a‖₊ := by - rw [← nnnorm_neg, neg_sub]; rfl - _ = spectralRadius ℂ (algebraMap ℂ A ‖star a * a‖ - star a * a) := by - refine (IsSelfAdjoint.spectralRadius_eq_nnnorm ?_).symm - rw [IsSelfAdjoint, star_sub, star_mul, star_star, ← algebraMap_star_comm] - congr! - exact RCLike.conj_ofReal _ - _ < ‖star a * a‖₊ := spectrum.spectralRadius_lt_of_forall_lt _ h₂) - -/-- For `x : A` which is invertible in `A`, the inverse lies in any unital C⋆-subalgebra `S` -containing `x`. -/ -theorem StarSubalgebra.isUnit_coe_inv_mem {S : StarSubalgebra ℂ A} (hS : IsClosed (S : Set A)) - {x : A} (h : IsUnit x) (hxS : x ∈ S) : ↑h.unit⁻¹ ∈ S := by - have hx := h.star.mul h - suffices this : (↑hx.unit⁻¹ : A) ∈ S by - rw [← one_mul (↑h.unit⁻¹ : A), ← hx.unit.inv_mul, mul_assoc, IsUnit.unit_spec, mul_assoc, - h.mul_val_inv, mul_one] - exact mul_mem this (star_mem hxS) - refine le_of_isClosed_of_mem ℂ hS (mul_mem (star_mem hxS) hxS) ?_ - haveI := (IsSelfAdjoint.star_mul_self x).isStarNormal - have hx' := elementalStarAlgebra.isUnit_of_isUnit_of_isStarNormal hx - convert (↑hx'.unit⁻¹ : elementalStarAlgebra ℂ (star x * x)).prop using 1 - refine left_inv_eq_right_inv hx.unit.inv_mul ?_ - exact (congr_arg ((↑) : _ → A) hx'.unit.mul_inv) - -/-- For a unital C⋆-subalgebra `S` of `A` and `x : S`, if `↑x : A` is invertible in `A`, then -`x` is invertible in `S`. -/ -theorem StarSubalgebra.coe_isUnit {S : StarSubalgebra ℂ A} (hS : IsClosed (S : Set A)) {x : S} : - IsUnit (x : A) ↔ IsUnit x := by - refine ⟨fun hx => - ⟨⟨x, ⟨(↑hx.unit⁻¹ : A), StarSubalgebra.isUnit_coe_inv_mem hS hx x.prop⟩, ?_, ?_⟩, rfl⟩, - fun hx => hx.map S.subtype⟩ - exacts [Subtype.coe_injective hx.mul_val_inv, Subtype.coe_injective hx.val_inv_mul] - -theorem StarSubalgebra.mem_spectrum_iff {S : StarSubalgebra ℂ A} (hS : IsClosed (S : Set A)) {x : S} - {z : ℂ} : z ∈ spectrum ℂ x ↔ z ∈ spectrum ℂ (x : A) := - not_iff_not.2 (StarSubalgebra.coe_isUnit hS).symm - -/-- **Spectral permanence.** The spectrum of an element is invariant of the (closed) -`StarSubalgebra` in which it is contained. -/ -theorem StarSubalgebra.spectrum_eq {S : StarSubalgebra ℂ A} (hS : IsClosed (S : Set A)) (x : S) : - spectrum ℂ x = spectrum ℂ (x : A) := - Set.ext fun _z => StarSubalgebra.mem_spectrum_iff hS - -variable (a) - /-- The natural map from `characterSpace ℂ (elementalStarAlgebra ℂ x)` to `spectrum ℂ x` given by evaluating `φ` at `x`. This is essentially just evaluation of the `gelfandTransform` of `x`, but because we want something in `spectrum ℂ x`, as opposed to @@ -212,9 +68,8 @@ noncomputable def elementalStarAlgebra.characterSpaceToSpectrum (x : A) (φ : characterSpace ℂ (elementalStarAlgebra ℂ x)) : spectrum ℂ x where val := φ ⟨x, self_mem ℂ x⟩ property := by - simpa only [StarSubalgebra.spectrum_eq (elementalStarAlgebra.isClosed ℂ x) - ⟨x, self_mem ℂ x⟩] using - AlgHom.apply_mem_spectrum φ ⟨x, self_mem ℂ x⟩ + simpa only [StarSubalgebra.spectrum_eq (hS := elementalStarAlgebra.isClosed ℂ x) + (a := ⟨x, self_mem ℂ x⟩)] using AlgHom.apply_mem_spectrum φ ⟨x, self_mem ℂ x⟩ #adaptation_note /-- nightly-2024-04-01 The simpNF linter now times out on this lemma. @@ -234,8 +89,8 @@ theorem elementalStarAlgebra.bijective_characterSpaceToSpectrum : · simpa only [elementalStarAlgebra.characterSpaceToSpectrum, Subtype.mk_eq_mk, ContinuousMap.coe_mk] using h · rintro ⟨z, hz⟩ - have hz' := (StarSubalgebra.spectrum_eq (elementalStarAlgebra.isClosed ℂ a) - ⟨a, self_mem ℂ a⟩).symm.subst hz + have hz' := (StarSubalgebra.spectrum_eq (hS := elementalStarAlgebra.isClosed ℂ a) + (a := ⟨a, self_mem ℂ a⟩).symm.subst hz) rw [CharacterSpace.mem_spectrum_iff_exists] at hz' obtain ⟨φ, rfl⟩ := hz' exact ⟨φ, rfl⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index cc87a12072105..d401e7f3e4c6d 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -169,7 +169,7 @@ instance IsStarNormal.instContinuousFunctionalCalculus {A : Type*} [NormedRing A case hom_map_spectrum => intro f simp only [StarAlgHom.comp_apply, StarAlgHom.coe_coe, StarSubalgebra.coe_subtype] - rw [← StarSubalgebra.spectrum_eq (elementalStarAlgebra.isClosed ℂ a), + rw [← StarSubalgebra.spectrum_eq (hS := elementalStarAlgebra.isClosed ℂ a), AlgEquiv.spectrum_eq (continuousFunctionalCalculus a), ContinuousMap.spectrum_eq_range] case predicate_hom => exact fun f ↦ ⟨by rw [← map_star]; exact Commute.all (star f) f |>.map _⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/Spectrum.lean b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean index ceac7a25b3c1e..80057039f51d6 100644 --- a/Mathlib/Analysis/CStarAlgebra/Spectrum.lean +++ b/Mathlib/Analysis/CStarAlgebra/Spectrum.lean @@ -3,16 +3,58 @@ Copyright (c) 2022 Jireh Loreaux. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jireh Loreaux -/ -import Mathlib.Analysis.Normed.Algebra.Spectrum +import Mathlib.Analysis.Complex.UpperHalfPlane.Topology import Mathlib.Analysis.CStarAlgebra.Unitization +import Mathlib.Analysis.Normed.Algebra.Spectrum import Mathlib.Analysis.SpecialFunctions.Exponential import Mathlib.Algebra.Star.StarAlgHom /-! # Spectral properties in C⋆-algebras + In this file, we establish various properties related to the spectrum of elements in C⋆-algebras. +In particular, we show that the spectrum of a unitary element is contained in the unit circle in +`ℂ`, the spectrum of a selfadjoint element is real, the spectral radius of a selfadjoint element +or normal element is its norm, among others. + +An essential feature of C⋆-algebras is **spectral permanence**. This is the property that the +spectrum of an element in a closed subalgebra is the same as the spectrum of the element in the +whole algebra. For Banach algebras more generally, and even for Banach ⋆-algebras, this fails. + +A consequence of spectral permanence is that one may always enlarge the C⋆-algebra (via a unital +embedding) while preserving the spectrum of any element. In addition, it allows us to make sense of +the spectrum of elements in non-unital C⋆-algebras by considering them as elements in the +`Unitization` of the C⋆-algebra, or indeed *any* unital C⋆-algebra. Of course, one may do this +(that is, consider the spectrum of an element in a non-unital by embedding it in a unital algebra) +for any Banach algebra, but the downside in that setting is that embedding in different unital +algebras results in varying spectra. + +In Mathlib, we don't *define* the spectrum of an element in a non-unital C⋆-algebra, and instead +simply consider the `quasispectrum` so as to avoid depending on a choice of unital algebra. However, +we can still establish a form of spectral permanence. + +## Main statements + ++ `unitary.spectrum_subset_circle`: The spectrum of a unitary element is contained in the unit + sphere in `ℂ`. ++ `IsSelfAdjoint.spectralRadius_eq_nnnorm`: The spectral radius of a selfadjoint element is equal + to its norm. ++ `IsStarNormal.spectralRadius_eq_nnnorm`: The spectral radius of a normal element is equal to + its norm. ++ `IsSelfAdjoint.mem_spectrum_eq_re`: Any element of the spectrum of a selfadjoint element is real. +* `StarSubalgebra.coe_isUnit`: for `x : S` in a C⋆-Subalgebra `S` of `A`, then `↑x : A` is a Unit + if and only if `x` is a unit. +* `StarSubalgebra.spectrum_eq`: **spectral_permanence** for `x : S`, where `S` is a C⋆-Subalgebra + of `A`, `spectrum ℂ x = spectrum ℂ (x : A)`. + +## TODO + ++ prove a variation of spectral permanence using `StarAlgHom` instead of `StarSubalgebra`. ++ prove a variation of spectral permanence for `quasispectrum`. + -/ +local notation "σ" => spectrum local postfix:max "⋆" => star section @@ -87,8 +129,10 @@ theorem IsStarNormal.spectralRadius_eq_nnnorm (a : A) [IsStarNormal a] : convert tendsto_nhds_unique h₂ (pow_nnnorm_pow_one_div_tendsto_nhds_spectralRadius (a⋆ * a)) rw [(IsSelfAdjoint.star_mul_self a).spectralRadius_eq_nnnorm, sq, nnnorm_star_mul_self, coe_mul] +variable [StarModule ℂ A] + /-- Any element of the spectrum of a selfadjoint is real. -/ -theorem IsSelfAdjoint.mem_spectrum_eq_re [StarModule ℂ A] {a : A} (ha : IsSelfAdjoint a) {z : ℂ} +theorem IsSelfAdjoint.mem_spectrum_eq_re {a : A} (ha : IsSelfAdjoint a) {z : ℂ} (hz : z ∈ spectrum ℂ a) : z = z.re := by have hu := exp_mem_unitary_of_mem_skewAdjoint ℂ (ha.smul_mem_skewAdjoint conj_I) let Iu := Units.mk0 I I_ne_zero @@ -101,22 +145,78 @@ theorem IsSelfAdjoint.mem_spectrum_eq_re [StarModule ℂ A] {a : A} (ha : IsSelf spectrum.subset_circle_of_unitary hu this /-- Any element of the spectrum of a selfadjoint is real. -/ -theorem selfAdjoint.mem_spectrum_eq_re [StarModule ℂ A] (a : selfAdjoint A) {z : ℂ} +theorem selfAdjoint.mem_spectrum_eq_re (a : selfAdjoint A) {z : ℂ} (hz : z ∈ spectrum ℂ (a : A)) : z = z.re := a.prop.mem_spectrum_eq_re hz +/-- Any element of the spectrum of a selfadjoint is real. -/ +theorem IsSelfAdjoint.im_eq_zero_of_mem_spectrum {a : A} (ha : IsSelfAdjoint a) + {z : ℂ} (hz : z ∈ spectrum ℂ a) : z.im = 0 := by + rw [ha.mem_spectrum_eq_re hz, ofReal_im] + /-- The spectrum of a selfadjoint is real -/ -theorem IsSelfAdjoint.val_re_map_spectrum [StarModule ℂ A] {a : A} (ha : IsSelfAdjoint a) : +theorem IsSelfAdjoint.val_re_map_spectrum {a : A} (ha : IsSelfAdjoint a) : spectrum ℂ a = ((↑) ∘ re '' spectrum ℂ a : Set ℂ) := le_antisymm (fun z hz => ⟨z, hz, (ha.mem_spectrum_eq_re hz).symm⟩) fun z => by rintro ⟨z, hz, rfl⟩ simpa only [(ha.mem_spectrum_eq_re hz).symm, Function.comp_apply] using hz /-- The spectrum of a selfadjoint is real -/ -theorem selfAdjoint.val_re_map_spectrum [StarModule ℂ A] (a : selfAdjoint A) : +theorem selfAdjoint.val_re_map_spectrum (a : selfAdjoint A) : spectrum ℂ (a : A) = ((↑) ∘ re '' spectrum ℂ (a : A) : Set ℂ) := a.property.val_re_map_spectrum +/-- The complement of the spectrum of a selfadjoint element in a C⋆-algebra is connected. -/ +lemma IsSelfAdjoint.isConnected_spectrum_compl {a : A} (ha : IsSelfAdjoint a) : + IsConnected (σ ℂ a)ᶜ := by + suffices IsConnected (((σ ℂ a)ᶜ ∩ {z | 0 ≤ z.im}) ∪ (σ ℂ a)ᶜ ∩ {z | z.im ≤ 0}) by + rw [← Set.inter_union_distrib_left, ← Set.setOf_or] at this + rw [← Set.inter_univ (σ ℂ a)ᶜ] + convert this using 2 + exact Eq.symm <| Set.eq_univ_of_forall (fun z ↦ le_total 0 z.im) + refine IsConnected.union ?nonempty ?upper ?lower + case nonempty => + have := Filter.NeBot.nonempty_of_mem inferInstance <| Filter.mem_map.mp <| + Complex.isometry_ofReal.antilipschitz.tendsto_cobounded (spectrum.isBounded a |>.compl) + exact this.image Complex.ofReal' |>.mono <| by simp + case' upper => apply Complex.isConnected_of_upperHalfPlane ?_ <| Set.inter_subset_right + case' lower => apply Complex.isConnected_of_lowerHalfPlane ?_ <| Set.inter_subset_right + all_goals + refine Set.subset_inter (fun z hz hz' ↦ ?_) (fun _ ↦ by simpa using le_of_lt) + rw [Set.mem_setOf_eq, ha.im_eq_zero_of_mem_spectrum hz'] at hz + simp_all + +namespace StarSubalgebra + +variable (S : StarSubalgebra ℂ A) [hS : IsClosed (S : Set A)] + +/-- For a unital C⋆-subalgebra `S` of `A` and `x : S`, if `↑x : A` is invertible in `A`, then +`x` is invertible in `S`. -/ +lemma coe_isUnit {a : S} : IsUnit (a : A) ↔ IsUnit a := by + refine ⟨fun ha ↦ ?_, IsUnit.map S.subtype⟩ + have ha₁ := ha.star.mul ha + have ha₂ := ha.mul ha.star + have spec_eq {x : S} (hx : IsSelfAdjoint x) : spectrum ℂ x = spectrum ℂ (x : A) := + Subalgebra.spectrum_eq_of_isPreconnected_compl S _ <| + (hx.starHom_apply S.subtype).isConnected_spectrum_compl.isPreconnected + rw [← StarMemClass.coe_star, ← MulMemClass.coe_mul, ← spectrum.zero_not_mem_iff ℂ, ← spec_eq, + spectrum.zero_not_mem_iff] at ha₁ ha₂ + · have h₁ : ha₁.unit⁻¹ * star a * a = 1 := mul_assoc _ _ a ▸ ha₁.val_inv_mul + have h₂ : a * (star a * ha₂.unit⁻¹) = 1 := (mul_assoc a _ _).symm ▸ ha₂.mul_val_inv + exact ⟨⟨a, ha₁.unit⁻¹ * star a, left_inv_eq_right_inv h₁ h₂ ▸ h₂, h₁⟩, rfl⟩ + · exact IsSelfAdjoint.mul_star_self a + · exact IsSelfAdjoint.star_mul_self a + +lemma mem_spectrum_iff {a : S} {z : ℂ} : z ∈ spectrum ℂ a ↔ z ∈ spectrum ℂ (a : A) := + not_iff_not.mpr S.coe_isUnit.symm + +/-- **Spectral permanence.** The spectrum of an element is invariant of the (closed) +`StarSubalgebra` in which it is contained. -/ +lemma spectrum_eq {a : S} : spectrum ℂ a = spectrum ℂ (a : A) := + Set.ext fun _ ↦ S.mem_spectrum_iff + +end StarSubalgebra + end ComplexScalars namespace NonUnitalStarAlgHom diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean index 28f5dd099e13f..794141a06f371 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Topology.lean @@ -131,3 +131,17 @@ lemma comp_ofComplex (f : ℍ → ℂ) (z : ℍ) : (↑ₕ f) z = f z := by rw [Function.comp_apply, ofComplex_apply] end UpperHalfPlane + +lemma Complex.isConnected_of_upperHalfPlane {s : Set ℂ} (hs₁ : {z | 0 < z.im} ⊆ s) + (hs₂ : s ⊆ {z | 0 ≤ z.im}) : IsConnected s := by + refine IsConnected.subset_closure ?_ hs₁ (by simpa using hs₂) + rw [isConnected_iff_connectedSpace] + exact inferInstanceAs (ConnectedSpace UpperHalfPlane) + +lemma Complex.isConnected_of_lowerHalfPlane {s : Set ℂ} (hs₁ : {z | z.im < 0} ⊆ s) + (hs₂ : s ⊆ {z | z.im ≤ 0}) : IsConnected s := by + rw [← Equiv.star.surjective.image_preimage s] + refine IsConnected.image (f := Equiv.star) ?_ continuous_star.continuousOn + apply Complex.isConnected_of_upperHalfPlane + · exact fun z hz ↦ hs₁ <| show star z ∈ _ by simpa + · exact fun z hz ↦ by simpa using show (star z).im ≤ 0 from hs₂ hz diff --git a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean index 6e710a4b5a224..d74a1d6c50562 100644 --- a/Mathlib/Analysis/Normed/Algebra/Spectrum.lean +++ b/Mathlib/Analysis/Normed/Algebra/Spectrum.lean @@ -573,11 +573,14 @@ section BoundarySpectrum local notation "σ" => spectrum -variable {𝕜 A SA : Type*} [NormedField 𝕜] [NormedRing A] [NormedAlgebra 𝕜 A] [CompleteSpace A] -variable [SetLike SA A] [SubringClass SA A] [instSMulMem : SMulMemClass SA 𝕜 A] -variable (S : SA) [hS : IsClosed (S : Set A)] (x : S) +variable {𝕜 A SA : Type*} [NormedRing A] [CompleteSpace A] [SetLike SA A] [SubringClass SA A] + +open Topology Filter Set -open Topology Filter +section NormedField + +variable [NormedField 𝕜] [NormedAlgebra 𝕜 A] [instSMulMem : SMulMemClass SA 𝕜 A] +variable (S : SA) [hS : IsClosed (S : Set A)] (x : S) open SubalgebraClass in include instSMulMem in @@ -602,9 +605,7 @@ lemma _root_.Subalgebra.isUnit_of_isUnit_val_of_eventually {l : Filter S} {a : S apply Units.mul_eq_one_iff_inv_eq.mp simpa [-IsUnit.mul_val_inv] using congr(($hx.mul_val_inv : A)) -open Set - -/-- If `S` is a closed subalgebra of a Banach algebra `A`, then for any +/-- If `S : Subalgebra 𝕜 A` is a closed subalgebra of a Banach algebra `A`, then for any `x : S`, the boundary of the spectrum of `x` relative to `S` is a subset of the spectrum of `↑x : A` relative to `A`. -/ lemma _root_.Subalgebra.frontier_spectrum : frontier (σ 𝕜 x) ⊆ σ 𝕜 (x : A) := by @@ -676,6 +677,25 @@ lemma Subalgebra.spectrum_isBounded_connectedComponentIn {z : 𝕜} (hz : z ∈ rw [spectrum_sUnion_connectedComponentIn S] exact subset_biUnion_of_mem (mem_diff_of_mem hz hz') |>.trans subset_union_right +end NormedField + +variable [NontriviallyNormedField 𝕜] [NormedAlgebra 𝕜 A] [SMulMemClass SA 𝕜 A] +variable (S : SA) [hS : IsClosed (S : Set A)] (x : S) + +/-- Let `S` be a closed subalgebra of a Banach algebra `A`. If for `x : S` the complement of the +spectrum of `↑x : A` is connected, then `spectrum 𝕜 x = spectrum 𝕜 (x : A)`. -/ +lemma Subalgebra.spectrum_eq_of_isPreconnected_compl (h : IsPreconnected (σ 𝕜 (x : A))ᶜ) : + σ 𝕜 x = σ 𝕜 (x : A) := by + nontriviality A + suffices σ 𝕜 x \ σ 𝕜 (x : A) = ∅ by + rw [spectrum_sUnion_connectedComponentIn, this] + simp + refine eq_empty_of_forall_not_mem fun z hz ↦ NormedSpace.unbounded_univ 𝕜 𝕜 ?_ + obtain ⟨hz, hz'⟩ := mem_diff _ |>.mp hz + have := (spectrum.isBounded (x : A)).union <| + h.connectedComponentIn hz' ▸ spectrum_isBounded_connectedComponentIn S x hz + simpa + end BoundarySpectrum namespace SpectrumRestricts From 61b9d5e77adcf9cc825a4993fc4f177ac448a892 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 16 Aug 2024 23:11:51 +0000 Subject: [PATCH 348/359] chore: avoid importing `Cardinal` in basic group theory (#15646) --- Mathlib.lean | 10 ++-- Mathlib/Algebra/Category/Grp/Colimits.lean | 2 +- Mathlib/Algebra/Category/Grp/EpiMono.lean | 2 +- Mathlib/Algebra/CharZero/Quotient.lean | 4 +- Mathlib/Algebra/ModEq.lean | 3 +- Mathlib/Algebra/Order/ToIntervalMod.lean | 1 - Mathlib/Algebra/Periodic.lean | 2 +- Mathlib/Algebra/Pointwise/Stabilizer.lean | 2 +- Mathlib/Data/ZMod/Basic.lean | 1 + Mathlib/GroupTheory/Abelianization.lean | 3 +- .../Basic.lean} | 22 +------- Mathlib/GroupTheory/Commutator/Finite.lean | 35 ++++++++++++ Mathlib/GroupTheory/CommutingProbability.lean | 1 + .../{Coset.lean => Coset/Basic.lean} | 37 +------------ Mathlib/GroupTheory/Coset/Card.lean | 52 ++++++++++++++++++ Mathlib/GroupTheory/Coxeter/Basic.lean | 1 + Mathlib/GroupTheory/Divisible.lean | 3 +- Mathlib/GroupTheory/DoubleCoset.lean | 2 +- Mathlib/GroupTheory/Finiteness.lean | 3 +- .../GroupTheory/GroupAction/CardCommute.lean | 32 +++++++++++ Mathlib/GroupTheory/GroupAction/Quotient.lean | 27 +--------- Mathlib/GroupTheory/Index.lean | 2 + Mathlib/GroupTheory/Nilpotent.lean | 1 + Mathlib/GroupTheory/OrderOfElement.lean | 1 + Mathlib/GroupTheory/PresentedGroup.lean | 2 +- .../Basic.lean} | 38 +------------ Mathlib/GroupTheory/QuotientGroup/Finite.lean | 53 +++++++++++++++++++ .../SpecificGroups/Alternating.lean | 2 +- .../GroupTheory/SpecificGroups/Dihedral.lean | 1 + .../LinearAlgebra/Alternating/DomCoprod.lean | 1 + Mathlib/LinearAlgebra/Quotient.lean | 3 +- .../MeasureTheory/MeasurableSpace/Basic.lean | 2 +- Mathlib/RingTheory/ClassGroup.lean | 1 - Mathlib/Topology/Algebra/Group/Basic.lean | 2 +- .../Topology/Algebra/InfiniteSum/Group.lean | 1 + .../Algebra/InfiniteSum/Nonarchimedean.lean | 1 + Mathlib/Topology/Algebra/Ring/Basic.lean | 1 + 37 files changed, 219 insertions(+), 138 deletions(-) rename Mathlib/GroupTheory/{Commutator.lean => Commutator/Basic.lean} (92%) create mode 100644 Mathlib/GroupTheory/Commutator/Finite.lean rename Mathlib/GroupTheory/{Coset.lean => Coset/Basic.lean} (94%) create mode 100644 Mathlib/GroupTheory/Coset/Card.lean create mode 100644 Mathlib/GroupTheory/GroupAction/CardCommute.lean rename Mathlib/GroupTheory/{QuotientGroup.lean => QuotientGroup/Basic.lean} (93%) create mode 100644 Mathlib/GroupTheory/QuotientGroup/Finite.lean diff --git a/Mathlib.lean b/Mathlib.lean index d1da137ddf3a0..c288e312d7466 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2698,7 +2698,8 @@ import Mathlib.GroupTheory.Abelianization import Mathlib.GroupTheory.Archimedean import Mathlib.GroupTheory.ClassEquation import Mathlib.GroupTheory.Commensurable -import Mathlib.GroupTheory.Commutator +import Mathlib.GroupTheory.Commutator.Basic +import Mathlib.GroupTheory.Commutator.Finite import Mathlib.GroupTheory.CommutingProbability import Mathlib.GroupTheory.Complement import Mathlib.GroupTheory.Congruence.Basic @@ -2706,7 +2707,8 @@ import Mathlib.GroupTheory.Congruence.BigOperators import Mathlib.GroupTheory.Congruence.Opposite import Mathlib.GroupTheory.Coprod.Basic import Mathlib.GroupTheory.CoprodI -import Mathlib.GroupTheory.Coset +import Mathlib.GroupTheory.Coset.Basic +import Mathlib.GroupTheory.Coset.Card import Mathlib.GroupTheory.CosetCover import Mathlib.GroupTheory.Coxeter.Basic import Mathlib.GroupTheory.Coxeter.Inversion @@ -2728,6 +2730,7 @@ import Mathlib.GroupTheory.FreeGroup.NielsenSchreier import Mathlib.GroupTheory.GroupAction.Basic import Mathlib.GroupTheory.GroupAction.BigOperators import Mathlib.GroupTheory.GroupAction.Blocks +import Mathlib.GroupTheory.GroupAction.CardCommute import Mathlib.GroupTheory.GroupAction.ConjAct import Mathlib.GroupTheory.GroupAction.DomAct.ActionHom import Mathlib.GroupTheory.GroupAction.DomAct.Basic @@ -2775,7 +2778,8 @@ import Mathlib.GroupTheory.Perm.Support import Mathlib.GroupTheory.Perm.ViaEmbedding import Mathlib.GroupTheory.PresentedGroup import Mathlib.GroupTheory.PushoutI -import Mathlib.GroupTheory.QuotientGroup +import Mathlib.GroupTheory.QuotientGroup.Basic +import Mathlib.GroupTheory.QuotientGroup.Finite import Mathlib.GroupTheory.Schreier import Mathlib.GroupTheory.SchurZassenhaus import Mathlib.GroupTheory.SemidirectProduct diff --git a/Mathlib/Algebra/Category/Grp/Colimits.lean b/Mathlib/Algebra/Category/Grp/Colimits.lean index 68b7e009d27e5..a53daac3791a2 100644 --- a/Mathlib/Algebra/Category/Grp/Colimits.lean +++ b/Mathlib/Algebra/Category/Grp/Colimits.lean @@ -4,10 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Mathlib.Algebra.Category.Grp.Preadditive -import Mathlib.GroupTheory.QuotientGroup import Mathlib.CategoryTheory.Limits.Shapes.Kernels import Mathlib.CategoryTheory.Limits.Shapes.FiniteLimits import Mathlib.CategoryTheory.ConcreteCategory.Elementwise +import Mathlib.GroupTheory.QuotientGroup.Basic /-! # The category of additive commutative groups has all colimits. diff --git a/Mathlib/Algebra/Category/Grp/EpiMono.lean b/Mathlib/Algebra/Category/Grp/EpiMono.lean index 31c09a8f4a5fe..472da6d5b3651 100644 --- a/Mathlib/Algebra/Category/Grp/EpiMono.lean +++ b/Mathlib/Algebra/Category/Grp/EpiMono.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import Mathlib.Algebra.Category.Grp.EquivalenceGroupAddGroup -import Mathlib.GroupTheory.QuotientGroup import Mathlib.CategoryTheory.ConcreteCategory.EpiMono import Mathlib.CategoryTheory.Limits.Constructions.EpiMono +import Mathlib.GroupTheory.QuotientGroup.Basic /-! # Monomorphisms and epimorphisms in `Group` diff --git a/Mathlib/Algebra/CharZero/Quotient.lean b/Mathlib/Algebra/CharZero/Quotient.lean index 309850f089606..8a44865763c7e 100644 --- a/Mathlib/Algebra/CharZero/Quotient.lean +++ b/Mathlib/Algebra/CharZero/Quotient.lean @@ -3,7 +3,9 @@ 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.GroupTheory.QuotientGroup +import Mathlib.Algebra.Field.Basic +import Mathlib.GroupTheory.QuotientGroup.Basic +import Mathlib.Algebra.Order.Group.Unbundled.Int /-! # Lemmas about quotients in characteristic zero diff --git a/Mathlib/Algebra/ModEq.lean b/Mathlib/Algebra/ModEq.lean index 5cd5445eab5d8..8c4bcc0fd4def 100644 --- a/Mathlib/Algebra/ModEq.lean +++ b/Mathlib/Algebra/ModEq.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Int.ModEq -import Mathlib.GroupTheory.QuotientGroup +import Mathlib.Algebra.Field.Basic +import Mathlib.GroupTheory.QuotientGroup.Basic /-! # Equality modulo an element diff --git a/Mathlib/Algebra/Order/ToIntervalMod.lean b/Mathlib/Algebra/Order/ToIntervalMod.lean index 1f303379050ac..938027eaa347f 100644 --- a/Mathlib/Algebra/Order/ToIntervalMod.lean +++ b/Mathlib/Algebra/Order/ToIntervalMod.lean @@ -8,7 +8,6 @@ import Mathlib.Algebra.Module.Defs import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Algebra.Periodic import Mathlib.Data.Int.SuccPred -import Mathlib.GroupTheory.QuotientGroup import Mathlib.Order.Circular import Mathlib.Data.List.TFAE import Mathlib.Data.Set.Lattice diff --git a/Mathlib/Algebra/Periodic.lean b/Mathlib/Algebra/Periodic.lean index 2a4c1ad2350fe..d3f5025d9b640 100644 --- a/Mathlib/Algebra/Periodic.lean +++ b/Mathlib/Algebra/Periodic.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.Group.Subgroup.ZPowers import Mathlib.Algebra.Group.Submonoid.Membership import Mathlib.Algebra.Ring.NegOnePow import Mathlib.Algebra.Order.Archimedean.Basic -import Mathlib.GroupTheory.Coset +import Mathlib.GroupTheory.Coset.Card /-! # Periodicity diff --git a/Mathlib/Algebra/Pointwise/Stabilizer.lean b/Mathlib/Algebra/Pointwise/Stabilizer.lean index a63e174396fa1..deab88e20079c 100644 --- a/Mathlib/Algebra/Pointwise/Stabilizer.lean +++ b/Mathlib/Algebra/Pointwise/Stabilizer.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Finset.Pointwise.Basic -import Mathlib.GroupTheory.QuotientGroup +import Mathlib.GroupTheory.QuotientGroup.Basic /-! # Stabilizer of a set under a pointwise action diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 19c83df818076..db8d87e23d3e4 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -8,6 +8,7 @@ import Mathlib.GroupTheory.GroupAction.SubMulAction import Mathlib.GroupTheory.OrderOfElement import Mathlib.Tactic.FinCases import Mathlib.Tactic.Linarith +import Mathlib.Data.Fintype.Units /-! diff --git a/Mathlib/GroupTheory/Abelianization.lean b/Mathlib/GroupTheory/Abelianization.lean index 0b848d11ae3ab..51b0cfca00034 100644 --- a/Mathlib/GroupTheory/Abelianization.lean +++ b/Mathlib/GroupTheory/Abelianization.lean @@ -3,8 +3,9 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Michael Howes -/ -import Mathlib.GroupTheory.Commutator +import Mathlib.Data.Finite.Card import Mathlib.GroupTheory.Finiteness +import Mathlib.GroupTheory.Commutator.Basic /-! # The abelianization of a group diff --git a/Mathlib/GroupTheory/Commutator.lean b/Mathlib/GroupTheory/Commutator/Basic.lean similarity index 92% rename from Mathlib/GroupTheory/Commutator.lean rename to Mathlib/GroupTheory/Commutator/Basic.lean index 9f7244cd843dc..8e4f605d4f565 100644 --- a/Mathlib/GroupTheory/Commutator.lean +++ b/Mathlib/GroupTheory/Commutator/Basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Jordan Brown, Thomas Browning, Patrick Lutz. All rights reser Released under Apache 2.0 license as described in the file LICENSE. Authors: Jordan Brown, Thomas Browning, Patrick Lutz -/ -import Mathlib.Algebra.Group.Subgroup.Finite import Mathlib.GroupTheory.Subgroup.Centralizer import Mathlib.Tactic.Group @@ -20,6 +19,8 @@ is the subgroup of `G` generated by the commutators `h₁ * h₂ * h₁⁻¹ * h * `⁅H₁, H₂⁆` : the commutator of the subgroups `H₁` and `H₂`. -/ +assert_not_exists Cardinal +assert_not_exists Multiset variable {G G' F : Type*} [Group G] [Group G'] [FunLike F G G'] [MonoidHomClass F G G'] variable (f : F) {g₁ g₂ g₃ g : G} @@ -186,25 +187,6 @@ theorem commutator_pi_pi_le {η : Type*} {Gs : η → Type*} [∀ i, Group (Gs i ⁅Subgroup.pi Set.univ H, Subgroup.pi Set.univ K⁆ ≤ Subgroup.pi Set.univ fun i => ⁅H i, K i⁆ := commutator_le.mpr fun _p hp _q hq i hi => commutator_mem_commutator (hp i hi) (hq i hi) -/-- The commutator of a finite direct product is contained in the direct product of the commutators. --/ -theorem commutator_pi_pi_of_finite {η : Type*} [Finite η] {Gs : η → Type*} [∀ i, Group (Gs i)] - (H K : ∀ i, Subgroup (Gs i)) : ⁅Subgroup.pi Set.univ H, Subgroup.pi Set.univ K⁆ = - Subgroup.pi Set.univ fun i => ⁅H i, K i⁆ := by - classical - apply le_antisymm (commutator_pi_pi_le H K) - rw [pi_le_iff] - intro i hi - rw [map_commutator] - apply commutator_mono <;> - · rw [le_pi_iff] - intro j _hj - rintro _ ⟨_, ⟨x, hx, rfl⟩, rfl⟩ - by_cases h : j = i - · subst h - simpa using hx - · simp [h, one_mem] - end Subgroup variable (G) diff --git a/Mathlib/GroupTheory/Commutator/Finite.lean b/Mathlib/GroupTheory/Commutator/Finite.lean new file mode 100644 index 0000000000000..0109ea7274cb1 --- /dev/null +++ b/Mathlib/GroupTheory/Commutator/Finite.lean @@ -0,0 +1,35 @@ +/- +Copyright (c) 2021 Jordan Brown, Thomas Browning, Patrick Lutz. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jordan Brown, Thomas Browning, Patrick Lutz +-/ +import Mathlib.Algebra.Group.Subgroup.Finite +import Mathlib.GroupTheory.Commutator.Basic + +/-! +The commutator of a finite direct product is contained in the direct product of the commutators. +-/ + + +namespace Subgroup + +/-- The commutator of a finite direct product is contained in the direct product of the commutators. +-/ +theorem commutator_pi_pi_of_finite {η : Type*} [Finite η] {Gs : η → Type*} [∀ i, Group (Gs i)] + (H K : ∀ i, Subgroup (Gs i)) : ⁅Subgroup.pi Set.univ H, Subgroup.pi Set.univ K⁆ = + Subgroup.pi Set.univ fun i => ⁅H i, K i⁆ := by + classical + apply le_antisymm (commutator_pi_pi_le H K) + rw [pi_le_iff] + intro i hi + rw [map_commutator] + apply commutator_mono <;> + · rw [le_pi_iff] + intro j _hj + rintro _ ⟨_, ⟨x, hx, rfl⟩, rfl⟩ + by_cases h : j = i + · subst h + simpa using hx + · simp [h, one_mem] + +end Subgroup diff --git a/Mathlib/GroupTheory/CommutingProbability.lean b/Mathlib/GroupTheory/CommutingProbability.lean index 95d545452fa22..a0a76793010e2 100644 --- a/Mathlib/GroupTheory/CommutingProbability.lean +++ b/Mathlib/GroupTheory/CommutingProbability.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ import Mathlib.GroupTheory.Abelianization +import Mathlib.GroupTheory.GroupAction.CardCommute import Mathlib.GroupTheory.SpecificGroups.Dihedral import Mathlib.Tactic.LinearCombination import Mathlib.Tactic.Qify diff --git a/Mathlib/GroupTheory/Coset.lean b/Mathlib/GroupTheory/Coset/Basic.lean similarity index 94% rename from Mathlib/GroupTheory/Coset.lean rename to Mathlib/GroupTheory/Coset/Basic.lean index a4c0c84b15acb..b25a17f71817f 100644 --- a/Mathlib/GroupTheory/Coset.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -7,7 +7,6 @@ import Mathlib.Algebra.Quotient import Mathlib.Algebra.Group.Subgroup.Actions import Mathlib.Algebra.Group.Subgroup.MulOpposite import Mathlib.GroupTheory.GroupAction.Basic -import Mathlib.SetTheory.Cardinal.Finite /-! # Cosets @@ -40,6 +39,7 @@ If instead `G` is an additive group, we can write (with `open scoped Pointwise` Properly merge with pointwise actions on sets, by renaming and deduplicating lemmas as appropriate. -/ +assert_not_exists Cardinal open Function MulOpposite Set open scoped Pointwise @@ -609,41 +609,6 @@ theorem quotientiInfEmbedding_apply_mk {ι : Type*} (f : ι → Subgroup α) (g quotientiInfEmbedding f (QuotientGroup.mk g) i = QuotientGroup.mk g := rfl -@[to_additive AddSubgroup.card_eq_card_quotient_mul_card_addSubgroup] -theorem card_eq_card_quotient_mul_card_subgroup (s : Subgroup α) : - Nat.card α = Nat.card (α ⧸ s) * Nat.card s := by - rw [← Nat.card_prod]; exact Nat.card_congr Subgroup.groupEquivQuotientProdSubgroup - -/-- **Lagrange's Theorem**: The order of a subgroup divides the order of its ambient group. -/ -@[to_additive "**Lagrange's Theorem**: The order of an additive subgroup divides the order of its - ambient additive group."] -theorem card_subgroup_dvd_card (s : Subgroup α) : Nat.card s ∣ Nat.card α := by - classical simp [card_eq_card_quotient_mul_card_subgroup s, @dvd_mul_left ℕ] - -@[to_additive] -theorem card_quotient_dvd_card (s : Subgroup α) : Nat.card (α ⧸ s) ∣ Nat.card α := by - simp [card_eq_card_quotient_mul_card_subgroup s, @dvd_mul_right ℕ] - -variable {H : Type*} [Group H] - -@[to_additive] -theorem card_dvd_of_injective (f : α →* H) (hf : Function.Injective f) : - Nat.card α ∣ Nat.card H := by - classical calc - Nat.card α = Nat.card (f.range : Subgroup H) := Nat.card_congr (Equiv.ofInjective f hf) - _ ∣ Nat.card H := card_subgroup_dvd_card _ - -@[to_additive] -theorem card_dvd_of_le {H K : Subgroup α} (hHK : H ≤ K) : Nat.card H ∣ Nat.card K := - card_dvd_of_injective (inclusion hHK) (inclusion_injective hHK) - -@[to_additive] -theorem card_comap_dvd_of_injective (K : Subgroup H) (f : α →* H) - (hf : Function.Injective f) : Nat.card (K.comap f) ∣ Nat.card K := - calc Nat.card (K.comap f) = Nat.card ((K.comap f).map f) := - Nat.card_congr (equivMapOfInjective _ _ hf).toEquiv - _ ∣ Nat.card K := card_dvd_of_le (map_comap_le _ _) - end Subgroup namespace MonoidHom diff --git a/Mathlib/GroupTheory/Coset/Card.lean b/Mathlib/GroupTheory/Coset/Card.lean new file mode 100644 index 0000000000000..79fd105a82f88 --- /dev/null +++ b/Mathlib/GroupTheory/Coset/Card.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2018 Mitchell Rowett. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mitchell Rowett, Scott Morrison +-/ +import Mathlib.GroupTheory.Coset.Basic +import Mathlib.SetTheory.Cardinal.Finite + +/-! +# Lagrange's theorem: the order of a subgroup divides the order of the group. +-/ + +namespace Subgroup + +variable {α : Type*} [Group α] + +@[to_additive AddSubgroup.card_eq_card_quotient_mul_card_addSubgroup] +theorem card_eq_card_quotient_mul_card_subgroup (s : Subgroup α) : + Nat.card α = Nat.card (α ⧸ s) * Nat.card s := by + rw [← Nat.card_prod]; exact Nat.card_congr Subgroup.groupEquivQuotientProdSubgroup + +/-- **Lagrange's Theorem**: The order of a subgroup divides the order of its ambient group. -/ +@[to_additive "**Lagrange's Theorem**: The order of an additive subgroup divides the order of its + ambient additive group."] +theorem card_subgroup_dvd_card (s : Subgroup α) : Nat.card s ∣ Nat.card α := by + classical simp [card_eq_card_quotient_mul_card_subgroup s, @dvd_mul_left ℕ] + +@[to_additive] +theorem card_quotient_dvd_card (s : Subgroup α) : Nat.card (α ⧸ s) ∣ Nat.card α := by + simp [card_eq_card_quotient_mul_card_subgroup s, @dvd_mul_right ℕ] + +variable {H : Type*} [Group H] + +@[to_additive] +theorem card_dvd_of_injective (f : α →* H) (hf : Function.Injective f) : + Nat.card α ∣ Nat.card H := by + classical calc + Nat.card α = Nat.card (f.range : Subgroup H) := Nat.card_congr (Equiv.ofInjective f hf) + _ ∣ Nat.card H := card_subgroup_dvd_card _ + +@[to_additive] +theorem card_dvd_of_le {H K : Subgroup α} (hHK : H ≤ K) : Nat.card H ∣ Nat.card K := + card_dvd_of_injective (inclusion hHK) (inclusion_injective hHK) + +@[to_additive] +theorem card_comap_dvd_of_injective (K : Subgroup H) (f : α →* H) + (hf : Function.Injective f) : Nat.card (K.comap f) ∣ Nat.card K := + calc Nat.card (K.comap f) = Nat.card ((K.comap f).map f) := + Nat.card_congr (equivMapOfInjective _ _ hf).toEquiv + _ ∣ Nat.card K := card_dvd_of_le (map_comap_le _ _) + +end Subgroup diff --git a/Mathlib/GroupTheory/Coxeter/Basic.lean b/Mathlib/GroupTheory/Coxeter/Basic.lean index 5fc6d5c513a62..fdbcf71b5c51d 100644 --- a/Mathlib/GroupTheory/Coxeter/Basic.lean +++ b/Mathlib/GroupTheory/Coxeter/Basic.lean @@ -8,6 +8,7 @@ import Mathlib.GroupTheory.PresentedGroup import Mathlib.GroupTheory.Coxeter.Matrix import Mathlib.Tactic.Ring import Mathlib.Tactic.Use +import Mathlib.Tactic.NormNum.DivMod /-! # Coxeter groups and Coxeter systems diff --git a/Mathlib/GroupTheory/Divisible.lean b/Mathlib/GroupTheory/Divisible.lean index 1dc47a2cc80bd..4276c6c3fd8b1 100644 --- a/Mathlib/GroupTheory/Divisible.lean +++ b/Mathlib/GroupTheory/Divisible.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import Mathlib.Algebra.Group.ULift -import Mathlib.GroupTheory.QuotientGroup +import Mathlib.GroupTheory.QuotientGroup.Basic +import Mathlib.Tactic.NormNum.Eq /-! # Divisible Group and rootable group diff --git a/Mathlib/GroupTheory/DoubleCoset.lean b/Mathlib/GroupTheory/DoubleCoset.lean index 4b564ca369672..1d710162e3e55 100644 --- a/Mathlib/GroupTheory/DoubleCoset.lean +++ b/Mathlib/GroupTheory/DoubleCoset.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Birkbeck -/ import Mathlib.Algebra.Group.Subgroup.Pointwise -import Mathlib.GroupTheory.Coset +import Mathlib.GroupTheory.Coset.Basic /-! # Double cosets diff --git a/Mathlib/GroupTheory/Finiteness.lean b/Mathlib/GroupTheory/Finiteness.lean index 86ceab0ad5626..6d7da785053a0 100644 --- a/Mathlib/GroupTheory/Finiteness.lean +++ b/Mathlib/GroupTheory/Finiteness.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ import Mathlib.Data.Set.Pointwise.Finite -import Mathlib.GroupTheory.QuotientGroup +import Mathlib.GroupTheory.QuotientGroup.Basic +import Mathlib.SetTheory.Cardinal.Finite /-! # Finitely generated monoids and groups diff --git a/Mathlib/GroupTheory/GroupAction/CardCommute.lean b/Mathlib/GroupTheory/GroupAction/CardCommute.lean new file mode 100644 index 0000000000000..4be202b059262 --- /dev/null +++ b/Mathlib/GroupTheory/GroupAction/CardCommute.lean @@ -0,0 +1,32 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Thomas Browning +-/ +import Mathlib.Algebra.Group.ConjFinite +import Mathlib.GroupTheory.GroupAction.Quotient + +/-! +# A consequence of Burnside's lemma + +See `Mathlib.GroupTheory.GroupAction.Quotient` for Burnside's lemma itself. +This lemma is separate because it requires `Nat.card` +and hence transitively the development of cardinals. +-/ + +open Fintype + +theorem card_comm_eq_card_conjClasses_mul_card (G : Type*) [Group G] : + Nat.card { p : G × G // Commute p.1 p.2 } = Nat.card (ConjClasses G) * Nat.card G := by + classical + rcases fintypeOrInfinite G; swap + · rw [mul_comm, Nat.card_eq_zero_of_infinite, Nat.card_eq_zero_of_infinite, zero_mul] + simp only [Nat.card_eq_fintype_card] + -- Porting note: Changed `calc` proof into a `rw` proof. + rw [card_congr (Equiv.subtypeProdEquivSigmaSubtype Commute), card_sigma, + sum_equiv ConjAct.toConjAct.toEquiv (fun a ↦ card { b // Commute a b }) + (fun g ↦ card (MulAction.fixedBy G g)) + fun g ↦ card_congr' <| congr_arg _ <| funext fun h ↦ mul_inv_eq_iff_eq_mul.symm.eq, + MulAction.sum_card_fixedBy_eq_card_orbits_mul_card_group] + congr 1; apply card_congr'; congr; ext + exact (Setoid.comm' _).trans isConj_iff.symm diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index a03534bfbc10f..8a484cbcfe58d 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -3,13 +3,12 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Thomas Browning -/ -import Mathlib.Algebra.Group.ConjFinite import Mathlib.Data.Fintype.BigOperators import Mathlib.Dynamics.PeriodicPts -import Mathlib.GroupTheory.Commutator -import Mathlib.GroupTheory.Coset import Mathlib.GroupTheory.GroupAction.ConjAct import Mathlib.GroupTheory.GroupAction.Hom +import Mathlib.GroupTheory.Coset.Basic +import Mathlib.GroupTheory.Commutator.Basic /-! # Properties of group actions involving quotient groups @@ -20,7 +19,6 @@ This file proves properties of group actions which use the quotient group constr * Burnside's lemma `sum_card_fixedBy_eq_card_orbits_mul_card_group` -/ - universe u v w variable {α : Type u} {β : Type v} {γ : Type w} @@ -422,24 +420,3 @@ theorem quotientCenterEmbedding_apply {S : Set G} (hS : closure S = ⊤) (g : G) rfl end Subgroup - -section conjClasses - -open Fintype - -theorem card_comm_eq_card_conjClasses_mul_card (G : Type*) [Group G] : - Nat.card { p : G × G // Commute p.1 p.2 } = Nat.card (ConjClasses G) * Nat.card G := by - classical - rcases fintypeOrInfinite G; swap - · rw [mul_comm, Nat.card_eq_zero_of_infinite, Nat.card_eq_zero_of_infinite, zero_mul] - simp only [Nat.card_eq_fintype_card] - -- Porting note: Changed `calc` proof into a `rw` proof. - rw [card_congr (Equiv.subtypeProdEquivSigmaSubtype Commute), card_sigma, - sum_equiv ConjAct.toConjAct.toEquiv (fun a ↦ card { b // Commute a b }) - (fun g ↦ card (MulAction.fixedBy G g)) - fun g ↦ card_congr' <| congr_arg _ <| funext fun h ↦ mul_inv_eq_iff_eq_mul.symm.eq, - MulAction.sum_card_fixedBy_eq_card_orbits_mul_card_group] - congr 1; apply card_congr'; congr; ext - exact (Setoid.comm' _).trans isConj_iff.symm - -end conjClasses diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index a389d1051072c..ac23f4a5dddfc 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -3,7 +3,9 @@ Copyright (c) 2021 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ +import Mathlib.Data.Finite.Card import Mathlib.Algebra.BigOperators.GroupWithZero.Finset +import Mathlib.GroupTheory.Coset.Card import Mathlib.GroupTheory.Finiteness import Mathlib.GroupTheory.GroupAction.Quotient diff --git a/Mathlib/GroupTheory/Nilpotent.lean b/Mathlib/GroupTheory/Nilpotent.lean index 290faebdaff6e..a4b435fb7cba6 100644 --- a/Mathlib/GroupTheory/Nilpotent.lean +++ b/Mathlib/GroupTheory/Nilpotent.lean @@ -6,6 +6,7 @@ Authors: Kevin Buzzard, Ines Wright, Joachim Breitner import Mathlib.GroupTheory.Solvable import Mathlib.GroupTheory.Sylow import Mathlib.Algebra.Group.Subgroup.Order +import Mathlib.GroupTheory.Commutator.Finite /-! diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index f50a3b8ac4190..9cb5066d29da6 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Julian Kuelshammer import Mathlib.Algebra.CharP.Defs import Mathlib.GroupTheory.Index import Mathlib.Order.Interval.Set.Infinite +import Mathlib.Algebra.Group.Subgroup.Finite /-! # Order of an element diff --git a/Mathlib/GroupTheory/PresentedGroup.lean b/Mathlib/GroupTheory/PresentedGroup.lean index 184b89f21a66b..492c52f336a3f 100644 --- a/Mathlib/GroupTheory/PresentedGroup.lean +++ b/Mathlib/GroupTheory/PresentedGroup.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Howes, Newell Jensen -/ import Mathlib.GroupTheory.FreeGroup.Basic -import Mathlib.GroupTheory.QuotientGroup +import Mathlib.GroupTheory.QuotientGroup.Basic /-! # Defining a group given by generators and relations diff --git a/Mathlib/GroupTheory/QuotientGroup.lean b/Mathlib/GroupTheory/QuotientGroup/Basic.lean similarity index 93% rename from Mathlib/GroupTheory/QuotientGroup.lean rename to Mathlib/GroupTheory/QuotientGroup/Basic.lean index b1570f17976f1..8989b6647b1d3 100644 --- a/Mathlib/GroupTheory/QuotientGroup.lean +++ b/Mathlib/GroupTheory/QuotientGroup/Basic.lean @@ -5,10 +5,9 @@ Authors: Kevin Buzzard, Patrick Massot This file is to a certain extent based on `quotient_module.lean` by Johannes Hölzl. -/ -import Mathlib.Algebra.Group.Subgroup.Finite import Mathlib.Algebra.Group.Subgroup.Pointwise import Mathlib.GroupTheory.Congruence.Basic -import Mathlib.GroupTheory.Coset +import Mathlib.GroupTheory.Coset.Basic /-! # Quotients of groups by normal subgroups @@ -611,38 +610,3 @@ theorem mk_int_mul (n : ℤ) (a : R) : ((n * a : R) : R ⧸ N) = n • ↑a := b rw [← zsmul_eq_mul, mk_zsmul N a n] end QuotientAddGroup - -namespace Group - -open scoped Classical - -open QuotientGroup Subgroup - -variable {F G H : Type u} [Group F] [Group G] [Group H] [Fintype F] [Fintype H] -variable (f : F →* G) (g : G →* H) - -/-- If `F` and `H` are finite such that `ker(G →* H) ≤ im(F →* G)`, then `G` is finite. -/ -@[to_additive "If `F` and `H` are finite such that `ker(G →+ H) ≤ im(F →+ G)`, then `G` is finite."] -noncomputable def fintypeOfKerLeRange (h : g.ker ≤ f.range) : Fintype G := - @Fintype.ofEquiv _ _ - (@instFintypeProd _ _ (Fintype.ofInjective _ <| kerLift_injective g) <| - Fintype.ofInjective _ <| inclusion_injective h) - groupEquivQuotientProdSubgroup.symm - -/-- If `F` and `H` are finite such that `ker(G →* H) = im(F →* G)`, then `G` is finite. -/ -@[to_additive "If `F` and `H` are finite such that `ker(G →+ H) = im(F →+ G)`, then `G` is finite."] -noncomputable def fintypeOfKerEqRange (h : g.ker = f.range) : Fintype G := - fintypeOfKerLeRange _ _ h.le - -/-- If `ker(G →* H)` and `H` are finite, then `G` is finite. -/ -@[to_additive "If `ker(G →+ H)` and `H` are finite, then `G` is finite."] -noncomputable def fintypeOfKerOfCodom [Fintype g.ker] : Fintype G := - fintypeOfKerLeRange ((topEquiv : _ ≃* G).toMonoidHom.comp <| inclusion le_top) g fun x hx => - ⟨⟨x, hx⟩, rfl⟩ - -/-- If `F` and `coker(F →* G)` are finite, then `G` is finite. -/ -@[to_additive "If `F` and `coker(F →+ G)` are finite, then `G` is finite."] -noncomputable def fintypeOfDomOfCoker [Normal f.range] [Fintype <| G ⧸ f.range] : Fintype G := - fintypeOfKerLeRange _ (mk' f.range) fun x => (eq_one_iff x).mp - -end Group diff --git a/Mathlib/GroupTheory/QuotientGroup/Finite.lean b/Mathlib/GroupTheory/QuotientGroup/Finite.lean new file mode 100644 index 0000000000000..83c6e3b290b14 --- /dev/null +++ b/Mathlib/GroupTheory/QuotientGroup/Finite.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2018 Kevin Buzzard, Patrick Massot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kevin Buzzard, Patrick Massot + +This file is to a certain extent based on `quotient_module.lean` by Johannes Hölzl. +-/ +import Mathlib.Algebra.Group.Subgroup.Finite +import Mathlib.GroupTheory.QuotientGroup.Basic + +/-! +# Deducing finiteness of a group. +-/ + +open Function +open scoped Pointwise + +universe u v w x + +namespace Group + +open scoped Classical + +open QuotientGroup Subgroup + +variable {F G H : Type u} [Group F] [Group G] [Group H] [Fintype F] [Fintype H] +variable (f : F →* G) (g : G →* H) + +/-- If `F` and `H` are finite such that `ker(G →* H) ≤ im(F →* G)`, then `G` is finite. -/ +@[to_additive "If `F` and `H` are finite such that `ker(G →+ H) ≤ im(F →+ G)`, then `G` is finite."] +noncomputable def fintypeOfKerLeRange (h : g.ker ≤ f.range) : Fintype G := + @Fintype.ofEquiv _ _ + (@instFintypeProd _ _ (Fintype.ofInjective _ <| kerLift_injective g) <| + Fintype.ofInjective _ <| inclusion_injective h) + groupEquivQuotientProdSubgroup.symm + +/-- If `F` and `H` are finite such that `ker(G →* H) = im(F →* G)`, then `G` is finite. -/ +@[to_additive "If `F` and `H` are finite such that `ker(G →+ H) = im(F →+ G)`, then `G` is finite."] +noncomputable def fintypeOfKerEqRange (h : g.ker = f.range) : Fintype G := + fintypeOfKerLeRange _ _ h.le + +/-- If `ker(G →* H)` and `H` are finite, then `G` is finite. -/ +@[to_additive "If `ker(G →+ H)` and `H` are finite, then `G` is finite."] +noncomputable def fintypeOfKerOfCodom [Fintype g.ker] : Fintype G := + fintypeOfKerLeRange ((topEquiv : _ ≃* G).toMonoidHom.comp <| inclusion le_top) g fun x hx => + ⟨⟨x, hx⟩, rfl⟩ + +/-- If `F` and `coker(F →* G)` are finite, then `G` is finite. -/ +@[to_additive "If `F` and `coker(F →+ G)` are finite, then `G` is finite."] +noncomputable def fintypeOfDomOfCoker [Normal f.range] [Fintype <| G ⧸ f.range] : Fintype G := + fintypeOfKerLeRange _ (mk' f.range) fun x => (eq_one_iff x).mp + +end Group diff --git a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean index 4836d6d64026d..200e6d06c39ac 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean @@ -3,10 +3,10 @@ Copyright (c) 2021 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ -import Mathlib.Algebra.Group.ConjFinite import Mathlib.GroupTheory.Perm.Fin import Mathlib.GroupTheory.Subgroup.Simple import Mathlib.Tactic.IntervalCases +import Mathlib.Data.Fintype.Units /-! # Alternating Groups diff --git a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean index 26ff5fc37728e..8cbe12005aff0 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean @@ -5,6 +5,7 @@ Authors: Shing Tak Lam -/ import Mathlib.Data.ZMod.Basic import Mathlib.GroupTheory.Exponent +import Mathlib.GroupTheory.GroupAction.CardCommute /-! # Dihedral Groups diff --git a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean index deb0d98dc3cb3..ce50529aa2ab6 100644 --- a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean +++ b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean @@ -6,6 +6,7 @@ Authors: Eric Wieser import Mathlib.LinearAlgebra.Alternating.Basic import Mathlib.LinearAlgebra.Multilinear.TensorProduct import Mathlib.GroupTheory.GroupAction.Quotient +import Mathlib.Algebra.Group.Subgroup.Finite /-! # Exterior product of alternating maps diff --git a/Mathlib/LinearAlgebra/Quotient.lean b/Mathlib/LinearAlgebra/Quotient.lean index f19649d0da829..61d6fd9576a2a 100644 --- a/Mathlib/LinearAlgebra/Quotient.lean +++ b/Mathlib/LinearAlgebra/Quotient.lean @@ -3,9 +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, Mario Carneiro, Kevin Buzzard, Yury Kudryashov -/ -import Mathlib.GroupTheory.QuotientGroup import Mathlib.LinearAlgebra.Span import Mathlib.Algebra.Module.Equiv.Basic +import Mathlib.GroupTheory.QuotientGroup.Basic +import Mathlib.SetTheory.Cardinal.Finite /-! # Quotients by submodules diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 7d3dad830959e..72f27d9b68033 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -5,12 +5,12 @@ Authors: Johannes Hölzl, Mario Carneiro -/ import Mathlib.Data.Finset.Update import Mathlib.Data.Prod.TProd -import Mathlib.GroupTheory.Coset import Mathlib.Logic.Equiv.Fin import Mathlib.MeasureTheory.MeasurableSpace.Instances import Mathlib.Order.LiminfLimsup import Mathlib.Data.Set.UnionLift import Mathlib.Order.Filter.SmallSets +import Mathlib.GroupTheory.Coset.Basic /-! # Measurable spaces and measurable functions diff --git a/Mathlib/RingTheory/ClassGroup.lean b/Mathlib/RingTheory/ClassGroup.lean index 9037779c7cd2a..1fee05b58a3b4 100644 --- a/Mathlib/RingTheory/ClassGroup.lean +++ b/Mathlib/RingTheory/ClassGroup.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ -import Mathlib.GroupTheory.QuotientGroup import Mathlib.RingTheory.DedekindDomain.Ideal /-! diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index 8b8f83b53c3bc..36c67204cd37c 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -5,10 +5,10 @@ Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import Mathlib.GroupTheory.GroupAction.ConjAct import Mathlib.GroupTheory.GroupAction.Quotient -import Mathlib.GroupTheory.QuotientGroup import Mathlib.Topology.Algebra.Monoid import Mathlib.Topology.Algebra.Constructions import Mathlib.Algebra.Order.Archimedean.Basic +import Mathlib.GroupTheory.QuotientGroup.Basic /-! # Topological groups diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean index b21eaf9c595e2..06be853f4ac9d 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ +import Mathlib.SetTheory.Cardinal.Finite import Mathlib.Topology.Algebra.InfiniteSum.Basic import Mathlib.Topology.Algebra.UniformGroup diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean b/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean index 953e8155df310..30d1d7e493bc0 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Nonarchimedean.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Mitchell Lee. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mitchell Lee -/ +import Mathlib.Algebra.Group.Subgroup.Finite import Mathlib.Topology.Algebra.InfiniteSum.Group import Mathlib.Topology.Algebra.Nonarchimedean.Basic diff --git a/Mathlib/Topology/Algebra/Ring/Basic.lean b/Mathlib/Topology/Algebra/Ring/Basic.lean index 84e24b6c25393..625d49d63a27c 100644 --- a/Mathlib/Topology/Algebra/Ring/Basic.lean +++ b/Mathlib/Topology/Algebra/Ring/Basic.lean @@ -25,6 +25,7 @@ of topological (semi)rings. - The indexed product of topological (semi)rings is a topological (semi)ring. -/ +assert_not_exists Cardinal open Set Filter TopologicalSpace Function Topology Filter From 45d8a0f51831d393d554abe71b02d972e488c190 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 16 Aug 2024 23:11:52 +0000 Subject: [PATCH 349/359] chore: clean up List/Indexes (#15825) There had been a bit of a mess here of shims connecting back to a Lean 3 era definition. Rip it all out, replacing the proofs that required these shims, and deprecate the now unneeded mess. --- Mathlib/Data/List/Basic.lean | 4 + Mathlib/Data/List/Indexes.lean | 186 ++++++++++++++++++++------------- 2 files changed, 115 insertions(+), 75 deletions(-) diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 8bce8077ffc81..46c4e404ea02c 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -41,6 +41,10 @@ theorem le_eq_not_gt [LT α] : ∀ l₁ l₂ : List α, (l₁ ≤ l₂) = ¬l₂ -- Porting note: Delete this attribute -- attribute [inline] List.head! +theorem getElem?_eq (l : List α) (i : Nat) : + l[i]? = if h : i < l.length then some l[i] else none := by + split <;> simp_all + /-- There is only one list of an empty type -/ instance uniqueOfIsEmpty [IsEmpty α] : Unique (List α) := { instInhabitedList with diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index 6da70619cfb4b..a9af61d933c06 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -25,34 +25,11 @@ variable {α : Type u} {β : Type v} section MapIdx --- Porting note: Add back old definition because it's easier for writing proofs. - -/-- Lean3 `map_with_index` helper function -/ -protected def oldMapIdxCore (f : ℕ → α → β) : ℕ → List α → List β - | _, [] => [] - | k, a :: as => f k a :: List.oldMapIdxCore f (k + 1) as - -/-- Given a function `f : ℕ → α → β` and `as : List α`, `as = [a₀, a₁, ...]`, returns the list -`[f 0 a₀, f 1 a₁, ...]`. -/ -protected def oldMapIdx (f : ℕ → α → β) (as : List α) : List β := - List.oldMapIdxCore f 0 as - @[simp] theorem mapIdx_nil {α β} (f : ℕ → α → β) : mapIdx f [] = [] := rfl -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 - · rfl - · rw [List.oldMapIdx] - simp only [List.oldMapIdxCore, hl, Nat.add_left_comm, Nat.add_comm, Nat.add_zero] --- Porting note: convert new definition to old definition. --- A few new theorems are added to achieve this --- 1. Prove that `oldMapIdxCore f (l ++ [e]) = oldMapIdxCore f l ++ [f l.length e]` --- 2. Prove that `oldMapIdx f (l ++ [e]) = oldMapIdx f l ++ [f l.length e]` --- 3. Prove list induction using `∀ l e, p [] → (p l → p (l ++ [e])) → p l` theorem list_reverse_induction (p : List α → Prop) (base : p []) (ind : ∀ (l : List α) (e : α), p l → p (l ++ [e])) : (∀ (l : List α), p l) := by let q := fun l ↦ p (reverse l) @@ -65,39 +42,6 @@ theorem list_reverse_induction (p : List α → Prop) (base : p []) · apply pq; simp only [reverse_nil, base] · apply pq; simp only [reverse_cons]; apply ind; apply qp; rw [reverse_reverse]; exact ih -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 - intros f n l₁ l₂ - generalize e : (l₁ ++ l₂).length = len - revert n l₁ l₂ - induction' len with len ih <;> intros n l₁ l₂ h - · have l₁_nil : l₁ = [] := by - cases l₁ - · rfl - · contradiction - have l₂_nil : l₂ = [] := by - cases l₂ - · rfl - · rw [List.length_append] at h; contradiction - simp only [l₁_nil, l₂_nil]; rfl - · cases' l₁ with head tail - · rfl - · simp only [List.oldMapIdxCore, List.append_eq, length_cons, cons_append,cons.injEq, true_and] - suffices n + Nat.succ (length tail) = n + 1 + tail.length by - rw [this] - apply ih (n + 1) _ _ _ - simp only [cons_append, length_cons, length_append, Nat.succ.injEq] at h - simp only [length_append, h] - rw [Nat.add_assoc]; simp only [Nat.add_comm] - -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 - unfold List.oldMapIdx - rw [List.oldMapIdxCore_append f 0 l [e]] - simp only [Nat.zero_add]; rfl - theorem mapIdxGo_append : ∀ (f : ℕ → α → β) (l₁ l₂ : List α) (arr : Array β), mapIdx.go f (l₁ ++ l₂) arr = mapIdx.go f l₂ (List.toArray (mapIdx.go f l₁ arr)) := by intros f l₁ l₂ arr @@ -136,14 +80,6 @@ theorem mapIdx_append_one : ∀ (f : ℕ → α → β) (l : List α) (e : α), simp only [mapIdx.go, Array.size_toArray, mapIdxGo_length, length_nil, Nat.add_zero, Array.toList_eq, Array.push_data, Array.data_toArray] -protected theorem new_def_eq_old_def : - ∀ (f : ℕ → α → β) (l : List α), l.mapIdx f = List.oldMapIdx f l := by - intro f - apply list_reverse_induction - · rfl - · intro l e h - rw [List.oldMapIdx_append, mapIdx_append_one, h] - @[local simp] theorem map_enumFrom_eq_zipWith : ∀ (l : List α) (n : ℕ) (f : ℕ → α → β), map (uncurry f) (enumFrom n l) = zipWith (fun i ↦ f (i + n)) (range (length l)) l := by @@ -168,13 +104,46 @@ theorem map_enumFrom_eq_zipWith : ∀ (l : List α) (n : ℕ) (f : ℕ → α simp only [comp, Nat.add_assoc, Nat.add_comm, Nat.add_succ] simp only [length_cons, Nat.succ.injEq] at e; exact e +theorem length_mapIdx_go (f : ℕ → α → β) : ∀ (l : List α) (arr : Array β), + (mapIdx.go f l arr).length = l.length + arr.size + | [], _ => by simp [mapIdx.go] + | a :: l, _ => by + simp only [mapIdx.go, length_cons] + rw [length_mapIdx_go] + simp + omega + +@[simp] theorem length_mapIdx (l : List α) (f : ℕ → α → β) : (l.mapIdx f).length = l.length := by + simp [mapIdx, length_mapIdx_go] + +theorem getElem?_mapIdx_go (f : ℕ → α → β) : ∀ (l : List α) (arr : Array β) (i : ℕ), + (mapIdx.go f l arr)[i]? = + if h : i < arr.size then some arr[i] else Option.map (f i) l[i - arr.size]? + | [], arr, i => by + simp [mapIdx.go, getElem?_eq, Array.getElem_eq_data_getElem] + | a :: l, arr, i => by + rw [mapIdx.go, getElem?_mapIdx_go] + simp only [Array.size_push] + split <;> split + · simp only [Option.some.injEq] + rw [Array.getElem_eq_data_getElem] + simp only [Array.push_data] + rw [getElem_append_left, Array.getElem_eq_data_getElem] + · have : i = arr.size := by omega + simp_all + · omega + · have : i - arr.size = i - (arr.size + 1) + 1 := by omega + simp_all + +@[simp] theorem getElem?_mapIdx (l : List α) (f : ℕ → α → β) (i : ℕ) : + (l.mapIdx f)[i]? = Option.map (f i) l[i]? := by + simp [mapIdx, getElem?_mapIdx_go] + theorem mapIdx_eq_enum_map (l : List α) (f : ℕ → α → β) : l.mapIdx f = l.enum.map (Function.uncurry f) := by - rw [List.new_def_eq_old_def] - induction' l with hd tl hl generalizing f - · rfl - · rw [List.oldMapIdx, List.oldMapIdxCore, List.oldMapIdxCore_eq, hl] - simp [map, enum_eq_zip_range, map_uncurry_zip_eq_zipWith] + ext1 i + simp only [getElem?_mapIdx, Option.map, getElem?_map, getElem?_enum] + split <;> simp @[simp] theorem mapIdx_cons (l : List α) (f : ℕ → α → β) (a : α) : @@ -188,12 +157,6 @@ theorem mapIdx_append (K L : List α) (f : ℕ → α → β) : · rfl · simp [IH fun i ↦ f (i + 1), Nat.add_assoc] -@[simp] -theorem length_mapIdx (l : List α) (f : ℕ → α → β) : (l.mapIdx f).length = l.length := by - induction' l with hd tl IH generalizing f - · rfl - · simp [IH] - @[simp] theorem mapIdx_eq_nil {f : ℕ → α → β} {l : List α} : List.mapIdx f l = [] ↔ l = [] := by rw [List.mapIdx_eq_enum_map, List.map_eq_nil, List.enum_eq_nil] @@ -211,6 +174,79 @@ theorem mapIdx_eq_ofFn (l : List α) (f : ℕ → α → β) : | nil => simp | cons _ _ IH => simp [IH] +section deprecated + +/-- Lean3 `map_with_index` helper function -/ +@[deprecated (since := "2024-08-15")] +protected def oldMapIdxCore (f : ℕ → α → β) : ℕ → List α → List β + | _, [] => [] + | k, a :: as => f k a :: List.oldMapIdxCore f (k + 1) as + +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")] +protected def oldMapIdx (f : ℕ → α → β) (as : List α) : List β := + List.oldMapIdxCore f 0 as + +set_option linter.deprecated false in +@[deprecated (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 + · rfl + · rw [List.oldMapIdx] + 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")] +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 + intros f n l₁ l₂ + generalize e : (l₁ ++ l₂).length = len + revert n l₁ l₂ + induction' len with len ih <;> intros n l₁ l₂ h + · have l₁_nil : l₁ = [] := by + cases l₁ + · rfl + · contradiction + have l₂_nil : l₂ = [] := by + cases l₂ + · rfl + · rw [List.length_append] at h; contradiction + simp only [l₁_nil, l₂_nil]; rfl + · cases' l₁ with head tail + · rfl + · simp only [List.oldMapIdxCore, List.append_eq, length_cons, cons_append,cons.injEq, true_and] + suffices n + Nat.succ (length tail) = n + 1 + tail.length by + rw [this] + apply ih (n + 1) _ _ _ + simp only [cons_append, length_cons, length_append, Nat.succ.injEq] at h + simp only [length_append, h] + rw [Nat.add_assoc]; simp only [Nat.add_comm] + +set_option linter.deprecated false in +@[deprecated (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 + unfold List.oldMapIdx + rw [List.oldMapIdxCore_append f 0 l [e]] + simp only [Nat.zero_add]; rfl + +set_option linter.deprecated false in +@[deprecated (since := "2024-08-15")] +protected theorem new_def_eq_old_def : + ∀ (f : ℕ → α → β) (l : List α), l.mapIdx f = List.oldMapIdx f l := by + intro f + apply list_reverse_induction + · rfl + · intro l e h + rw [List.oldMapIdx_append, mapIdx_append_one, h] + +end deprecated + end MapIdx section FoldrIdx From e8679e49ac8a467f2439ff3f145dc9266a3154da Mon Sep 17 00:00:00 2001 From: grunweg Date: Fri, 16 Aug 2024 23:11:54 +0000 Subject: [PATCH 350/359] chore(tests): fix linter warnings (#15872) Fix long lines, remove unused tactics and add missing end commands. The remaining changes are silencing linters in `tests`, which cannot be backported. Split out from #15845. --- Mathlib/Data/Option/Basic.lean | 2 -- Mathlib/Tactic/Convert.lean | 3 ++- Mathlib/Tactic/FunProp/Core.lean | 3 ++- Mathlib/Tactic/FunProp/Types.lean | 4 ++-- Mathlib/Tactic/Linter/TextBased.lean | 3 ++- test/Clean.lean | 2 ++ test/DefEqTransformations.lean | 2 ++ test/InferParam.lean | 2 ++ test/TermCongr.lean | 2 ++ test/Use.lean | 2 ++ test/convert.lean | 2 ++ test/notation3.lean | 2 ++ 12 files changed, 22 insertions(+), 7 deletions(-) diff --git a/Mathlib/Data/Option/Basic.lean b/Mathlib/Data/Option/Basic.lean index db3660d4695ba..533abe5270de7 100644 --- a/Mathlib/Data/Option/Basic.lean +++ b/Mathlib/Data/Option/Basic.lean @@ -355,7 +355,6 @@ theorem elim_apply {f : γ → α → β} {x : α → β} {i : Option γ} {y : @[simp] lemma bnot_isSome (a : Option α) : (! a.isSome) = a.isNone := by - funext cases a <;> simp @[simp] @@ -365,7 +364,6 @@ lemma bnot_comp_isSome : (! ·) ∘ @Option.isSome α = Option.isNone := by @[simp] lemma bnot_isNone (a : Option α) : (! a.isNone) = a.isSome := by - funext cases a <;> simp @[simp] diff --git a/Mathlib/Tactic/Convert.lean b/Mathlib/Tactic/Convert.lean index bd9219b70db22..aa2d570e1a2a9 100644 --- a/Mathlib/Tactic/Convert.lean +++ b/Mathlib/Tactic/Convert.lean @@ -46,7 +46,8 @@ def Lean.MVarId.convertLocalDecl (g : MVarId) (fvarId : FVarId) (typeNew : Expr) (patterns : List (TSyntax `rcasesPat) := []) : MetaM (MVarId × List MVarId) := g.withContext do let typeOld ← fvarId.getType - let v ← mkFreshExprMVar (← mkAppM ``Eq (if symm then #[typeNew, typeOld] else #[typeOld, typeNew])) + let v ← mkFreshExprMVar (← mkAppM ``Eq + (if symm then #[typeNew, typeOld] else #[typeOld, typeNew])) let pf ← if symm then mkEqSymm v else pure v let res ← g.replaceLocalDecl fvarId typeNew pf let gs ← v.mvarId!.congrN! depth config patterns diff --git a/Mathlib/Tactic/FunProp/Core.lean b/Mathlib/Tactic/FunProp/Core.lean index be950365215c5..03d0b6b5c6a57 100644 --- a/Mathlib/Tactic/FunProp/Core.lean +++ b/Mathlib/Tactic/FunProp/Core.lean @@ -39,7 +39,8 @@ sythesized value{indentExpr val}\nis not definitionally equal to{indentExpr x}" return false -/-- Synthesize arguments `xs` either with typeclass synthesis, with `fun_prop` or with discharger. -/ +/-- Synthesize arguments `xs` either with typeclass synthesis, +with `fun_prop` or with a discharger. -/ def synthesizeArgs (thmId : Origin) (xs : Array Expr) (bis : Array BinderInfo) (funProp : Expr → FunPropM (Option Result)) : FunPropM Bool := do diff --git a/Mathlib/Tactic/FunProp/Types.lean b/Mathlib/Tactic/FunProp/Types.lean index 3d8330020fa34..27036b17bce46 100644 --- a/Mathlib/Tactic/FunProp/Types.lean +++ b/Mathlib/Tactic/FunProp/Types.lean @@ -152,8 +152,8 @@ Messages are logged only when `transitionDepth = 0` i.e. when `fun_prop` is **no function property like continuity from another property like differentiability. The main reason is that if the user forgets to add a continuity theorem for function `foo` then `fun_prop` should report that there is a continuity theorem for `foo` missing. If we would log -messages `transitionDepth > 0` then user will see messages saying that there is a missing theorem for -differentiability, smoothness, ... for `foo`. -/ +messages `transitionDepth > 0` then user will see messages saying that there is a missing theorem +for differentiability, smoothness, ... for `foo`. -/ def logError (msg : String) : FunPropM Unit := do if (← read).transitionDepth = 0 then modify fun s => diff --git a/Mathlib/Tactic/Linter/TextBased.lean b/Mathlib/Tactic/Linter/TextBased.lean index 544f44eb1d63c..3da71f0a028bd 100644 --- a/Mathlib/Tactic/Linter/TextBased.lean +++ b/Mathlib/Tactic/Linter/TextBased.lean @@ -410,7 +410,8 @@ def lintFile (path : FilePath) (sizeLimit : Option ℕ) (exceptions : Array Erro let allOutput := (Array.map (fun lint ↦ (Array.map (fun (e, n) ↦ ErrorContext.mk e n path)) (lint lines))) allLinters -- This this list is not sorted: for github, this is fine. - errors := errors.append (allOutput.flatten.filter (fun e ↦ (e.find?_comparable exceptions).isNone)) + errors := errors.append + (allOutput.flatten.filter (fun e ↦ (e.find?_comparable exceptions).isNone)) return errors /-- Lint a collection of modules for style violations. diff --git a/test/Clean.lean b/test/Clean.lean index 8e6bab55aaed1..ec96dc00e6179 100644 --- a/test/Clean.lean +++ b/test/Clean.lean @@ -51,3 +51,5 @@ example : True := by guard_hyp z :ₛ Nat := let_fun x := 1; x + x trivial + +end Tests diff --git a/test/DefEqTransformations.lean b/test/DefEqTransformations.lean index 9ae2b7dc6e3f5..dc01755e7f38c 100644 --- a/test/DefEqTransformations.lean +++ b/test/DefEqTransformations.lean @@ -182,3 +182,5 @@ example (n : Fin 5) : n = ⟨n.val2, n.prop2⟩ := by eta_struct guard_target =ₛ n = n rfl + +end Tests diff --git a/test/InferParam.lean b/test/InferParam.lean index d8f9a84f68a5c..835574cb0dec9 100644 --- a/test/InferParam.lean +++ b/test/InferParam.lean @@ -19,3 +19,5 @@ example : 0 ≤ 2 + 2 := by example : 0 ≤ 2 + 2 := by apply zero_le_add' infer_param + +end InferParamTest diff --git a/test/TermCongr.lean b/test/TermCongr.lean index 0f3376e1efaa0..c429d5c4d916a 100644 --- a/test/TermCongr.lean +++ b/test/TermCongr.lean @@ -160,3 +160,5 @@ example {s t : α → Prop} (h : s = t) (p : α → Prop) : congr(∀ (n : Subtype $h), p n) end limitations + +end Tests diff --git a/test/Use.lean b/test/Use.lean index 9dd4b75d1c145..a82114dfd2b85 100644 --- a/test/Use.lean +++ b/test/Use.lean @@ -227,3 +227,5 @@ example (h1 : 1 > 0) : ∃ (n : Nat) (_h : n > 0), n = n := by example : let P : Nat → Prop := fun _x => ∃ _n : Nat, True; P 1 := by intro P use 1 + +end UseTests diff --git a/test/convert.lean b/test/convert.lean index d260bb7ec431c..f00e2c3f65a52 100644 --- a/test/convert.lean +++ b/test/convert.lean @@ -125,3 +125,5 @@ example (x y z : Nat) (h : x + y = z) : y + x = z := by convert_to y + x = _ at h · rw [Nat.add_comm] exact h + +end Tests diff --git a/test/notation3.lean b/test/notation3.lean index 44116cca2ce50..eecf73b2e31a5 100644 --- a/test/notation3.lean +++ b/test/notation3.lean @@ -202,3 +202,5 @@ notation3 "δNat" => (default : Nat) #guard_msgs in #check (default : Nat) /-- info: δNat : ℕ -/ #guard_msgs in #check @default Nat (Inhabited.mk 5) + +end Test From 0dd5858ed899ab9af9621480ca63655b01548c29 Mon Sep 17 00:00:00 2001 From: grhkm21 Date: Sat, 17 Aug 2024 00:16:30 +0000 Subject: [PATCH 351/359] docs(Algebra/Group/Subgroup): Add note about other induction paths on groups (#15903) I thought `Subgroup.closure_induction_left` didn't exist but it did (in a separate file), so I add a note about it. --- Mathlib/Algebra/Group/Subgroup/Basic.lean | 10 ++++++++-- 1 file changed, 8 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index c307e84d8c382..5ca49efc06b4f 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -877,11 +877,17 @@ theorem closure_eq_of_le (h₁ : k ⊆ K) (h₂ : K ≤ closure k) : closure k = /-- An induction principle for closure membership. If `p` holds for `1` and all elements of `k`, and is preserved under multiplication and inverse, then `p` holds for all elements of the closure -of `k`. -/ +of `k`. + +See also `Subgroup.closure_induction_left` and `Subgroup.closure_induction_right` for versions that +only require showing `p` is preserved by multiplication by elements in `k`. -/ @[to_additive (attr := elab_as_elim) "An induction principle for additive closure membership. If `p` holds for `0` and all elements of `k`, and is preserved under addition and inverses, then `p` - holds for all elements of the additive closure of `k`."] + holds for all elements of the additive closure of `k`. + + See also `AddSubgroup.closure_induction_left` and `AddSubgroup.closure_induction_left` for + versions that only require showing `p` is preserved by addition by elements in `k`."] theorem closure_induction {p : G → Prop} {x} (h : x ∈ closure k) (mem : ∀ x ∈ k, p x) (one : p 1) (mul : ∀ x y, p x → p y → p (x * y)) (inv : ∀ x, p x → p x⁻¹) : p x := (@closure_le _ _ ⟨⟨⟨setOf p, fun {x y} ↦ mul x y⟩, one⟩, fun {x} ↦ inv x⟩ k).2 mem h From fd0dba7f3a926163728a44e4a29507cf14474e3c Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Sat, 17 Aug 2024 01:15:00 +0000 Subject: [PATCH 352/359] feat(ModelTheory): More general ways of deriving `DecidableEq` for logic symbols (#15901) Replaces two `DecidableEq` instances in `ModelTheory/Ring/Basic` with more general instances. Co-authored-by: Aaron Anderson <65780815+awainverse@users.noreply.github.com> --- Mathlib/ModelTheory/Algebra/Ring/Basic.lean | 10 ++++++---- Mathlib/ModelTheory/Basic.lean | 18 ++++++++++++++++++ 2 files changed, 24 insertions(+), 4 deletions(-) diff --git a/Mathlib/ModelTheory/Algebra/Ring/Basic.lean b/Mathlib/ModelTheory/Algebra/Ring/Basic.lean index 7df3ce6ac5bdb..34291af0a1546 100644 --- a/Mathlib/ModelTheory/Algebra/Ring/Basic.lean +++ b/Mathlib/ModelTheory/Algebra/Ring/Basic.lean @@ -64,11 +64,13 @@ namespace Ring open ringFunc Language -instance (n : ℕ) : DecidableEq (Language.ring.Functions n) := by - dsimp [Language.ring]; infer_instance +/-- This instance does not get inferred without `instDecidableEqFunctions` in +`ModelTheory/Basic`. -/ +example (n : ℕ) : DecidableEq (Language.ring.Functions n) := inferInstance -instance (n : ℕ) : DecidableEq (Language.ring.Relations n) := by - dsimp [Language.ring]; infer_instance +/-- This instance does not get inferred without `instDecidableEqRelations` in +`ModelTheory/Basic`. -/ +example (n : ℕ) : DecidableEq (Language.ring.Relations n) := inferInstance /-- `RingFunc.add`, but with the defeq type `Language.ring.Functions 2` instead of `RingFunc 2` -/ diff --git a/Mathlib/ModelTheory/Basic.lean b/Mathlib/ModelTheory/Basic.lean index 58d599986ed9e..e5610d623e310 100644 --- a/Mathlib/ModelTheory/Basic.lean +++ b/Mathlib/ModelTheory/Basic.lean @@ -81,6 +81,12 @@ instance inhabited₂ [h : Inhabited a₂] : Inhabited (Sequence₂ a₀ a₁ a instance {n : ℕ} : IsEmpty (Sequence₂ a₀ a₁ a₂ (n + 3)) := inferInstanceAs (IsEmpty PEmpty) +instance [DecidableEq a₀] [DecidableEq a₁] [DecidableEq a₂] {n : ℕ} : + DecidableEq (Sequence₂ a₀ a₁ a₂ n) := + match n with + | 0 | 1 | 2 => ‹_› + | _ + 3 => inferInstance + @[simp] theorem lift_mk {i : ℕ} : Cardinal.lift.{v,u} #(Sequence₂ a₀ a₁ a₂ i) @@ -240,6 +246,18 @@ theorem card_mk₂ (c f₁ f₂ : Type u) (r₁ r₂ : Type v) : Cardinal.lift.{u} #r₁ + Cardinal.lift.{u} #r₂ := by simp [card_eq_card_functions_add_card_relations, add_assoc] +/-- Passes a `DecidableEq` instance on a type of function symbols through the `Language` +constructor. Despite the fact that this is proven by `inferInstance`, it is still needed - +see the `example`s in `ModelTheory/Ring/Basic`. -/ +instance instDecidableEqFunctions {f : ℕ → Type*} {R : ℕ → Type*} (n : ℕ) [DecidableEq (f n)] : + DecidableEq ((⟨f, R⟩ : Language).Functions n) := inferInstance + +/-- Passes a `DecidableEq` instance on a type of relation symbols through the `Language` +constructor. Despite the fact that this is proven by `inferInstance`, it is still needed - +see the `example`s in `ModelTheory/Ring/Basic`. -/ +instance instDecidableEqRelations {f : ℕ → Type*} {R : ℕ → Type*} (n : ℕ) [DecidableEq (R n)] : + DecidableEq ((⟨f, R⟩ : Language).Relations n) := inferInstance + variable (L) (M : Type w) /-- A first-order structure on a type `M` consists of interpretations of all the symbols in a given From f2a46e0f999a4bc9772e939b6889d3180e024d31 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 17 Aug 2024 03:44:13 +0000 Subject: [PATCH 353/359] chore: bump dependencies (#15908) --- .../Enumerative/Composition.lean | 2 +- Mathlib/Data/List/Basic.lean | 36 ++----------------- lake-manifest.json | 2 +- 3 files changed, 5 insertions(+), 35 deletions(-) diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index 05f69e27409d1..cbe6dc1551920 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -629,7 +629,7 @@ theorem getElem_splitWrtCompositionAux (l : List α) (ns : List ℕ) {i : ℕ} · rw [Nat.add_zero, List.take_zero, sum_nil] simp · simp only [splitWrtCompositionAux, getElem_cons_succ, IH, take, - sum_cons, Nat.add_eq, add_zero, splitAt_eq_take_drop, drop_take, drop_drop] + sum_cons, Nat.add_eq, add_zero, splitAt_eq, drop_take, drop_drop] rw [add_comm (sum _) n, Nat.add_sub_add_left] /-- The `i`-th sublist in the splitting of a list `l` along a composition `c`, is the slice of `l` diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 46c4e404ea02c..e47f6853ce6e7 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1609,39 +1609,9 @@ section SplitAtOn variable (p : α → Bool) (xs ys : List α) (ls : List (List α)) (f : List α → List α) -@[simp] -theorem splitAt_eq_take_drop (n : ℕ) (l : List α) : splitAt n l = (take n l, drop n l) := by - by_cases h : n < l.length <;> rw [splitAt, go_eq_take_drop] - · rw [if_pos h]; rfl - · rw [if_neg h, take_of_length_le <| le_of_not_lt h, drop_eq_nil_of_le <| le_of_not_lt h] -where - go_eq_take_drop (n : ℕ) (l xs : List α) (acc : Array α) : splitAt.go l xs n acc = - if n < xs.length then (acc.toList ++ take n xs, drop n xs) else (l, []) := by - split_ifs with h - · induction n generalizing xs acc with - | zero => - rw [splitAt.go, take, drop, append_nil] - · intros h₁; rw [h₁] at h; contradiction - · intros; contradiction - | succ _ ih => - cases xs with - | nil => contradiction - | cons hd tl => - rw [length] at h - rw [splitAt.go, take, drop, append_cons, Array.toList_eq, ← Array.push_data, - ← Array.toList_eq] - exact ih _ _ <| (by omega) - · induction n generalizing xs acc with - | zero => - replace h : xs.length = 0 := by omega - rw [eq_nil_of_length_eq_zero h, splitAt.go] - | succ _ ih => - cases xs with - | nil => rw [splitAt.go] - | cons hd tl => - rw [length] at h - rw [splitAt.go] - exact ih _ _ <| not_imp_not.mpr (Nat.add_lt_add_right · 1) h +attribute [simp] splitAt_eq + +@[deprecated (since := "2024-08-17")] alias splitAt_eq_take_drop := splitAt_eq @[simp] theorem splitOn_nil [DecidableEq α] (a : α) : [].splitOn a = [[]] := diff --git a/lake-manifest.json b/lake-manifest.json index fdeb2fb3a7531..b1bab701ac2be 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "5e5e54c4028f7b6bd086ebb72e5822794c64c35d", + "rev": "65f464e6a5f30b2c2a58569f530c4677b371cec6", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From c1cb845a81543663e44a4e096bfb297bc0489e5c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Sat, 17 Aug 2024 05:53:55 +0000 Subject: [PATCH 354/359] feat: absolute continuity of `Measure.compProd` (#11273) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For `μ ν : Measure α` and `κ η : Kernel α β`, - If `μ ≪ ν` and `∀ᵐ a ∂μ, κ a ≪ η a`, then `μ ⊗ₘ κ ≪ ν ⊗ₘ η` - If `μ ⊗ₘ κ ≪ ν ⊗ₘ η`, and `κ` is everywhere nonzero then `μ ≪ ν` Also remove a few s-finite assumptions in the `Probability/Kernel/MeasureCompProd` file. --- Mathlib/Probability/Kernel/Composition.lean | 7 +- .../Probability/Kernel/MeasureCompProd.lean | 85 ++++++++++++++++--- 2 files changed, 78 insertions(+), 14 deletions(-) diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 7c4d7a7363b7e..6b68eb2887189 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -301,9 +301,14 @@ theorem ae_ae_of_ae_compProd {p : β × γ → Prop} (h : ∀ᵐ bc ∂(κ ⊗ ∀ᵐ b ∂κ a, ∀ᵐ c ∂η (a, b), p (b, c) := ae_null_of_compProd_null h -lemma ae_compProd_of_ae_ae {p : β × γ → Prop} (hp : MeasurableSet {x | p x}) +lemma ae_compProd_of_ae_ae {κ : Kernel α β} {η : Kernel (α × β) γ} + {p : β × γ → Prop} (hp : MeasurableSet {x | p x}) (h : ∀ᵐ b ∂κ a, ∀ᵐ c ∂η (a, b), p (b, c)) : ∀ᵐ bc ∂(κ ⊗ₖ η) a, p bc := by + by_cases hκ : IsSFiniteKernel κ + swap; · simp [compProd_of_not_isSFiniteKernel_left _ _ hκ] + by_cases hη : IsSFiniteKernel η + swap; · simp [compProd_of_not_isSFiniteKernel_right _ _ hη] simp_rw [ae_iff] at h ⊢ rw [compProd_null] · exact h diff --git a/Mathlib/Probability/Kernel/MeasureCompProd.lean b/Mathlib/Probability/Kernel/MeasureCompProd.lean index 99269703362e9..8e89b9365934c 100644 --- a/Mathlib/Probability/Kernel/MeasureCompProd.lean +++ b/Mathlib/Probability/Kernel/MeasureCompProd.lean @@ -30,7 +30,7 @@ open ProbabilityTheory namespace MeasureTheory.Measure variable {α β : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} - {μ : Measure α} {κ η : Kernel α β} + {μ ν : Measure α} {κ η : Kernel α β} /-- The composition-product of a measure and a kernel. -/ noncomputable @@ -67,14 +67,16 @@ lemma compProd_apply_prod [SFinite μ] [IsSFiniteKernel κ] rw [Set.indicator_apply] split_ifs with ha <;> simp [ha] -lemma compProd_congr [SFinite μ] [IsSFiniteKernel κ] [IsSFiniteKernel η] +lemma compProd_congr [IsSFiniteKernel κ] [IsSFiniteKernel η] (h : κ =ᵐ[μ] η) : μ ⊗ₘ κ = μ ⊗ₘ η := by - ext s hs - have : (fun a ↦ κ a (Prod.mk a ⁻¹' s)) =ᵐ[μ] fun a ↦ η a (Prod.mk a ⁻¹' s) := by - filter_upwards [h] with a ha using by rw [ha] - rw [compProd_apply hs, lintegral_congr_ae this, compProd_apply hs] - -lemma ae_compProd_of_ae_ae [SFinite μ] [IsSFiniteKernel κ] {p : α × β → Prop} + by_cases hμ : SFinite μ + · ext s hs + have : (fun a ↦ κ a (Prod.mk a ⁻¹' s)) =ᵐ[μ] fun a ↦ η a (Prod.mk a ⁻¹' s) := by + filter_upwards [h] with a ha using by rw [ha] + rw [compProd_apply hs, lintegral_congr_ae this, compProd_apply hs] + · simp [compProd_of_not_sfinite _ _ hμ] + +lemma ae_compProd_of_ae_ae {p : α × β → Prop} (hp : MeasurableSet {x | p x}) (h : ∀ᵐ a ∂μ, ∀ᵐ b ∂(κ a), p (a, b)) : ∀ᵐ x ∂(μ ⊗ₘ κ), p x := Kernel.ae_compProd_of_ae_ae hp h @@ -89,15 +91,20 @@ lemma ae_compProd_iff [SFinite μ] [IsSFiniteKernel κ] {p : α × β → Prop} (∀ᵐ x ∂(μ ⊗ₘ κ), p x) ↔ ∀ᵐ a ∂μ, ∀ᵐ b ∂(κ a), p (a, b) := Kernel.ae_compProd_iff hp -lemma compProd_add_left (μ ν : Measure α) [SFinite μ] [SFinite ν] (κ : Kernel α β) - [IsSFiniteKernel κ] : +lemma compProd_add_left (μ ν : Measure α) [SFinite μ] [SFinite ν] (κ : Kernel α β) : (μ + ν) ⊗ₘ κ = μ ⊗ₘ κ + ν ⊗ₘ κ := by - rw [Measure.compProd, Kernel.const_add, Kernel.compProd_add_left]; rfl + by_cases hκ : IsSFiniteKernel κ + · rw [Measure.compProd, Kernel.const_add, Kernel.compProd_add_left] + rfl + · simp [compProd_of_not_isSFiniteKernel _ _ hκ] -lemma compProd_add_right (μ : Measure α) [SFinite μ] (κ η : Kernel α β) +lemma compProd_add_right (μ : Measure α) (κ η : Kernel α β) [IsSFiniteKernel κ] [IsSFiniteKernel η] : μ ⊗ₘ (κ + η) = μ ⊗ₘ κ + μ ⊗ₘ η := by - rw [Measure.compProd, Kernel.prodMkLeft_add, Kernel.compProd_add_right]; rfl + by_cases hμ : SFinite μ + · rw [Measure.compProd, Kernel.prodMkLeft_add, Kernel.compProd_add_right] + rfl + · simp [compProd_of_not_sfinite _ _ hμ] section Integral @@ -172,4 +179,56 @@ instance [IsFiniteMeasure μ] [IsFiniteKernel κ] : IsFiniteMeasure (μ ⊗ₘ instance [IsProbabilityMeasure μ] [IsMarkovKernel κ] : IsProbabilityMeasure (μ ⊗ₘ κ) := by rw [compProd]; infer_instance +section AbsolutelyContinuous + +lemma absolutelyContinuous_compProd_left [SFinite ν] (hμν : μ ≪ ν) (κ : Kernel α β) : + μ ⊗ₘ κ ≪ ν ⊗ₘ κ := by + by_cases hκ : IsSFiniteKernel κ + · have : SFinite μ := sFinite_of_absolutelyContinuous hμν + refine Measure.AbsolutelyContinuous.mk fun s hs hs_zero ↦ ?_ + rw [Measure.compProd_apply hs, lintegral_eq_zero_iff (Kernel.measurable_kernel_prod_mk_left hs)] + at hs_zero ⊢ + exact hμν.ae_eq hs_zero + · simp [compProd_of_not_isSFiniteKernel _ _ hκ] + +lemma absolutelyContinuous_compProd_right [SFinite μ] [IsSFiniteKernel η] + (hκη : ∀ᵐ a ∂μ, κ a ≪ η a) : + μ ⊗ₘ κ ≪ μ ⊗ₘ η := by + by_cases hκ : IsSFiniteKernel κ + · refine Measure.AbsolutelyContinuous.mk fun s hs hs_zero ↦ ?_ + rw [Measure.compProd_apply hs, lintegral_eq_zero_iff (Kernel.measurable_kernel_prod_mk_left hs)] + at hs_zero ⊢ + filter_upwards [hs_zero, hκη] with a ha_zero ha_ac using ha_ac ha_zero + · simp [compProd_of_not_isSFiniteKernel _ _ hκ] + +lemma absolutelyContinuous_compProd [SFinite ν] [IsSFiniteKernel η] + (hμν : μ ≪ ν) (hκη : ∀ᵐ a ∂μ, κ a ≪ η a) : + μ ⊗ₘ κ ≪ ν ⊗ₘ η := + have : SFinite μ := sFinite_of_absolutelyContinuous hμν + (Measure.absolutelyContinuous_compProd_right hκη).trans + (Measure.absolutelyContinuous_compProd_left hμν _) + +lemma absolutelyContinuous_of_compProd [SFinite μ] [IsSFiniteKernel κ] [h_zero : ∀ a, NeZero (κ a)] + (h : μ ⊗ₘ κ ≪ ν ⊗ₘ η) : + μ ≪ ν := by + refine Measure.AbsolutelyContinuous.mk (fun s hs hs0 ↦ ?_) + have h1 : (ν ⊗ₘ η) (s ×ˢ Set.univ) = 0 := by + by_cases hν : SFinite ν + swap; · simp [compProd_of_not_sfinite _ _ hν] + by_cases hη : IsSFiniteKernel η + swap; · simp [compProd_of_not_isSFiniteKernel _ _ hη] + rw [Measure.compProd_apply_prod hs MeasurableSet.univ] + exact setLIntegral_measure_zero _ _ hs0 + have h2 : (μ ⊗ₘ κ) (s ×ˢ Set.univ) = 0 := h h1 + rw [Measure.compProd_apply_prod hs MeasurableSet.univ, lintegral_eq_zero_iff] at h2 + swap; · exact Kernel.measurable_coe _ MeasurableSet.univ + by_contra hμs + have : Filter.NeBot (ae (μ.restrict s)) := by simp [hμs] + obtain ⟨a, ha⟩ : ∃ a, κ a Set.univ = 0 := h2.exists + refine absurd ha ?_ + simp only [Measure.measure_univ_eq_zero] + exact (h_zero a).out + +end AbsolutelyContinuous + end MeasureTheory.Measure From 3805ed6e8b8a81617ebfe7db6b0dbe9862075594 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sat, 17 Aug 2024 07:09:37 +0000 Subject: [PATCH 355/359] chore: update Mathlib dependencies 2024-08-17 (#15912) 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 b1bab701ac2be..717ed46f42568 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "65f464e6a5f30b2c2a58569f530c4677b371cec6", + "rev": "a975dea2c4d8258a55b4f9861c537e2bb0f9ef63", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From b2af0a88e4928979df8926e973e7175f7bcecc52 Mon Sep 17 00:00:00 2001 From: Jeremy Tan Jie Rui Date: Sat, 17 Aug 2024 09:33:35 +0000 Subject: [PATCH 356/359] style: replace `$` with `<|` (#15913) The mathlib style guide requires using `<|` instead of `$`. The linter in #15909 will enforce this. --- LongestPole/Main.lean | 2 +- Mathlib/Algebra/Associated/Basic.lean | 4 +- .../Algebra/BigOperators/Group/Finset.lean | 10 ++--- Mathlib/Algebra/Field/Basic.lean | 8 ++-- Mathlib/Algebra/Field/Opposite.lean | 4 +- Mathlib/Algebra/Group/Action/Defs.lean | 2 +- Mathlib/Algebra/Group/Basic.lean | 4 +- Mathlib/Algebra/Group/Hom/Defs.lean | 8 ++-- Mathlib/Algebra/GroupWithZero/Basic.lean | 2 +- Mathlib/Algebra/GroupWithZero/Center.lean | 2 +- Mathlib/Algebra/GroupWithZero/Hom.lean | 4 +- .../GroupWithZero/NonZeroDivisors.lean | 2 +- .../Homology/ShortComplex/ExactFunctor.lean | 20 +++++----- Mathlib/Algebra/MvPolynomial/Degrees.lean | 2 +- Mathlib/Algebra/Order/Antidiag/Pi.lean | 8 ++-- Mathlib/Algebra/Order/CauSeq/Basic.lean | 2 +- .../Algebra/Order/CauSeq/BigOperators.lean | 9 +++-- Mathlib/Algebra/Order/Floor/Div.lean | 16 ++++---- Mathlib/Algebra/Order/Group/Basic.lean | 2 +- Mathlib/Algebra/Order/Group/PosPart.lean | 4 +- Mathlib/Algebra/Order/Module/Algebra.lean | 2 +- Mathlib/Algebra/Order/Module/Rat.lean | 4 +- .../Algebra/Order/Monoid/Canonical/Defs.lean | 2 +- Mathlib/Algebra/Order/Ring/Abs.lean | 4 +- Mathlib/Algebra/Order/Ring/Basic.lean | 6 +-- .../Algebra/Order/Ring/Unbundled/Basic.lean | 2 +- Mathlib/Algebra/Order/Ring/WithTop.lean | 4 +- .../Polynomial/Degree/TrailingDegree.lean | 4 +- Mathlib/Algebra/Polynomial/EraseLead.lean | 2 +- Mathlib/Algebra/Polynomial/RingDivision.lean | 2 +- Mathlib/Algebra/Ring/Hom/Defs.lean | 4 +- Mathlib/Algebra/Ring/Parity.lean | 2 +- Mathlib/AlgebraicGeometry/Modules/Tilde.lean | 4 +- Mathlib/AlgebraicGeometry/Stalk.lean | 4 +- Mathlib/Analysis/CStarAlgebra/Multiplier.lean | 4 +- Mathlib/Analysis/Calculus/Deriv/Shift.lean | 2 +- Mathlib/Analysis/Convex/Combination.lean | 8 ++-- Mathlib/Analysis/Convex/Cone/Closure.lean | 2 +- Mathlib/Analysis/Convex/Deriv.lean | 4 +- .../Analysis/InnerProductSpace/TwoDim.lean | 2 +- Mathlib/Analysis/MeanInequalities.lean | 4 +- Mathlib/Analysis/Normed/Field/Basic.lean | 2 +- Mathlib/Analysis/Normed/Order/UpperLower.lean | 14 +++---- Mathlib/Analysis/RCLike/Basic.lean | 2 +- .../SpecialFunctions/Complex/Arg.lean | 4 +- .../Trigonometric/Bounds.lean | 26 ++++++------- .../CategoryTheory/Category/PartialFun.lean | 2 +- .../WithAlgebraicStructures.lean | 4 +- .../Combinatorics/Additive/AP/Three/Defs.lean | 8 ++-- .../Combinatorics/Additive/Corner/Defs.lean | 4 +- .../Combinatorics/Additive/Corner/Roth.lean | 18 ++++----- .../Combinatorics/Additive/Dissociation.lean | 6 +-- Mathlib/Combinatorics/Additive/Energy.lean | 2 +- .../Additive/ErdosGinzburgZiv.lean | 16 ++++---- .../Combinatorics/Additive/FreimanHom.lean | 4 +- Mathlib/Combinatorics/Colex.lean | 17 +++++---- .../SetFamily/AhlswedeZhang.lean | 28 +++++++------- Mathlib/Combinatorics/SimpleGraph/Basic.lean | 2 +- .../Combinatorics/SimpleGraph/Coloring.lean | 4 +- .../Connectivity/WalkCounting.lean | 2 +- .../SimpleGraph/Hamiltonian.lean | 8 ++-- .../SimpleGraph/Regularity/Uniform.lean | 12 +++--- .../SimpleGraph/Triangle/Basic.lean | 16 ++++---- .../SimpleGraph/Triangle/Counting.lean | 18 ++++----- .../SimpleGraph/Triangle/Removal.lean | 2 +- .../SimpleGraph/Triangle/Tripartite.lean | 4 +- Mathlib/Computability/Primrec.lean | 16 ++++---- Mathlib/Data/Complex/Abs.lean | 2 +- Mathlib/Data/Fin/Basic.lean | 38 +++++++++---------- Mathlib/Data/Fin/Tuple/Curry.lean | 2 +- Mathlib/Data/Fin/Tuple/Finset.lean | 2 +- Mathlib/Data/Finset/Basic.lean | 2 +- Mathlib/Data/Finset/Card.lean | 2 +- Mathlib/Data/Finset/Density.lean | 6 +-- Mathlib/Data/Finset/Powerset.lean | 10 ++--- Mathlib/Data/Fintype/BigOperators.lean | 4 +- Mathlib/Data/Int/Star.lean | 2 +- Mathlib/Data/List/Perm.lean | 2 +- Mathlib/Data/List/Range.lean | 2 +- Mathlib/Data/Matrix/Composition.lean | 2 +- Mathlib/Data/Multiset/Basic.lean | 4 +- Mathlib/Data/Multiset/Bind.lean | 2 +- Mathlib/Data/Multiset/Fintype.lean | 2 +- Mathlib/Data/NNRat/Defs.lean | 2 +- Mathlib/Data/Nat/Cast/Basic.lean | 2 +- Mathlib/Data/Nat/Cast/Order/Field.lean | 2 +- Mathlib/Data/Nat/Choose/Lucas.lean | 2 +- Mathlib/Data/Nat/Defs.lean | 26 ++++++------- Mathlib/Data/Nat/Factorial/Basic.lean | 2 +- Mathlib/Data/Nat/Factorization/Root.lean | 4 +- Mathlib/Data/Nat/Find.lean | 2 +- Mathlib/Data/Nat/Log.lean | 2 +- Mathlib/Data/Real/ConjExponents.lean | 8 ++-- Mathlib/Data/Real/Sqrt.lean | 4 +- Mathlib/Data/Set/Function.lean | 2 +- Mathlib/Data/Set/Image.lean | 2 +- .../IsAlgClosed/AlgebraicClosure.lean | 6 +-- .../SplittingField/Construction.lean | 6 +-- .../Euclidean/Angle/Unoriented/Affine.lean | 2 +- Mathlib/GroupTheory/FiniteAbelian.lean | 2 +- Mathlib/GroupTheory/OrderOfElement.lean | 8 ++-- Mathlib/GroupTheory/Perm/Basic.lean | 4 +- .../AffineSpace/FiniteDimensional.lean | 6 +-- Mathlib/LinearAlgebra/Matrix/ZPow.lean | 2 +- Mathlib/LinearAlgebra/Prod.lean | 2 +- Mathlib/LinearAlgebra/Ray.lean | 2 +- Mathlib/Logic/Equiv/Defs.lean | 14 +++---- Mathlib/Logic/Godel/GodelBetaFunction.lean | 6 +-- .../Constructions/BorelSpace/Metric.lean | 2 +- .../Constructions/EventuallyMeasurable.lean | 2 +- .../Function/Intersectivity.lean | 30 +++++++-------- .../MeasurableSpace/CountablyGenerated.lean | 16 ++++---- .../MeasureTheory/Measure/FiniteMeasure.lean | 2 +- .../MeasureTheory/Measure/Haar/Quotient.lean | 2 +- .../Measure/ProbabilityMeasure.lean | 2 +- Mathlib/MeasureTheory/Order/UpperLower.lean | 21 +++++----- Mathlib/MeasureTheory/OuterMeasure/AE.lean | 2 +- .../NumberTheory/NumberField/Units/Basic.lean | 2 +- Mathlib/Order/Antichain.lean | 4 +- Mathlib/Order/Basic.lean | 2 +- Mathlib/Order/BooleanAlgebra.lean | 2 +- Mathlib/Order/Category/BoolAlg.lean | 2 +- Mathlib/Order/Directed.lean | 4 +- Mathlib/Order/Filter/AtTopBot.lean | 2 +- .../Order/Filter/CountableSeparatingOn.lean | 4 +- Mathlib/Order/Filter/Curry.lean | 4 +- Mathlib/Order/Interval/Finset/Nat.lean | 4 +- Mathlib/Order/Interval/Set/Image.lean | 8 ++-- Mathlib/Order/Part.lean | 4 +- Mathlib/Order/RelSeries.lean | 2 +- Mathlib/Order/SuccPred/Basic.lean | 2 +- Mathlib/Order/SupClosed.lean | 8 ++-- Mathlib/Order/UpperLower/Basic.lean | 4 +- Mathlib/RingTheory/Flat/Basic.lean | 4 +- Mathlib/RingTheory/Flat/CategoryTheory.lean | 16 ++++---- Mathlib/RingTheory/HahnSeries/Summable.lean | 2 +- Mathlib/RingTheory/Ideal/Operations.lean | 2 +- Mathlib/RingTheory/Polynomial/Bernstein.lean | 2 +- .../Polynomial/Cyclotomic/Expand.lean | 2 +- Mathlib/RingTheory/TwoSidedIdeal/Basic.lean | 2 +- Mathlib/RingTheory/TwoSidedIdeal/Lattice.lean | 12 +++--- .../WittVector/DiscreteValuationRing.lean | 2 +- Mathlib/RingTheory/WittVector/InitTail.lean | 2 +- Mathlib/SetTheory/Game/Basic.lean | 4 +- Mathlib/SetTheory/Game/PGame.lean | 4 +- Mathlib/Tactic/Cases.lean | 2 +- Mathlib/Tactic/FunProp/RefinedDiscrTree.lean | 20 +++++----- Mathlib/Tactic/Positivity/Basic.lean | 8 ++-- Mathlib/Tactic/Widget/InteractiveUnfold.lean | 4 +- Mathlib/Topology/Basic.lean | 2 +- Mathlib/Topology/Inseparable.lean | 2 +- Mathlib/Topology/Order/Bornology.lean | 4 +- Mathlib/Topology/Order/IntermediateValue.lean | 2 +- 153 files changed, 435 insertions(+), 432 deletions(-) diff --git a/LongestPole/Main.lean b/LongestPole/Main.lean index 1ad6f2a7b33cd..4c21dc471afe4 100644 --- a/LongestPole/Main.lean +++ b/LongestPole/Main.lean @@ -22,7 +22,7 @@ open Lean Meta /-- Runs a terminal command and retrieves its output -/ def runCmd (cmd : String) (args : Array String) (throwFailure := true) : IO String := do let out ← IO.Process.output { cmd := cmd, args := args } - if out.exitCode != 0 && throwFailure then throw $ IO.userError out.stderr + if out.exitCode != 0 && throwFailure then throw <| IO.userError out.stderr else return out.stdout def runCurl (args : Array String) (throwFailure := true) : IO String := do diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index 8ea3c00ea7004..9ff65fedb0e76 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -851,8 +851,8 @@ instance uniqueUnits : Unique (Associates α)ˣ where uniq := by rintro ⟨a, b, hab, hba⟩ revert hab hba - exact Quotient.inductionOn₂ a b $ fun a b hab hba ↦ Units.ext $ Quotient.sound $ - associated_one_of_associated_mul_one $ Quotient.exact hab + exact Quotient.inductionOn₂ a b <| fun a b hab hba ↦ Units.ext <| Quotient.sound <| + associated_one_of_associated_mul_one <| Quotient.exact hab @[deprecated (since := "2024-07-22")] alias mul_eq_one_iff := mul_eq_one @[deprecated (since := "2024-07-22")] protected alias units_eq_one := Subsingleton.elim diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index eaeaf7b7f6f44..602f1900802b6 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -456,7 +456,7 @@ of `s`, and over all subsets of `s` to which one adds `x`. -/ of `s`, and over all subsets of `s` to which one adds `x`."] lemma prod_powerset_cons (ha : a ∉ s) (f : Finset α → β) : ∏ t ∈ (s.cons a ha).powerset, f t = (∏ t ∈ s.powerset, f t) * - ∏ t ∈ s.powerset.attach, f (cons a t $ not_mem_mono (mem_powerset.1 t.2) ha) := by + ∏ t ∈ s.powerset.attach, f (cons a t <| not_mem_mono (mem_powerset.1 t.2) ha) := by classical simp_rw [cons_eq_insert] rw [prod_powerset_insert ha, prod_attach _ fun t ↦ f (insert a t)] @@ -721,7 +721,7 @@ but differ in the type of their element, `univ.pi t` is a `Finset (Π a ∈ univ lemma prod_univ_pi [DecidableEq ι] [Fintype ι] {κ : ι → Type*} (t : ∀ i, Finset (κ i)) (f : (∀ i ∈ (univ : Finset ι), κ i) → β) : ∏ x ∈ univ.pi t, f x = ∏ x ∈ Fintype.piFinset t, f fun a _ ↦ x a := by - apply prod_nbij' (fun x i ↦ x i $ mem_univ _) (fun x i _ ↦ x i) <;> simp + apply prod_nbij' (fun x i ↦ x i <| mem_univ _) (fun x i _ ↦ x i) <;> simp @[to_additive (attr := simp)] lemma prod_diag [DecidableEq α] (s : Finset α) (f : α × α → β) : @@ -1751,7 +1751,7 @@ variable [DecidableEq ι] [CancelCommMonoid α] {s t : Finset ι} {f : ι → α @[to_additive] lemma prod_sdiff_eq_prod_sdiff_iff : ∏ i ∈ s \ t, f i = ∏ i ∈ t \ s, f i ↔ ∏ i ∈ s, f i = ∏ i ∈ t, f i := - eq_comm.trans $ eq_iff_eq_of_mul_eq_mul $ by + eq_comm.trans <| eq_iff_eq_of_mul_eq_mul <| by rw [← prod_union disjoint_sdiff_self_left, ← prod_union disjoint_sdiff_self_left, sdiff_union_self_eq_union, sdiff_union_self_eq_union, union_comm] @@ -1985,7 +1985,7 @@ theorem prod_subtype_mul_prod_subtype {α β : Type*} [Fintype α] [CommMonoid @[to_additive] lemma prod_subset {s : Finset ι} {f : ι → α} (h : ∀ i, f i ≠ 1 → i ∈ s) : ∏ i ∈ s, f i = ∏ i, f i := - Finset.prod_subset s.subset_univ $ by simpa [not_imp_comm (a := _ ∈ s)] + Finset.prod_subset s.subset_univ <| by simpa [not_imp_comm (a := _ ∈ s)] @[to_additive] lemma prod_ite_eq_ite_exists (p : ι → Prop) [DecidablePred p] (h : ∀ i j, p i → p j → i = j) @@ -2032,7 +2032,7 @@ variable [CommMonoid α] @[to_additive (attr := simp)] lemma prod_attach_univ [Fintype ι] (f : {i // i ∈ @univ ι _} → α) : ∏ i ∈ univ.attach, f i = ∏ i, f ⟨i, mem_univ _⟩ := - Fintype.prod_equiv (Equiv.subtypeUnivEquiv mem_univ) _ _ $ by simp + Fintype.prod_equiv (Equiv.subtypeUnivEquiv mem_univ) _ _ <| by simp @[to_additive] theorem prod_erase_attach [DecidableEq ι] {s : Finset ι} (f : ι → α) (i : ↑s) : diff --git a/Mathlib/Algebra/Field/Basic.lean b/Mathlib/Algebra/Field/Basic.lean index b0dabde1c5792..fa1a788fb74a7 100644 --- a/Mathlib/Algebra/Field/Basic.lean +++ b/Mathlib/Algebra/Field/Basic.lean @@ -256,9 +256,9 @@ protected abbrev divisionSemiring [DivisionSemiring β] (zero : f 0 = 0) (one : (natCast : ∀ n : ℕ, f n = n) (nnratCast : ∀ q : ℚ≥0, f q = q) : DivisionSemiring α where toSemiring := hf.semiring f zero one add mul nsmul npow natCast __ := hf.groupWithZero f zero one mul inv div npow zpow - nnratCast_def q := hf $ by rw [nnratCast, NNRat.cast_def, div, natCast, natCast] + nnratCast_def q := hf <| by rw [nnratCast, NNRat.cast_def, div, natCast, natCast] nnqsmul := (· • ·) - nnqsmul_def q a := hf $ by rw [nnqsmul, NNRat.smul_def, mul, nnratCast] + nnqsmul_def q a := hf <| by rw [nnqsmul, NNRat.smul_def, mul, nnratCast] /-- Pullback a `DivisionSemiring` along an injective function. -/ -- See note [reducible non-instances] @@ -274,9 +274,9 @@ protected abbrev divisionRing [DivisionRing β] (zero : f 0 = 0) (one : f 1 = 1) toRing := hf.ring f zero one add mul neg sub nsmul zsmul npow natCast intCast __ := hf.groupWithZero f zero one mul inv div npow zpow __ := hf.divisionSemiring f zero one add mul inv div nsmul nnqsmul npow zpow natCast nnratCast - ratCast_def q := hf $ by erw [ratCast, div, intCast, natCast, Rat.cast_def] + ratCast_def q := hf <| by erw [ratCast, div, intCast, natCast, Rat.cast_def] qsmul := (· • ·) - qsmul_def q a := hf $ by erw [qsmul, mul, Rat.smul_def, ratCast] + qsmul_def q a := hf <| by erw [qsmul, mul, Rat.smul_def, ratCast] /-- Pullback a `Field` along an injective function. -/ -- See note [reducible non-instances] diff --git a/Mathlib/Algebra/Field/Opposite.lean b/Mathlib/Algebra/Field/Opposite.lean index cde950bef7a7c..99f0c22d73388 100644 --- a/Mathlib/Algebra/Field/Opposite.lean +++ b/Mathlib/Algebra/Field/Opposite.lean @@ -35,7 +35,7 @@ instance instDivisionSemiring [DivisionSemiring α] : DivisionSemiring αᵐᵒ __ := instGroupWithZero nnqsmul := _ nnqsmul_def := fun q a => rfl - nnratCast_def q := unop_injective $ by rw [unop_nnratCast, unop_div, unop_natCast, unop_natCast, + nnratCast_def q := unop_injective <| by rw [unop_nnratCast, unop_div, unop_natCast, unop_natCast, NNRat.cast_def, div_eq_mul_inv, Nat.cast_comm] instance instDivisionRing [DivisionRing α] : DivisionRing αᵐᵒᵖ where @@ -63,7 +63,7 @@ instance instDivisionSemiring [DivisionSemiring α] : DivisionSemiring αᵃᵒ __ := instGroupWithZero nnqsmul := _ nnqsmul_def := fun q a => rfl - nnratCast_def q := unop_injective $ by rw [unop_nnratCast, unop_div, unop_natCast, unop_natCast, + nnratCast_def q := unop_injective <| by rw [unop_nnratCast, unop_div, unop_natCast, unop_natCast, NNRat.cast_def, div_eq_mul_inv] instance instDivisionRing [DivisionRing α] : DivisionRing αᵃᵒᵖ where diff --git a/Mathlib/Algebra/Group/Action/Defs.lean b/Mathlib/Algebra/Group/Action/Defs.lean index b0519ca6ce79c..220aa2d2f5e5c 100644 --- a/Mathlib/Algebra/Group/Action/Defs.lean +++ b/Mathlib/Algebra/Group/Action/Defs.lean @@ -489,7 +489,7 @@ end Mul variable [Group H] [MulAction G H] [SMulCommClass G H H] [IsScalarTower G H H] lemma smul_inv (g : G) (a : H) : (g • a)⁻¹ = g⁻¹ • a⁻¹ := - inv_eq_of_mul_eq_one_right $ by rw [smul_mul_smul, mul_inv_cancel, mul_inv_cancel, one_smul] + inv_eq_of_mul_eq_one_right <| by rw [smul_mul_smul, mul_inv_cancel, mul_inv_cancel, one_smul] lemma smul_zpow (g : G) (a : H) (n : ℤ) : (g • a) ^ n = g ^ n • a ^ n := by cases n <;> simp [smul_pow, smul_inv] diff --git a/Mathlib/Algebra/Group/Basic.lean b/Mathlib/Algebra/Group/Basic.lean index c487707b09042..9bd58b84c981b 100644 --- a/Mathlib/Algebra/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Basic.lean @@ -189,11 +189,11 @@ lemma pow_sub_mul_pow (a : M) (h : m ≤ n) : a ^ (n - m) * a ^ m = a ^ n := by @[to_additive sub_one_nsmul_add] lemma mul_pow_sub_one (hn : n ≠ 0) (a : M) : a * a ^ (n - 1) = a ^ n := by - rw [← pow_succ', Nat.sub_add_cancel $ Nat.one_le_iff_ne_zero.2 hn] + rw [← pow_succ', Nat.sub_add_cancel <| Nat.one_le_iff_ne_zero.2 hn] @[to_additive add_sub_one_nsmul] lemma pow_sub_one_mul (hn : n ≠ 0) (a : M) : a ^ (n - 1) * a = a ^ n := by - rw [← pow_succ, Nat.sub_add_cancel $ Nat.one_le_iff_ne_zero.2 hn] + rw [← pow_succ, Nat.sub_add_cancel <| Nat.one_le_iff_ne_zero.2 hn] /-- If `x ^ n = 1`, then `x ^ m` is the same as `x ^ (m % n)` -/ @[to_additive nsmul_eq_mod_nsmul "If `n • x = 0`, then `m • x` is the same as `(m % n) • x`"] diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index 6c96b0483e734..b75baec7c85f7 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -837,8 +837,8 @@ instance : Monoid (Monoid.End M) where mul_assoc _ _ _ := MonoidHom.comp_assoc _ _ _ mul_one := MonoidHom.comp_id one_mul := MonoidHom.id_comp - npow n f := (npowRec n f).copy f^[n] $ by induction n <;> simp [npowRec, *] <;> rfl - npow_succ n f := DFunLike.coe_injective $ Function.iterate_succ _ _ + npow n f := (npowRec n f).copy f^[n] <| by induction n <;> simp [npowRec, *] <;> rfl + npow_succ n f := DFunLike.coe_injective <| Function.iterate_succ _ _ instance : Inhabited (Monoid.End M) := ⟨1⟩ @@ -878,8 +878,8 @@ instance monoid : Monoid (AddMonoid.End A) where mul_assoc _ _ _ := AddMonoidHom.comp_assoc _ _ _ mul_one := AddMonoidHom.comp_id one_mul := AddMonoidHom.id_comp - npow n f := (npowRec n f).copy (Nat.iterate f n) $ by induction n <;> simp [npowRec, *] <;> rfl - npow_succ n f := DFunLike.coe_injective $ Function.iterate_succ _ _ + npow n f := (npowRec n f).copy (Nat.iterate f n) <| by induction n <;> simp [npowRec, *] <;> rfl + npow_succ n f := DFunLike.coe_injective <| Function.iterate_succ _ _ @[simp, norm_cast] lemma coe_pow (f : AddMonoid.End A) (n : ℕ) : (↑(f ^ n) : A → A) = f^[n] := rfl diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index b21f2ccf96142..78525d46a0706 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -150,7 +150,7 @@ lemma pow_eq_zero_of_le : ∀ {m n} (hmn : m ≤ n) (ha : a ^ m = 0), a ^ n = 0 | _, _, Nat.le.refl, ha => ha | _, _, Nat.le.step hmn, ha => by rw [pow_succ, pow_eq_zero_of_le hmn ha, zero_mul] -lemma ne_zero_pow (hn : n ≠ 0) (ha : a ^ n ≠ 0) : a ≠ 0 := by rintro rfl; exact ha $ zero_pow hn +lemma ne_zero_pow (hn : n ≠ 0) (ha : a ^ n ≠ 0) : a ≠ 0 := by rintro rfl; exact ha <| zero_pow hn @[simp] lemma zero_pow_eq_zero [Nontrivial M₀] : (0 : M₀) ^ n = 0 ↔ n ≠ 0 := diff --git a/Mathlib/Algebra/GroupWithZero/Center.lean b/Mathlib/Algebra/GroupWithZero/Center.lean index d8366c9759d48..3f721aaa7d178 100644 --- a/Mathlib/Algebra/GroupWithZero/Center.lean +++ b/Mathlib/Algebra/GroupWithZero/Center.lean @@ -38,7 +38,7 @@ lemma center_units_subset : center G₀ˣ ⊆ ((↑) : G₀ˣ → G₀) ⁻¹' c intro u hu a obtain rfl | ha := eq_or_ne a 0 · rw [zero_mul, mul_zero] - · exact congr_arg Units.val $ hu $ Units.mk0 a ha + · exact congr_arg Units.val <| hu <| Units.mk0 a ha /-- In a group with zero, the center of the units is the preimage of the center. -/ lemma center_units_eq : center G₀ˣ = ((↑) : G₀ˣ → G₀) ⁻¹' center G₀ := diff --git a/Mathlib/Algebra/GroupWithZero/Hom.lean b/Mathlib/Algebra/GroupWithZero/Hom.lean index 7f2b2a294119a..75d6481cf412f 100644 --- a/Mathlib/Algebra/GroupWithZero/Hom.lean +++ b/Mathlib/Algebra/GroupWithZero/Hom.lean @@ -175,11 +175,11 @@ lemma comp_assoc (f : α →*₀ β) (g : β →*₀ γ) (h : γ →*₀ δ) : lemma cancel_right {g₁ g₂ : β →*₀ γ} {f : α →*₀ β} (hf : Surjective f) : g₁.comp f = g₂.comp f ↔ g₁ = g₂ := - ⟨fun h ↦ ext $ hf.forall.2 (DFunLike.ext_iff.1 h), fun h ↦ h ▸ rfl⟩ + ⟨fun h ↦ ext <| hf.forall.2 (DFunLike.ext_iff.1 h), fun h ↦ h ▸ rfl⟩ lemma cancel_left {g : β →*₀ γ} {f₁ f₂ : α →*₀ β} (hg : Injective g) : g.comp f₁ = g.comp f₂ ↔ f₁ = f₂ := - ⟨fun h ↦ ext fun x ↦ hg $ by rw [← comp_apply, h, + ⟨fun h ↦ ext fun x ↦ hg <| by rw [← comp_apply, h, comp_apply], fun h ↦ h ▸ rfl⟩ lemma toMonoidHom_injective : Injective (toMonoidHom : (α →*₀ β) → α →* β) := diff --git a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean index 0ed3c88eff1ab..3c09328f414ca 100644 --- a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean +++ b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean @@ -283,7 +283,7 @@ def unitsNonZeroDivisorsEquiv : M₀⁰ˣ ≃* M₀ˣ where right_inv _ := rfl @[simp, norm_cast] lemma nonZeroDivisors.associated_coe : Associated (a : M₀) b ↔ Associated a b := - unitsNonZeroDivisorsEquiv.symm.exists_congr_left.trans $ by simp [Associated]; norm_cast + unitsNonZeroDivisorsEquiv.symm.exists_congr_left.trans <| by simp [Associated]; norm_cast end MonoidWithZero diff --git a/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean b/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean index feb861d1e0060..974a1c8685585 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ExactFunctor.lean @@ -112,8 +112,8 @@ lemma preservesFiniteLimits_tfae : List.TFAE [ ∀ (S : ShortComplex C), S.ShortExact → (S.map F).Exact ∧ Mono (F.map S.f), ∀ (S : ShortComplex C), S.Exact ∧ Mono S.f → (S.map F).Exact ∧ Mono (F.map S.f), - ∀ ⦃X Y : C⦄ (f : X ⟶ Y), Nonempty $ PreservesLimit (parallelPair f 0) F, - Nonempty $ PreservesFiniteLimits F + ∀ ⦃X Y : C⦄ (f : X ⟶ Y), Nonempty <| PreservesLimit (parallelPair f 0) F, + Nonempty <| PreservesFiniteLimits F ] := by tfae_have 1 → 2 · rintro hF S ⟨hS, hf⟩ @@ -123,7 +123,7 @@ lemma preservesFiniteLimits_tfae : List.TFAE let φ : T.map F ⟶ S.map F := { τ₁ := 𝟙 _ τ₂ := 𝟙 _ - τ₃ := F.map $ Abelian.factorThruCoimage S.g + τ₃ := F.map <| Abelian.factorThruCoimage S.g comm₂₃ := show 𝟙 _ ≫ F.map _ = F.map (cokernel.π _) ≫ _ by rw [Category.id_comp, ← F.map_comp, cokernel.π_desc] } exact (exact_iff_of_epi_of_isIso_of_mono φ).1 (hF T ⟨(S.exact_iff_exact_coimage_π).1 hS⟩).1 @@ -171,8 +171,8 @@ lemma preservesFiniteColimits_tfae : List.TFAE [ ∀ (S : ShortComplex C), S.ShortExact → (S.map F).Exact ∧ Epi (F.map S.g), ∀ (S : ShortComplex C), S.Exact ∧ Epi S.g → (S.map F).Exact ∧ Epi (F.map S.g), - ∀ ⦃X Y : C⦄ (f : X ⟶ Y), Nonempty $ PreservesColimit (parallelPair f 0) F, - Nonempty $ PreservesFiniteColimits F + ∀ ⦃X Y : C⦄ (f : X ⟶ Y), Nonempty <| PreservesColimit (parallelPair f 0) F, + Nonempty <| PreservesFiniteColimits F ] := by tfae_have 1 → 2 · rintro hF S ⟨hS, hf⟩ @@ -180,7 +180,7 @@ lemma preservesFiniteColimits_tfae : List.TFAE refine ⟨?_, inferInstance⟩ let T := ShortComplex.mk (Abelian.image.ι S.f) S.g (Abelian.image_ι_comp_eq_zero S.zero) let φ : S.map F ⟶ T.map F := - { τ₁ := F.map $ Abelian.factorThruImage S.f + { τ₁ := F.map <| Abelian.factorThruImage S.f τ₂ := 𝟙 _ τ₃ := 𝟙 _ comm₁₂ := show _ ≫ F.map (kernel.ι _) = F.map _ ≫ 𝟙 _ by @@ -235,10 +235,10 @@ lemma exact_tfae : List.TFAE tfae_have 2 → 1 · intro hF S hS - have : Mono (S.map F).f := exact_iff_mono _ (by simp) |>.1 $ - hF (.mk (0 : 0 ⟶ S.X₁) S.f $ by simp) (exact_iff_mono _ (by simp) |>.2 hS.mono_f) - have : Epi (S.map F).g := exact_iff_epi _ (by simp) |>.1 $ - hF (.mk S.g (0 : S.X₃ ⟶ 0) $ by simp) (exact_iff_epi _ (by simp) |>.2 hS.epi_g) + have : Mono (S.map F).f := exact_iff_mono _ (by simp) |>.1 <| + hF (.mk (0 : 0 ⟶ S.X₁) S.f <| by simp) (exact_iff_mono _ (by simp) |>.2 hS.mono_f) + have : Epi (S.map F).g := exact_iff_epi _ (by simp) |>.1 <| + hF (.mk S.g (0 : S.X₃ ⟶ 0) <| by simp) (exact_iff_epi _ (by simp) |>.2 hS.epi_g) exact ⟨hF S hS.exact⟩ tfae_have 3 → 4 diff --git a/Mathlib/Algebra/MvPolynomial/Degrees.lean b/Mathlib/Algebra/MvPolynomial/Degrees.lean index 388b7941515bf..93924cb3df6c2 100644 --- a/Mathlib/Algebra/MvPolynomial/Degrees.lean +++ b/Mathlib/Algebra/MvPolynomial/Degrees.lean @@ -435,7 +435,7 @@ theorem totalDegree_finset_sum {ι : Type*} (s : Finset ι) (f : ι → MvPolyno lemma totalDegree_finsetSum_le {ι : Type*} {s : Finset ι} {f : ι → MvPolynomial σ R} {d : ℕ} (hf : ∀ i ∈ s, (f i).totalDegree ≤ d) : (s.sum f).totalDegree ≤ d := - (totalDegree_finset_sum ..).trans $ Finset.sup_le hf + (totalDegree_finset_sum ..).trans <| Finset.sup_le hf lemma degreeOf_le_totalDegree (f : MvPolynomial σ R) (i : σ) : f.degreeOf i ≤ f.totalDegree := degreeOf_le_iff.mpr fun d hd ↦ (eq_or_ne (d i) 0).elim (·.trans_le zero_le') fun h ↦ diff --git a/Mathlib/Algebra/Order/Antidiag/Pi.lean b/Mathlib/Algebra/Order/Antidiag/Pi.lean index 3bbf985a82d41..2573fd6a28d1f 100644 --- a/Mathlib/Algebra/Order/Antidiag/Pi.lean +++ b/Mathlib/Algebra/Order/Antidiag/Pi.lean @@ -75,7 +75,7 @@ where (fun ab => (aux d ab.2).1.map { toFun := Fin.cons (ab.1) inj' := Fin.cons_right_injective _ }) - (fun i _hi j _hj hij => Finset.disjoint_left.2 fun t hti htj => hij $ by + (fun i _hi j _hj hij => Finset.disjoint_left.2 fun t hti htj => hij <| by simp_rw [Finset.mem_map, Embedding.coeFn_mk] at hti htj obtain ⟨ai, hai, hij'⟩ := hti obtain ⟨aj, haj, rfl⟩ := htj @@ -106,7 +106,7 @@ choice. /-- The finset of functions `ι → μ` with support contained in `s` and sum `n`. -/ def piAntidiag (s : Finset ι) (n : μ) : Finset (ι → μ) := by - refine (Fintype.truncEquivFinOfCardEq $ Fintype.card_coe s).lift + refine (Fintype.truncEquivFinOfCardEq <| Fintype.card_coe s).lift (fun e ↦ (finAntidiagonal s.card n).map ⟨fun f i ↦ if hi : i ∈ s then f (e ⟨i, hi⟩) else 0, ?_⟩) fun e₁ e₂ ↦ ?_ · rintro f g hfg @@ -114,7 +114,7 @@ def piAntidiag (s : Finset ι) (n : μ) : Finset (ι → μ) := by simpa using congr_fun hfg (e.symm i) · ext f simp only [mem_map, mem_finAntidiagonal] - refine Equiv.exists_congr ((e₁.symm.trans e₂).arrowCongr $ .refl _) fun g ↦ ?_ + refine Equiv.exists_congr ((e₁.symm.trans e₂).arrowCongr <| .refl _) fun g ↦ ?_ have := Fintype.sum_equiv (e₂.symm.trans e₁) _ g fun _ ↦ rfl aesop @@ -162,7 +162,7 @@ lemma pairwiseDisjoint_piAntidiag_map_addRightEmbedding (hi : i ∉ s) (n : μ) simp only [ne_eq, antidiagonal_congr' hab hcd, disjoint_left, mem_map, mem_piAntidiag, addRightEmbedding_apply, not_exists, not_and, and_imp, forall_exists_index] rintro hfg _ f rfl - rfl g rfl - hgf - exact hfg $ by simpa [sum_add_distrib, hi] using congr_arg (∑ j ∈ s, · j) hgf.symm + exact hfg <| by simpa [sum_add_distrib, hi] using congr_arg (∑ j ∈ s, · j) hgf.symm lemma piAntidiag_cons (hi : i ∉ s) (n : μ) : piAntidiag (cons i s hi) n = (antidiagonal n).disjiUnion (fun p : μ × μ ↦ diff --git a/Mathlib/Algebra/Order/CauSeq/Basic.lean b/Mathlib/Algebra/Order/CauSeq/Basic.lean index aa0df8061099e..461ef07e12d62 100644 --- a/Mathlib/Algebra/Order/CauSeq/Basic.lean +++ b/Mathlib/Algebra/Order/CauSeq/Basic.lean @@ -116,7 +116,7 @@ lemma bounded (hf : IsCauSeq abv f) : ∃ r, ∀ i, abv (f i) < r := by refine ⟨R i + 1, fun j ↦ ?_⟩ obtain hji | hij := le_total j i · exact (this i _ hji).trans_lt (lt_add_one _) - · simpa using (abv_add abv _ _).trans_lt $ add_lt_add_of_le_of_lt (this i _ le_rfl) (h _ hij) + · simpa using (abv_add abv _ _).trans_lt <| add_lt_add_of_le_of_lt (this i _ le_rfl) (h _ hij) lemma bounded' (hf : IsCauSeq abv f) (x : α) : ∃ r > x, ∀ i, abv (f i) < r := let ⟨r, h⟩ := hf.bounded diff --git a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean index 129e4899f3ace..2ac291cf0c7a6 100644 --- a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean +++ b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean @@ -94,7 +94,7 @@ theorem _root_.cauchy_product (ha : IsCauSeq abs fun m ↦ ∑ n ∈ range m, ab abv (f i) * abv ((∑ k ∈ range (K - i), g k) - ∑ k ∈ range K, g k)) ≤ ∑ i ∈ range (max N M + 1), abv (f i) * (ε / (2 * P)) := by gcongr with m hmJ - refine le_of_lt $ hN (K - m) (le_tsub_of_add_le_left $ hK.trans' ?_) K hKN.le + refine le_of_lt <| hN (K - m) (le_tsub_of_add_le_left <| hK.trans' ?_) K hKN.le rw [two_mul] gcongr · exact (mem_range.1 hmJ).le @@ -133,8 +133,9 @@ theorem _root_.cauchy_product (ha : IsCauSeq abs fun m ↦ ∑ n ∈ range m, ab rw [← sum_mul, ← sum_range_sub_sum_range (le_of_lt hNMK)] have := lt_of_le_of_lt (abv_nonneg _ _) (hQ 0) gcongr - exact (le_abs_self _).trans_lt $ hM _ ((Nat.le_succ_of_le (le_max_right _ _)).trans hNMK.le) - _ $ Nat.le_succ_of_le $ le_max_right _ _ + exact (le_abs_self _).trans_lt <| + hM _ ((Nat.le_succ_of_le (le_max_right _ _)).trans hNMK.le) _ <| + Nat.le_succ_of_le <| le_max_right _ _ variable [Archimedean α] @@ -172,7 +173,7 @@ lemma of_decreasing_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ lemma of_mono_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ m, |f n| ≤ a) (hnm : ∀ n ≥ m, f n ≤ f n.succ) : IsCauSeq abs f := - (of_decreasing_bounded (-f) (a := a) (m := m) (by simpa using ham) $ by simpa using hnm).of_neg + (of_decreasing_bounded (-f) (a := a) (m := m) (by simpa using ham) <| by simpa using hnm).of_neg lemma geo_series [Nontrivial β] (x : β) (hx1 : abv x < 1) : IsCauSeq abv fun n ↦ ∑ m ∈ range n, x ^ m := by diff --git a/Mathlib/Algebra/Order/Floor/Div.lean b/Mathlib/Algebra/Order/Floor/Div.lean index 33bc2b55c54a5..35d76dbf57197 100644 --- a/Mathlib/Algebra/Order/Floor/Div.lean +++ b/Mathlib/Algebra/Order/Floor/Div.lean @@ -124,7 +124,7 @@ variable [LinearOrderedAddCommMonoid α] [OrderedAddCommMonoid β] [SMulZeroClas lemma floorDiv_le_ceilDiv : b ⌊/⌋ a ≤ b ⌈/⌉ a := by obtain ha | ha := le_or_lt a 0 · simp [ha] - · exact le_of_smul_le_smul_left ((smul_floorDiv_le ha).trans $ le_smul_ceilDiv ha) ha + · exact le_of_smul_le_smul_left ((smul_floorDiv_le ha).trans <| le_smul_ceilDiv ha) ha end LinearOrderedAddCommMonoid @@ -135,11 +135,11 @@ section FloorDiv variable [FloorDiv α β] {a : α} @[simp] lemma floorDiv_one [Nontrivial α] (b : β) : b ⌊/⌋ (1 : α) = b := - eq_of_forall_le_iff $ fun c ↦ by simp [zero_lt_one' α] + eq_of_forall_le_iff <| fun c ↦ by simp [zero_lt_one' α] @[simp] lemma smul_floorDiv [PosSMulMono α β] [PosSMulReflectLE α β] (ha : 0 < a) (b : β) : a • b ⌊/⌋ a = b := - eq_of_forall_le_iff $ by simp [smul_le_smul_iff_of_pos_left, ha] + eq_of_forall_le_iff <| by simp [smul_le_smul_iff_of_pos_left, ha] end FloorDiv @@ -147,11 +147,11 @@ section CeilDiv variable [CeilDiv α β] {a : α} @[simp] lemma ceilDiv_one [Nontrivial α] (b : β) : b ⌈/⌉ (1 : α) = b := - eq_of_forall_ge_iff $ fun c ↦ by simp [zero_lt_one' α] + eq_of_forall_ge_iff <| fun c ↦ by simp [zero_lt_one' α] @[simp] lemma smul_ceilDiv [PosSMulMono α β] [PosSMulReflectLE α β] (ha : 0 < a) (b : β) : a • b ⌈/⌉ a = b := - eq_of_forall_ge_iff $ by simp [smul_le_smul_iff_of_pos_left, ha] + eq_of_forall_ge_iff <| by simp [smul_le_smul_iff_of_pos_left, ha] end CeilDiv @@ -177,14 +177,14 @@ namespace Nat instance instFloorDiv : FloorDiv ℕ ℕ where floorDiv := HDiv.hDiv floorDiv_gc a ha := by simpa [mul_comm] using Nat.galoisConnection_mul_div ha - floorDiv_nonpos a ha b := by rw [ha.antisymm $ zero_le _, Nat.div_zero] + floorDiv_nonpos a ha b := by rw [ha.antisymm <| zero_le _, Nat.div_zero] zero_floorDiv := Nat.zero_div instance instCeilDiv : CeilDiv ℕ ℕ where ceilDiv a b := (a + b - 1) / b ceilDiv_gc a ha b c := by - simp [div_le_iff_le_mul_add_pred ha, add_assoc, tsub_add_cancel_of_le $ succ_le_iff.2 ha] - ceilDiv_nonpos a ha b := by simp_rw [ha.antisymm $ zero_le _, Nat.div_zero] + simp [div_le_iff_le_mul_add_pred ha, add_assoc, tsub_add_cancel_of_le <| succ_le_iff.2 ha] + ceilDiv_nonpos a ha b := by simp_rw [ha.antisymm <| zero_le _, Nat.div_zero] zero_ceilDiv a := by cases a <;> simp [Nat.div_eq_zero_iff] @[simp] lemma floorDiv_eq_div (a b : ℕ) : a ⌊/⌋ b = a / b := rfl diff --git a/Mathlib/Algebra/Order/Group/Basic.lean b/Mathlib/Algebra/Order/Group/Basic.lean index 09aeb718a03a4..a5778d8524f10 100644 --- a/Mathlib/Algebra/Order/Group/Basic.lean +++ b/Mathlib/Algebra/Order/Group/Basic.lean @@ -89,7 +89,7 @@ that here because importing that definition would create import cycles."] lemma zpow_left_injective (hn : n ≠ 0) : Injective ((· ^ n) : α → α) := by obtain hn | hn := hn.lt_or_lt · refine fun a b (hab : a ^ n = b ^ n) ↦ - (zpow_strictMono_left _ $ Int.neg_pos_of_neg hn).injective ?_ + (zpow_strictMono_left _ <| Int.neg_pos_of_neg hn).injective ?_ rw [zpow_neg, zpow_neg, hab] · exact (zpow_strictMono_left _ hn).injective diff --git a/Mathlib/Algebra/Order/Group/PosPart.lean b/Mathlib/Algebra/Order/Group/PosPart.lean index 8afd1f0c6ea35..078e76ef92cb4 100644 --- a/Mathlib/Algebra/Order/Group/PosPart.lean +++ b/Mathlib/Algebra/Order/Group/PosPart.lean @@ -206,7 +206,7 @@ variable [LinearOrder α] [Group α] {a : α} rw [oneLePart, ← maxDefault, ← sup_eq_maxDefault]; simp_rw [sup_comm] @[to_additive (attr := simp) posPart_pos_iff] lemma one_lt_oneLePart_iff : 1 < a⁺ᵐ ↔ 1 < a := - lt_iff_lt_of_le_iff_le $ (one_le_oneLePart _).le_iff_eq.trans oneLePart_eq_one + lt_iff_lt_of_le_iff_le <| (one_le_oneLePart _).le_iff_eq.trans oneLePart_eq_one @[to_additive posPart_eq_of_posPart_pos] lemma oneLePart_of_one_lt_oneLePart (ha : 1 < a⁺ᵐ) : a⁺ᵐ = a := by @@ -219,7 +219,7 @@ variable [CovariantClass α α (· * ·) (· ≤ ·)] simp_rw [← one_le_inv']; rw [leOnePart, ← maxDefault, ← sup_eq_maxDefault]; simp_rw [sup_comm] @[to_additive (attr := simp) negPart_pos_iff] lemma one_lt_ltOnePart_iff : 1 < a⁻ᵐ ↔ a < 1 := - lt_iff_lt_of_le_iff_le $ (one_le_leOnePart _).le_iff_eq.trans leOnePart_eq_one + lt_iff_lt_of_le_iff_le <| (one_le_leOnePart _).le_iff_eq.trans leOnePart_eq_one end covariantmul end LinearOrder diff --git a/Mathlib/Algebra/Order/Module/Algebra.lean b/Mathlib/Algebra/Order/Module/Algebra.lean index 079e1ce13f6fa..083743a9ceaa4 100644 --- a/Mathlib/Algebra/Order/Module/Algebra.lean +++ b/Mathlib/Algebra/Order/Module/Algebra.lean @@ -91,7 +91,7 @@ def evalAlgebraMap : PositivityExt where eval {u β} _zβ _pβ e := do let _instβring ← synthInstanceQ q(OrderedSemiring $β) let _instαβsmul ← synthInstanceQ q(SMulPosMono $α $β) assertInstancesCommute - return .nonnegative q(algebraMap_nonneg $β $ le_of_lt $pa) + return .nonnegative q(algebraMap_nonneg $β <| le_of_lt $pa) | .nonnegative pa => let _instαring ← synthInstanceQ q(OrderedCommSemiring $α) let _instβring ← synthInstanceQ q(OrderedSemiring $β) diff --git a/Mathlib/Algebra/Order/Module/Rat.lean b/Mathlib/Algebra/Order/Module/Rat.lean index 4837629b8c8b4..0368ae68efa89 100644 --- a/Mathlib/Algebra/Order/Module/Rat.lean +++ b/Mathlib/Algebra/Order/Module/Rat.lean @@ -33,7 +33,7 @@ variable [LinearOrderedSemifield α] instance LinearOrderedSemifield.toPosSMulStrictMono_rat : PosSMulStrictMono ℚ≥0 α where elim q hq a b hab := by - rw [NNRat.smul_def, NNRat.smul_def]; exact mul_lt_mul_of_pos_left hab $ NNRat.cast_pos.2 hq + rw [NNRat.smul_def, NNRat.smul_def]; exact mul_lt_mul_of_pos_left hab <| NNRat.cast_pos.2 hq end LinearOrderedSemifield @@ -42,6 +42,6 @@ variable [LinearOrderedField α] instance LinearOrderedField.toPosSMulStrictMono_rat : PosSMulStrictMono ℚ α where elim q hq a b hab := by - rw [Rat.smul_def, Rat.smul_def]; exact mul_lt_mul_of_pos_left hab $ Rat.cast_pos.2 hq + rw [Rat.smul_def, Rat.smul_def]; exact mul_lt_mul_of_pos_left hab <| Rat.cast_pos.2 hq end LinearOrderedField diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean index 0646b0b825b56..87f6d53965914 100644 --- a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean @@ -112,7 +112,7 @@ theorem bot_eq_one : (⊥ : α) = 1 := le_antisymm bot_le (one_le ⊥) @[to_additive] instance CanonicallyOrderedCommMonoid.toUniqueUnits : Unique αˣ where - uniq a := Units.ext ((mul_eq_one_iff_of_one_le (α := α) (one_le _) $ one_le _).1 a.mul_inv).1 + uniq a := Units.ext ((mul_eq_one_iff_of_one_le (α := α) (one_le _) <| one_le _).1 a.mul_inv).1 @[deprecated (since := "2024-07-24")] alias mul_eq_one_iff := mul_eq_one @[deprecated (since := "2024-07-24")] alias add_eq_zero_iff := add_eq_zero diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index f834a11db01cb..a7a48f42842b7 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -111,13 +111,13 @@ lemma abs_lt_of_sq_lt_sq (h : a ^ 2 < b ^ 2) (hb : 0 ≤ b) : |a| < b := by rwa [← abs_of_nonneg hb, ← sq_lt_sq] lemma abs_lt_of_sq_lt_sq' (h : a ^ 2 < b ^ 2) (hb : 0 ≤ b) : -b < a ∧ a < b := - abs_lt.1 $ abs_lt_of_sq_lt_sq h hb + abs_lt.1 <| abs_lt_of_sq_lt_sq h hb lemma abs_le_of_sq_le_sq (h : a ^ 2 ≤ b ^ 2) (hb : 0 ≤ b) : |a| ≤ b := by rwa [← abs_of_nonneg hb, ← sq_le_sq] lemma abs_le_of_sq_le_sq' (h : a ^ 2 ≤ b ^ 2) (hb : 0 ≤ b) : -b ≤ a ∧ a ≤ b := - abs_le.1 $ abs_le_of_sq_le_sq h hb + abs_le.1 <| abs_le_of_sq_le_sq h hb lemma sq_eq_sq_iff_abs_eq_abs (a b : α) : a ^ 2 = b ^ 2 ↔ |a| = |b| := by simp only [le_antisymm_iff, sq_le_sq] diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index a4780f5c89c08..84c8ec4672fa0 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -269,12 +269,12 @@ lemma add_pow_le (ha : 0 ≤ a) (hb : 0 ≤ b) : ∀ n, (a + b) ^ n ≤ 2 ^ (n - rw [pow_succ] calc _ ≤ 2 ^ n * (a ^ (n + 1) + b ^ (n + 1)) * (a + b) := - mul_le_mul_of_nonneg_right (add_pow_le ha hb (n + 1)) $ add_nonneg ha hb + mul_le_mul_of_nonneg_right (add_pow_le ha hb (n + 1)) <| add_nonneg ha hb _ = 2 ^ n * (a ^ (n + 2) + b ^ (n + 2) + (a ^ (n + 1) * b + b ^ (n + 1) * a)) := by rw [mul_assoc, mul_add, add_mul, add_mul, ← pow_succ, ← pow_succ, add_comm _ (b ^ _), add_add_add_comm, add_comm (_ * a)] _ ≤ 2 ^ n * (a ^ (n + 2) + b ^ (n + 2) + (a ^ (n + 1) * a + b ^ (n + 1) * b)) := - mul_le_mul_of_nonneg_left (add_le_add_left ?_ _) $ pow_nonneg (zero_le_two (α := R)) _ + mul_le_mul_of_nonneg_left (add_le_add_left ?_ _) <| pow_nonneg (zero_le_two (α := R)) _ _ = _ := by simp only [← pow_succ, ← two_mul, ← mul_assoc]; rfl · obtain hab | hba := le_total a b · exact mul_add_mul_le_mul_add_mul (pow_le_pow_left ha hab _) hab @@ -289,7 +289,7 @@ protected lemma Even.add_pow_le (hn : Even n) : _ = 2 ^ n * (a ^ 2 + b ^ 2) ^ n := by -- TODO: Should be `Nat.cast_commute` rw [Commute.mul_pow]; simp [Commute, SemiconjBy, two_mul, mul_two] _ ≤ 2 ^ n * (2 ^ (n - 1) * ((a ^ 2) ^ n + (b ^ 2) ^ n)) := mul_le_mul_of_nonneg_left - (add_pow_le (sq_nonneg _) (sq_nonneg _) _) $ pow_nonneg (zero_le_two (α := R)) _ + (add_pow_le (sq_nonneg _) (sq_nonneg _) _) <| pow_nonneg (zero_le_two (α := R)) _ _ = _ := by simp only [← mul_assoc, ← pow_add, ← pow_mul] cases n diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean index 66f0f84ca3589..91e49342408c0 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean @@ -151,7 +151,7 @@ lemma pow_le_pow_of_le_one [ZeroLEOneClass α] [PosMulMono α] [MulPosMono α] | _, _, Nat.le.refl => le_rfl | _, _, Nat.le.step h => by rw [pow_succ'] - exact (mul_le_of_le_one_left (pow_nonneg ha₀ _) ha₁).trans $ pow_le_pow_of_le_one ha₀ ha₁ h + exact (mul_le_of_le_one_left (pow_nonneg ha₀ _) ha₁).trans <| pow_le_pow_of_le_one ha₀ ha₁ h lemma pow_le_of_le_one [ZeroLEOneClass α] [PosMulMono α] [MulPosMono α] (h₀ : 0 ≤ a) (h₁ : a ≤ 1) {n : ℕ} (hn : n ≠ 0) : a ^ n ≤ a := diff --git a/Mathlib/Algebra/Order/Ring/WithTop.lean b/Mathlib/Algebra/Order/Ring/WithTop.lean index bcea8dcf42272..8cea38308e3b5 100644 --- a/Mathlib/Algebra/Order/Ring/WithTop.lean +++ b/Mathlib/Algebra/Order/Ring/WithTop.lean @@ -30,10 +30,10 @@ instance instMulZeroClass : MulZeroClass (WithTop α) where | ⊤, (b : α) => if b = 0 then 0 else ⊤ | ⊤, ⊤ => ⊤ mul_zero a := match a with - | (a : α) => congr_arg some $ mul_zero _ + | (a : α) => congr_arg some <| mul_zero _ | ⊤ => if_pos rfl zero_mul b := match b with - | (b : α) => congr_arg some $ zero_mul _ + | (b : α) => congr_arg some <| zero_mul _ | ⊤ => if_pos rfl @[simp, norm_cast] lemma coe_mul (a b : α) : (↑(a * b) : WithTop α) = a * b := rfl diff --git a/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean b/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean index af0a168b4849f..78ee6ce26e4a0 100644 --- a/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean +++ b/Mathlib/Algebra/Polynomial/Degree/TrailingDegree.lean @@ -122,7 +122,7 @@ theorem natTrailingDegree_le_of_ne_zero (h : coeff p n ≠ 0) : natTrailingDegre constructor · rintro h by_contra hp - obtain ⟨n, hpn, hn⟩ := by simpa using min_mem_image_coe $ support_nonempty.2 hp + obtain ⟨n, hpn, hn⟩ := by simpa using min_mem_image_coe <| support_nonempty.2 hp obtain rfl := (trailingDegree_eq_iff_natTrailingDegree_eq hp).1 hn.symm exact hpn h · rintro rfl @@ -139,7 +139,7 @@ lemma trailingDegree_eq_zero : trailingDegree p = 0 ↔ coeff p 0 ≠ 0 := simp [natTrailingDegree, or_comm] lemma natTrailingDegree_ne_zero : natTrailingDegree p ≠ 0 ↔ p ≠ 0 ∧ coeff p 0 = 0 := - natTrailingDegree_eq_zero.not.trans $ by rw [not_or, not_ne_iff] + natTrailingDegree_eq_zero.not.trans <| by rw [not_or, not_ne_iff] lemma trailingDegree_ne_zero : trailingDegree p ≠ 0 ↔ coeff p 0 = 0 := trailingDegree_eq_zero.not_left diff --git a/Mathlib/Algebra/Polynomial/EraseLead.lean b/Mathlib/Algebra/Polynomial/EraseLead.lean index 5c50263b65651..8697d21e6eb6a 100644 --- a/Mathlib/Algebra/Polynomial/EraseLead.lean +++ b/Mathlib/Algebra/Polynomial/EraseLead.lean @@ -202,7 +202,7 @@ theorem eraseLead_natDegree_le (f : R[X]) : (eraseLead f).natDegree ≤ f.natDeg lemma natDegree_eraseLead (h : f.nextCoeff ≠ 0) : f.eraseLead.natDegree = f.natDegree - 1 := by have := natDegree_pos_of_nextCoeff_ne_zero h - refine f.eraseLead_natDegree_le.antisymm $ le_natDegree_of_ne_zero ?_ + refine f.eraseLead_natDegree_le.antisymm <| le_natDegree_of_ne_zero ?_ rwa [eraseLead_coeff_of_ne _ (tsub_lt_self _ _).ne, ← nextCoeff_of_natDegree_pos] all_goals positivity diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index 5732a7c7bd649..c3fd1bd83e3f6 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -130,7 +130,7 @@ theorem natDegree_pow (p : R[X]) (n : ℕ) : natDegree (p ^ n) = n * natDegree p classical obtain rfl | hp := eq_or_ne p 0 · obtain rfl | hn := eq_or_ne n 0 <;> simp [*] - exact natDegree_pow' $ by + exact natDegree_pow' <| by rw [← leadingCoeff_pow, Ne, leadingCoeff_eq_zero]; exact pow_ne_zero _ hp theorem degree_le_mul_left (p : R[X]) (hq : q ≠ 0) : degree p ≤ degree (p * q) := by diff --git a/Mathlib/Algebra/Ring/Hom/Defs.lean b/Mathlib/Algebra/Ring/Hom/Defs.lean index e24b81bf70f81..6ffa90e13d5da 100644 --- a/Mathlib/Algebra/Ring/Hom/Defs.lean +++ b/Mathlib/Algebra/Ring/Hom/Defs.lean @@ -598,8 +598,8 @@ instance instMonoid : Monoid (α →+* α) where mul_one := comp_id one_mul := id_comp mul_assoc f g h := comp_assoc _ _ _ - npow n f := (npowRec n f).copy f^[n] $ by induction n <;> simp [npowRec, *] - npow_succ n f := DFunLike.coe_injective $ Function.iterate_succ _ _ + npow n f := (npowRec n f).copy f^[n] <| by induction n <;> simp [npowRec, *] + npow_succ n f := DFunLike.coe_injective <| Function.iterate_succ _ _ @[simp, norm_cast] lemma coe_pow (f : α →+* α) (n : ℕ) : ⇑(f ^ n) = f^[n] := rfl diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index 3efcda0789876..ef9ba083a8241 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -143,7 +143,7 @@ lemma Odd.pow_add_pow_eq_zero [IsCancelAdd α] (hn : Odd n) (hab : a + b = 0) : obtain ⟨k, rfl⟩ := hn induction' k with k ih · simpa - have : a ^ 2 = b ^ 2 := add_right_cancel $ + have : a ^ 2 = b ^ 2 := add_right_cancel <| calc a ^ 2 + a * b = 0 := by rw [sq, ← mul_add, hab, mul_zero] _ = b ^ 2 + a * b := by rw [sq, ← add_mul, add_comm, hab, zero_mul] diff --git a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean index 39f6878821861..6233014ed00a7 100644 --- a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean +++ b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean @@ -136,7 +136,7 @@ def tildeInType : Sheaf (Type u) (PrimeSpectrum.Top R) := instance (U : (Opens (PrimeSpectrum.Top R))ᵒᵖ) : AddCommGroup (M.tildeInType.1.obj U) := - inferInstanceAs $ AddCommGroup (Tilde.sectionsSubmodule M U) + inferInstanceAs <| AddCommGroup (Tilde.sectionsSubmodule M U) /-- `M^~` as a presheaf of abelian groups over `Spec R` @@ -159,7 +159,7 @@ def tildeInAddCommGrp : Sheaf AddCommGrp (PrimeSpectrum.Top R) := noncomputable instance (U : (Opens (PrimeSpectrum.Top R))ᵒᵖ) : Module ((Spec (.of R)).ringCatSheaf.1.obj U) (M.tildeInAddCommGrp.1.obj U) := - inferInstanceAs $ Module _ (Tilde.sectionsSubmodule M U) + inferInstanceAs <| Module _ (Tilde.sectionsSubmodule M U) /-- `M^~` as a sheaf of `𝒪_{Spec R}`-modules diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index 00a1dd8d05033..f1b9f8a1215c8 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -37,12 +37,12 @@ theorem IsAffineOpen.fromSpecStalk_eq {X : Scheme} (x : X) {U V : X.Opens} Opens.isBasis_iff_nbhd.mp (isBasis_affine_open X) (show x ∈ U ⊓ V from ⟨hxU, hxV⟩) transitivity fromSpecStalk h₁ h₂ · delta fromSpecStalk - rw [← hU.map_fromSpec h₁ (homOfLE $ h₃.trans inf_le_left).op] + rw [← hU.map_fromSpec h₁ (homOfLE <| h₃.trans inf_le_left).op] erw [← Scheme.Spec_map (X.presheaf.map _).op, ← Scheme.Spec_map (X.presheaf.germ ⟨x, h₂⟩).op] rw [← Functor.map_comp_assoc, ← op_comp, TopCat.Presheaf.germ_res, Scheme.Spec_map, Quiver.Hom.unop_op] · delta fromSpecStalk - rw [← hV.map_fromSpec h₁ (homOfLE $ h₃.trans inf_le_right).op] + rw [← hV.map_fromSpec h₁ (homOfLE <| h₃.trans inf_le_right).op] erw [← Scheme.Spec_map (X.presheaf.map _).op, ← Scheme.Spec_map (X.presheaf.germ ⟨x, h₂⟩).op] rw [← Functor.map_comp_assoc, ← op_comp, TopCat.Presheaf.germ_res, Scheme.Spec_map, Quiver.Hom.unop_op] diff --git a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean index 06bd224d0b51a..248d3fa9e31ac 100644 --- a/Mathlib/Analysis/CStarAlgebra/Multiplier.lean +++ b/Mathlib/Analysis/CStarAlgebra/Multiplier.lean @@ -559,8 +559,8 @@ theorem norm_fst_eq_snd (a : 𝓜(𝕜, A)) : ‖a.fst‖ = ‖a.snd‖ := by intro b convert mul_le_mul_right' (mul_le_mul_left' (f.le_opNNNorm b) C) ‖b‖₊ using 1 ring - have := NNReal.div_le_of_le_mul $ f.opNNNorm_le_bound _ $ by - simpa only [sqrt_sq, sqrt_mul] using fun b ↦ sqrt_le_sqrt.2 $ (h b).trans (h1 b) + have := NNReal.div_le_of_le_mul <| f.opNNNorm_le_bound _ <| by + simpa only [sqrt_sq, sqrt_mul] using fun b ↦ sqrt_le_sqrt.2 <| (h b).trans (h1 b) convert NNReal.rpow_le_rpow this two_pos.le · simp only [NNReal.rpow_two, div_pow, sq_sqrt] simp only [sq, mul_self_div_self] diff --git a/Mathlib/Analysis/Calculus/Deriv/Shift.lean b/Mathlib/Analysis/Calculus/Deriv/Shift.lean index ae0d346bba6f0..45f45b6335f23 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Shift.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Shift.lean @@ -57,7 +57,7 @@ lemma deriv_comp_add_const : deriv (fun x ↦ f (x + a)) x = deriv f (x + a) := simpa [add_comm] using deriv_comp_const_add f a x lemma deriv_comp_const_sub : deriv (fun x ↦ f (a - x)) x = -deriv f (a - x) := by - simp_rw [sub_eq_add_neg, deriv_comp_neg (f $ a + ·), deriv_comp_const_add] + simp_rw [sub_eq_add_neg, deriv_comp_neg (f <| a + ·), deriv_comp_const_add] lemma deriv_comp_sub_const : deriv (fun x ↦ f (x - a)) x = deriv f (x - a) := by simp_rw [sub_eq_add_neg, deriv_comp_add_const] diff --git a/Mathlib/Analysis/Convex/Combination.lean b/Mathlib/Analysis/Convex/Combination.lean index a1b480de2d161..20fcc69a29fb2 100644 --- a/Mathlib/Analysis/Convex/Combination.lean +++ b/Mathlib/Analysis/Convex/Combination.lean @@ -376,7 +376,7 @@ lemma Finset.mem_convexHull' {s : Finset E} {x : E} : x ∈ convexHull R (s : Set E) ↔ ∃ w : E → R, (∀ y ∈ s, 0 ≤ w y) ∧ ∑ y ∈ s, w y = 1 ∧ ∑ y ∈ s, w y • y = x := by rw [mem_convexHull] - refine exists_congr fun w ↦ and_congr_right' $ and_congr_right fun hw ↦ ?_ + refine exists_congr fun w ↦ and_congr_right' <| and_congr_right fun hw ↦ ?_ simp_rw [centerMass_eq_of_sum_1 _ _ hw, id_eq] theorem Set.Finite.convexHull_eq {s : Set E} (hs : s.Finite) : convexHull R s = @@ -517,13 +517,13 @@ variable {s t t₁ t₂ : Finset E} lemma AffineIndependent.convexHull_inter (hs : AffineIndependent R ((↑) : s → E)) (ht₁ : t₁ ⊆ s) (ht₂ : t₂ ⊆ s) : convexHull R (t₁ ∩ t₂ : Set E) = convexHull R t₁ ∩ convexHull R t₂ := by - refine (Set.subset_inter (convexHull_mono inf_le_left) $ + refine (Set.subset_inter (convexHull_mono inf_le_left) <| convexHull_mono inf_le_right).antisymm ?_ simp_rw [Set.subset_def, mem_inter_iff, Set.inf_eq_inter, ← coe_inter, mem_convexHull'] rintro x ⟨⟨w₁, h₁w₁, h₂w₁, h₃w₁⟩, w₂, -, h₂w₂, h₃w₂⟩ let w (x : E) : R := (if x ∈ t₁ then w₁ x else 0) - if x ∈ t₂ then w₂ x else 0 have h₁w : ∑ i ∈ s, w i = 0 := by simp [w, Finset.inter_eq_right.2, *] - replace hs := hs.eq_zero_of_sum_eq_zero_subtype h₁w $ by + replace hs := hs.eq_zero_of_sum_eq_zero_subtype h₁w <| by simp only [w, sub_smul, zero_smul, ite_smul, Finset.sum_sub_distrib, ← Finset.sum_filter, h₃w₁, Finset.filter_mem_eq_inter, Finset.inter_eq_right.2 ht₁, Finset.inter_eq_right.2 ht₂, h₃w₂, sub_self] @@ -593,7 +593,7 @@ lemma mem_convexHull_pi (h : ∀ i ∈ s, x i ∈ convexHull 𝕜 (t i)) : x ∈ @[simp] lemma convexHull_pi (s : Set ι) (t : Π i, Set (E i)) : convexHull 𝕜 (s.pi t) = s.pi (fun i ↦ convexHull 𝕜 (t i)) := - Set.Subset.antisymm (convexHull_min (Set.pi_mono fun _ _ ↦ subset_convexHull _ _) $ convex_pi $ + Set.Subset.antisymm (convexHull_min (Set.pi_mono fun _ _ ↦ subset_convexHull _ _) <| convex_pi <| fun _ _ ↦ convex_convexHull _ _) fun _ ↦ mem_convexHull_pi end pi diff --git a/Mathlib/Analysis/Convex/Cone/Closure.lean b/Mathlib/Analysis/Convex/Cone/Closure.lean index a467f4ce5c3a5..2e5daf051b7a5 100644 --- a/Mathlib/Analysis/Convex/Cone/Closure.lean +++ b/Mathlib/Analysis/Convex/Cone/Closure.lean @@ -51,7 +51,7 @@ variable {E : Type*} [AddCommMonoid E] [TopologicalSpace E] [ContinuousAdd E] [M [ContinuousConstSMul 𝕜 E] lemma toConvexCone_closure_pointed (K : PointedCone 𝕜 E) : (K : ConvexCone 𝕜 E).closure.Pointed := - subset_closure $ PointedCone.toConvexCone_pointed _ + subset_closure <| PointedCone.toConvexCone_pointed _ /-- The closure of a pointed cone inside a topological space as a pointed cone. This construction is mainly used for defining maps between proper cones. -/ diff --git a/Mathlib/Analysis/Convex/Deriv.lean b/Mathlib/Analysis/Convex/Deriv.lean index 20066fe9abaa8..6954ba6c48f2b 100644 --- a/Mathlib/Analysis/Convex/Deriv.lean +++ b/Mathlib/Analysis/Convex/Deriv.lean @@ -239,7 +239,7 @@ lemma convexOn_of_hasDerivWithinAt2_nonneg {D : Set ℝ} (hD : Convex ℝ D) {f convert hf''₀ _ hx using 1 dsimp rw [deriv_eqOn isOpen_interior (fun y hy ↦ ?_) hx] - exact (hf'' _ hy).congr this $ by rw [this hy] + exact (hf'' _ hy).congr this <| by rw [this hy] /-- If a function `f` is continuous on a convex set `D ⊆ ℝ`, is twice differentiable on its interior, and `f''` is nonpositive on the interior, then `f` is concave on `D`. -/ @@ -255,7 +255,7 @@ lemma concaveOn_of_hasDerivWithinAt2_nonpos {D : Set ℝ} (hD : Convex ℝ D) {f convert hf''₀ _ hx using 1 dsimp rw [deriv_eqOn isOpen_interior (fun y hy ↦ ?_) hx] - exact (hf'' _ hy).congr this $ by rw [this hy] + exact (hf'' _ hy).congr this <| by rw [this hy] /-- If a function `f` is continuous on a convex set `D ⊆ ℝ` and `f''` is strictly positive on the interior, then `f` is strictly convex on `D`. diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index 8792d83e9e577..5f43d060381e7 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -418,7 +418,7 @@ theorem nonneg_inner_and_areaForm_eq_zero_iff_sameRay (x y : E) : have hx' : 0 < ‖x‖ := by simpa using hx have ha' : 0 ≤ a := nonneg_of_mul_nonneg_left ha (by positivity) have hb' : b = 0 := eq_zero_of_ne_zero_of_mul_right_eq_zero (pow_ne_zero 2 hx'.ne') hb - exact (SameRay.sameRay_nonneg_smul_right x ha').add_right $ by simp [hb'] + exact (SameRay.sameRay_nonneg_smul_right x ha').add_right <| by simp [hb'] · intro h obtain ⟨r, hr, rfl⟩ := h.exists_nonneg_left hx simp only [inner_smul_right, real_inner_self_eq_norm_sq, LinearMap.map_smulₛₗ, diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index 0f52b6382904c..521077e3726ed 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -823,10 +823,10 @@ lemma inner_le_weight_mul_Lp_of_nonneg (s : Finset ι) {p : ℝ} (hp : 1 ≤ p) · cases' H' with H' H' <;> simp [H', -one_div, -sum_eq_zero_iff, -rpow_eq_zero_iff, H] replace H' : (∀ i ∈ s, w i ≠ ⊤) ∧ ∀ i ∈ s, w i * f i ^ p ≠ ⊤ := by simpa [rpow_eq_top_iff,hp₀, hp₁, hp₀.not_lt, hp₁.not_lt, sum_eq_top_iff, not_or] using H' - have := coe_le_coe.2 $ NNReal.inner_le_weight_mul_Lp s hp.le (fun i ↦ ENNReal.toNNReal (w i)) + have := coe_le_coe.2 <| NNReal.inner_le_weight_mul_Lp s hp.le (fun i ↦ ENNReal.toNNReal (w i)) fun i ↦ ENNReal.toNNReal (f i) rw [coe_mul] at this - simp_rw [← coe_rpow_of_nonneg _ $ inv_nonneg.2 hp₀.le, coe_finset_sum, ENNReal.toNNReal_rpow, + simp_rw [← coe_rpow_of_nonneg _ <| inv_nonneg.2 hp₀.le, coe_finset_sum, ENNReal.toNNReal_rpow, ← ENNReal.toNNReal_mul, sum_congr rfl fun i hi ↦ coe_toNNReal (H'.2 i hi)] at this simp [← ENNReal.coe_rpow_of_nonneg, hp₀.le, hp₁.le] at this convert this using 2 with i hi diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index c401ebaa128e2..c0e8c5ba1564c 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -843,7 +843,7 @@ instance (priority := 100) NormedDivisionRing.to_topologicalDivisionRing : TopologicalDivisionRing α where protected lemma IsOfFinOrder.norm_eq_one (ha : IsOfFinOrder a) : ‖a‖ = 1 := - ((normHom : α →*₀ ℝ).toMonoidHom.isOfFinOrder ha).eq_one $ norm_nonneg _ + ((normHom : α →*₀ ℝ).toMonoidHom.isOfFinOrder ha).eq_one <| norm_nonneg _ example [Monoid β] (φ : β →* α) {x : β} {k : ℕ+} (h : x ^ (k : ℕ) = 1) : ‖φ x‖ = 1 := (φ.isOfFinOrder <| isOfFinOrder_iff_pow_eq_one.2 ⟨_, k.2, h⟩).norm_eq_one diff --git a/Mathlib/Analysis/Normed/Order/UpperLower.lean b/Mathlib/Analysis/Normed/Order/UpperLower.lean index 824536ea40231..b30f015561105 100644 --- a/Mathlib/Analysis/Normed/Order/UpperLower.lean +++ b/Mathlib/Analysis/Normed/Order/UpperLower.lean @@ -142,8 +142,8 @@ lemma dist_anti_right_pi : AntitoneOn (dist x) (Iic x) := by lemma dist_le_dist_of_le_pi (ha : a₂ ≤ a₁) (h₁ : a₁ ≤ b₁) (hb : b₁ ≤ b₂) : dist a₁ b₁ ≤ dist a₂ b₂ := - (dist_mono_right_pi h₁ (h₁.trans hb) hb).trans $ - dist_anti_left_pi (ha.trans $ h₁.trans hb) (h₁.trans hb) ha + (dist_mono_right_pi h₁ (h₁.trans hb) hb).trans <| + dist_anti_left_pi (ha.trans <| h₁.trans hb) (h₁.trans hb) ha theorem IsUpperSet.exists_subset_ball (hs : IsUpperSet s) (hx : x ∈ closure s) (hδ : 0 < δ) : ∃ y, closedBall y (δ / 4) ⊆ closedBall x δ ∧ closedBall y (δ / 4) ⊆ interior s := by @@ -213,7 +213,7 @@ protected lemma IsClosed.lowerClosure_pi (hs : IsClosed s) (hs' : BddAbove s) : haveI : BoundedGENhdsClass ℝ := by infer_instance obtain ⟨a, ha⟩ := hx.bddBelow_range obtain ⟨b, hb, φ, hφ, hbf⟩ := tendsto_subseq_of_bounded (hs'.isBounded_inter bddBelow_Ici) fun n ↦ - ⟨hg n, (ha $ mem_range_self _).trans $ hfg _⟩ + ⟨hg n, (ha <| mem_range_self _).trans <| hfg _⟩ exact ⟨b, closure_minimal inter_subset_left hs hb, le_of_tendsto_of_tendsto' (hx.comp hφ.tendsto_atTop) hbf fun _ ↦ hfg _⟩ @@ -225,14 +225,14 @@ protected lemma IsClopen.lowerClosure_pi (hs : IsClopen s) (hs' : BddAbove s) : lemma closure_upperClosure_comm_pi (hs : BddBelow s) : closure (upperClosure s : Set (ι → ℝ)) = upperClosure (closure s) := - (closure_minimal (upperClosure_anti subset_closure) $ - isClosed_closure.upperClosure_pi hs.closure).antisymm $ + (closure_minimal (upperClosure_anti subset_closure) <| + isClosed_closure.upperClosure_pi hs.closure).antisymm <| upperClosure_min (closure_mono subset_upperClosure) (upperClosure s).upper.closure lemma closure_lowerClosure_comm_pi (hs : BddAbove s) : closure (lowerClosure s : Set (ι → ℝ)) = lowerClosure (closure s) := - (closure_minimal (lowerClosure_mono subset_closure) $ - isClosed_closure.lowerClosure_pi hs.closure).antisymm $ + (closure_minimal (lowerClosure_mono subset_closure) <| + isClosed_closure.lowerClosure_pi hs.closure).antisymm <| lowerClosure_min (closure_mono subset_lowerClosure) (lowerClosure s).lower.closure end Finite diff --git a/Mathlib/Analysis/RCLike/Basic.lean b/Mathlib/Analysis/RCLike/Basic.lean index 5fc1882b59565..8031b85a253d9 100644 --- a/Mathlib/Analysis/RCLike/Basic.lean +++ b/Mathlib/Analysis/RCLike/Basic.lean @@ -418,7 +418,7 @@ theorem mul_conj (z : K) : z * conj z = ‖z‖ ^ 2 := by theorem conj_mul (z : K) : conj z * z = ‖z‖ ^ 2 := by rw [mul_comm, mul_conj] lemma inv_eq_conj (hz : ‖z‖ = 1) : z⁻¹ = conj z := - inv_eq_of_mul_eq_one_left $ by simp_rw [conj_mul, hz, algebraMap.coe_one, one_pow] + inv_eq_of_mul_eq_one_left <| by simp_rw [conj_mul, hz, algebraMap.coe_one, one_pow] theorem normSq_sub (z w : K) : normSq (z - w) = normSq z + normSq w - 2 * re (z * conj w) := by simp only [normSq_add, sub_eq_add_neg, map_neg, mul_neg, normSq_neg, map_neg] diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean index 567ae939e14d0..57b6747ca5744 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean @@ -107,7 +107,7 @@ theorem arg_cos_add_sin_mul_I {θ : ℝ} (hθ : θ ∈ Set.Ioc (-π) π) : arg ( lemma arg_exp_mul_I (θ : ℝ) : arg (exp (θ * I)) = toIocMod (mul_pos two_pos Real.pi_pos) (-π) θ := by convert arg_cos_add_sin_mul_I (θ := toIocMod (mul_pos two_pos Real.pi_pos) (-π) θ) _ using 2 - · rw [← exp_mul_I, eq_sub_of_add_eq $ toIocMod_add_toIocDiv_zsmul _ _ θ, ofReal_sub, + · rw [← exp_mul_I, eq_sub_of_add_eq <| toIocMod_add_toIocDiv_zsmul _ _ θ, ofReal_sub, ofReal_zsmul, ofReal_mul, ofReal_ofNat, exp_mul_I_periodic.sub_zsmul_eq] · convert toIocMod_mem_Ioc _ _ _ ring @@ -312,7 +312,7 @@ lemma abs_eq_one_iff' : abs x = 1 ↔ ∃ θ ∈ Set.Ioc (-π) π, exp (θ * I) refine ⟨toIocMod (mul_pos two_pos Real.pi_pos) (-π) θ, ?_, ?_⟩ · convert toIocMod_mem_Ioc _ _ _ ring - · rw [eq_sub_of_add_eq $ toIocMod_add_toIocDiv_zsmul _ _ θ, ofReal_sub, + · rw [eq_sub_of_add_eq <| toIocMod_add_toIocDiv_zsmul _ _ θ, ofReal_sub, ofReal_zsmul, ofReal_mul, ofReal_ofNat, exp_mul_I_periodic.sub_zsmul_eq] · rintro ⟨θ, _, rfl⟩ exact ⟨θ, rfl⟩ diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean index cc9089213f44b..a67d9921c5ba3 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean @@ -56,7 +56,7 @@ lemma le_sin (hx : x ≤ 0) : x ≤ sin x := by simpa using sin_le <| neg_nonneg lemma one_sub_sq_div_two_le_cos : 1 - x ^ 2 / 2 ≤ cos x := by wlog hx₀ : 0 ≤ x - · simpa using this $ neg_nonneg.2 $ le_of_not_le hx₀ + · simpa using this <| neg_nonneg.2 <| le_of_not_le hx₀ suffices MonotoneOn (fun x ↦ cos x + x ^ 2 / 2) (Ici 0) by simpa using this left_mem_Ici hx₀ hx₀ refine monotoneOn_of_hasDerivWithinAt_nonneg @@ -69,12 +69,12 @@ lemma one_sub_sq_div_two_le_cos : 1 - x ^ 2 / 2 ≤ cos x := by lemma two_div_pi_mul_le_sin (hx₀ : 0 ≤ x) (hx : x ≤ π / 2) : 2 / π * x ≤ sin x := by rw [← sub_nonneg] suffices ConcaveOn ℝ (Icc 0 (π / 2)) (fun x ↦ sin x - 2 / π * x) by - refine (le_min ?_ ?_).trans $ this.min_le_of_mem_Icc ⟨hx₀, hx⟩ <;> field_simp + refine (le_min ?_ ?_).trans <| this.min_le_of_mem_Icc ⟨hx₀, hx⟩ <;> field_simp exact concaveOn_of_hasDerivWithinAt2_nonpos (convex_Icc ..) - (Continuous.continuousOn $ by fun_prop) - (fun x _ ↦ ((hasDerivAt_sin ..).sub $ (hasDerivAt_id ..).const_mul (2 / π)).hasDerivWithinAt) + (Continuous.continuousOn <| by fun_prop) + (fun x _ ↦ ((hasDerivAt_sin ..).sub <| (hasDerivAt_id ..).const_mul (2 / π)).hasDerivWithinAt) (fun x _ ↦ (hasDerivAt_cos ..).hasDerivWithinAt.sub_const _) - fun x hx ↦ neg_nonpos.2 $ sin_nonneg_of_mem_Icc $ Icc_subset_Icc_right (by linarith) $ + fun x hx ↦ neg_nonpos.2 <| sin_nonneg_of_mem_Icc <| Icc_subset_Icc_right (by linarith) <| interior_subset hx /-- **Jordan's inequality** for negative values. -/ @@ -88,11 +88,11 @@ lemma one_sub_two_div_pi_mul_le_cos (hx₀ : 0 ≤ x) (hx : x ≤ π / 2) : 1 - lemma cos_quadratic_upper_bound (hx : |x| ≤ π) : cos x ≤ 1 - 2 / π ^ 2 * x ^ 2 := by wlog hx₀ : 0 ≤ x - · simpa using this (by rwa [abs_neg]) $ neg_nonneg.2 $ le_of_not_le hx₀ + · simpa using this (by rwa [abs_neg]) <| neg_nonneg.2 <| le_of_not_le hx₀ rw [abs_of_nonneg hx₀] at hx -- TODO: `compute_deriv` tactic? have hderiv (x) : HasDerivAt (fun x ↦ 1 - 2 / π ^ 2 * x ^ 2 - cos x) _ x := - (((hasDerivAt_pow ..).const_mul _).const_sub _).sub $ hasDerivAt_cos _ + (((hasDerivAt_pow ..).const_mul _).const_sub _).sub <| hasDerivAt_cos _ simp only [Nat.cast_ofNat, Nat.succ_sub_succ_eq_sub, tsub_zero, pow_one, ← neg_sub', neg_sub, ← mul_assoc] at hderiv have hmono : MonotoneOn (fun x ↦ 1 - 2 / π ^ 2 * x ^ 2 - cos x) (Icc 0 (π / 2)) := by @@ -100,7 +100,7 @@ lemma cos_quadratic_upper_bound (hx : |x| ≤ π) : cos x ≤ 1 - 2 / π ^ 2 * x set_option tactic.skipAssignedInstances false in refine monotoneOn_of_hasDerivWithinAt_nonneg (convex_Icc ..) - (Continuous.continuousOn $ by fun_prop) + (Continuous.continuousOn <| by fun_prop) (fun x _ ↦ (hderiv _).hasDerivWithinAt) fun x hx ↦ sub_nonneg.2 ?_ have ⟨hx₀, hx⟩ := interior_subset hx @@ -113,20 +113,20 @@ lemma cos_quadratic_upper_bound (hx : |x| ≤ π) : cos x ≤ 1 - 2 / π ^ 2 * x -- Compiles without this option, but somewhat slower. set_option tactic.skipAssignedInstances false in refine concaveOn_of_hasDerivWithinAt2_nonpos (convex_Icc ..) - (Continuous.continuousOn $ by fun_prop) (fun x _ ↦ (hderiv _).hasDerivWithinAt) - (fun x _ ↦ ((hasDerivAt_sin ..).sub $ (hasDerivAt_id ..).const_mul _).hasDerivWithinAt) + (Continuous.continuousOn <| by fun_prop) (fun x _ ↦ (hderiv _).hasDerivWithinAt) + (fun x _ ↦ ((hasDerivAt_sin ..).sub <| (hasDerivAt_id ..).const_mul _).hasDerivWithinAt) fun x hx ↦ ?_ have ⟨hx, hx'⟩ := interior_subset hx calc _ ≤ (0 : ℝ) - 0 := by gcongr - · exact cos_nonpos_of_pi_div_two_le_of_le hx $ hx'.trans $ by linarith + · exact cos_nonpos_of_pi_div_two_le_of_le hx <| hx'.trans <| by linarith · positivity _ = 0 := sub_zero _ rw [← sub_nonneg] obtain hx' | hx' := le_total x (π / 2) - · simpa using hmono (left_mem_Icc.2 $ by positivity) ⟨hx₀, hx'⟩ hx₀ - · refine (le_min ?_ ?_).trans $ hconc.min_le_of_mem_Icc ⟨hx', hx⟩ <;> field_simp <;> norm_num + · simpa using hmono (left_mem_Icc.2 <| by positivity) ⟨hx₀, hx'⟩ hx₀ + · refine (le_min ?_ ?_).trans <| hconc.min_le_of_mem_Icc ⟨hx', hx⟩ <;> field_simp <;> norm_num /-- For 0 < x ≤ 1 we have x - x ^ 3 / 4 < sin x. diff --git a/Mathlib/CategoryTheory/Category/PartialFun.lean b/Mathlib/CategoryTheory/Category/PartialFun.lean index b0c1ab026a13c..099e4390c830c 100644 --- a/Mathlib/CategoryTheory/Category/PartialFun.lean +++ b/Mathlib/CategoryTheory/Category/PartialFun.lean @@ -145,7 +145,7 @@ noncomputable def partialFunEquivPointed : PartialFun.{u} ≌ Pointed := · intro h split_ifs at h with ha rw [some_inj] at h - exact ⟨b, ⟨ha, h.symm⟩, rfl⟩) $ + exact ⟨b, ⟨ha, h.symm⟩, rfl⟩) <| NatIso.ofComponents (fun X ↦ Pointed.Iso.mk (by classical exact Equiv.optionSubtypeNe X.point) (by rfl)) fun {X Y} f ↦ Pointed.Hom.ext <| funext fun a ↦ by diff --git a/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean index 8a6c50484bdcd..dd4c4c861c62c 100644 --- a/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean +++ b/Mathlib/CategoryTheory/Limits/ConcreteCategory/WithAlgebraicStructures.lean @@ -70,8 +70,8 @@ lemma colimit_no_zero_smul_divisor obtain ⟨j'', H⟩ := H simpa [elementwise_of% (colimit.w F), map_zero] using congr(colimit.ι F _ $(H (IsFiltered.sup {j, j', j''} { ⟨j, j', by simp, by simp, i⟩ }) - (IsFiltered.toSup _ _ $ by simp) - (F.map (IsFiltered.toSup _ _ $ by simp) x) + (IsFiltered.toSup _ _ <| by simp) + (F.map (IsFiltered.toSup _ _ <| by simp) x) (by rw [← IsFiltered.toSup_commutes (f := i) (mY := by simp) (mf := by simp), F.map_comp, comp_apply, ← map_smul, ← map_smul, h, map_zero]))) diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean index 0046f90cc5f9e..30c00a8a00135 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean @@ -141,7 +141,7 @@ lemma IsMulFreimanIso.threeGPFree_congr (hf : IsMulFreimanIso 2 s t f) : obtain ⟨a, ha, rfl⟩ := hf.bijOn.surjOn hfa obtain ⟨b, hb, rfl⟩ := hf.bijOn.surjOn hfb obtain ⟨c, hc, rfl⟩ := hf.bijOn.surjOn hfc - exact congr_arg f $ hs ha hb hc $ (hf.mul_eq_mul ha hc hb hb).1 habc + exact congr_arg f <| hs ha hb hc <| (hf.mul_eq_mul ha hc hb hb).1 habc @[to_additive] theorem ThreeGPFree.image' [FunLike F α β] [MulHomClass F α β] (f : F) (hf : (s * s).InjOn f) @@ -184,7 +184,7 @@ lemma ThreeGPFree.eq_right (hs : ThreeGPFree s) : @[to_additive] theorem ThreeGPFree.smul_set (hs : ThreeGPFree s) : ThreeGPFree (a • s) := by rintro _ ⟨b, hb, rfl⟩ _ ⟨c, hc, rfl⟩ _ ⟨d, hd, rfl⟩ h - exact congr_arg (a • ·) $ hs hb hc hd $ by simpa [mul_mul_mul_comm _ _ a] using h + exact congr_arg (a • ·) <| hs hb hc hd <| by simpa [mul_mul_mul_comm _ _ a] using h @[to_additive] lemma threeGPFree_smul_set : ThreeGPFree (a • s) ↔ ThreeGPFree s where mp hs b hb c hc d hd h := mul_left_cancel @@ -214,7 +214,7 @@ variable [CancelCommMonoidWithZero α] [NoZeroDivisors α] {s : Set α} {a : α} lemma ThreeGPFree.smul_set₀ (hs : ThreeGPFree s) (ha : a ≠ 0) : ThreeGPFree (a • s) := by rintro _ ⟨b, hb, rfl⟩ _ ⟨c, hc, rfl⟩ _ ⟨d, hd, rfl⟩ h - exact congr_arg (a • ·) $ hs hb hc hd $ by simpa [mul_mul_mul_comm _ _ a, ha] using h + exact congr_arg (a • ·) <| hs hb hc hd <| by simpa [mul_mul_mul_comm _ _ a, ha] using h theorem threeGPFree_smul_set₀ (ha : a ≠ 0) : ThreeGPFree (a • s) ↔ ThreeGPFree s := ⟨fun hs b hb c hc d hd h ↦ @@ -441,7 +441,7 @@ theorem addRothNumber_Ico (a b : ℕ) : addRothNumber (Ico a b) = rothNumberNat lemma Fin.addRothNumber_eq_rothNumberNat (hkn : 2 * k ≤ n) : addRothNumber (Iio k : Finset (Fin n.succ)) = rothNumberNat k := - IsAddFreimanIso.addRothNumber_congr $ mod_cast isAddFreimanIso_Iio two_ne_zero hkn + IsAddFreimanIso.addRothNumber_congr <| mod_cast isAddFreimanIso_Iio two_ne_zero hkn lemma Fin.addRothNumber_le_rothNumberNat (k n : ℕ) (hkn : k ≤ n) : addRothNumber (Iio k : Finset (Fin n.succ)) ≤ rothNumberNat k := by diff --git a/Mathlib/Combinatorics/Additive/Corner/Defs.lean b/Mathlib/Combinatorics/Additive/Corner/Defs.lean index e9cd2b2dd2ed4..08295c48887f6 100644 --- a/Mathlib/Combinatorics/Additive/Corner/Defs.lean +++ b/Mathlib/Combinatorics/Additive/Corner/Defs.lean @@ -55,7 +55,7 @@ lemma IsCorner.mono (hAB : A ⊆ B) (hA : IsCorner A x₁ y₁ x₂ y₂) : IsCo add_eq_add := hA.add_eq_add lemma IsCornerFree.mono (hAB : A ⊆ B) (hB : IsCornerFree B) : IsCornerFree A := - fun _x₁ _y₁ _x₂ _y₂ hxyd ↦ hB $ hxyd.mono hAB + fun _x₁ _y₁ _x₂ _y₂ hxyd ↦ hB <| hxyd.mono hAB @[simp] lemma not_isCorner_empty : ¬ IsCorner ∅ x₁ y₁ x₂ y₂ := by simp [isCorner_iff] @@ -76,7 +76,7 @@ lemma IsCorner.image (hf : IsAddFreimanHom 2 s t f) (hAs : (A : Set (G × G)) lemma IsCornerFree.of_image (hf : IsAddFreimanHom 2 s t f) (hf' : s.InjOn f) (hAs : (A : Set (G × G)) ⊆ s ×ˢ s) (hA : IsCornerFree (Prod.map f f '' A)) : IsCornerFree A := fun _x₁ _y₁ _x₂ _y₂ hxy ↦ - hf' (hAs hxy.fst_fst_mem).1 (hAs hxy.snd_fst_mem).1 $ hA $ hxy.image hf hAs + hf' (hAs hxy.fst_fst_mem).1 (hAs hxy.snd_fst_mem).1 <| hA <| hxy.image hf hAs lemma isCorner_image (hf : IsAddFreimanIso 2 s t f) (hAs : A ⊆ s ×ˢ s) (hx₁ : x₁ ∈ s) (hy₁ : y₁ ∈ s) (hx₂ : x₂ ∈ s) (hy₂ : y₂ ∈ s) : diff --git a/Mathlib/Combinatorics/Additive/Corner/Roth.lean b/Mathlib/Combinatorics/Additive/Corner/Roth.lean index cc3a7145d5a4d..69eec794bef92 100644 --- a/Mathlib/Combinatorics/Additive/Corner/Roth.lean +++ b/Mathlib/Combinatorics/Additive/Corner/Roth.lean @@ -53,7 +53,7 @@ private lemma noAccidental (hs : IsCornerFree (A : Set (G × G))) : NoAccidental (triangleIndices A) where eq_or_eq_or_eq a a' b b' c c' ha hb hc := by simp only [mk_mem_triangleIndices] at ha hb hc - exact .inl $ hs ⟨hc.1, hb.1, ha.1, hb.2.symm.trans ha.2⟩ + exact .inl <| hs ⟨hc.1, hb.1, ha.1, hb.2.symm.trans ha.2⟩ private lemma farFromTriangleFree_graph [Fintype G] [DecidableEq G] (hε : ε * card G ^ 2 ≤ A.card) : (graph <| triangleIndices A).FarFromTriangleFree (ε / 9) := by @@ -96,7 +96,7 @@ theorem corners_theorem (ε : ℝ) (hε : 0 < ε) (hG : cornersTheoremBound ε classical have h₁ := (farFromTriangleFree_graph hAε).le_card_cliqueFinset rw [card_triangles, card_triangleIndices] at h₁ - convert h₁.trans (Nat.cast_le.2 $ card_le_univ _) using 1 <;> simp <;> ring + convert h₁.trans (Nat.cast_le.2 <| card_le_univ _) using 1 <;> simp <;> ring /-- The **corners theorem** for `ℕ`. @@ -118,8 +118,8 @@ theorem corners_theorem_nat (hε : 0 < ε) (hn : cornersTheoremBound (ε / 9) omega rw [this] at hA have := Fin.isAddFreimanIso_Iio two_ne_zero (le_refl (2 * n)) - have := hA.of_image this.isAddFreimanHom Fin.val_injective.injOn $ by - refine Set.image_subset_iff.2 $ hAn.trans fun x hx ↦ ?_ + have := hA.of_image this.isAddFreimanHom Fin.val_injective.injOn <| by + refine Set.image_subset_iff.2 <| hAn.trans fun x hx ↦ ?_ simp only [coe_range, Set.mem_prod, Set.mem_Iio] at hx exact ⟨Fin.natCast_strictMono (by omega) hx.1, Fin.natCast_strictMono (by omega) hx.2⟩ rw [← coe_image] at this @@ -155,10 +155,10 @@ theorem roth_3ap_theorem (ε : ℝ) (hε : 0 < ε) (hG : cornersTheoremBound ε ∃ x₁ y₁ x₂ y₂, y₁ - x₁ ∈ A ∧ y₂ - x₁ ∈ A ∧ y₁ - x₂ ∈ A ∧ x₁ + y₂ = x₂ + y₁ ∧ x₁ ≠ x₂ := by simpa [IsCornerFree, isCorner_iff, B, -exists_and_left, -exists_and_right] using corners_theorem ε hε hG B this - have := hA hx₂y₁ hx₁y₁ hx₁y₂ $ by -- TODO: This really ought to just be `by linear_combination h` + have := hA hx₂y₁ hx₁y₁ hx₁y₂ <| by -- TODO: This really ought to just be `by linear_combination h` rw [sub_add_sub_comm, add_comm, add_sub_add_comm, add_right_cancel_iff, sub_eq_sub_iff_add_eq_add, add_comm, hxy, add_comm] - exact hx₁x₂ $ by simpa using this.symm + exact hx₁x₂ <| by simpa using this.symm /-- **Roth's theorem** for `ℕ`. @@ -176,8 +176,8 @@ theorem roth_3ap_theorem_nat (ε : ℝ) (hε : 0 < ε) (hG : cornersTheoremBound omega rw [this] at hA have := Fin.isAddFreimanIso_Iio two_ne_zero (le_refl (2 * n)) - have := hA.of_image this.isAddFreimanHom Fin.val_injective.injOn $ Set.image_subset_iff.2 $ - hAn.trans fun x hx ↦ Fin.natCast_strictMono (by omega) $ by + have := hA.of_image this.isAddFreimanHom Fin.val_injective.injOn <| Set.image_subset_iff.2 <| + hAn.trans fun x hx ↦ Fin.natCast_strictMono (by omega) <| by simpa only [coe_range, Set.mem_Iio] using hx rw [← coe_image] at this refine roth_3ap_theorem (ε / 3) (by positivity) (by simp; omega) _ ?_ this @@ -188,7 +188,7 @@ theorem roth_3ap_theorem_nat (ε : ℝ) (hε : 0 < ε) (hG : cornersTheoremBound _ ≤ A.card := hAε _ = _ := by rw [card_image_of_injOn] - exact (CharP.natCast_injOn_Iio (Fin (2 * n).succ) (2 * n).succ).mono $ hAn.trans $ by + exact (CharP.natCast_injOn_Iio (Fin (2 * n).succ) (2 * n).succ).mono <| hAn.trans <| by simp; omega open Asymptotics Filter diff --git a/Mathlib/Combinatorics/Additive/Dissociation.lean b/Mathlib/Combinatorics/Additive/Dissociation.lean index 153c28f9d1b21..bfb19d5419f00 100644 --- a/Mathlib/Combinatorics/Additive/Dissociation.lean +++ b/Mathlib/Combinatorics/Additive/Dissociation.lean @@ -39,7 +39,7 @@ def MulDissociated (s : Set α) : Prop := {t : Finset α | ↑t ⊆ s}.InjOn ( @[to_additive] lemma mulDissociated_iff_sum_eq_subsingleton : MulDissociated s ↔ ∀ a, {t : Finset α | ↑t ⊆ s ∧ ∏ x in t, x = a}.Subsingleton := - ⟨fun hs _ _t ht _u hu ↦ hs ht.1 hu.1 $ ht.2.trans hu.2.symm, + ⟨fun hs _ _t ht _u hu ↦ hs ht.1 hu.1 <| ht.2.trans hu.2.symm, fun hs _t ht _u hu htu ↦ hs _ ⟨ht, htu⟩ ⟨hu, rfl⟩⟩ @[to_additive] lemma MulDissociated.subset {t : Set α} (hst : s ⊆ t) (ht : MulDissociated t) : @@ -137,13 +137,13 @@ lemma exists_subset_mulSpan_card_le_of_forall_mulDissociated by_cases ha' : a ∈ s' · exact subset_mulSpan ha' obtain ⟨t, u, ht, hu, htu⟩ := not_mulDissociated_iff_exists_disjoint.1 fun h ↦ - hs'max _ (insert_subset_iff.2 ⟨ha, hs'.1⟩) h $ ssubset_insert ha' + hs'max _ (insert_subset_iff.2 ⟨ha, hs'.1⟩) h <| ssubset_insert ha' by_cases hat : a ∈ t · have : a = (∏ b in u, b) / ∏ b in t.erase a, b := by rw [prod_erase_eq_div hat, htu.2.2, div_div_self'] rw [this] exact prod_div_prod_mem_mulSpan - ((subset_insert_iff_of_not_mem $ disjoint_left.1 htu.1 hat).1 hu) (subset_insert_iff.1 ht) + ((subset_insert_iff_of_not_mem <| disjoint_left.1 htu.1 hat).1 hu) (subset_insert_iff.1 ht) rw [coe_subset, subset_insert_iff_of_not_mem hat] at ht by_cases hau : a ∈ u · have : a = (∏ b in t, b) / ∏ b in u.erase a, b := by diff --git a/Mathlib/Combinatorics/Additive/Energy.lean b/Mathlib/Combinatorics/Additive/Energy.lean index fdc2bce807bcc..341b4509c4465 100644 --- a/Mathlib/Combinatorics/Additive/Energy.lean +++ b/Mathlib/Combinatorics/Additive/Energy.lean @@ -132,7 +132,7 @@ variable {s t} @[to_additive] lemma mulEnergy_eq_sum_sq [Fintype α] (s t : Finset α) : Eₘ[s, t] = ∑ a, ((s ×ˢ t).filter fun (x, y) ↦ x * y = a).card ^ 2 := by rw [mulEnergy_eq_sum_sq'] - exact Fintype.sum_subset $ by aesop (add simp [filter_eq_empty_iff, mul_mem_mul]) + exact Fintype.sum_subset <| by aesop (add simp [filter_eq_empty_iff, mul_mem_mul]) @[to_additive card_sq_le_card_mul_addEnergy] lemma card_sq_le_card_mul_mulEnergy (s t u : Finset α) : diff --git a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean index 6359c2a1ca925..a3bad041f207e 100644 --- a/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean +++ b/Mathlib/Combinatorics/Additive/ErdosGinzburgZiv.lean @@ -67,14 +67,14 @@ private theorem ZMod.erdos_ginzburg_ziv_prime (a : ι → ZMod p) (hs : s.card = have hpN : p ∣ N := char_dvd_card_solutions_of_add_lt p (totalDegree_f₁_add_totalDegree_f₂.trans_eq hs') -- Hence, `2 ≤ p ≤ N` and we can make a common root `x ≠ 0`. - obtain ⟨x, hx⟩ := Fintype.exists_ne_of_one_lt_card ((Fact.out : p.Prime).one_lt.trans_le $ + obtain ⟨x, hx⟩ := Fintype.exists_ne_of_one_lt_card ((Fact.out : p.Prime).one_lt.trans_le <| Nat.le_of_dvd hN₀ hpN) zero_sol -- This common root gives us the required subsequence, namely the `i ∈ s` such that `x i ≠ 0`. refine ⟨(s.attach.filter fun a ↦ x.1 a ≠ 0).map ⟨(↑), Subtype.val_injective⟩, ?_, ?_, ?_⟩ · simp (config := { contextual := true }) [subset_iff] -- From `f₁ x = 0`, we get that `p` divides the number of `a` such that `x a ≠ 0`. · rw [card_map] - refine Nat.eq_of_dvd_of_lt_two_mul (Finset.card_pos.2 ?_).ne' ?_ $ + refine Nat.eq_of_dvd_of_lt_two_mul (Finset.card_pos.2 ?_).ne' ?_ <| (Finset.card_filter_le _ _).trans_lt ?_ -- This number is nonzero because `x ≠ 0`. · rw [← Subtype.coe_ne_coe, Function.ne_iff] at hx @@ -135,11 +135,11 @@ theorem Int.erdos_ginzburg_ziv (a : ι → ℤ) (hs : 2 * n - 1 ≤ s.card) : -- `t ∈ ℬ` of `(∑ i ∈ t, a i) / n` is divisible by `m`. obtain ⟨ℬ, hℬ𝒜, hℬcard, hℬ⟩ := ihm (fun t ↦ (∑ i ∈ t, a i) / n) h𝒜card.ge -- We are done. - refine ⟨ℬ.biUnion fun x ↦ x, biUnion_subset.2 fun t ht ↦ (h𝒜 $ hℬ𝒜 ht).1, ?_, ?_⟩ - · rw [card_biUnion (h𝒜disj.mono hℬ𝒜), sum_const_nat fun t ht ↦ (h𝒜 $ hℬ𝒜 ht).2.1, hℬcard] + refine ⟨ℬ.biUnion fun x ↦ x, biUnion_subset.2 fun t ht ↦ (h𝒜 <| hℬ𝒜 ht).1, ?_, ?_⟩ + · rw [card_biUnion (h𝒜disj.mono hℬ𝒜), sum_const_nat fun t ht ↦ (h𝒜 <| hℬ𝒜 ht).2.1, hℬcard] rwa [sum_biUnion, natCast_mul, mul_comm, ← Int.dvd_div_iff_mul_dvd, Int.sum_div] - · exact fun t ht ↦ (h𝒜 $ hℬ𝒜 ht).2.2 - · exact dvd_sum fun t ht ↦ (h𝒜 $ hℬ𝒜 ht).2.2 + · exact fun t ht ↦ (h𝒜 <| hℬ𝒜 ht).2.2 + · exact dvd_sum fun t ht ↦ (h𝒜 <| hℬ𝒜 ht).2.2 · exact h𝒜disj.mono hℬ𝒜 -- Now, let's find those `2 * m - 1` sets. rintro k hk @@ -165,12 +165,12 @@ theorem Int.erdos_ginzburg_ziv (a : ι → ℤ) (hs : 2 * n - 1 ≤ s.card) : have : t₀ ∉ 𝒜 := by rintro h obtain rfl : n = 0 := by - simpa [← card_eq_zero, ht₀card] using sdiff_disjoint.mono ht₀ $ subset_biUnion_of_mem id h + simpa [← card_eq_zero, ht₀card] using sdiff_disjoint.mono ht₀ <| subset_biUnion_of_mem id h omega refine ⟨𝒜.cons t₀ this, by rw [card_cons, h𝒜card], ?_, ?_⟩ · simp only [cons_eq_insert, coe_insert, Set.pairwise_insert_of_symmetric symmetric_disjoint, mem_coe, ne_eq] - exact ⟨h𝒜disj, fun t ht _ ↦ sdiff_disjoint.mono ht₀ $ subset_biUnion_of_mem id ht⟩ + exact ⟨h𝒜disj, fun t ht _ ↦ sdiff_disjoint.mono ht₀ <| subset_biUnion_of_mem id ht⟩ · simp only [cons_eq_insert, mem_insert, forall_eq_or_imp, and_assoc] exact ⟨ht₀.trans sdiff_subset, ht₀card, ht₀sum, h𝒜⟩ diff --git a/Mathlib/Combinatorics/Additive/FreimanHom.lean b/Mathlib/Combinatorics/Additive/FreimanHom.lean index 3b9fd4e291330..ceb7bc28e54f4 100644 --- a/Mathlib/Combinatorics/Additive/FreimanHom.lean +++ b/Mathlib/Combinatorics/Additive/FreimanHom.lean @@ -319,7 +319,7 @@ namespace Fin variable {k m n : ℕ} private lemma aux (hm : m ≠ 0) (hkmn : m * k ≤ n) : k < (n + 1) := - Nat.lt_succ_iff.2 $ le_trans (Nat.le_mul_of_pos_left _ hm.bot_lt) hkmn + Nat.lt_succ_iff.2 <| le_trans (Nat.le_mul_of_pos_left _ hm.bot_lt) hkmn /-- **No wrap-around principle**. @@ -335,7 +335,7 @@ lemma isAddFreimanIso_Iic (hm : m ≠ 0) (hkmn : m * k ≤ n) : have (u : Multiset (Fin (n + 1))) : Nat.castRingHom _ (u.map val).sum = u.sum := by simp rw [← this, ← this] have {u : Multiset (Fin (n + 1))} (huk : ∀ x ∈ u, x ≤ k) (hu : card u = m) : - (u.map val).sum < (n + 1) := Nat.lt_succ_iff.2 $ hkmn.trans' $ by + (u.map val).sum < (n + 1) := Nat.lt_succ_iff.2 <| hkmn.trans' <| by rw [← hu, ← card_map] refine sum_le_card_nsmul (u.map val) k ?_ simpa [le_iff_val_le_val, -val_fin_le, Nat.mod_eq_of_lt, aux hm hkmn] using huk diff --git a/Mathlib/Combinatorics/Colex.lean b/Mathlib/Combinatorics/Colex.lean index e4b1b34ce802f..1f8a439741e9c 100644 --- a/Mathlib/Combinatorics/Colex.lean +++ b/Mathlib/Combinatorics/Colex.lean @@ -345,15 +345,15 @@ lemma lt_iff_exists_filter_lt : have mem_u {w : α} : w ∈ u ↔ w ∈ t ∧ w ∉ s ∧ ∀ a ∈ s, a ∉ t → a < w := by simp [u, and_assoc] have hu : u.Nonempty := h.imp fun _ ↦ mem_u.2 let m := max' _ hu - have ⟨hmt, hms, hm⟩ : m ∈ t ∧ m ∉ s ∧ ∀ a ∈ s, a ∉ t → a < m := mem_u.1 $ max'_mem _ _ + have ⟨hmt, hms, hm⟩ : m ∈ t ∧ m ∉ s ∧ ∀ a ∈ s, a ∉ t → a < m := mem_u.1 <| max'_mem _ _ refine ⟨m, hmt, hms, fun a hma ↦ ⟨fun has ↦ not_imp_comm.1 (hm _ has) hma.asymm, fun hat ↦ ?_⟩⟩ by_contra has have hau : a ∈ u := mem_u.2 ⟨hat, has, fun b hbs hbt ↦ (hm _ hbs hbt).trans hma⟩ - exact hma.not_le $ le_max' _ _ hau + exact hma.not_le <| le_max' _ _ hau · rintro ⟨w, hwt, hws, hw⟩ refine ⟨w, hwt, hws, fun a has hat ↦ ?_⟩ by_contra! hwa - exact hat $ (hw $ hwa.lt_of_ne $ ne_of_mem_of_not_mem hwt hat).1 has + exact hat <| (hw <| hwa.lt_of_ne <| ne_of_mem_of_not_mem hwt hat).1 has /-- If `s ≤ t` in colex and `s.card ≤ t.card`, then `s \ {a} ≤ t \ {min t}` for any `a ∈ s`. -/ lemma erase_le_erase_min' (hst : toColex s ≤ toColex t) (hcard : s.card ≤ t.card) (ha : a ∈ s) : @@ -364,9 +364,9 @@ lemma erase_le_erase_min' (hst : toColex s ≤ toColex t) (hcard : s.card ≤ t. -- Case on whether `s = t` obtain rfl | h' := eq_or_ne s t -- If `s = t`, then `s \ {a} ≤ s \ {m}` because `m ≤ a` - · exact (erase_le_erase ha $ min'_mem _ _).2 $ min'_le _ _ $ ha + · exact (erase_le_erase ha <| min'_mem _ _).2 <| min'_le _ _ <| ha -- If `s ≠ t`, call `w` the colex witness. Case on whether `w < a` or `a < w` - replace hst := hst.lt_of_ne $ toColex_inj.not.2 h' + replace hst := hst.lt_of_ne <| toColex_inj.not.2 h' simp only [lt_iff_exists_filter_lt, mem_sdiff, filter_inj, and_assoc] at hst obtain ⟨w, hwt, hws, hw⟩ := hst obtain hwa | haw := (ne_of_mem_of_not_mem ha hws).symm.lt_or_lt @@ -379,14 +379,15 @@ lemma erase_le_erase_min' (hst : toColex s ≤ toColex t) (hcard : s.card ≤ t. obtain rfl | hbt := hbt · assumption · by_contra! hab - exact hbt $ (hw $ hwa.trans_le hab).1 $ mem_of_mem_erase hbs + exact hbt <| (hw <| hwa.trans_le hab).1 <| mem_of_mem_erase hbs -- If `a < w`, case on whether `m < w` or `m = w` obtain rfl | hmw : m = w ∨ m < w := (min'_le _ _ hwt).eq_or_lt -- If `m = w`, then `s \ {a} = t \ {m}` · have : erase t m ⊆ erase s a := by rintro b hb rw [mem_erase] at hb ⊢ - exact ⟨(haw.trans_le $ min'_le _ _ hb.2).ne', (hw $ hb.1.lt_of_le' $ min'_le _ _ hb.2).2 hb.2⟩ + exact ⟨(haw.trans_le <| min'_le _ _ hb.2).ne', + (hw <| hb.1.lt_of_le' <| min'_le _ _ hb.2).2 hb.2⟩ rw [eq_of_subset_of_card_le this] rw [card_erase_of_mem ha, card_erase_of_mem (min'_mem _ _)] exact tsub_le_tsub_right hcard _ @@ -398,7 +399,7 @@ lemma erase_le_erase_min' (hst : toColex s ≤ toColex t) (hcard : s.card ≤ t. obtain rfl | hbt := hbt · assumption · by_contra! hwb - exact hbt $ (hw $ hwb.lt_of_ne $ ne_of_mem_of_not_mem hwt hbt).1 $ mem_of_mem_erase hbs + exact hbt <| (hw <| hwb.lt_of_ne <| ne_of_mem_of_not_mem hwt hbt).1 <| mem_of_mem_erase hbs /-- Strictly monotone functions preserve the colex ordering. -/ lemma toColex_image_le_toColex_image (hf : StrictMono f) : diff --git a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean index d7ff51b86ec91..e27ea7862f429 100644 --- a/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean +++ b/Mathlib/Combinatorics/SetFamily/AhlswedeZhang.lean @@ -87,9 +87,9 @@ private lemma Fintype.sum_div_mul_card_choose_card : have (n) (hn : n ∈ range (card α + 1)) : ((card α).choose n / ((card α - n) * (card α).choose n) : ℚ) = (card α - n : ℚ)⁻¹ := by rw [div_mul_cancel_right₀] - exact cast_ne_zero.2 (choose_pos $ mem_range_succ_iff.1 hn).ne' + exact cast_ne_zero.2 (choose_pos <| mem_range_succ_iff.1 hn).ne' simp only [sum_congr rfl this, mul_eq_mul_left_iff, cast_eq_zero] - convert Or.inl $ sum_range_reflect _ _ with a ha + convert Or.inl <| sum_range_reflect _ _ with a ha rw [add_tsub_cancel_right, cast_sub (mem_range_succ_iff.mp ha)] end @@ -124,7 +124,7 @@ lemma truncatedSup_of_mem (h : a ∈ lowerClosure s) : lemma truncatedSup_of_not_mem (h : a ∉ lowerClosure s) : truncatedSup s a = ⊤ := dif_neg h -@[simp] lemma truncatedSup_empty (a : α) : truncatedSup ∅ a = ⊤ := truncatedSup_of_not_mem $ by simp +@[simp] lemma truncatedSup_empty (a : α) : truncatedSup ∅ a = ⊤ := truncatedSup_of_not_mem (by simp) @[simp] lemma truncatedSup_singleton (b a : α) : truncatedSup {b} a = if a ≤ b then b else ⊤ := by simp [truncatedSup]; split_ifs <;> simp [Finset.filter_true_of_mem, *] @@ -133,7 +133,7 @@ lemma le_truncatedSup : a ≤ truncatedSup s a := by rw [truncatedSup] split_ifs with h · obtain ⟨ℬ, hb, h⟩ := h - exact h.trans $ le_sup' id $ mem_filter.2 ⟨hb, h⟩ + exact h.trans <| le_sup' id <| mem_filter.2 ⟨hb, h⟩ · exact le_top lemma map_truncatedSup [@DecidableRel β (· ≤ ·)] (e : α ≃o β) (s : Finset α) (a : α) : @@ -201,10 +201,10 @@ lemma truncatedInf_le : truncatedInf s a ≤ a := by unfold truncatedInf split_ifs with h · obtain ⟨b, hb, hba⟩ := h - exact hba.trans' $ inf'_le id $ mem_filter.2 ⟨hb, ‹_›⟩ + exact hba.trans' <| inf'_le id <| mem_filter.2 ⟨hb, ‹_›⟩ · exact bot_le -@[simp] lemma truncatedInf_empty (a : α) : truncatedInf ∅ a = ⊥ := truncatedInf_of_not_mem $ by simp +@[simp] lemma truncatedInf_empty (a : α) : truncatedInf ∅ a = ⊥ := truncatedInf_of_not_mem (by simp) @[simp] lemma truncatedInf_singleton (b a : α) : truncatedInf {b} a = if b ≤ a then b else ⊥ := by simp only [truncatedInf, coe_singleton, upperClosure_singleton, UpperSet.mem_Ici_iff, @@ -245,7 +245,7 @@ lemma truncatedInf_union_right (hs : a ∉ upperClosure s) (ht : a ∈ upperClos lemma truncatedInf_union_of_not_mem (hs : a ∉ upperClosure s) (ht : a ∉ upperClosure t) : truncatedInf (s ∪ t) a = ⊥ := - truncatedInf_of_not_mem $ by rw [coe_union, upperClosure_union]; exact fun h ↦ h.elim hs ht + truncatedInf_of_not_mem <| by rw [coe_union, upperClosure_union]; exact fun h ↦ h.elim hs ht end SemilatticeInf @@ -277,11 +277,11 @@ lemma truncatedInf_sups (hs : a ∈ upperClosure s) (ht : a ∈ upperClosure t) lemma truncatedSup_infs_of_not_mem (ha : a ∉ lowerClosure s ⊓ lowerClosure t) : truncatedSup (s ⊼ t) a = ⊤ := - truncatedSup_of_not_mem $ by rwa [coe_infs, lowerClosure_infs] + truncatedSup_of_not_mem <| by rwa [coe_infs, lowerClosure_infs] lemma truncatedInf_sups_of_not_mem (ha : a ∉ upperClosure s ⊔ upperClosure t) : truncatedInf (s ⊻ t) a = ⊥ := - truncatedInf_of_not_mem $ by rwa [coe_sups, upperClosure_sups] + truncatedInf_of_not_mem <| by rwa [coe_sups, upperClosure_sups] end DistribLattice @@ -301,8 +301,8 @@ variable [DecidableEq α] [Fintype α] lemma card_truncatedSup_union_add_card_truncatedSup_infs (𝒜 ℬ : Finset (Finset α)) (s : Finset α) : (truncatedSup (𝒜 ∪ ℬ) s).card + (truncatedSup (𝒜 ⊼ ℬ) s).card = (truncatedSup 𝒜 s).card + (truncatedSup ℬ s).card := by - by_cases h𝒜 : s ∈ lowerClosure (𝒜 : Set $ Finset α) <;> - by_cases hℬ : s ∈ lowerClosure (ℬ : Set $ Finset α) + by_cases h𝒜 : s ∈ lowerClosure (𝒜 : Set <| Finset α) <;> + by_cases hℬ : s ∈ lowerClosure (ℬ : Set <| Finset α) · rw [truncatedSup_union h𝒜 hℬ, truncatedSup_infs h𝒜 hℬ] exact card_union_add_card_inter _ _ · rw [truncatedSup_union_left h𝒜 hℬ, truncatedSup_of_not_mem hℬ, @@ -315,8 +315,8 @@ lemma card_truncatedSup_union_add_card_truncatedSup_infs (𝒜 ℬ : Finset (Fin lemma card_truncatedInf_union_add_card_truncatedInf_sups (𝒜 ℬ : Finset (Finset α)) (s : Finset α) : (truncatedInf (𝒜 ∪ ℬ) s).card + (truncatedInf (𝒜 ⊻ ℬ) s).card = (truncatedInf 𝒜 s).card + (truncatedInf ℬ s).card := by - by_cases h𝒜 : s ∈ upperClosure (𝒜 : Set $ Finset α) <;> - by_cases hℬ : s ∈ upperClosure (ℬ : Set $ Finset α) + by_cases h𝒜 : s ∈ upperClosure (𝒜 : Set <| Finset α) <;> + by_cases hℬ : s ∈ upperClosure (ℬ : Set <| Finset α) · rw [truncatedInf_union h𝒜 hℬ, truncatedInf_sups h𝒜 hℬ] exact card_inter_add_card_union _ _ · rw [truncatedInf_union_left h𝒜 hℬ, truncatedInf_of_not_mem hℬ, @@ -365,7 +365,7 @@ lemma IsAntichain.le_infSum (h𝒜 : IsAntichain (· ⊆ ·) (𝒜 : Set (Finset _ ≤ _ := sum_le_univ_sum_of_nonneg fun s ↦ by positivity refine sum_congr rfl fun s hs ↦ ?_ rw [truncatedInf_of_isAntichain h𝒜 hs, div_mul_cancel_left₀] - have := (nonempty_iff_ne_empty.2 $ ne_of_mem_of_not_mem hs h𝒜₀).card_pos + have := (nonempty_iff_ne_empty.2 <| ne_of_mem_of_not_mem hs h𝒜₀).card_pos positivity variable [Nonempty α] diff --git a/Mathlib/Combinatorics/SimpleGraph/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Basic.lean index fff296e425ccc..6aaf4e25f4463 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Basic.lean @@ -816,7 +816,7 @@ theorem deleteEdges_deleteEdges (s s' : Set (Sym2 V)) : lemma deleteEdges_le (s : Set (Sym2 V)) : G.deleteEdges s ≤ G := sdiff_le lemma deleteEdges_anti (h : s₁ ⊆ s₂) : G.deleteEdges s₂ ≤ G.deleteEdges s₁ := - sdiff_le_sdiff_left $ fromEdgeSet_mono h + sdiff_le_sdiff_left <| fromEdgeSet_mono h lemma deleteEdges_mono (h : G ≤ H) : G.deleteEdges s ≤ H.deleteEdges s := sdiff_le_sdiff_right h diff --git a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean index 455f7edb1ac0e..20c6684e3f809 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Coloring.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Coloring.lean @@ -370,8 +370,8 @@ lemma chromaticNumber_eq_iff_forall_surjective (hG : G.Colorable n) : rw [← hG.chromaticNumber_le.ge_iff_eq, le_chromaticNumber_iff_forall_surjective] theorem chromaticNumber_bot [Nonempty V] : (⊥ : SimpleGraph V).chromaticNumber = 1 := by - have : (⊥ : SimpleGraph V).Colorable 1 := ⟨.mk 0 $ by simp⟩ - exact this.chromaticNumber_le.antisymm $ ENat.one_le_iff_pos.2 $ chromaticNumber_pos this + have : (⊥ : SimpleGraph V).Colorable 1 := ⟨.mk 0 <| by simp⟩ + exact this.chromaticNumber_le.antisymm <| ENat.one_le_iff_pos.2 <| chromaticNumber_pos this @[simp] theorem chromaticNumber_top [Fintype V] : (⊤ : SimpleGraph V).chromaticNumber = Fintype.card V := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean index bb84cf7219624..e24349799722e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity/WalkCounting.lean @@ -169,7 +169,7 @@ instance : Decidable G.Connected := by infer_instance instance instDecidableMemSupp (c : G.ConnectedComponent) (v : V) : Decidable (v ∈ c.supp) := - c.recOn (fun w ↦ decidable_of_iff (G.Reachable v w) $ by simp) + c.recOn (fun w ↦ decidable_of_iff (G.Reachable v w) <| by simp) (fun _ _ _ _ ↦ Subsingleton.elim _ _) lemma odd_card_iff_odd_components : Odd (Nat.card V) ↔ diff --git a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean index 1f9b4d949285c..769c8115eecfa 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean @@ -59,7 +59,7 @@ lemma IsHamiltonian.support_toFinset (hp : p.IsHamiltonian) : p.support.toFinset /-- The length of a hamiltonian path is one less than the number of vertices of the graph. -/ lemma IsHamiltonian.length_eq (hp : p.IsHamiltonian) : p.length = Fintype.card α - 1 := - eq_tsub_of_add_eq $ by + eq_tsub_of_add_eq <| by rw [← length_support, ← List.sum_toFinset_count_eq_length, Finset.sum_congr rfl fun _ _ ↦ hp _, ← card_eq_sum_ones, hp.support_toFinset, card_univ] @@ -128,7 +128,7 @@ def IsHamiltonian (G : SimpleGraph α) : Prop := lemma IsHamiltonian.mono {H : SimpleGraph α} (hGH : G ≤ H) (hG : G.IsHamiltonian) : H.IsHamiltonian := - fun hα ↦ let ⟨_, p, hp⟩ := hG hα; ⟨_, p.map $ .ofLE hGH, hp.map _ bijective_id⟩ + fun hα ↦ let ⟨_, p, hp⟩ := hG hα; ⟨_, p.map <| .ofLE hGH, hp.map _ bijective_id⟩ lemma IsHamiltonian.connected (hG : G.IsHamiltonian) : G.Connected where preconnected a b := by @@ -138,7 +138,7 @@ lemma IsHamiltonian.connected (hG : G.IsHamiltonian) : G.Connected where obtain ⟨_, p, hp⟩ := hG Fintype.one_lt_card.ne' have a_mem := hp.mem_support a have b_mem := hp.mem_support b - exact ((p.takeUntil a a_mem).reverse.append $ p.takeUntil b b_mem).reachable - nonempty := not_isEmpty_iff.1 fun _ ↦ by simpa using hG $ by simp [@Fintype.card_eq_zero] + exact ((p.takeUntil a a_mem).reverse.append <| p.takeUntil b b_mem).reachable + nonempty := not_isEmpty_iff.1 fun _ ↦ by simpa using hG <| by simp [@Fintype.card_eq_zero] end SimpleGraph diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean index eaf8ae304e6c0..34b41bc57eb0c 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Uniform.lean @@ -86,7 +86,7 @@ lemma isUniform_one : G.IsUniform (1 : 𝕜) s t := by variable {G} lemma IsUniform.pos (hG : G.IsUniform ε s t) : 0 < ε := - not_le.1 fun hε ↦ (hε.trans $ abs_nonneg _).not_lt $ hG (empty_subset _) (empty_subset _) + not_le.1 fun hε ↦ (hε.trans <| abs_nonneg _).not_lt <| hG (empty_subset _) (empty_subset _) (by simpa using mul_nonpos_of_nonneg_of_nonpos (Nat.cast_nonneg _) hε) (by simpa using mul_nonpos_of_nonneg_of_nonpos (Nat.cast_nonneg _) hε) @@ -302,7 +302,7 @@ lemma IsEquipartition.card_interedges_sparsePairs_le (hP : P.IsEquipartition) (h private lemma aux {i j : ℕ} (hj : 0 < j) : j * (j - 1) * (i / j + 1) ^ 2 < (i + j) ^ 2 := by have : j * (j - 1) < j ^ 2 := by rw [sq]; exact Nat.mul_lt_mul_of_pos_left (Nat.sub_lt hj zero_lt_one) hj - apply (Nat.mul_lt_mul_of_pos_right this $ pow_pos Nat.succ_pos' _).trans_le + apply (Nat.mul_lt_mul_of_pos_right this <| pow_pos Nat.succ_pos' _).trans_le rw [← mul_pow] exact Nat.pow_le_pow_of_le_left (add_le_add_right (Nat.mul_div_le i j) _) _ @@ -332,10 +332,10 @@ lemma IsEquipartition.card_biUnion_offDiag_le (hε : 0 < ε) (hP : P.IsEquiparti rw [div_le_iff (Nat.cast_pos.2 (P.parts_nonempty hA.ne_empty).card_pos)] have : (A.card : 𝕜) + P.parts.card ≤ 2 * A.card := by rw [two_mul]; exact add_le_add_left (Nat.cast_le.2 P.card_parts_le_card) _ - refine (mul_le_mul_of_nonneg_left this $ by positivity).trans ?_ + refine (mul_le_mul_of_nonneg_left this <| by positivity).trans ?_ suffices 1 ≤ ε/4 * P.parts.card by rw [mul_left_comm, ← sq] - convert mul_le_mul_of_nonneg_left this (mul_nonneg zero_le_two $ sq_nonneg (A.card : 𝕜)) + convert mul_le_mul_of_nonneg_left this (mul_nonneg zero_le_two <| sq_nonneg (A.card : 𝕜)) using 1 <;> ring rwa [← div_le_iff', one_div_div] positivity @@ -347,7 +347,7 @@ lemma IsEquipartition.sum_nonUniforms_lt' (hA : A.Nonempty) (hε : 0 < ε) (hP : _ ≤ (P.nonUniforms G ε).card • (↑(A.card / P.parts.card + 1) : 𝕜) ^ 2 := sum_le_card_nsmul _ _ _ ?_ _ = _ := nsmul_eq_mul _ _ - _ ≤ _ := mul_le_mul_of_nonneg_right hG $ by positivity + _ ≤ _ := mul_le_mul_of_nonneg_right hG <| by positivity _ < _ := ?_ · simp only [Prod.forall, Finpartition.mk_mem_nonUniforms, and_imp] rintro U V hU hV - - @@ -418,7 +418,7 @@ lemma unreduced_edges_subset : obtain rfl | hUV := eq_or_ne U V · exact Or.inr (Or.inl ⟨U, hU, hx, hy, G.ne_of_adj h⟩) by_cases h₂ : G.IsUniform (ε/8) U V - · exact Or.inr $ Or.inr ⟨U, V, hU, hV, hUV, h' _ hU _ hV hx hy hUV h₂, hx, hy, h⟩ + · exact Or.inr <| Or.inr ⟨U, V, hU, hV, hUV, h' _ hU _ hV hx hy hUV h₂, hx, hy, h⟩ · exact Or.inl ⟨U, V, hU, hV, hUV, h₂, hx, hy⟩ end SimpleGraph diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean index 9eb128cce610b..62b63093df2d8 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean @@ -52,7 +52,7 @@ protected lemma LocallyLinear.edgeDisjointTriangles : G.LocallyLinear → G.Edge And.left nonrec lemma EdgeDisjointTriangles.mono (h : G ≤ H) (hH : H.EdgeDisjointTriangles) : - G.EdgeDisjointTriangles := hH.mono $ cliqueSet_mono h + G.EdgeDisjointTriangles := hH.mono <| cliqueSet_mono h @[simp] lemma edgeDisjointTriangles_bot : (⊥ : SimpleGraph α).EdgeDisjointTriangles := by simp [EdgeDisjointTriangles] @@ -133,7 +133,7 @@ alias ⟨EdgeDisjointTriangles.mem_sym2_subsingleton, _⟩ := variable [DecidableEq α] [Fintype α] [DecidableRel G.Adj] instance EdgeDisjointTriangles.instDecidable : Decidable G.EdgeDisjointTriangles := - decidable_of_iff ((G.cliqueFinset 3 : Set (Finset α)).Pairwise fun x y ↦ ((x ∩ y).card ≤ 1)) $ by + decidable_of_iff ((G.cliqueFinset 3 : Set (Finset α)).Pairwise fun x y ↦ ((x ∩ y).card ≤ 1)) <| by simp only [coe_cliqueFinset, EdgeDisjointTriangles, Finset.card_le_one, ← coe_inter]; rfl instance LocallyLinear.instDecidable : Decidable G.LocallyLinear := And.decidable @@ -151,7 +151,7 @@ lemma EdgeDisjointTriangles.card_edgeFinset_le (hG : G.EdgeDisjointTriangles) : simp [insert_subset, *] · simpa only [card_le_one, mem_bipartiteBelow, and_imp, Set.Subsingleton, Set.mem_setOf_eq, mem_cliqueFinset_iff, mem_cliqueSet_iff] - using hG.mem_sym2_subsingleton (G.not_isDiag_of_mem_edgeSet $ mem_edgeFinset.1 he) + using hG.mem_sym2_subsingleton (G.not_isDiag_of_mem_edgeSet <| mem_edgeFinset.1 he) lemma LocallyLinear.card_edgeFinset (hG : G.LocallyLinear) : G.edgeFinset.card = 3 * (G.cliqueFinset 3).card := by @@ -165,7 +165,7 @@ lemma LocallyLinear.card_edgeFinset (hG : G.LocallyLinear) : rintro _ a b c hab hac hbc rfl calc _ ≤ ({s(a, b), s(a, c), s(b, c)} : Finset _).card := card_le_card ?_ - _ ≤ 3 := (card_insert_le _ _).trans (succ_le_succ $ (card_insert_le _ _).trans_eq $ by + _ ≤ 3 := (card_insert_le _ _).trans (succ_le_succ <| (card_insert_le _ _).trans_eq <| by rw [card_singleton]) simp only [subset_iff, Sym2.forall, mem_sym2_iff, le_eq_subset, mem_bipartiteBelow, mem_insert, mem_edgeFinset, mem_singleton, and_imp, mem_edgeSet, Sym2.mem_iff, forall_eq_or_imp, @@ -213,7 +213,7 @@ private lemma farFromTriangleFree_of_disjoint_triangles_aux {tris : Finset (Fins by_contra! h refine hH t ?_ simp only [not_and, mem_sdiff, not_not, mem_edgeFinset, mem_edgeSet] at h - obtain ⟨x, y, z, xy, xz, yz, rfl⟩ := is3Clique_iff.1 (mem_cliqueFinset_iff.1 $ htris ht) + obtain ⟨x, y, z, xy, xz, yz, rfl⟩ := is3Clique_iff.1 (mem_cliqueFinset_iff.1 <| htris ht) rw [is3Clique_triple_iff] refine ⟨h _ _ ?_ ?_ xy.ne xy, h _ _ ?_ ?_ xz.ne xz, h _ _ ?_ ?_ yz.ne yz⟩ <;> simp choose fx fy hfx hfy hfne fmem using this @@ -237,9 +237,9 @@ lemma farFromTriangleFree_of_disjoint_triangles (tris : Finset (Finset α)) G.FarFromTriangleFree ε := by rw [farFromTriangleFree_iff] intros H _ hG hH - rw [← Nat.cast_sub (card_le_card $ edgeFinset_mono hG)] + rw [← Nat.cast_sub (card_le_card <| edgeFinset_mono hG)] exact tris_big.trans - (Nat.cast_le.2 $ farFromTriangleFree_of_disjoint_triangles_aux htris pd hG hH) + (Nat.cast_le.2 <| farFromTriangleFree_of_disjoint_triangles_aux htris pd hG hH) protected lemma EdgeDisjointTriangles.farFromTriangleFree (hG : G.EdgeDisjointTriangles) (tris_big : ε * (card α ^ 2 : ℕ) ≤ (G.cliqueFinset 3).card) : @@ -275,7 +275,7 @@ lemma FarFromTriangleFree.lt_half (hG : G.FarFromTriangleFree ε) : ε < 2⁻¹ apply tsub_lt_self <;> positivity lemma FarFromTriangleFree.lt_one (hG : G.FarFromTriangleFree ε) : ε < 1 := - hG.lt_half.trans $ inv_lt_one one_lt_two + hG.lt_half.trans <| inv_lt_one one_lt_two theorem FarFromTriangleFree.nonpos (h₀ : G.FarFromTriangleFree ε) (h₁ : G.CliqueFree 3) : ε ≤ 0 := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean index f0f5b628e442c..41df6e05b9aba 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean @@ -29,13 +29,13 @@ namespace SimpleGraph /-- The vertices of `s` whose density in `t` is `ε` less than expected. -/ private noncomputable def badVertices (ε : ℝ) (s t : Finset α) : Finset α := - s.filter fun x ↦ (t.filter $ G.Adj x).card < (G.edgeDensity s t - ε) * t.card + s.filter fun x ↦ (t.filter <| G.Adj x).card < (G.edgeDensity s t - ε) * t.card private lemma card_interedges_badVertices_le : (Rel.interedges G.Adj (badVertices G ε s t) t).card ≤ (badVertices G ε s t).card * t.card * (G.edgeDensity s t - ε) := by classical - refine (Nat.cast_le.2 $ (card_le_card $ subset_of_eq (Rel.interedges_eq_biUnion _)).trans + refine (Nat.cast_le.2 <| (card_le_card <| subset_of_eq (Rel.interedges_eq_biUnion _)).trans card_biUnion_le).trans ?_ simp_rw [Nat.cast_sum, card_map, ← nsmul_eq_mul, smul_mul_assoc, mul_comm (t.card : ℝ)] exact sum_le_card_nsmul _ _ _ fun x hx ↦ (mem_filter.1 hx).2.le @@ -44,14 +44,14 @@ private lemma edgeDensity_badVertices_le (hε : 0 ≤ ε) (dst : 2 * ε ≤ G.ed G.edgeDensity (badVertices G ε s t) t ≤ G.edgeDensity s t - ε := by rw [edgeDensity_def] push_cast - refine div_le_of_nonneg_of_le_mul (by positivity) (sub_nonneg_of_le $ by linarith) ?_ + refine div_le_of_nonneg_of_le_mul (by positivity) (sub_nonneg_of_le <| by linarith) ?_ rw [mul_comm] exact G.card_interedges_badVertices_le private lemma card_badVertices_le (dst : 2 * ε ≤ G.edgeDensity s t) (hst : G.IsUniform ε s t) : (badVertices G ε s t).card ≤ s.card * ε := by have hε : ε ≤ 1 := (le_mul_of_one_le_of_le_of_nonneg (by norm_num) le_rfl hst.pos.le).trans - (dst.trans $ by exact_mod_cast edgeDensity_le_one _ _ _) + (dst.trans <| by exact_mod_cast edgeDensity_le_one _ _ _) by_contra! h have : |(G.edgeDensity (badVertices G ε s t) t - G.edgeDensity s t : ℝ)| < ε := hst (filter_subset _ _) Subset.rfl h.le (mul_le_of_le_one_right (Nat.cast_nonneg _) hε) @@ -61,7 +61,7 @@ private lemma card_badVertices_le (dst : 2 * ε ≤ G.edgeDensity s t) (hst : G. /-- A subset of the triangles constructed in a weird way to make them easy to count. -/ private lemma triangle_split_helper [DecidableEq α] : (s \ (badVertices G ε s t ∪ badVertices G ε s u)).biUnion - (fun x ↦ (G.interedges (t.filter $ G.Adj x) (u.filter $ G.Adj x)).image (x, ·)) ⊆ + (fun x ↦ (G.interedges (t.filter <| G.Adj x) (u.filter <| G.Adj x)).image (x, ·)) ⊆ (s ×ˢ t ×ˢ u).filter (fun (x, y, z) ↦ G.Adj x y ∧ G.Adj x z ∧ G.Adj y z) := by rintro ⟨x, y, z⟩ simp only [mem_filter, mem_product, mem_biUnion, mem_sdiff, exists_prop, mem_union, @@ -109,18 +109,18 @@ lemma triangle_counting' let X' := s \ (badVertices G ε s t ∪ badVertices G ε s u) have : X'.biUnion _ ⊆ (s ×ˢ t ×ˢ u).filter fun (a, b, c) ↦ G.Adj a b ∧ G.Adj a c ∧ G.Adj b c := triangle_split_helper _ - refine le_trans ?_ (Nat.cast_le.2 $ card_le_card this) + refine le_trans ?_ (Nat.cast_le.2 <| card_le_card this) rw [card_biUnion, Nat.cast_sum] - · apply le_trans _ (card_nsmul_le_sum X' _ _ $ G.good_vertices_triangle_card dst dsu dtu utu) + · apply le_trans _ (card_nsmul_le_sum X' _ _ <| G.good_vertices_triangle_card dst dsu dtu utu) rw [nsmul_eq_mul] have := hst.pos.le suffices hX' : (1 - 2 * ε) * s.card ≤ X'.card by - exact Eq.trans_le (by ring) (mul_le_mul_of_nonneg_right hX' $ by positivity) + exact Eq.trans_le (by ring) (mul_le_mul_of_nonneg_right hX' <| by positivity) have i : badVertices G ε s t ∪ badVertices G ε s u ⊆ s := union_subset (filter_subset _ _) (filter_subset _ _) rw [sub_mul, one_mul, card_sdiff i, Nat.cast_sub (card_le_card i), sub_le_sub_iff_left, mul_assoc, mul_comm ε, two_mul] - refine (Nat.cast_le.2 $ card_union_le _ _).trans ?_ + refine (Nat.cast_le.2 <| card_union_le _ _).trans ?_ rw [Nat.cast_add] exact add_le_add h₁ h₂ rintro a _ b _ t diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean index 96c9a38f6795c..7af960554a9a8 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Removal.lean @@ -61,7 +61,7 @@ private lemma card_bound (hP₁ : P.IsEquipartition) (hP₃ : P.parts.card ≤ b calc _ ≤ card α / (2 * P.parts.card : ℝ) := by gcongr _ ≤ ↑(card α / P.parts.card) := - (div_le_iff' (by positivity)).2 $ mod_cast (aux ‹_› P.card_parts_le_card).le + (div_le_iff' (by positivity)).2 <| mod_cast (aux ‹_› P.card_parts_le_card).le _ ≤ (s.card : ℝ) := mod_cast hP₁.average_le_card_part hX private lemma triangle_removal_aux (hε : 0 < ε) (hP₁ : P.IsEquipartition) diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean index dac9907861a7d..ceb4e3902e9d9 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean @@ -211,7 +211,7 @@ section Fintype variable [Fintype α] [Fintype β] [Fintype γ] lemma cliqueFinset_eq_image [NoAccidental t] : (graph t).cliqueFinset 3 = t.image toTriangle := - coe_injective $ by push_cast; exact cliqueSet_eq_image _ + coe_injective <| by push_cast; exact cliqueSet_eq_image _ lemma cliqueFinset_eq_map [NoAccidental t] : (graph t).cliqueFinset 3 = t.map toTriangle := by simp [cliqueFinset_eq_image, map_eq_image] @@ -224,7 +224,7 @@ lemma farFromTriangleFree [ExplicitDisjoint t] {ε : 𝕜} (graph t).FarFromTriangleFree ε := farFromTriangleFree_of_disjoint_triangles (t.map toTriangle) (map_subset_iff_subset_preimage.2 fun x hx ↦ by simpa using toTriangle_is3Clique hx) - (map_toTriangle_disjoint t) $ by simpa [add_assoc] using ht + (map_toTriangle_disjoint t) <| by simpa [add_assoc] using ht end Fintype end DecidableEq diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index 55c72b43472b4..0d8527495b4d2 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -1001,9 +1001,9 @@ theorem nat_strong_rec (f : α → ℕ → σ) {g : α → List σ → Option σ theorem listLookup [DecidableEq α] : Primrec₂ (List.lookup : α → List (α × β) → Option β) := (to₂ <| list_rec snd (const none) <| to₂ <| - cond (Primrec.beq.comp (fst.comp fst) (fst.comp $ fst.comp snd)) - (option_some.comp $ snd.comp $ fst.comp snd) - (snd.comp $ snd.comp snd)).of_eq + cond (Primrec.beq.comp (fst.comp fst) (fst.comp <| fst.comp snd)) + (option_some.comp <| snd.comp <| fst.comp snd) + (snd.comp <| snd.comp snd)).of_eq fun a ps => by induction' ps with p ps ih <;> simp [List.lookup, *] cases ha : a == p.1 <;> simp [ha] @@ -1016,7 +1016,7 @@ theorem nat_omega_rec' (f : β → σ) {m : β → ℕ} {l : β → List β} {g let mapGraph (M : List (β × σ)) (bs : List β) : List σ := bs.bind (Option.toList <| M.lookup ·) let bindList (b : β) : ℕ → List β := fun n ↦ n.rec [b] fun _ bs ↦ bs.bind l let graph (b : β) : ℕ → List (β × σ) := fun i ↦ i.rec [] fun i ih ↦ - (bindList b (m b - i)).filterMap fun b' ↦ (g b' $ mapGraph ih (l b')).map (b', ·) + (bindList b (m b - i)).filterMap fun b' ↦ (g b' <| mapGraph ih (l b')).map (b', ·) have mapGraph_primrec : Primrec₂ mapGraph := to₂ <| list_bind snd <| optionToList.comp₂ <| listLookup.comp₂ .right (fst.comp₂ .left) have bindList_primrec : Primrec₂ (bindList) := @@ -1028,9 +1028,9 @@ theorem nat_omega_rec' (f : β → σ) {m : β → ℕ} {l : β → List β} {g to₂ <| listFilterMap (bindList_primrec.comp (fst.comp fst) - (nat_sub.comp (hm.comp $ fst.comp fst) (fst.comp snd))) <| + (nat_sub.comp (hm.comp <| fst.comp fst) (fst.comp snd))) <| to₂ <| option_map - (hg.comp snd (mapGraph_primrec.comp (snd.comp $ snd.comp fst) (hl.comp snd))) + (hg.comp snd (mapGraph_primrec.comp (snd.comp <| snd.comp fst) (hl.comp snd))) (Primrec₂.pair.comp₂ (snd.comp₂ .left) .right) have : Primrec (fun b => ((graph b (m b + 1)).get? 0).map Prod.snd) := option_map (list_get?.comp (graph_primrec.comp Primrec.id (succ.comp hm)) (const 0)) @@ -1043,14 +1043,14 @@ theorem nat_omega_rec' (f : β → σ) {m : β → ℕ} {l : β → List β} {g induction' k with k ih <;> simp [bindList] intro a₂ a₁ ha₁ ha₂ have : k ≤ m b := - Nat.lt_succ.mp (by simpa using Nat.add_lt_of_lt_sub $ Nat.zero_lt_of_lt (ih a₁ ha₁)) + Nat.lt_succ.mp (by simpa using Nat.add_lt_of_lt_sub <| Nat.zero_lt_of_lt (ih a₁ ha₁)) have : m a₁ ≤ m b - k := Nat.lt_succ.mp (by rw [← Nat.succ_sub this]; simpa using ih a₁ ha₁) exact lt_of_lt_of_le (Ord a₁ a₂ ha₂) this List.eq_nil_iff_forall_not_mem.mpr (by intro b' ha'; by_contra; simpa using bindList_m_lt (m b + 1) b' ha') have mapGraph_graph {bs bs' : List β} (has : bs' ⊆ bs) : - mapGraph (bs.map $ fun x => (x, f x)) bs' = bs'.map f := by + mapGraph (bs.map <| fun x => (x, f x)) bs' = bs'.map f := by induction' bs' with b bs' ih <;> simp [mapGraph] · have : b ∈ bs ∧ bs' ⊆ bs := by simpa using has rcases this with ⟨ha, has'⟩ diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 3fdd10b7dd853..dbca5f2e696e6 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -232,7 +232,7 @@ theorem isCauSeq_re (f : CauSeq ℂ Complex.abs) : IsCauSeq abs' fun n => (f n). theorem isCauSeq_im (f : CauSeq ℂ Complex.abs) : IsCauSeq abs' fun n => (f n).im := fun ε ε0 => (f.cauchy ε0).imp fun i H j ij ↦ by - simpa only [← ofReal_sub, abs_ofReal, sub_re] using (abs_im_le_abs _).trans_lt $ H _ ij + simpa only [← ofReal_sub, abs_ofReal, sub_re] using (abs_im_le_abs _).trans_lt <| H _ ij /-- The real part of a complex Cauchy sequence, as a real Cauchy sequence. -/ noncomputable def cauSeqRe (f : CauSeq ℂ Complex.abs) : CauSeq ℝ abs' := diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index ec2e6c8e48d03..1eef0858bcbcf 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -111,10 +111,10 @@ protected lemma lt_of_lt_of_le : a < b → b ≤ c → a < c := Nat.lt_of_lt_of_ protected lemma le_rfl : a ≤ a := Nat.le_refl _ protected lemma lt_iff_le_and_ne : a < b ↔ a ≤ b ∧ a ≠ b := by rw [← val_ne_iff]; exact Nat.lt_iff_le_and_ne -protected lemma lt_or_lt_of_ne (h : a ≠ b) : a < b ∨ b < a := Nat.lt_or_lt_of_ne $ val_ne_iff.2 h +protected lemma lt_or_lt_of_ne (h : a ≠ b) : a < b ∨ b < a := Nat.lt_or_lt_of_ne <| val_ne_iff.2 h protected lemma lt_or_le (a b : Fin n) : a < b ∨ b ≤ a := Nat.lt_or_ge _ _ protected lemma le_or_lt (a b : Fin n) : a ≤ b ∨ b < a := (b.lt_or_le a).symm -protected lemma le_of_eq (hab : a = b) : a ≤ b := Nat.le_of_eq $ congr_arg val hab +protected lemma le_of_eq (hab : a = b) : a ≤ b := Nat.le_of_eq <| congr_arg val hab protected lemma ge_of_eq (hab : a = b) : b ≤ a := Fin.le_of_eq hab.symm protected lemma eq_or_lt_of_le : a ≤ b → a = b ∨ a < b := by rw [Fin.ext_iff]; exact Nat.eq_or_lt_of_le @@ -127,10 +127,10 @@ lemma lt_last_iff_ne_last {a : Fin (n + 1)} : a < last n ↔ a ≠ last n := by simp [Fin.lt_iff_le_and_ne, le_last] lemma ne_zero_of_lt {a b : Fin (n + 1)} (hab : a < b) : b ≠ 0 := - Fin.ne_of_gt $ Fin.lt_of_le_of_lt a.zero_le hab + Fin.ne_of_gt <| Fin.lt_of_le_of_lt a.zero_le hab lemma ne_last_of_lt {a b : Fin (n + 1)} (hab : a < b) : a ≠ last n := - Fin.ne_of_lt $ Fin.lt_of_lt_of_le hab b.le_last + Fin.ne_of_lt <| Fin.lt_of_lt_of_le hab b.le_last /-- Equivalence between `Fin n` and `{ i // i < n }`. -/ @[simps apply symm_apply] @@ -979,7 +979,7 @@ lemma succAbove_castSucc_of_le (p i : Fin n) (h : p ≤ i) : succAbove p.castSuc succAbove_castSucc_of_le _ _ Fin.le_rfl lemma succAbove_pred_of_lt (p i : Fin (n + 1)) (h : p < i) - (hi := Fin.ne_of_gt $ Fin.lt_of_le_of_lt p.zero_le h) : succAbove p (i.pred hi) = i := by + (hi := Fin.ne_of_gt <| Fin.lt_of_le_of_lt p.zero_le h) : succAbove p (i.pred hi) = i := by rw [succAbove_of_lt_succ _ _ (succ_pred _ _ ▸ h), succ_pred] lemma succAbove_pred_of_le (p i : Fin (n + 1)) (h : i ≤ p) (hi : i ≠ 0) : @@ -989,7 +989,7 @@ lemma succAbove_pred_of_le (p i : Fin (n + 1)) (h : i ≤ p) (hi : i ≠ 0) : succAbove p (p.pred h) = (p.pred h).castSucc := succAbove_pred_of_le _ _ Fin.le_rfl h lemma succAbove_castPred_of_lt (p i : Fin (n + 1)) (h : i < p) - (hi := Fin.ne_of_lt $ Nat.lt_of_lt_of_le h p.le_last) : succAbove p (i.castPred hi) = i := by + (hi := Fin.ne_of_lt <| Nat.lt_of_lt_of_le h p.le_last) : succAbove p (i.castPred hi) = i := by rw [succAbove_of_castSucc_lt _ _ (castSucc_castPred _ _ ▸ h), castSucc_castPred] lemma succAbove_castPred_of_le (p i : Fin (n + 1)) (h : p ≤ i) (hi : i ≠ last n) : @@ -1028,9 +1028,9 @@ lemma succAbove_right_injective : Injective p.succAbove := by split_ifs at hij with hi hj hj · exact castSucc_injective _ hij · rw [hij] at hi - cases hj $ Nat.lt_trans j.castSucc_lt_succ hi + cases hj <| Nat.lt_trans j.castSucc_lt_succ hi · rw [← hij] at hj - cases hi $ Nat.lt_trans i.castSucc_lt_succ hj + cases hi <| Nat.lt_trans i.castSucc_lt_succ hj · exact succ_injective _ hij /-- Given a fixed pivot `p : Fin (n + 1)`, `p.succAbove` is injective. -/ @@ -1089,7 +1089,7 @@ lemma succAbove_lt_iff_castSucc_lt (p : Fin (n + 1)) (i : Fin n) : cases' castSucc_lt_or_lt_succ p i with H H · rwa [iff_true_right H, succAbove_of_castSucc_lt _ _ H] · rw [castSucc_lt_iff_succ_le, iff_false_right (Fin.not_le.2 H), succAbove_of_lt_succ _ _ H] - exact Fin.not_lt.2 $ Fin.le_of_lt H + exact Fin.not_lt.2 <| Fin.le_of_lt H lemma succAbove_lt_iff_succ_le (p : Fin (n + 1)) (i : Fin n) : p.succAbove i < p ↔ succ i ≤ p := by @@ -1101,7 +1101,7 @@ lemma lt_succAbove_iff_le_castSucc (p : Fin (n + 1)) (i : Fin n) : p < p.succAbove i ↔ p ≤ castSucc i := by cases' castSucc_lt_or_lt_succ p i with H H · rw [iff_false_right (Fin.not_le.2 H), succAbove_of_castSucc_lt _ _ H] - exact Fin.not_lt.2 $ Fin.le_of_lt H + exact Fin.not_lt.2 <| Fin.le_of_lt H · rwa [succAbove_of_lt_succ _ _ H, iff_true_left H, le_castSucc_iff] lemma lt_succAbove_iff_lt_castSucc (p : Fin (n + 1)) (i : Fin n) : @@ -1114,12 +1114,12 @@ lemma succAbove_pos [NeZero n] (p : Fin (n + 1)) (i : Fin n) (h : 0 < i) : 0 < p · simp [succAbove_of_le_castSucc _ _ (Fin.not_lt.1 H)] lemma castPred_succAbove (x : Fin n) (y : Fin (n + 1)) (h : castSucc x < y) - (h' := Fin.ne_last_of_lt $ (succAbove_lt_iff_castSucc_lt ..).2 h) : + (h' := Fin.ne_last_of_lt <| (succAbove_lt_iff_castSucc_lt ..).2 h) : (y.succAbove x).castPred h' = x := by rw [castPred_eq_iff_eq_castSucc, succAbove_of_castSucc_lt _ _ h] lemma pred_succAbove (x : Fin n) (y : Fin (n + 1)) (h : y ≤ castSucc x) - (h' := Fin.ne_zero_of_lt $ (lt_succAbove_iff_le_castSucc ..).2 h) : + (h' := Fin.ne_zero_of_lt <| (lt_succAbove_iff_le_castSucc ..).2 h) : (y.succAbove x).pred h' = x := by simp only [succAbove_of_le_castSucc _ _ h, pred_succ] lemma exists_succAbove_eq {x y : Fin (n + 1)} (h : x ≠ y) : ∃ z, y.succAbove z = x := by @@ -1208,11 +1208,11 @@ section PredAbove def predAbove (p : Fin n) (i : Fin (n + 1)) : Fin n := if h : castSucc p < i then pred i (Fin.ne_zero_of_lt h) - else castPred i (Fin.ne_of_lt $ Fin.lt_of_le_of_lt (Fin.not_lt.1 h) (castSucc_lt_last _)) + else castPred i (Fin.ne_of_lt <| Fin.lt_of_le_of_lt (Fin.not_lt.1 h) (castSucc_lt_last _)) lemma predAbove_of_le_castSucc (p : Fin n) (i : Fin (n + 1)) (h : i ≤ castSucc p) - (hi := Fin.ne_of_lt $ Fin.lt_of_le_of_lt h $ castSucc_lt_last _) : - p.predAbove i = i.castPred hi := dif_neg $ Fin.not_lt.2 h + (hi := Fin.ne_of_lt <| Fin.lt_of_le_of_lt h <| castSucc_lt_last _) : + p.predAbove i = i.castPred hi := dif_neg <| Fin.not_lt.2 h lemma predAbove_of_lt_succ (p : Fin n) (i : Fin (n + 1)) (h : i < succ p) (hi := Fin.ne_last_of_lt h) : p.predAbove i = i.castPred hi := @@ -1222,7 +1222,7 @@ lemma predAbove_of_castSucc_lt (p : Fin n) (i : Fin (n + 1)) (h : castSucc p < i (hi := Fin.ne_zero_of_lt h) : p.predAbove i = i.pred hi := dif_pos h lemma predAbove_of_succ_le (p : Fin n) (i : Fin (n + 1)) (h : succ p ≤ i) - (hi := Fin.ne_of_gt $ Fin.lt_of_lt_of_le (succ_pos _) h) : + (hi := Fin.ne_of_gt <| Fin.lt_of_lt_of_le (succ_pos _) h) : p.predAbove i = i.pred hi := predAbove_of_castSucc_lt _ _ (castSucc_lt_iff_succ_le.mpr h) lemma predAbove_succ_of_lt (p i : Fin n) (h : i < p) (hi := succ_ne_last_of_lt h) : @@ -1250,7 +1250,7 @@ lemma predAbove_pred_of_lt (p i : Fin (n + 1)) (h : i < p) (hp := Fin.ne_zero_of rw [predAbove_of_lt_succ _ _ (succ_pred _ _ ▸ h)] lemma predAbove_pred_of_le (p i : Fin (n + 1)) (h : p ≤ i) (hp : p ≠ 0) - (hi := Fin.ne_of_gt $ Fin.lt_of_lt_of_le (Fin.pos_iff_ne_zero.2 hp) h) : + (hi := Fin.ne_of_gt <| Fin.lt_of_lt_of_le (Fin.pos_iff_ne_zero.2 hp) h) : (pred p hp).predAbove i = pred i hi := by rw [predAbove_of_succ_le _ _ (succ_pred _ _ ▸ h)] lemma predAbove_pred_self (p : Fin (n + 1)) (hp : p ≠ 0) : (pred p hp).predAbove p = pred p hp := @@ -1261,7 +1261,7 @@ lemma predAbove_castPred_of_lt (p i : Fin (n + 1)) (h : p < i) (hp := Fin.ne_las rw [predAbove_of_castSucc_lt _ _ (castSucc_castPred _ _ ▸ h)] lemma predAbove_castPred_of_le (p i : Fin (n + 1)) (h : i ≤ p) (hp : p ≠ last n) - (hi := Fin.ne_of_lt $ Fin.lt_of_le_of_lt h $ Fin.lt_last_iff_ne_last.2 hp) : + (hi := Fin.ne_of_lt <| Fin.lt_of_le_of_lt h <| Fin.lt_last_iff_ne_last.2 hp) : (castPred p hp).predAbove i = castPred i hi := by rw [predAbove_of_le_castSucc _ _ (castSucc_castPred _ _ ▸ h)] @@ -1335,7 +1335,7 @@ then back to `Fin n` by subtracting one from anything above `p` is the identity. lemma predAbove_succAbove (p : Fin n) (i : Fin n) : p.predAbove ((castSucc p).succAbove i) = i := by obtain h | h := p.le_or_lt i · rw [succAbove_castSucc_of_le _ _ h, predAbove_succ_of_le _ _ h] - · rw [succAbove_castSucc_of_lt _ _ h, predAbove_castSucc_of_le _ _ $ Fin.le_of_lt h] + · rw [succAbove_castSucc_of_lt _ _ h, predAbove_castSucc_of_le _ _ <| Fin.le_of_lt h] /-- `succ` commutes with `predAbove`. -/ @[simp] lemma succ_predAbove_succ (a : Fin n) (b : Fin (n + 1)) : diff --git a/Mathlib/Data/Fin/Tuple/Curry.lean b/Mathlib/Data/Fin/Tuple/Curry.lean index 1b752e325ecb5..65e49ebf5c77c 100644 --- a/Mathlib/Data/Fin/Tuple/Curry.lean +++ b/Mathlib/Data/Fin/Tuple/Curry.lean @@ -73,7 +73,7 @@ variable {n : ℕ} {p : Fin n → Type u} {τ : Type u} theorem curry_uncurry (f : Function.FromTypes p τ) : curry (uncurry f) = f := by induction n with | zero => rfl - | succ n ih => exact funext (ih $ f ·) + | succ n ih => exact funext (ih <| f ·) @[simp] theorem uncurry_curry (f : ((i : Fin n) → p i) → τ) : diff --git a/Mathlib/Data/Fin/Tuple/Finset.lean b/Mathlib/Data/Fin/Tuple/Finset.lean index 5a102d77f270b..103c2df7fa9f0 100644 --- a/Mathlib/Data/Fin/Tuple/Finset.lean +++ b/Mathlib/Data/Fin/Tuple/Finset.lean @@ -29,7 +29,7 @@ lemma cons_mem_piFinset_cons {x₀ : α 0} {x : ∀ i : Fin n, α i.succ} simp_rw [mem_piFinset_succ, cons_zero, tail_cons] lemma snoc_mem_piFinset_snoc {x : ∀ i : Fin n, α i.castSucc} {xₙ : α (.last n)} - {s : ∀ i : Fin n, Finset (α i.castSucc)} {sₙ : Finset (α $ .last n)} : + {s : ∀ i : Fin n, Finset (α i.castSucc)} {sₙ : Finset (α <| .last n)} : snoc x xₙ ∈ piFinset (snoc s sₙ) ↔ x ∈ piFinset s ∧ xₙ ∈ sₙ := by simp_rw [mem_piFinset_succ', init_snoc, snoc_last] diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index bf99110752d4a..bd963644daca7 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -1592,7 +1592,7 @@ theorem erase_empty (a : α) : erase ∅ a = ∅ := rfl protected lemma Nontrivial.erase_nonempty (hs : s.Nontrivial) : (s.erase a).Nonempty := - (hs.exists_ne a).imp $ by aesop + (hs.exists_ne a).imp <| by aesop @[simp] lemma erase_nonempty (ha : a ∈ s) : (s.erase a).Nonempty ↔ s.Nontrivial := by simp only [Finset.Nonempty, mem_erase, and_comm (b := _ ∈ _)] diff --git a/Mathlib/Data/Finset/Card.lean b/Mathlib/Data/Finset/Card.lean index 36650f6030a8e..ee90c07098ad8 100644 --- a/Mathlib/Data/Finset/Card.lean +++ b/Mathlib/Data/Finset/Card.lean @@ -426,7 +426,7 @@ lemma surj_on_of_inj_on_of_card_le (f : ∀ a ∈ s, β) (hf : ∀ a ha, f a ha intro ⟨_, _⟩ ⟨_, _⟩ h exact Subtype.eq <| hinj _ _ _ _ h obtain rfl : image (fun a : { a // a ∈ s } => f a a.prop) s.attach = t := - eq_of_subset_of_card_le (image_subset_iff.2 $ by simpa) (by simp [hst, h]) + eq_of_subset_of_card_le (image_subset_iff.2 <| by simpa) (by simp [hst, h]) simp only [mem_image, mem_attach, true_and, Subtype.exists, forall_exists_index] exact fun b a ha hb ↦ ⟨a, ha, hb.symm⟩ diff --git a/Mathlib/Data/Finset/Density.lean b/Mathlib/Data/Finset/Density.lean index c486e085f6987..30ed02c7deba6 100644 --- a/Mathlib/Data/Finset/Density.lean +++ b/Mathlib/Data/Finset/Density.lean @@ -82,10 +82,10 @@ protected alias ⟨_, Nonempty.dens_pos⟩ := dens_pos protected alias ⟨_, Nonempty.dens_ne_zero⟩ := dens_ne_zero lemma dens_le_dens (h : s ⊆ t) : dens s ≤ dens t := - div_le_div_of_nonneg_right (mod_cast card_mono h) $ by positivity + div_le_div_of_nonneg_right (mod_cast card_mono h) <| by positivity lemma dens_lt_dens (h : s ⊂ t) : dens s < dens t := - div_lt_div_of_pos_right (mod_cast card_strictMono h) $ by + div_lt_div_of_pos_right (mod_cast card_strictMono h) <| by cases isEmpty_or_nonempty α · simp [Subsingleton.elim s t, ssubset_irrfl] at h · exact mod_cast Fintype.card_pos @@ -134,7 +134,7 @@ lemma dens_sdiff_add_dens (s t : Finset α) : dens (s \ t) + dens t = (s ∪ t). rw [← dens_union_of_disjoint sdiff_disjoint, sdiff_union_self_eq_union] lemma dens_sdiff_comm (h : card s = card t) : dens (s \ t) = dens (t \ s) := - add_left_injective (dens t) $ by + add_left_injective (dens t) <| by simp_rw [dens_sdiff_add_dens, union_comm s, ← dens_sdiff_add_dens, dens, h] @[simp] diff --git a/Mathlib/Data/Finset/Powerset.lean b/Mathlib/Data/Finset/Powerset.lean index 7fca063a429b2..ca7a34c9e0cbb 100644 --- a/Mathlib/Data/Finset/Powerset.lean +++ b/Mathlib/Data/Finset/Powerset.lean @@ -112,12 +112,12 @@ instance decidableForallOfDecidableSubsets {s : Finset α} {p : ∀ t ⊆ s, Pro /-- For predicate `p` decidable on subsets, it is decidable whether `p` holds for any subset. -/ instance decidableExistsOfDecidableSubsets' {s : Finset α} {p : Finset α → Prop} [∀ t, Decidable (p t)] : Decidable (∃ t ⊆ s, p t) := - decidable_of_iff (∃ (t : _) (_h : t ⊆ s), p t) $ by simp + decidable_of_iff (∃ (t : _) (_h : t ⊆ s), p t) <| by simp /-- For predicate `p` decidable on subsets, it is decidable whether `p` holds for every subset. -/ instance decidableForallOfDecidableSubsets' {s : Finset α} {p : Finset α → Prop} [∀ t, Decidable (p t)] : Decidable (∀ t ⊆ s, p t) := - decidable_of_iff (∀ (t : _) (_h : t ⊆ s), p t) $ by simp + decidable_of_iff (∀ (t : _) (_h : t ⊆ s), p t) <| by simp end Powerset @@ -195,7 +195,7 @@ theorem powersetCard_zero (s : Finset α) : s.powersetCard 0 = {∅} := by exact ⟨empty_subset s, rfl⟩⟩ lemma powersetCard_empty_subsingleton (n : ℕ) : - (powersetCard n (∅ : Finset α) : Set $ Finset α).Subsingleton := by + (powersetCard n (∅ : Finset α) : Set <| Finset α).Subsingleton := by simp [Set.Subsingleton, subset_empty] @[simp] @@ -209,9 +209,9 @@ theorem powersetCard_one (s : Finset α) : @[simp] lemma powersetCard_eq_empty : powersetCard n s = ∅ ↔ s.card < n := by - refine ⟨?_, fun h ↦ card_eq_zero.1 $ by rw [card_powersetCard, Nat.choose_eq_zero_of_lt h]⟩ + refine ⟨?_, fun h ↦ card_eq_zero.1 <| by rw [card_powersetCard, Nat.choose_eq_zero_of_lt h]⟩ contrapose! - exact fun h ↦ nonempty_iff_ne_empty.1 $ (exists_subset_card_eq h).imp $ by simp + exact fun h ↦ nonempty_iff_ne_empty.1 <| (exists_subset_card_eq h).imp <| by simp @[simp] lemma powersetCard_card_add (s : Finset α) (hn : 0 < n) : s.powersetCard (s.card + n) = ∅ := by simpa diff --git a/Mathlib/Data/Fintype/BigOperators.lean b/Mathlib/Data/Fintype/BigOperators.lean index 02838b7dad046..2cef64927c55b 100644 --- a/Mathlib/Data/Fintype/BigOperators.lean +++ b/Mathlib/Data/Fintype/BigOperators.lean @@ -148,7 +148,7 @@ lemma card_filter_piFinset_eq_of_mem [∀ i, DecidableEq (α i)] lemma card_filter_piFinset_const_eq_of_mem (s : Finset κ) (i : ι) {x : κ} (hx : x ∈ s) : ((piFinset fun _ ↦ s).filter fun f ↦ f i = x).card = s.card ^ (card ι - 1) := - (card_filter_piFinset_eq_of_mem _ _ hx).trans $ by + (card_filter_piFinset_eq_of_mem _ _ hx).trans <| by rw [prod_const s.card, card_erase_of_mem (mem_univ _), card_univ] lemma card_filter_piFinset_eq [∀ i, DecidableEq (α i)] (s : ∀ i, Finset (α i)) (i : ι) (a : α i) : @@ -161,7 +161,7 @@ lemma card_filter_piFinset_eq [∀ i, DecidableEq (α i)] (s : ∀ i, Finset (α lemma card_filter_piFinset_const (s : Finset κ) (i : ι) (j : κ) : ((piFinset fun _ ↦ s).filter fun f ↦ f i = j).card = if j ∈ s then s.card ^ (card ι - 1) else 0 := - (card_filter_piFinset_eq _ _ _).trans $ by + (card_filter_piFinset_eq _ _ _).trans <| by rw [prod_const s.card, card_erase_of_mem (mem_univ _), card_univ] end Fintype diff --git a/Mathlib/Data/Int/Star.lean b/Mathlib/Data/Int/Star.lean index c654adbb78ad6..9492664978ba1 100644 --- a/Mathlib/Data/Int/Star.lean +++ b/Mathlib/Data/Int/Star.lean @@ -23,7 +23,7 @@ namespace Int refine le_antisymm (closure_le.2 <| range_subset_iff.2 hn.pow_nonneg) fun x hx ↦ ?_ have : x = x.natAbs • 1 ^ n := by simpa [eq_comm (a := x)] using hx rw [this] - exact nsmul_mem (subset_closure $ mem_range_self _) _ + exact nsmul_mem (subset_closure <| mem_range_self _) _ @[simp] lemma addSubmonoid_closure_range_mul_self : closure (range fun x : ℤ ↦ x * x) = nonneg _ := by diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index f804a9590f642..4aaffa789fd29 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -570,7 +570,7 @@ theorem nodup_permutations'Aux_iff {s : List α} {x : α} : Nodup (permutations' intro H obtain ⟨k, hk, hk'⟩ := nthLe_of_mem H rw [nodup_iff_nthLe_inj] at h - refine k.succ_ne_self.symm $ h k (k + 1) ?_ ?_ ?_ + refine k.succ_ne_self.symm <| h k (k + 1) ?_ ?_ ?_ · simpa [Nat.lt_succ_iff] using hk.le · simpa using hk rw [nthLe_permutations'Aux, nthLe_permutations'Aux] diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index 4f1ef5b6229a7..65b273ac1e98b 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -169,7 +169,7 @@ lemma mem_mem_ranges_iff_lt_natSum (l : List ℕ) {n : ℕ} : /-- The members of `l.ranges` have no duplicate -/ theorem ranges_nodup {l s : List ℕ} (hs : s ∈ ranges l) : s.Nodup := - (List.pairwise_join.mp $ by rw [ranges_join']; exact nodup_range _).1 s hs + (List.pairwise_join.mp <| by rw [ranges_join']; exact nodup_range _).1 s hs end Ranges diff --git a/Mathlib/Data/Matrix/Composition.lean b/Mathlib/Data/Matrix/Composition.lean index ae65a787e8669..31de61ff426dd 100644 --- a/Mathlib/Data/Matrix/Composition.lean +++ b/Mathlib/Data/Matrix/Composition.lean @@ -57,7 +57,7 @@ def compRingEquiv : Matrix I I (Matrix J J R) ≃+* Matrix (I × J) (I × J) R w __ := Matrix.compAddEquiv I I J J R map_mul' _ _ := by ext _ _ - exact (Matrix.sum_apply _ _ _ _).trans $ Eq.symm Fintype.sum_prod_type + exact (Matrix.sum_apply _ _ _ _).trans <| Eq.symm Fintype.sum_prod_type end Semiring diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 523808b4def1c..6889f2116d1eb 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1468,7 +1468,7 @@ theorem sub_eq_fold_erase (s t : Multiset α) : s - t = foldl erase erase_comm s @[simp] theorem card_sub {s t : Multiset α} (h : t ≤ s) : card (s - t) = card s - card t := - Nat.eq_sub_of_add_eq $ by rw [← card_add, tsub_add_cancel_of_le h] + Nat.eq_sub_of_add_eq <| by rw [← card_add, tsub_add_cancel_of_le h] /-! ### Union -/ @@ -2206,7 +2206,7 @@ theorem ext' {s t : Multiset α} : (∀ a, count a s = count a t) → s = t := ext.2 lemma count_injective : Injective fun (s : Multiset α) a ↦ s.count a := - fun _s _t hst ↦ ext' $ congr_fun hst + fun _s _t hst ↦ ext' <| congr_fun hst @[simp] theorem coe_inter (s t : List α) : (s ∩ t : Multiset α) = (s.bagInter t : List α) := by ext; simp diff --git a/Mathlib/Data/Multiset/Bind.lean b/Mathlib/Data/Multiset/Bind.lean index 5dbf593b2a902..3fe645c8e220e 100644 --- a/Mathlib/Data/Multiset/Bind.lean +++ b/Mathlib/Data/Multiset/Bind.lean @@ -191,7 +191,7 @@ theorem le_bind {α β : Type*} {f : α → Multiset β} (S : Multiset α) {x : f x ≤ S.bind f := by classical refine le_iff_count.2 fun a ↦ ?_ - obtain ⟨m', hm'⟩ := exists_cons_of_mem $ mem_map_of_mem (fun b ↦ count a (f b)) hx + obtain ⟨m', hm'⟩ := exists_cons_of_mem <| mem_map_of_mem (fun b ↦ count a (f b)) hx rw [count_bind, hm', sum_cons] exact Nat.le_add_right _ _ diff --git a/Mathlib/Data/Multiset/Fintype.lean b/Mathlib/Data/Multiset/Fintype.lean index 9f48dea81db32..eef7f8002adbf 100644 --- a/Mathlib/Data/Multiset/Fintype.lean +++ b/Mathlib/Data/Multiset/Fintype.lean @@ -130,7 +130,7 @@ namespace Multiset have : card (s.1.filter fun x ↦ a = x.1) ≤ card (s.1.filter fun x ↦ a = x.1) - 1 := by simpa [Finset.card, eq_comm] using Finset.card_mono h omega - exact Nat.le_of_pred_lt (han.trans_lt $ by simpa using hsm hn) + exact Nat.le_of_pred_lt (han.trans_lt <| by simpa using hsm hn) end Multiset diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index 7273bddabfad4..dc43e243e47d4 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -343,7 +343,7 @@ lemma divNat_inj (h₁ : d₁ ≠ 0) (h₂ : d₂ ≠ 0) : divNat n₁ d₁ = di @[simp] lemma divNat_zero (n : ℕ) : divNat n 0 = 0 := by simp [divNat]; rfl @[simp] lemma num_divNat_den (q : ℚ≥0) : divNat q.num q.den = q := - ext $ by rw [← (q : ℚ).mkRat_num_den']; simp [num_coe, den_coe] + ext <| by rw [← (q : ℚ).mkRat_num_den']; simp [num_coe, den_coe] lemma natCast_eq_divNat (n : ℕ) : (n : ℚ≥0) = divNat n 1 := (num_divNat_den _).symm diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index 8e33f0b37f992..82a0796b55644 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -130,7 +130,7 @@ theorem map_ofNat' {A} [AddMonoidWithOne A] [FunLike F A B] [AddMonoidHomClass F { toFun := fun n ↦ n • (1 : A) map_zero' := zero_nsmul _ map_add' := add_nsmul _ } - exact eq_natCast' f $ by simp [f] + exact eq_natCast' f <| by simp [f] end AddMonoidHomClass diff --git a/Mathlib/Data/Nat/Cast/Order/Field.lean b/Mathlib/Data/Nat/Cast/Order/Field.lean index 0ae2a744f0146..74674ae19eeb2 100644 --- a/Mathlib/Data/Nat/Cast/Order/Field.lean +++ b/Mathlib/Data/Nat/Cast/Order/Field.lean @@ -22,7 +22,7 @@ variable {α : Type*} [LinearOrderedSemifield α] lemma cast_inv_le_one : ∀ n : ℕ, (n⁻¹ : α) ≤ 1 | 0 => by simp - | n + 1 => inv_le_one $ by simp [Nat.cast_nonneg] + | n + 1 => inv_le_one <| by simp [Nat.cast_nonneg] /-- Natural division is always less than division in the field. -/ theorem cast_div_le {m n : ℕ} : ((m / n : ℕ) : α) ≤ m / n := by diff --git a/Mathlib/Data/Nat/Choose/Lucas.lean b/Mathlib/Data/Nat/Choose/Lucas.lean index c896d134554d5..ed55883414890 100644 --- a/Mathlib/Data/Nat/Choose/Lucas.lean +++ b/Mathlib/Data/Nat/Choose/Lucas.lean @@ -47,7 +47,7 @@ theorem choose_modEq_choose_mod_mul_choose_div : rw [Prod.mk.injEq] constructor <;> intro h · simp only [mem_product, mem_range] at hx - have h' : x₁ < p := lt_of_lt_of_le hx.left $ mod_lt _ Fin.size_pos' + have h' : x₁ < p := lt_of_lt_of_le hx.left <| mod_lt _ Fin.size_pos' rw [h, add_mul_mod_self_left, add_mul_div_left _ _ Fin.size_pos', eq_comm (b := x₂)] exact ⟨mod_eq_of_lt h', self_eq_add_left.mpr (div_eq_of_lt h')⟩ · rw [← h.left, ← h.right, mod_add_div] diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 4670f85b6aeb7..31452114335f3 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -139,7 +139,7 @@ lemma one_lt_iff_ne_zero_and_ne_one : ∀ {n : ℕ}, 1 < n ↔ n ≠ 0 ∧ n ≠ lemma le_one_iff_eq_zero_or_eq_one : ∀ {n : ℕ}, n ≤ 1 ↔ n = 0 ∨ n = 1 := by simp [le_succ_iff] -@[simp] lemma lt_one_iff : n < 1 ↔ n = 0 := Nat.lt_succ_iff.trans $ by rw [le_zero_eq] +@[simp] lemma lt_one_iff : n < 1 ↔ n = 0 := Nat.lt_succ_iff.trans <| by rw [le_zero_eq] lemma one_le_of_lt (h : a < b) : 1 ≤ b := Nat.lt_of_le_of_lt (Nat.zero_le _) h @@ -316,9 +316,9 @@ lemma mul_eq_left (ha : a ≠ 0) : a * b = a ↔ b = 1 := by simpa using Nat.mul lemma mul_eq_right (hb : b ≠ 0) : a * b = b ↔ a = 1 := by simpa using Nat.mul_left_inj hb (c := 1) -- TODO: Deprecate -lemma mul_right_eq_self_iff (ha : 0 < a) : a * b = a ↔ b = 1 := mul_eq_left $ ne_of_gt ha +lemma mul_right_eq_self_iff (ha : 0 < a) : a * b = a ↔ b = 1 := mul_eq_left <| ne_of_gt ha -lemma mul_left_eq_self_iff (hb : 0 < b) : a * b = b ↔ a = 1 := mul_eq_right $ ne_of_gt hb +lemma mul_left_eq_self_iff (hb : 0 < b) : a * b = b ↔ a = 1 := mul_eq_right <| ne_of_gt hb protected lemma le_of_mul_le_mul_right (h : a * c ≤ b * c) (hc : 0 < c) : a ≤ b := Nat.le_of_mul_le_mul_left (by simpa [Nat.mul_comm]) hc @@ -361,7 +361,7 @@ lemma succ_mul_pos (m : ℕ) (hn : 0 < n) : 0 < succ m * n := Nat.mul_pos m.succ lemma mul_self_le_mul_self (h : m ≤ n) : m * m ≤ n * n := Nat.mul_le_mul h h lemma mul_lt_mul'' (hac : a < c) (hbd : b < d) : a * b < c * d := - Nat.mul_lt_mul_of_lt_of_le hac (Nat.le_of_lt hbd) $ by omega + Nat.mul_lt_mul_of_lt_of_le hac (Nat.le_of_lt hbd) <| by omega lemma mul_self_lt_mul_self (h : m < n) : m * m < n * n := mul_lt_mul'' h h @@ -387,7 +387,7 @@ lemma add_sub_one_le_mul (ha : a ≠ 0) (hb : b ≠ 0) : a + b - 1 ≤ a * b := cases a · cases ha rfl · rw [succ_add, Nat.add_one_sub_one, succ_mul] - exact Nat.add_le_add_right (Nat.le_mul_of_pos_right _ $ Nat.pos_iff_ne_zero.2 hb) _ + exact Nat.add_le_add_right (Nat.le_mul_of_pos_right _ <| Nat.pos_iff_ne_zero.2 hb) _ protected lemma add_le_mul {a : ℕ} (ha : 2 ≤ a) : ∀ {b : ℕ} (_ : 2 ≤ b), a + b ≤ a * b | 2, _ => by omega @@ -421,10 +421,10 @@ protected lemma div_le_div_right (h : a ≤ b) : a / c ≤ b / c := (le_div_iff_mul_le' hc).2 <| Nat.le_trans (Nat.div_mul_le_self _ _) h lemma lt_of_div_lt_div (h : a / c < b / c) : a < b := - Nat.lt_of_not_le fun hab ↦ Nat.not_le_of_lt h $ Nat.div_le_div_right hab + Nat.lt_of_not_le fun hab ↦ Nat.not_le_of_lt h <| Nat.div_le_div_right hab protected lemma div_pos (hba : b ≤ a) (hb : 0 < b) : 0 < a / b := - Nat.pos_of_ne_zero fun h ↦ Nat.lt_irrefl a $ + Nat.pos_of_ne_zero fun h ↦ Nat.lt_irrefl a <| calc a = a % b := by simpa [h] using (mod_add_div a b).symm _ < b := mod_lt a hb @@ -504,7 +504,7 @@ protected lemma div_le_of_le_mul' (h : m ≤ k * n) : m / k ≤ n := by protected lemma div_le_div_of_mul_le_mul (hd : d ≠ 0) (hdc : d ∣ c) (h : a * d ≤ c * b) : a / b ≤ c / d := - Nat.div_le_of_le_mul' $ by + Nat.div_le_of_le_mul' <| by rwa [← Nat.mul_div_assoc _ hdc, Nat.le_div_iff_mul_le (Nat.pos_iff_ne_zero.2 hd), b.mul_comm] protected lemma div_le_self' (m n : ℕ) : m / n ≤ m := by @@ -641,7 +641,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 @@ -1038,7 +1038,7 @@ lemma eq_of_dvd_of_lt_two_mul (ha : a ≠ 0) (hdvd : b ∣ a) (hlt : a < 2 * b) cases Nat.not_le_of_lt hlt (Nat.mul_le_mul_right _ (by omega)) lemma mod_eq_iff_lt (hn : n ≠ 0) : m % n = m ↔ m < n := - ⟨fun h ↦ by rw [← h]; exact mod_lt _ $ Nat.pos_iff_ne_zero.2 hn, mod_eq_of_lt⟩ + ⟨fun h ↦ by rw [← h]; exact mod_lt _ <| Nat.pos_iff_ne_zero.2 hn, mod_eq_of_lt⟩ @[simp] lemma mod_succ_eq_iff_lt : m % n.succ = m ↔ m < n.succ := @@ -1098,12 +1098,12 @@ protected lemma dvd_add_right (h : a ∣ b) : a ∣ b + c ↔ a ∣ c := (Nat.dv /-- special case of `mul_dvd_mul_iff_left` for `ℕ`. Duplicated here to keep simple imports for this file. -/ protected lemma mul_dvd_mul_iff_left (ha : 0 < a) : a * b ∣ a * c ↔ b ∣ c := - exists_congr fun d ↦ by rw [Nat.mul_assoc, Nat.mul_right_inj $ ne_of_gt ha] + exists_congr fun d ↦ by rw [Nat.mul_assoc, Nat.mul_right_inj <| ne_of_gt ha] /-- special case of `mul_dvd_mul_iff_right` for `ℕ`. Duplicated here to keep simple imports for this file. -/ protected lemma mul_dvd_mul_iff_right (hc : 0 < c) : a * c ∣ b * c ↔ a ∣ b := - exists_congr fun d ↦ by rw [Nat.mul_right_comm, Nat.mul_left_inj $ ne_of_gt hc] + exists_congr fun d ↦ by rw [Nat.mul_right_comm, Nat.mul_left_inj <| ne_of_gt hc] -- Moved to Batteries @@ -1474,7 +1474,7 @@ lemma not_exists_sq' : m ^ 2 < n → n < (m + 1) ^ 2 → ¬∃ t, t ^ 2 = n := b instance decidableLoHi (lo hi : ℕ) (P : ℕ → Prop) [H : DecidablePred P] : Decidable (∀ x, lo ≤ x → x < hi → P x) := - decidable_of_iff (∀ x < hi - lo, P (lo + x)) $ by + decidable_of_iff (∀ x < hi - lo, P (lo + x)) <| by refine ⟨fun al x hl hh ↦ ?_, fun al x h ↦ al _ (Nat.le_add_right _ _) (Nat.lt_sub_iff_add_lt'.1 h)⟩ have := al (x - lo) ((Nat.sub_lt_sub_iff_right hl).2 hh) diff --git a/Mathlib/Data/Nat/Factorial/Basic.lean b/Mathlib/Data/Nat/Factorial/Basic.lean index 49a834165d323..f41001d9b85d8 100644 --- a/Mathlib/Data/Nat/Factorial/Basic.lean +++ b/Mathlib/Data/Nat/Factorial/Basic.lean @@ -86,7 +86,7 @@ theorem factorial_lt (hn : 0 < n) : n ! < m ! ↔ n < m := by exact Nat.mul_pos hk k.factorial_pos induction h generalizing hn with | refl => exact this hn - | step hnk ih => exact lt_trans (ih hn) $ this <| lt_trans hn <| lt_of_succ_le hnk + | step hnk ih => exact lt_trans (ih hn) <| this <| lt_trans hn <| lt_of_succ_le hnk @[gcongr] lemma factorial_lt_of_lt {m n : ℕ} (hn : 0 < n) (h : n < m) : n ! < m ! := (factorial_lt hn).mpr h diff --git a/Mathlib/Data/Nat/Factorization/Root.lean b/Mathlib/Data/Nat/Factorization/Root.lean index d0aa141fa8fcd..7e0c2c1f19bee 100644 --- a/Mathlib/Data/Nat/Factorization/Root.lean +++ b/Mathlib/Data/Nat/Factorization/Root.lean @@ -68,7 +68,7 @@ lemma floorRoot_ne_zero : floorRoot n a ≠ 0 ↔ n ≠ 0 ∧ a ≠ 0 := by simp (config := { contextual := true }) [floorRoot, not_imp_not, not_or] @[simp] lemma floorRoot_eq_zero : floorRoot n a = 0 ↔ n = 0 ∨ a = 0 := - floorRoot_ne_zero.not_right.trans $ by simp only [not_and_or, ne_eq, not_not] + floorRoot_ne_zero.not_right.trans <| by simp only [not_and_or, ne_eq, not_not] @[simp] lemma factorization_floorRoot (n a : ℕ) : (floorRoot n a).factorization = a.factorization ⌊/⌋ n := by @@ -130,7 +130,7 @@ lemma ceilRoot_ne_zero : ceilRoot n a ≠ 0 ↔ n ≠ 0 ∧ a ≠ 0 := by simp (config := { contextual := true }) [ceilRoot_def, not_imp_not, not_or] @[simp] lemma ceilRoot_eq_zero : ceilRoot n a = 0 ↔ n = 0 ∨ a = 0 := - ceilRoot_ne_zero.not_right.trans $ by simp only [not_and_or, ne_eq, not_not] + ceilRoot_ne_zero.not_right.trans <| by simp only [not_and_or, ne_eq, not_not] @[simp] lemma factorization_ceilRoot (n a : ℕ) : (ceilRoot n a).factorization = a.factorization ⌈/⌉ n := by diff --git a/Mathlib/Data/Nat/Find.lean b/Mathlib/Data/Nat/Find.lean index 30b550bb0fc2d..014c1729bbfe4 100644 --- a/Mathlib/Data/Nat/Find.lean +++ b/Mathlib/Data/Nat/Find.lean @@ -200,7 +200,7 @@ lemma findGreatest_mono_right (P : ℕ → Prop) [DecidablePred P] {m n} (hmn : · simp rw [findGreatest_succ] split_ifs - · exact le_trans ih $ le_trans (findGreatest_le _) (le_succ _) + · exact le_trans ih <| le_trans (findGreatest_le _) (le_succ _) · exact ih lemma findGreatest_mono_left [DecidablePred Q] (hPQ : ∀ n, P n → Q n) (n : ℕ) : diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index fef1cd6a9e339..0e33dd6f02898 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -60,7 +60,7 @@ theorem log_of_one_lt_of_le {b n : ℕ} (h : 1 < b) (hn : b ≤ n) : log b n = l rw [log] exact if_pos ⟨hn, h⟩ -@[simp] lemma log_zero_left : ∀ n, log 0 n = 0 := log_of_left_le_one $ Nat.zero_le _ +@[simp] lemma log_zero_left : ∀ n, log 0 n = 0 := log_of_left_le_one <| Nat.zero_le _ @[simp] theorem log_zero_right (b : ℕ) : log b 0 = 0 := diff --git a/Mathlib/Data/Real/ConjExponents.lean b/Mathlib/Data/Real/ConjExponents.lean index 6b6586723f5c1..5323170031fd4 100644 --- a/Mathlib/Data/Real/ConjExponents.lean +++ b/Mathlib/Data/Real/ConjExponents.lean @@ -106,10 +106,10 @@ theorem inv_add_inv_conj_ennreal : (ENNReal.ofReal p)⁻¹ + (ENNReal.ofReal q) end protected lemma inv_inv (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) : a⁻¹.IsConjExponent b⁻¹ := - ⟨one_lt_inv ha $ by linarith, by simpa only [inv_inv]⟩ + ⟨one_lt_inv ha <| by linarith, by simpa only [inv_inv]⟩ lemma inv_one_sub_inv (ha₀ : 0 < a) (ha₁ : a < 1) : a⁻¹.IsConjExponent (1 - a)⁻¹ := - .inv_inv ha₀ (sub_pos_of_lt ha₁) $ add_tsub_cancel_of_le ha₁.le + .inv_inv ha₀ (sub_pos_of_lt ha₁) <| add_tsub_cancel_of_le ha₁.le lemma one_sub_inv_inv (ha₀ : 0 < a) (ha₁ : a < 1) : (1 - a)⁻¹.IsConjExponent a⁻¹ := (inv_one_sub_inv ha₀ ha₁).symm @@ -195,11 +195,11 @@ end protected lemma inv_inv (ha : a ≠ 0) (hb : b ≠ 0) (hab : a + b = 1) : a⁻¹.IsConjExponent b⁻¹ := - ⟨one_lt_inv ha.bot_lt $ by rw [← hab]; exact lt_add_of_pos_right _ hb.bot_lt, by + ⟨one_lt_inv ha.bot_lt <| by rw [← hab]; exact lt_add_of_pos_right _ hb.bot_lt, by simpa only [inv_inv] using hab⟩ lemma inv_one_sub_inv (ha₀ : a ≠ 0) (ha₁ : a < 1) : a⁻¹.IsConjExponent (1 - a)⁻¹ := - .inv_inv ha₀ (tsub_pos_of_lt ha₁).ne' $ add_tsub_cancel_of_le ha₁.le + .inv_inv ha₀ (tsub_pos_of_lt ha₁).ne' <| add_tsub_cancel_of_le ha₁.le lemma one_sub_inv_inv (ha₀ : a ≠ 0) (ha₁ : a < 1) : (1 - a)⁻¹.IsConjExponent a⁻¹ := (inv_one_sub_inv ha₀ ha₁).symm diff --git a/Mathlib/Data/Real/Sqrt.lean b/Mathlib/Data/Real/Sqrt.lean index 3c3768bf2fd11..d94c45c1cc6cd 100644 --- a/Mathlib/Data/Real/Sqrt.lean +++ b/Mathlib/Data/Real/Sqrt.lean @@ -267,7 +267,7 @@ alias ⟨_, sqrt_pos_of_pos⟩ := sqrt_pos lemma sqrt_le_sqrt_iff' (hx : 0 < x) : √x ≤ √y ↔ x ≤ y := by obtain hy | hy := le_total y 0 - · exact iff_of_false ((sqrt_eq_zero_of_nonpos hy).trans_lt $ sqrt_pos.2 hx).not_le + · exact iff_of_false ((sqrt_eq_zero_of_nonpos hy).trans_lt <| sqrt_pos.2 hx).not_le (hy.trans_lt hx).not_le · exact sqrt_le_sqrt_iff hy @@ -433,7 +433,7 @@ open Finset /-- **Cauchy-Schwarz inequality** for finsets using square roots in `ℝ≥0`. -/ lemma sum_mul_le_sqrt_mul_sqrt (s : Finset ι) (f g : ι → ℝ≥0) : ∑ i ∈ s, f i * g i ≤ sqrt (∑ i ∈ s, f i ^ 2) * sqrt (∑ i ∈ s, g i ^ 2) := - (le_sqrt_iff_sq_le.2 $ sum_mul_sq_le_sq_mul_sq _ _ _).trans_eq <| sqrt_mul _ _ + (le_sqrt_iff_sq_le.2 <| sum_mul_sq_le_sq_mul_sq _ _ _).trans_eq <| sqrt_mul _ _ /-- **Cauchy-Schwarz inequality** for finsets using square roots in `ℝ≥0`. -/ lemma sum_sqrt_mul_sqrt_le (s : Finset ι) (f g : ι → ℝ≥0) : diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 141de079ea67a..937769bbf1c82 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -929,7 +929,7 @@ theorem BijOn.image_eq (h : BijOn f s t) : f '' s = t := h.surjOn.image_eq_of_mapsTo h.mapsTo lemma BijOn.forall {p : β → Prop} (hf : BijOn f s t) : (∀ b ∈ t, p b) ↔ ∀ a ∈ s, p (f a) where - mp h a ha := h _ $ hf.mapsTo ha + mp h a ha := h _ <| hf.mapsTo ha mpr h b hb := by obtain ⟨a, ha, rfl⟩ := hf.surjOn hb; exact h _ ha lemma BijOn.exists {p : β → Prop} (hf : BijOn f s t) : (∃ b ∈ t, p b) ↔ ∃ a ∈ s, p (f a) where diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 09886a1171c8b..482de2b723d5d 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -191,7 +191,7 @@ theorem _root_.Function.Injective.mem_set_image {f : α → β} (hf : Injective lemma preimage_subset_of_surjOn {t : Set β} (hf : Injective f) (h : SurjOn f s t) : f ⁻¹' t ⊆ s := fun _ hx ↦ - hf.mem_set_image.1 $ h hx + hf.mem_set_image.1 <| h hx theorem forall_mem_image {f : α → β} {s : Set α} {p : β → Prop} : (∀ y ∈ f '' s, p y) ↔ ∀ ⦃x⦄, x ∈ s → p (f x) := by simp diff --git a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index 662deae08d25c..bef0b0e2de8f4 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -379,7 +379,7 @@ instance instGroupWithZero : GroupWithZero (AlgebraicClosure k) := let e := algEquivAlgebraicClosureAux k { inv := fun a ↦ e.symm (e a)⁻¹ inv_zero := by simp - mul_inv_cancel := fun a ha ↦ e.injective $ by simp [(AddEquivClass.map_ne_zero_iff _).2 ha] + mul_inv_cancel := fun a ha ↦ e.injective <| by simp [(AddEquivClass.map_ne_zero_iff _).2 ha] __ := e.surjective.nontrivial } instance instField : Field (AlgebraicClosure k) where @@ -392,9 +392,9 @@ instance instField : Field (AlgebraicClosure k) where nnratCast_def q := by change algebraMap k _ _ = _; simp_rw [NNRat.cast_def, map_div₀, map_natCast] ratCast_def q := by change algebraMap k _ _ = _; rw [Rat.cast_def, map_div₀, map_intCast, map_natCast] - nnqsmul_def q x := Quotient.inductionOn x fun p ↦ congr_arg Quotient.mk'' $ by + nnqsmul_def q x := Quotient.inductionOn x fun p ↦ congr_arg Quotient.mk'' <| by ext; simp [MvPolynomial.algebraMap_eq, NNRat.smul_def] - qsmul_def q x := Quotient.inductionOn x fun p ↦ congr_arg Quotient.mk'' $ by + qsmul_def q x := Quotient.inductionOn x fun p ↦ congr_arg Quotient.mk'' <| by ext; simp [MvPolynomial.algebraMap_eq, Rat.smul_def] instance isAlgClosed : IsAlgClosed (AlgebraicClosure k) := diff --git a/Mathlib/FieldTheory/SplittingField/Construction.lean b/Mathlib/FieldTheory/SplittingField/Construction.lean index 0a11bdfa13be0..fb0572b56c8b9 100644 --- a/Mathlib/FieldTheory/SplittingField/Construction.lean +++ b/Mathlib/FieldTheory/SplittingField/Construction.lean @@ -245,7 +245,7 @@ instance instGroupWithZero : GroupWithZero (SplittingField f) := let e := algEquivSplittingFieldAux f { inv := fun a ↦ e.symm (e a)⁻¹ inv_zero := by simp - mul_inv_cancel := fun a ha ↦ e.injective $ by simp [(AddEquivClass.map_ne_zero_iff _).2 ha] + mul_inv_cancel := fun a ha ↦ e.injective <| by simp [(AddEquivClass.map_ne_zero_iff _).2 ha] __ := e.surjective.nontrivial } instance instField : Field (SplittingField f) where @@ -258,9 +258,9 @@ instance instField : Field (SplittingField f) where nnratCast_def q := by change algebraMap K _ _ = _; simp_rw [NNRat.cast_def, map_div₀, map_natCast] ratCast_def q := by change algebraMap K _ _ = _; rw [Rat.cast_def, map_div₀, map_intCast, map_natCast] - nnqsmul_def q x := Quotient.inductionOn x fun p ↦ congr_arg Quotient.mk'' $ by + nnqsmul_def q x := Quotient.inductionOn x fun p ↦ congr_arg Quotient.mk'' <| by ext; simp [MvPolynomial.algebraMap_eq, NNRat.smul_def] - qsmul_def q x := Quotient.inductionOn x fun p ↦ congr_arg Quotient.mk'' $ by + qsmul_def q x := Quotient.inductionOn x fun p ↦ congr_arg Quotient.mk'' <| by ext; simp [MvPolynomial.algebraMap_eq, Rat.smul_def] instance instCharZero [CharZero K] : CharZero (SplittingField f) := diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean index 1bdfc60311837..9dcdde2ba5730 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/Affine.lean @@ -134,7 +134,7 @@ nonrec theorem angle_le_pi (p1 p2 p3 : P) : ∠ p1 p2 p3 ≤ π := @[simp] lemma angle_self_right (p₀ p : P) : ∠ p p₀ p₀ = π / 2 := by rw [angle_comm, angle_self_left] /-- The angle ∠ABA at a point is `0`, unless `A = B`. -/ -theorem angle_self_of_ne (h : p ≠ p₀) : ∠ p p₀ p = 0 := angle_self $ vsub_ne_zero.2 h +theorem angle_self_of_ne (h : p ≠ p₀) : ∠ p p₀ p = 0 := angle_self <| vsub_ne_zero.2 h @[deprecated (since := "2024-02-14")] alias angle_eq_left := angle_self_left @[deprecated (since := "2024-02-14")] alias angle_eq_right := angle_self_right diff --git a/Mathlib/GroupTheory/FiniteAbelian.lean b/Mathlib/GroupTheory/FiniteAbelian.lean index 5c79378d634fc..0e74393883b9e 100644 --- a/Mathlib/GroupTheory/FiniteAbelian.lean +++ b/Mathlib/GroupTheory/FiniteAbelian.lean @@ -75,7 +75,7 @@ private def directSumNeZeroMulEquiv (ι : Type) [DecidableEq ι] (p : ι → ℕ · rw [map_zero, map_zero] · rw [DirectSum.toAddMonoid_of] split_ifs with h - · simp [(ZMod.subsingleton_iff.2 $ by rw [h, pow_zero]).elim x 0] + · simp [(ZMod.subsingleton_iff.2 <| by rw [h, pow_zero]).elim x 0] · simp_rw [directSumNeZeroMulHom, DirectSum.toAddMonoid_of] · rw [map_add, map_add, hx, hy] map_add' := map_add (directSumNeZeroMulHom p n) diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 9cb5066d29da6..fa52929c1c355 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -902,12 +902,12 @@ theorem zpow_mod_card (a : G) (n : ℤ) : a ^ (n % Fintype.card G : ℤ) = a ^ n @[to_additive (attr := simp) mod_natCard_nsmul] lemma pow_mod_natCard {G} [Group G] (a : G) (n : ℕ) : a ^ (n % Nat.card G) = a ^ n := by - rw [eq_comm, ← pow_mod_orderOf, ← Nat.mod_mod_of_dvd n $ orderOf_dvd_natCard _, pow_mod_orderOf] + rw [eq_comm, ← pow_mod_orderOf, ← Nat.mod_mod_of_dvd n <| orderOf_dvd_natCard _, pow_mod_orderOf] @[to_additive (attr := simp) mod_natCard_zsmul] lemma zpow_mod_natCard {G} [Group G] (a : G) (n : ℤ) : a ^ (n % Nat.card G : ℤ) = a ^ n := by - rw [eq_comm, ← zpow_mod_orderOf, - ← Int.emod_emod_of_dvd n $ Int.natCast_dvd_natCast.2 $ orderOf_dvd_natCard _, zpow_mod_orderOf] + rw [eq_comm, ← zpow_mod_orderOf, ← Int.emod_emod_of_dvd n <| + Int.natCast_dvd_natCast.2 <| orderOf_dvd_natCard _, zpow_mod_orderOf] /-- If `gcd(|G|,n)=1` then the `n`th power map is a bijection -/ @[to_additive (attr := simps) "If `gcd(|G|,n)=1` then the smul by `n` is a bijection"] @@ -1114,7 +1114,7 @@ lemma charP_of_prime_pow_injective (R) [Ring R] [Fintype R] (p n : ℕ) [hp : Fa obtain ⟨c, hc⟩ := CharP.exists R have hcpn : c ∣ p ^ n := by rw [← CharP.cast_eq_zero_iff R c, ← hn, Nat.cast_card_eq_zero] obtain ⟨i, hi, rfl⟩ : ∃ i ≤ n, c = p ^ i := by rwa [Nat.dvd_prime_pow hp.1] at hcpn - obtain rfl : i = n := hR i hi $ by rw [← Nat.cast_pow, CharP.cast_eq_zero] + obtain rfl : i = n := hR i hi <| by rw [← Nat.cast_pow, CharP.cast_eq_zero] assumption namespace SemiconjBy diff --git a/Mathlib/GroupTheory/Perm/Basic.lean b/Mathlib/GroupTheory/Perm/Basic.lean index fbe49dd2f60d0..4ff40a3a4a271 100644 --- a/Mathlib/GroupTheory/Perm/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Basic.lean @@ -36,9 +36,9 @@ instance permGroup : Group (Perm α) where mul_one := refl_trans inv_mul_cancel := self_trans_symm npow n f := f ^ n - npow_succ n f := coe_fn_injective $ Function.iterate_succ _ _ + npow_succ n f := coe_fn_injective <| Function.iterate_succ _ _ zpow := zpowRec fun n f ↦ f ^ n - zpow_succ' n f := coe_fn_injective $ Function.iterate_succ _ _ + zpow_succ' n f := coe_fn_injective <| Function.iterate_succ _ _ @[simp] theorem default_eq : (default : Perm α) = 1 := diff --git a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean index 5be189e9b0611..e59a26e0c1ce5 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean @@ -156,8 +156,8 @@ theorem finrank_vectorSpan_range_le [Fintype ι] (p : ι → P) {n : ℕ} (hc : /-- The `vectorSpan` of an indexed family of `n + 1` points has dimension at most `n`. -/ lemma finrank_vectorSpan_range_add_one_le [Fintype ι] [Nonempty ι] (p : ι → P) : finrank k (vectorSpan k (Set.range p)) + 1 ≤ Fintype.card ι := - (le_tsub_iff_right $ Nat.succ_le_iff.2 Fintype.card_pos).1 $ finrank_vectorSpan_range_le _ _ - (tsub_add_cancel_of_le $ Nat.succ_le_iff.2 Fintype.card_pos).symm + (le_tsub_iff_right <| Nat.succ_le_iff.2 Fintype.card_pos).1 <| finrank_vectorSpan_range_le _ _ + (tsub_add_cancel_of_le <| Nat.succ_le_iff.2 Fintype.card_pos).symm /-- `n + 1` points are affinely independent if and only if their `vectorSpan` has dimension `n`. -/ @@ -241,7 +241,7 @@ lemma AffineIndependent.card_lt_card_of_affineSpan_lt_affineSpan {s t : Finset V · simp [Set.subset_empty_iff] at hst have := hs'.to_subtype have := ht'.to_set.to_subtype - have dir_lt := AffineSubspace.direction_lt_of_nonempty (k := k) hst $ hs'.to_set.affineSpan k + have dir_lt := AffineSubspace.direction_lt_of_nonempty (k := k) hst <| hs'.to_set.affineSpan k rw [direction_affineSpan, direction_affineSpan, ← @Subtype.range_coe _ (s : Set V), ← @Subtype.range_coe _ (t : Set V)] at dir_lt have finrank_lt := add_lt_add_right (Submodule.finrank_lt_finrank_of_lt dir_lt) 1 diff --git a/Mathlib/LinearAlgebra/Matrix/ZPow.lean b/Mathlib/LinearAlgebra/Matrix/ZPow.lean index 5523292767d36..d28a1a1448f30 100644 --- a/Mathlib/LinearAlgebra/Matrix/ZPow.lean +++ b/Mathlib/LinearAlgebra/Matrix/ZPow.lean @@ -216,7 +216,7 @@ theorem zpow_add_one_of_ne_neg_one {A : M} : ∀ n : ℤ, n ≠ -1 → A ^ (n + rcases nonsing_inv_cancel_or_zero A with (⟨h, _⟩ | h) · apply zpow_add_one (isUnit_det_of_left_inverse h) · show A ^ (-((n + 1 : ℕ) : ℤ)) = A ^ (-((n + 2 : ℕ) : ℤ)) * A - simp_rw [zpow_neg_natCast, ← inv_pow', h, zero_pow $ Nat.succ_ne_zero _, zero_mul] + simp_rw [zpow_neg_natCast, ← inv_pow', h, zero_pow <| Nat.succ_ne_zero _, zero_mul] theorem zpow_mul (A : M) (h : IsUnit A.det) : ∀ m n : ℤ, A ^ (m * n) = (A ^ m) ^ n | (m : ℕ), (n : ℕ) => by rw [zpow_natCast, zpow_natCast, ← pow_mul, ← zpow_natCast, Int.ofNat_mul] diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index f20dd2b8bdbdf..8a58b466b4060 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -871,7 +871,7 @@ theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) @[deprecated (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)) ≤ + tailing f i n ⊔ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) ≤ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i n) := by dsimp [tailing, tunnel, tunnel', tunnelAux] erw [← Submodule.map_sup, sup_comm, Submodule.fst_sup_snd, Submodule.map_comp, Submodule.map_comp] diff --git a/Mathlib/LinearAlgebra/Ray.lean b/Mathlib/LinearAlgebra/Ray.lean index 04a2786fe4bde..02d99609f467d 100644 --- a/Mathlib/LinearAlgebra/Ray.lean +++ b/Mathlib/LinearAlgebra/Ray.lean @@ -106,7 +106,7 @@ lemma sameRay_nonneg_smul_right (v : M) (h : 0 ≤ a) : SameRay R v (a • v) := obtain h | h := (algebraMap_nonneg R h).eq_or_gt · rw [← algebraMap_smul R a v, h, zero_smul] exact zero_right _ - · refine Or.inr $ Or.inr ⟨algebraMap S R a, 1, h, by nontriviality R; exact zero_lt_one, ?_⟩ + · refine Or.inr <| Or.inr ⟨algebraMap S R a, 1, h, by nontriviality R; exact zero_lt_one, ?_⟩ rw [algebraMap_smul, one_smul] /-- A nonnegative multiple of a vector is in the same ray as that vector. -/ diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index b42323a8b7ae0..07b977deedc18 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -721,13 +721,13 @@ protected lemma exists_congr_left : (∃ a, p a) ↔ ∃ b, p (e.symm b) := e.symm.exists_congr_right.symm protected lemma exists_congr (h : ∀ a, p a ↔ q (e a)) : (∃ a, p a) ↔ ∃ b, q b := - e.exists_congr_left.trans $ by simp [h] + e.exists_congr_left.trans <| by simp [h] protected lemma exists_congr' (h : ∀ b, p (e.symm b) ↔ q b) : (∃ a, p a) ↔ ∃ b, q b := - e.exists_congr_left.trans $ by simp [h] + e.exists_congr_left.trans <| by simp [h] protected lemma existsUnique_congr_right : (∃! a, q (e a)) ↔ ∃! b, q b := - e.exists_congr $ by simpa using fun _ _ ↦ e.forall_congr (by simp) + e.exists_congr <| by simpa using fun _ _ ↦ e.forall_congr (by simp) protected lemma existsUnique_congr_left : (∃! a, p a) ↔ ∃! b, p (e.symm b) := e.symm.existsUnique_congr_right.symm @@ -739,12 +739,12 @@ alias exists_unique_congr_left := Equiv.existsUnique_congr_right alias exists_unique_congr_left' := Equiv.existsUnique_congr_left protected lemma existsUnique_congr (h : ∀ a, p a ↔ q (e a)) : (∃! a, p a) ↔ ∃! b, q b := - e.existsUnique_congr_left.trans $ by simp [h] + e.existsUnique_congr_left.trans <| by simp [h] @[deprecated (since := "2024-06-11")] alias exists_unique_congr := Equiv.existsUnique_congr protected lemma existsUnique_congr' (h : ∀ b, p (e.symm b) ↔ q b) : (∃! a, p a) ↔ ∃! b, q b := - e.existsUnique_congr_left.trans $ by simp [h] + e.existsUnique_congr_left.trans <| by simp [h] -- We next build some higher arity versions of `Equiv.forall_congr`. -- Although they appear to just be repeated applications of `Equiv.forall_congr`, @@ -756,7 +756,7 @@ protected lemma existsUnique_congr' (h : ∀ b, p (e.symm b) ↔ q b) : (∃! a, protected theorem forall₂_congr {α₁ α₂ β₁ β₂ : Sort*} {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (h : ∀ {x y}, p x y ↔ q (eα x) (eβ y)) : (∀ x y, p x y) ↔ ∀ x y, q x y := - eα.forall_congr fun _ ↦ eβ.forall_congr $ @h _ + eα.forall_congr fun _ ↦ eβ.forall_congr <| @h _ protected theorem forall₂_congr' {α₁ α₂ β₁ β₂ : Sort*} {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (h : ∀ {x y}, p (eα.symm x) (eβ.symm y) ↔ q x y) : @@ -766,7 +766,7 @@ protected theorem forall₃_congr {α₁ α₂ β₁ β₂ γ₁ γ₂ : Sort*} {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p x y z ↔ q (eα x) (eβ y) (eγ z)) : (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := - Equiv.forall₂_congr _ _ <| Equiv.forall_congr _ $ @h _ _ + Equiv.forall₂_congr _ _ <| Equiv.forall_congr _ <| @h _ _ protected theorem forall₃_congr' {α₁ α₂ β₁ β₂ γ₁ γ₂ : Sort*} {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} diff --git a/Mathlib/Logic/Godel/GodelBetaFunction.lean b/Mathlib/Logic/Godel/GodelBetaFunction.lean index a504860f1408a..d83abadb18246 100644 --- a/Mathlib/Logic/Godel/GodelBetaFunction.lean +++ b/Mathlib/Logic/Godel/GodelBetaFunction.lean @@ -46,7 +46,7 @@ lemma coprime_mul_succ {n m a} (h : n ≤ m) (ha : m - n ∣ a) : Coprime (n * a Nat.coprime_of_dvd fun p pp hn hm => by have : p ∣ (m - n) * a := by simpa [Nat.succ_sub_succ, ← Nat.mul_sub_right_distrib] using - Nat.dvd_sub (Nat.succ_le_succ $ Nat.mul_le_mul_right a h) hm hn + Nat.dvd_sub (Nat.succ_le_succ <| Nat.mul_le_mul_right a h) hm hn have : p ∣ a := by rcases (Nat.Prime.dvd_mul pp).mp this with (hp | hp) · exact Nat.dvd_trans hp ha @@ -62,7 +62,7 @@ private def coprimes (a : Fin m → ℕ) : Fin m → ℕ := fun i => (i + 1) * ( lemma coprimes_lt (a : Fin m → ℕ) (i) : a i < coprimes a i := by have h₁ : a i < supOfSeq a := - Nat.lt_add_one_iff.mpr (le_max_of_le_right $ Finset.le_sup (by simp)) + Nat.lt_add_one_iff.mpr (le_max_of_le_right <| Finset.le_sup (by simp)) have h₂ : supOfSeq a ≤ (i + 1) * (supOfSeq a)! + 1 := le_trans (self_le_factorial _) (le_trans (Nat.le_mul_of_pos_left (supOfSeq a)! (succ_pos i)) (le_add_right _ _)) @@ -75,7 +75,7 @@ private lemma pairwise_coprime_coprimes (a : Fin m → ℕ) : Pairwise (Coprime unfold Function.onFun coprimes have hja : j < supOfSeq a := lt_of_lt_of_le j.prop (le_step (le_max_left _ _)) exact coprime_mul_succ - (Nat.succ_le_succ $ le_of_lt ltij) + (Nat.succ_le_succ <| le_of_lt ltij) (Nat.dvd_factorial (by simp [Nat.succ_sub_succ, ltij]) (by simpa only [Nat.succ_sub_succ] using le_of_lt (lt_of_le_of_lt (sub_le j i) hja))) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Metric.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Metric.lean index c18fa56992e55..8abe04e25e044 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Metric.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Metric.lean @@ -196,7 +196,7 @@ theorem exists_borelSpace_of_countablyGenerated_of_separatesPoints (α : Type*) ∃ τ : TopologicalSpace α, SecondCountableTopology α ∧ T4Space α ∧ BorelSpace α := by rcases measurableEquiv_nat_bool_of_countablyGenerated α with ⟨s, ⟨f⟩⟩ letI := induced f inferInstance - let F := f.toEquiv.toHomeomorphOfInducing $ inducing_induced _ + let F := f.toEquiv.toHomeomorphOfInducing <| inducing_induced _ exact ⟨inferInstance, F.secondCountableTopology, F.symm.t4Space, MeasurableEmbedding.borelSpace f.measurableEmbedding F.inducing⟩ diff --git a/Mathlib/MeasureTheory/Constructions/EventuallyMeasurable.lean b/Mathlib/MeasureTheory/Constructions/EventuallyMeasurable.lean index 38354c1a4bf18..fe7b968db5833 100644 --- a/Mathlib/MeasureTheory/Constructions/EventuallyMeasurable.lean +++ b/Mathlib/MeasureTheory/Constructions/EventuallyMeasurable.lean @@ -75,7 +75,7 @@ namespace EventuallyMeasurableSpace instance measurableSingleton [MeasurableSingletonClass α] : @MeasurableSingletonClass α (EventuallyMeasurableSpace m l) := - @MeasurableSingletonClass.mk _ (_) $ fun x => (MeasurableSet.singleton x).eventuallyMeasurableSet + @MeasurableSingletonClass.mk _ (_) <| fun x => (MeasurableSet.singleton x).eventuallyMeasurableSet end EventuallyMeasurableSpace diff --git a/Mathlib/MeasureTheory/Function/Intersectivity.lean b/Mathlib/MeasureTheory/Function/Intersectivity.lean index a9113bf78d52b..5ea828b2fcf5b 100644 --- a/Mathlib/MeasureTheory/Function/Intersectivity.lean +++ b/Mathlib/MeasureTheory/Function/Intersectivity.lean @@ -62,27 +62,27 @@ lemma bergelson' {s : ℕ → Set α} (hs : ∀ n, MeasurableSet (s n)) (hr₀ : simp only [f, Pi.natCast_def, Pi.smul_apply, Pi.inv_apply, Finset.sum_apply, eq_self_iff_true, forall_const, imp_true_iff, smul_eq_mul] have hf n : Measurable (f n) := Measurable.mul' (@measurable_const ℝ≥0∞ _ _ _ (↑(n + 1))⁻¹) - (Finset.measurable_sum' _ fun i _ ↦ measurable_one.indicator $ hs i) + (Finset.measurable_sum' _ fun i _ ↦ measurable_one.indicator <| hs i) have hf₁ n : f n ≤ 1 := by rintro a rw [hfapp, ← ENNReal.div_eq_inv_mul] - refine (ENNReal.div_le_iff_le_mul (Or.inl $ Nat.cast_ne_zero.2 n.succ_ne_zero) $ + refine (ENNReal.div_le_iff_le_mul (Or.inl <| Nat.cast_ne_zero.2 n.succ_ne_zero) <| Or.inr one_ne_zero).2 ?_ rw [mul_comm, ← nsmul_eq_mul, ← Finset.card_range n.succ] exact Finset.sum_le_card_nsmul _ _ _ fun _ _ ↦ indicator_le (fun _ _ ↦ le_rfl) _ -- By assumption, `f n` has integral at least `r`. have hrf n : r ≤ ∫⁻ a, f n a ∂μ := by simp_rw [hfapp] - rw [lintegral_const_mul _ (Finset.measurable_sum _ fun _ _ ↦ measurable_one.indicator $ hs _), + rw [lintegral_const_mul _ (Finset.measurable_sum _ fun _ _ ↦ measurable_one.indicator <| hs _), lintegral_finset_sum _ fun _ _ ↦ measurable_one.indicator (hs _)] simp only [lintegral_indicator_one (hs _)] rw [← ENNReal.div_eq_inv_mul, ENNReal.le_div_iff_mul_le (by simp) (by simp), ← nsmul_eq_mul'] simpa using Finset.card_nsmul_le_sum (Finset.range (n + 1)) _ _ fun _ _ ↦ hr _ -- Collect some basic fact - have hμ : μ ≠ 0 := by rintro rfl; exact hr₀ $ le_bot_iff.1 $ hr 0 + have hμ : μ ≠ 0 := by rintro rfl; exact hr₀ <| le_bot_iff.1 <| hr 0 have : ∫⁻ x, limsup (f · x) atTop ∂μ ≤ μ univ := by rw [← lintegral_one] - exact lintegral_mono fun a ↦ limsup_le_of_le ⟨0, fun R _ ↦ bot_le⟩ $ + exact lintegral_mono fun a ↦ limsup_le_of_le ⟨0, fun R _ ↦ bot_le⟩ <| Eventually.of_forall fun n ↦ hf₁ _ _ -- By the first moment method, there exists some `x ∉ N` such that `limsup f n x` is at least `r`. obtain ⟨x, hxN, hx⟩ := exists_not_mem_null_laverage_le hμ @@ -90,23 +90,23 @@ lemma bergelson' {s : ℕ → Set α} (hs : ∀ n, MeasurableSet (s n)) (hr₀ : replace hx : r / μ univ ≤ limsup (f · x) atTop := calc _ ≤ limsup (⨍⁻ x, f · x ∂μ) atTop := le_limsup_of_le ⟨1, eventually_map.2 ?_⟩ fun b hb ↦ ?_ - _ ≤ ⨍⁻ x, limsup (f · x) atTop ∂μ := limsup_lintegral_le 1 hf (ae_of_all _ $ hf₁ ·) (by simp) + _ ≤ ⨍⁻ x, limsup (f · x) atTop ∂μ := limsup_lintegral_le 1 hf (ae_of_all _ <| hf₁ ·) (by simp) _ ≤ limsup (f · x) atTop := hx -- This exactly means that the `s n` containing `x` have all their finite intersection non-null. · refine ⟨{n | x ∈ s n}, fun hxs ↦ ?_, fun u hux hu ↦ ?_⟩ -- This next block proves that a set of strictly positive natural density is infinite, mixed -- with the fact that `{n | x ∈ s n}` has strictly positive natural density. -- TODO: Separate it out to a lemma once we have a natural density API. - · refine ENNReal.div_ne_zero.2 ⟨hr₀, measure_ne_top _ _⟩ $ eq_bot_mono hx $ Tendsto.limsup_eq $ - tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds + · refine ENNReal.div_ne_zero.2 ⟨hr₀, measure_ne_top _ _⟩ <| eq_bot_mono hx <| + Tendsto.limsup_eq <| tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds (h := fun n ↦ (n.succ : ℝ≥0∞)⁻¹ * hxs.toFinset.card) ?_ bot_le fun n ↦ mul_le_mul_left' ?_ _ - · simpa using ENNReal.Tendsto.mul_const (ENNReal.tendsto_inv_nat_nhds_zero.comp $ - tendsto_add_atTop_nat 1) (.inr $ ENNReal.natCast_ne_top _) + · simpa using ENNReal.Tendsto.mul_const (ENNReal.tendsto_inv_nat_nhds_zero.comp <| + tendsto_add_atTop_nat 1) (.inr <| ENNReal.natCast_ne_top _) · classical simpa only [Finset.sum_apply, indicator_apply, Pi.one_apply, Finset.sum_boole, Nat.cast_le] using Finset.card_le_card fun m hm ↦ hxs.mem_toFinset.2 (Finset.mem_filter.1 hm).2 · simp_rw [← hu.mem_toFinset] - exact hN₁ _ ⟨x, mem_iInter₂.2 fun n hn ↦ hux $ hu.mem_toFinset.1 hn, hxN⟩ + exact hN₁ _ ⟨x, mem_iInter₂.2 fun n hn ↦ hux <| hu.mem_toFinset.1 hn, hxN⟩ · refine Eventually.of_forall fun n ↦ ?_ obtain rfl | _ := eq_zero_or_neZero μ · simp @@ -121,9 +121,9 @@ measure at least `r` has an infinite subset whose finite intersections all have lemma bergelson [Infinite ι] {s : ι → Set α} (hs : ∀ i, MeasurableSet (s i)) (hr₀ : r ≠ 0) (hr : ∀ i, r ≤ μ (s i)) : ∃ t : Set ι, t.Infinite ∧ ∀ ⦃u⦄, u ⊆ t → u.Finite → 0 < μ (⋂ i ∈ u, s i) := by - obtain ⟨t, ht, h⟩ := bergelson' (fun n ↦ hs $ Infinite.natEmbedding _ n) hr₀ (fun n ↦ hr _) - refine ⟨_, ht.image $ (Infinite.natEmbedding _).injective.injOn, fun u hut hu ↦ - (h (preimage_subset_of_surjOn (Infinite.natEmbedding _).injective hut) $ hu.preimage - (Embedding.injective _).injOn).trans_le $ measure_mono $ subset_iInter₂ fun i hi ↦ ?_⟩ + obtain ⟨t, ht, h⟩ := bergelson' (fun n ↦ hs <| Infinite.natEmbedding _ n) hr₀ (fun n ↦ hr _) + refine ⟨_, ht.image <| (Infinite.natEmbedding _).injective.injOn, fun u hut hu ↦ + (h (preimage_subset_of_surjOn (Infinite.natEmbedding _).injective hut) <| hu.preimage + (Embedding.injective _).injOn).trans_le <| measure_mono <| subset_iInter₂ fun i hi ↦ ?_⟩ obtain ⟨n, -, rfl⟩ := hut hi exact iInter₂_subset n hi diff --git a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean index 9e2582523ac5f..f2a385de89049 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/CountablyGenerated.lean @@ -89,7 +89,7 @@ lemma generateFrom_natGeneratingSequence (α : Type*) [m : MeasurableSpace α] lemma measurableSet_natGeneratingSequence [MeasurableSpace α] [CountablyGenerated α] (n : ℕ) : MeasurableSet (natGeneratingSequence α n) := - measurableSet_countableGeneratingSet $ Set.enumerateCountable_mem _ + measurableSet_countableGeneratingSet <| Set.enumerateCountable_mem _ empty_mem_countableGeneratingSet n theorem CountablyGenerated.comap [m : MeasurableSpace β] [h : CountablyGenerated β] (f : α → β) : @@ -159,7 +159,7 @@ theorem separating_of_generateFrom (S : Set (Set α)) letI := generateFrom S intros x y hxy rw [← forall_generateFrom_mem_iff_mem_iff] at hxy - exact separatesPoints_def $ fun _ hs ↦ (hxy _ hs).mp + exact separatesPoints_def <| fun _ hs ↦ (hxy _ hs).mp theorem SeparatesPoints.mono {m m' : MeasurableSpace α} [hsep : @SeparatesPoints _ m] (h : m ≤ m') : @SeparatesPoints _ m' := @SeparatesPoints.mk _ m' fun _ _ hxy ↦ @@ -194,12 +194,12 @@ theorem CountablySeparated.subtype_iff [MeasurableSpace α] {s : Set α} : instance (priority := 100) Subtype.separatesPoints [MeasurableSpace α] [h : SeparatesPoints α] {s : Set α} : SeparatesPoints s := - ⟨fun _ _ hxy ↦ Subtype.val_injective $ h.1 _ _ fun _ ht ↦ hxy _ $ measurable_subtype_coe ht⟩ + ⟨fun _ _ hxy ↦ Subtype.val_injective <| h.1 _ _ fun _ ht ↦ hxy _ <| measurable_subtype_coe ht⟩ instance (priority := 100) Subtype.countablySeparated [MeasurableSpace α] [h : CountablySeparated α] {s : Set α} : CountablySeparated s := by rw [CountablySeparated.subtype_iff] - exact h.countably_separated.mono (fun s ↦ id) $ subset_univ _ + exact h.countably_separated.mono (fun s ↦ id) <| subset_univ _ instance (priority := 100) separatesPoints_of_measurableSingletonClass [MeasurableSpace α] [MeasurableSingletonClass α] : SeparatesPoints α := by @@ -235,7 +235,7 @@ theorem exists_countablyGenerated_le_of_countablySeparated [m : MeasurableSpace refine ⟨generateFrom b, ?_, ?_, generateFrom_le hbm⟩ · use b rw [@separatesPoints_iff] - exact fun x y hxy ↦ hb _ trivial _ trivial fun _ hs ↦ hxy _ $ measurableSet_generateFrom hs + exact fun x y hxy ↦ hb _ trivial _ trivial fun _ hs ↦ hxy _ <| measurableSet_generateFrom hs open Function @@ -270,9 +270,9 @@ the Cantor Space. -/ theorem measurableEquiv_nat_bool_of_countablyGenerated [MeasurableSpace α] [CountablyGenerated α] [SeparatesPoints α] : ∃ s : Set (ℕ → Bool), Nonempty (α ≃ᵐ s) := by - use range (mapNatBool α), Equiv.ofInjective _ $ + use range (mapNatBool α), Equiv.ofInjective _ <| injective_mapNatBool _, - Measurable.subtype_mk $ measurable_mapNatBool _ + Measurable.subtype_mk <| measurable_mapNatBool _ simp_rw [← generateFrom_natGeneratingSequence α] apply measurable_generateFrom rintro _ ⟨n, rfl⟩ @@ -299,7 +299,7 @@ theorem measurableSingletonClass_of_countablySeparated rcases measurable_injection_nat_bool_of_countablySeparated α with ⟨f, fmeas, finj⟩ refine ⟨fun x ↦ ?_⟩ rw [← finj.preimage_image {x}, image_singleton] - exact fmeas $ MeasurableSet.singleton _ + exact fmeas <| MeasurableSet.singleton _ end SeparatesPoints diff --git a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean index 049c09311b973..f10ffdcda79a3 100644 --- a/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/FiniteMeasure.lean @@ -130,7 +130,7 @@ theorem toMeasure_injective : Function.Injective ((↑) : FiniteMeasure Ω → M instance instFunLike : FunLike (FiniteMeasure Ω) (Set Ω) ℝ≥0 where coe μ s := ((μ : Measure Ω) s).toNNReal - coe_injective' μ ν h := toMeasure_injective $ Measure.ext fun s _ ↦ by + coe_injective' μ ν h := toMeasure_injective <| Measure.ext fun s _ ↦ by simpa [ENNReal.toNNReal_eq_toNNReal_iff, measure_ne_top] using congr_fun h s lemma coeFn_def (μ : FiniteMeasure Ω) : μ = fun s ↦ ((μ : Measure Ω) s).toNNReal := rfl diff --git a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean index 31dbff7cf6c84..a63821bbd70b0 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Quotient.lean @@ -228,7 +228,7 @@ theorem MeasureTheory.QuotientMeasureEqMeasurePreimage.haarMeasure_quotient [Loc MeasureTheory.QuotientMeasureEqMeasurePreimage.mulInvariantMeasure_quotient ν rw [haarMeasure_unique μ K'] have finiteCovol : covolume Γ.op G ν ≠ ⊤ := - ne_top_of_lt $ QuotientMeasureEqMeasurePreimage.covolume_ne_top μ (ν := ν) + ne_top_of_lt <| QuotientMeasureEqMeasurePreimage.covolume_ne_top μ (ν := ν) obtain ⟨s, fund_dom_s⟩ := i rw [fund_dom_s.covolume_eq_volume] at finiteCovol -- TODO: why `rw` fails? diff --git a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean index d908cc8f3afd6..62aa6c8957d30 100644 --- a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean @@ -128,7 +128,7 @@ theorem toMeasure_injective : Function.Injective ((↑) : ProbabilityMeasure Ω instance instFunLike : FunLike (ProbabilityMeasure Ω) (Set Ω) ℝ≥0 where coe μ s := ((μ : Measure Ω) s).toNNReal - coe_injective' μ ν h := toMeasure_injective $ Measure.ext fun s _ ↦ by + coe_injective' μ ν h := toMeasure_injective <| Measure.ext fun s _ ↦ by simpa [ENNReal.toNNReal_eq_toNNReal_iff, measure_ne_top] using congr_fun h s lemma coeFn_def (μ : ProbabilityMeasure Ω) : μ = fun s ↦ ((μ : Measure Ω) s).toNNReal := rfl diff --git a/Mathlib/MeasureTheory/Order/UpperLower.lean b/Mathlib/MeasureTheory/Order/UpperLower.lean index ef8fa8580acfc..5e56bec648560 100644 --- a/Mathlib/MeasureTheory/Order/UpperLower.lean +++ b/Mathlib/MeasureTheory/Order/UpperLower.lean @@ -64,7 +64,7 @@ private lemma aux₀ intro H obtain ⟨ε, -, hε', hε₀⟩ := exists_seq_strictAnti_tendsto_nhdsWithin (0 : ℝ) refine not_eventually.2 - (Frequently.of_forall fun _ ↦ lt_irrefl $ ENNReal.ofReal $ 4⁻¹ ^ Fintype.card ι) + (Frequently.of_forall fun _ ↦ lt_irrefl <| ENNReal.ofReal <| 4⁻¹ ^ Fintype.card ι) ((Filter.Tendsto.eventually_lt (H.comp hε₀) tendsto_const_nhds ?_).mono fun n ↦ lt_of_le_of_lt ?_) on_goal 2 => @@ -72,7 +72,8 @@ private lemma aux₀ ENNReal.ofReal (4⁻¹ ^ Fintype.card ι) = volume (closedBall (f (ε n) (hε' n)) (ε n / 4)) / volume (closedBall x (ε n)) := ?_ _ ≤ volume (closure s ∩ closedBall x (ε n)) / volume (closedBall x (ε n)) := by - gcongr; exact subset_inter ((hf₁ _ $ hε' n).trans interior_subset_closure) $ hf₀ _ $ hε' n + gcongr + exact subset_inter ((hf₁ _ <| hε' n).trans interior_subset_closure) <| hf₀ _ <| hε' n have := hε' n rw [Real.volume_pi_closedBall, Real.volume_pi_closedBall, ← ENNReal.ofReal_div_of_pos, ← div_pow, mul_div_mul_left _ _ (two_ne_zero' ℝ), div_right_comm, div_self, one_div] @@ -92,14 +93,14 @@ private lemma aux₁ intro H obtain ⟨ε, -, hε', hε₀⟩ := exists_seq_strictAnti_tendsto_nhdsWithin (0 : ℝ) refine not_eventually.2 - (Frequently.of_forall fun _ ↦ lt_irrefl $ 1 - ENNReal.ofReal (4⁻¹ ^ Fintype.card ι)) - ((Filter.Tendsto.eventually_lt tendsto_const_nhds (H.comp hε₀) $ + (Frequently.of_forall fun _ ↦ lt_irrefl <| 1 - ENNReal.ofReal (4⁻¹ ^ Fintype.card ι)) + ((Filter.Tendsto.eventually_lt tendsto_const_nhds (H.comp hε₀) <| ENNReal.sub_lt_self ENNReal.one_ne_top one_ne_zero ?_).mono fun n ↦ lt_of_le_of_lt' ?_) on_goal 2 => calc volume (closure s ∩ closedBall x (ε n)) / volume (closedBall x (ε n)) - ≤ volume (closedBall x (ε n) \ closedBall (f (ε n) $ hε' n) (ε n / 4)) / + ≤ volume (closedBall x (ε n) \ closedBall (f (ε n) <| hε' n) (ε n / 4)) / volume (closedBall x (ε n)) := by gcongr rw [diff_eq_compl_inter] @@ -124,9 +125,9 @@ theorem IsUpperSet.null_frontier (hs : IsUpperSet s) : volume (frontier s) = 0 : by_cases h : x ∈ closure s <;> simp only [mem_compl_iff, mem_setOf, h, not_false_eq_true, indicator_of_not_mem, indicator_of_mem, Pi.one_apply] - · refine aux₁ fun _ ↦ hs.compl.exists_subset_ball $ frontier_subset_closure ?_ + · refine aux₁ fun _ ↦ hs.compl.exists_subset_ball <| frontier_subset_closure ?_ rwa [frontier_compl] - · exact aux₀ fun _ ↦ hs.exists_subset_ball $ frontier_subset_closure hx + · exact aux₀ fun _ ↦ hs.exists_subset_ball <| frontier_subset_closure hx theorem IsLowerSet.null_frontier (hs : IsLowerSet s) : volume (frontier s) = 0 := by refine measure_mono_null (fun x hx ↦ ?_) @@ -135,13 +136,13 @@ theorem IsLowerSet.null_frontier (hs : IsLowerSet s) : volume (frontier s) = 0 : by_cases h : x ∈ closure s <;> simp only [mem_compl_iff, mem_setOf, h, not_false_eq_true, indicator_of_not_mem, indicator_of_mem, Pi.one_apply] - · refine aux₁ fun _ ↦ hs.compl.exists_subset_ball $ frontier_subset_closure ?_ + · refine aux₁ fun _ ↦ hs.compl.exists_subset_ball <| frontier_subset_closure ?_ rwa [frontier_compl] - · exact aux₀ fun _ ↦ hs.exists_subset_ball $ frontier_subset_closure hx + · exact aux₀ fun _ ↦ hs.exists_subset_ball <| frontier_subset_closure hx theorem Set.OrdConnected.null_frontier (hs : s.OrdConnected) : volume (frontier s) = 0 := by rw [← hs.upperClosure_inter_lowerClosure] - exact measure_mono_null (frontier_inter_subset _ _) $ measure_union_null + exact measure_mono_null (frontier_inter_subset _ _) <| measure_union_null (measure_inter_null_of_null_left _ (UpperSet.upper _).null_frontier) (measure_inter_null_of_null_right _ (LowerSet.lower _).null_frontier) diff --git a/Mathlib/MeasureTheory/OuterMeasure/AE.lean b/Mathlib/MeasureTheory/OuterMeasure/AE.lean index 2b8af44469a76..bf766c468a301 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/AE.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/AE.lean @@ -121,7 +121,7 @@ theorem ae_eq_trans {f g h : α → β} (h₁ : f =ᵐ[μ] g) (h₂ : g =ᵐ[μ] refine ⟨fun h a ha ↦ by simpa [ha] using (h {a}ᶜ).1, fun h s ↦ ⟨fun hs ↦ ?_, ?_⟩⟩ · rw [← compl_empty_iff, ← not_nonempty_iff_eq_empty] rintro ⟨a, ha⟩ - exact h _ $ measure_mono_null (singleton_subset_iff.2 ha) hs + exact h _ <| measure_mono_null (singleton_subset_iff.2 ha) hs · rintro rfl simp diff --git a/Mathlib/NumberTheory/NumberField/Units/Basic.lean b/Mathlib/NumberTheory/NumberField/Units/Basic.lean index 84ab8c9274a4e..bacedd52bdeb2 100644 --- a/Mathlib/NumberTheory/NumberField/Units/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Units/Basic.lean @@ -102,7 +102,7 @@ def torsion : Subgroup (𝓞 K)ˣ := CommGroup.torsion (𝓞 K)ˣ theorem mem_torsion {x : (𝓞 K)ˣ} [NumberField K] : x ∈ torsion K ↔ ∀ w : InfinitePlace K, w x = 1 := by rw [eq_iff_eq (x : K) 1, torsion, CommGroup.mem_torsion] - refine ⟨fun hx φ ↦ (((φ.comp $ algebraMap (𝓞 K) K).toMonoidHom.comp $ + refine ⟨fun hx φ ↦ (((φ.comp <| algebraMap (𝓞 K) K).toMonoidHom.comp <| Units.coeHom _).isOfFinOrder hx).norm_eq_one, fun h ↦ isOfFinOrder_iff_pow_eq_one.2 ?_⟩ obtain ⟨n, hn, hx⟩ := Embeddings.pow_eq_one_of_norm_eq_one K ℂ x.val.isIntegral_coe h exact ⟨n, hn, by ext; rw [NumberField.RingOfIntegers.coe_eq_algebraMap, coe_pow, hx, diff --git a/Mathlib/Order/Antichain.lean b/Mathlib/Order/Antichain.lean index 7350341af87c7..cdfcbf4009f47 100644 --- a/Mathlib/Order/Antichain.lean +++ b/Mathlib/Order/Antichain.lean @@ -283,11 +283,11 @@ variable [PartialOrder α] [PartialOrder β] {f : α → β} {s : Set α} lemma IsAntichain.of_strictMonoOn_antitoneOn (hf : StrictMonoOn f s) (hf' : AntitoneOn f s) : IsAntichain (· ≤ ·) s := - fun _a ha _b hb hab' hab ↦ (hf ha hb $ hab.lt_of_ne hab').not_le (hf' ha hb hab) + fun _a ha _b hb hab' hab ↦ (hf ha hb <| hab.lt_of_ne hab').not_le (hf' ha hb hab) lemma IsAntichain.of_monotoneOn_strictAntiOn (hf : MonotoneOn f s) (hf' : StrictAntiOn f s) : IsAntichain (· ≤ ·) s := - fun _a ha _b hb hab' hab ↦ (hf ha hb hab).not_lt (hf' ha hb $ hab.lt_of_ne hab') + fun _a ha _b hb hab' hab ↦ (hf ha hb hab).not_lt (hf' ha hb <| hab.lt_of_ne hab') end General diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 8f3d90b665353..6e4c780550a19 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -400,7 +400,7 @@ lemma exists_forall_ge_and [LinearOrder α] {p q : α → Prop} : (∃ i, ∀ j ≥ i, p j) → (∃ i, ∀ j ≥ i, q j) → ∃ i, ∀ j ≥ i, p j ∧ q j | ⟨a, ha⟩, ⟨b, hb⟩ => let ⟨c, hac, hbc⟩ := exists_ge_of_linear a b - ⟨c, fun _d hcd ↦ ⟨ha _ $ hac.trans hcd, hb _ $ hbc.trans hcd⟩⟩ + ⟨c, fun _d hcd ↦ ⟨ha _ <| hac.trans hcd, hb _ <| hbc.trans hcd⟩⟩ theorem lt_imp_lt_of_le_imp_le {β} [LinearOrder α] [Preorder β] {a b : α} {c d : β} (H : a ≤ b → c ≤ d) (h : d < c) : b < a := diff --git a/Mathlib/Order/BooleanAlgebra.lean b/Mathlib/Order/BooleanAlgebra.lean index 942a70f8f6bec..509ef845db9a3 100644 --- a/Mathlib/Order/BooleanAlgebra.lean +++ b/Mathlib/Order/BooleanAlgebra.lean @@ -767,7 +767,7 @@ protected abbrev Function.Injective.booleanAlgebra [Sup α] [Inf α] [Top α] [B sdiff_eq a b := by refine hf ((map_sdiff _ _).trans (sdiff_eq.trans ?_)) rw [map_inf, map_compl] - himp_eq a b := hf $ (map_himp _ _).trans $ himp_eq.trans $ by rw [map_sup, map_compl] + himp_eq a b := hf <| (map_himp _ _).trans <| himp_eq.trans <| by rw [map_sup, map_compl] end lift diff --git a/Mathlib/Order/Category/BoolAlg.lean b/Mathlib/Order/Category/BoolAlg.lean index 3afeb442490cf..e90e6e55aed90 100644 --- a/Mathlib/Order/Category/BoolAlg.lean +++ b/Mathlib/Order/Category/BoolAlg.lean @@ -106,6 +106,6 @@ theorem boolAlg_dual_comp_forget_to_bddDistLat : /-- The powerset functor. `Set` as a contravariant functor. -/ @[simps] def typeToBoolAlgOp : Type u ⥤ BoolAlgᵒᵖ where - obj X := op $ BoolAlg.of (Set X) + obj X := op <| BoolAlg.of (Set X) map {X Y} f := Quiver.Hom.op (CompleteLatticeHom.setPreimage f : BoundedLatticeHom (Set Y) (Set X)) diff --git a/Mathlib/Order/Directed.lean b/Mathlib/Order/Directed.lean index 22dbc59d0d72e..8ab4e604c1f95 100644 --- a/Mathlib/Order/Directed.lean +++ b/Mathlib/Order/Directed.lean @@ -269,14 +269,14 @@ variable {f : α → β} {s : Set α} lemma constant_of_monotone_antitone [IsDirected α (· ≤ ·)] (hf : Monotone f) (hf' : Antitone f) (a b : α) : f a = f b := by obtain ⟨c, hac, hbc⟩ := exists_ge_ge a b - exact le_antisymm ((hf hac).trans $ hf' hbc) ((hf hbc).trans $ hf' hac) + exact le_antisymm ((hf hac).trans <| hf' hbc) ((hf hbc).trans <| hf' hac) /-- If `f` is monotone and antitone on a directed set `s`, then `f` is constant on `s`. -/ lemma constant_of_monotoneOn_antitoneOn (hf : MonotoneOn f s) (hf' : AntitoneOn f s) (hs : DirectedOn (· ≤ ·) s) : ∀ ⦃a⦄, a ∈ s → ∀ ⦃b⦄, b ∈ s → f a = f b := by rintro a ha b hb obtain ⟨c, hc, hac, hbc⟩ := hs _ ha _ hb - exact le_antisymm ((hf ha hc hac).trans $ hf' hb hc hbc) ((hf hb hc hbc).trans $ hf' ha hc hac) + exact le_antisymm ((hf ha hc hac).trans <| hf' hb hc hbc) ((hf hb hc hbc).trans <| hf' ha hc hac) end Preorder diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index aaa43927f9f6c..f2357f5cf20d7 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -793,7 +793,7 @@ end OrderedSemiring theorem zero_pow_eventuallyEq [MonoidWithZero α] : (fun n : ℕ => (0 : α) ^ n) =ᶠ[atTop] fun _ => 0 := - eventually_atTop.2 ⟨1, fun _n hn ↦ zero_pow $ Nat.one_le_iff_ne_zero.1 hn⟩ + eventually_atTop.2 ⟨1, fun _n hn ↦ zero_pow <| Nat.one_le_iff_ne_zero.1 hn⟩ section OrderedRing diff --git a/Mathlib/Order/Filter/CountableSeparatingOn.lean b/Mathlib/Order/Filter/CountableSeparatingOn.lean index affcf2f203ee0..b1f268cda8419 100644 --- a/Mathlib/Order/Filter/CountableSeparatingOn.lean +++ b/Mathlib/Order/Filter/CountableSeparatingOn.lean @@ -127,13 +127,13 @@ theorem HasCountableSeparatingOn.subtype_iff {α : Type*} {p : Set α → Prop} HasCountableSeparatingOn t (fun u ↦ ∃ v, p v ∧ (↑) ⁻¹' v = u) univ ↔ HasCountableSeparatingOn α p t := by constructor <;> intro h - · exact h.of_subtype $ fun s ↦ id + · exact h.of_subtype <| fun s ↦ id rcases h with ⟨S, Sct, Sp, hS⟩ use {Subtype.val ⁻¹' s | s ∈ S}, Sct.image _, ?_, ?_ · rintro u ⟨t, tS, rfl⟩ exact ⟨t, Sp _ tS, rfl⟩ rintro x - y - hxy - exact Subtype.val_injective $ hS _ (Subtype.coe_prop _) _ (Subtype.coe_prop _) + exact Subtype.val_injective <| hS _ (Subtype.coe_prop _) _ (Subtype.coe_prop _) fun s hs ↦ hxy (Subtype.val ⁻¹' s) ⟨s, hs, rfl⟩ namespace Filter diff --git a/Mathlib/Order/Filter/Curry.lean b/Mathlib/Order/Filter/Curry.lean index 4652e9981ebdf..561f48bd033b3 100644 --- a/Mathlib/Order/Filter/Curry.lean +++ b/Mathlib/Order/Filter/Curry.lean @@ -79,11 +79,11 @@ theorem frequently_curry_prod_iff {α β : Type*} {l : Filter α} {m : Filter β refine ⟨fun h => ?_, fun ⟨hs, ht⟩ => ?_⟩ · exact frequently_prod_and.mp (Frequently.filter_mono h curry_le_prod) rw [frequently_curry_iff] - exact Frequently.mono hs $ fun x hx => Frequently.mono ht (by simp[hx]) + exact Frequently.mono hs <| fun x hx => Frequently.mono ht (by simp [hx]) theorem prod_mem_curry {α β : Type*} {l : Filter α} {m : Filter β} {s : Set α} {t : Set β} (hs : s ∈ l) (ht : t ∈ m) : s ×ˢ t ∈ l.curry m := - curry_le_prod $ prod_mem_prod hs ht + curry_le_prod <| prod_mem_prod hs ht theorem eventually_curry_prod_iff {α β : Type*} {l : Filter α} {m : Filter β} [NeBot l] [NeBot m] (s : Set α) (t : Set β) : diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index 8d85a50b90b7e..133d9233ac466 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -83,7 +83,7 @@ theorem card_Ioo : (Ioo a b).card = b - a - 1 := @[simp] theorem card_uIcc : (uIcc a b).card = (b - a : ℤ).natAbs + 1 := - (card_Icc _ _).trans $ by rw [← Int.natCast_inj, sup_eq_max, inf_eq_min, Int.ofNat_sub] <;> omega + (card_Icc _ _).trans <| by rw [← Int.natCast_inj, sup_eq_max, inf_eq_min, Int.ofNat_sub] <;> omega @[simp] lemma card_Iic : (Iic b).card = b + 1 := by rw [Iic_eq_Icc, card_Icc, Nat.bot_eq_zero, Nat.sub_zero] @@ -199,7 +199,7 @@ theorem mod_injOn_Ico (n a : ℕ) : Set.InjOn (· % a) (Finset.Ico n (n + a)) := rw [Ico_succ_left_eq_erase_Ico, succ_add, succ_eq_add_one, Ico_succ_right_eq_insert_Ico (by omega)] rintro k hk l hl (hkl : k % a = l % a) - have ha : 0 < a := Nat.pos_iff_ne_zero.2 $ by rintro rfl; simp at hk + have ha : 0 < a := Nat.pos_iff_ne_zero.2 <| by rintro rfl; simp at hk simp only [Finset.mem_coe, Finset.mem_insert, Finset.mem_erase] at hk hl rcases hk with ⟨hkn, rfl | hk⟩ <;> rcases hl with ⟨hln, rfl | hl⟩ · rfl diff --git a/Mathlib/Order/Interval/Set/Image.lean b/Mathlib/Order/Interval/Set/Image.lean index 98810e2ae65b4..0b194b81b848d 100644 --- a/Mathlib/Order/Interval/Set/Image.lean +++ b/Mathlib/Order/Interval/Set/Image.lean @@ -431,18 +431,18 @@ lemma directedOn_le_Iic (b : α) : DirectedOn (· ≤ ·) (Iic b) := fun _x hx _y hy ↦ ⟨b, le_rfl, hx, hy⟩ lemma directedOn_le_Icc (a b : α) : DirectedOn (· ≤ ·) (Icc a b) := - fun _x hx _y hy ↦ ⟨b, right_mem_Icc.2 $ hx.1.trans hx.2, hx.2, hy.2⟩ + fun _x hx _y hy ↦ ⟨b, right_mem_Icc.2 <| hx.1.trans hx.2, hx.2, hy.2⟩ lemma directedOn_le_Ioc (a b : α) : DirectedOn (· ≤ ·) (Ioc a b) := - fun _x hx _y hy ↦ ⟨b, right_mem_Ioc.2 $ hx.1.trans_le hx.2, hx.2, hy.2⟩ + fun _x hx _y hy ↦ ⟨b, right_mem_Ioc.2 <| hx.1.trans_le hx.2, hx.2, hy.2⟩ lemma directedOn_ge_Ici (a : α) : DirectedOn (· ≥ ·) (Ici a) := fun _x hx _y hy ↦ ⟨a, le_rfl, hx, hy⟩ lemma directedOn_ge_Icc (a b : α) : DirectedOn (· ≥ ·) (Icc a b) := - fun _x hx _y hy ↦ ⟨a, left_mem_Icc.2 $ hx.1.trans hx.2, hx.1, hy.1⟩ + fun _x hx _y hy ↦ ⟨a, left_mem_Icc.2 <| hx.1.trans hx.2, hx.1, hy.1⟩ lemma directedOn_ge_Ico (a b : α) : DirectedOn (· ≥ ·) (Ico a b) := - fun _x hx _y hy ↦ ⟨a, left_mem_Ico.2 $ hx.1.trans_lt hx.2, hx.1, hy.1⟩ + fun _x hx _y hy ↦ ⟨a, left_mem_Ico.2 <| hx.1.trans_lt hx.2, hx.1, hy.1⟩ end Preorder diff --git a/Mathlib/Order/Part.lean b/Mathlib/Order/Part.lean index 7b1dbb20f678d..095e4622ab550 100644 --- a/Mathlib/Order/Part.lean +++ b/Mathlib/Order/Part.lean @@ -47,10 +47,10 @@ section seq variable {β γ : Type _} {f : α → Part (β → γ)} {g : α → Part β} lemma Monotone.partSeq (hf : Monotone f) (hg : Monotone g) : Monotone fun x ↦ f x <*> g x := by - simpa only [seq_eq_bind_map] using hf.partBind $ Monotone.of_apply₂ fun _ ↦ hg.partMap + simpa only [seq_eq_bind_map] using hf.partBind <| Monotone.of_apply₂ fun _ ↦ hg.partMap lemma Antitone.partSeq (hf : Antitone f) (hg : Antitone g) : Antitone fun x ↦ f x <*> g x := by - simpa only [seq_eq_bind_map] using hf.partBind $ Antitone.of_apply₂ fun _ ↦ hg.partMap + simpa only [seq_eq_bind_map] using hf.partBind <| Antitone.of_apply₂ fun _ ↦ hg.partMap end seq diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index f1b12c5c24bcf..95ce68bb9c968 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -184,7 +184,7 @@ theorem subsingleton_of_length_eq_zero (hs : s.length = 0) : {x | x ∈ s}.Subsi exact finCongr (by rw [hs, zero_add]) |>.injective <| Subsingleton.elim (α := Fin 1) _ _ theorem length_ne_zero_of_nontrivial (h : {x | x ∈ s}.Nontrivial) : s.length ≠ 0 := - fun hs ↦ h.not_subsingleton $ subsingleton_of_length_eq_zero hs + fun hs ↦ h.not_subsingleton <| subsingleton_of_length_eq_zero hs theorem length_pos_of_nontrivial (h : {x | x ∈ s}.Nontrivial) : 0 < s.length := Nat.pos_iff_ne_zero.mpr <| length_ne_zero_of_nontrivial h diff --git a/Mathlib/Order/SuccPred/Basic.lean b/Mathlib/Order/SuccPred/Basic.lean index 8f35e8e5a9135..3f18d2007775b 100644 --- a/Mathlib/Order/SuccPred/Basic.lean +++ b/Mathlib/Order/SuccPred/Basic.lean @@ -266,7 +266,7 @@ theorem succ_mono : Monotone (succ : α → α) := fun _ _ => succ_le_succ lemma le_succ_of_wcovBy (h : a ⩿ b) : b ≤ succ a := by obtain hab | ⟨-, hba⟩ := h.covBy_or_le_and_le · by_contra hba - exact h.2 (lt_succ_of_not_isMax hab.lt.not_isMax) $ hab.lt.succ_le.lt_of_not_le hba + exact h.2 (lt_succ_of_not_isMax hab.lt.not_isMax) <| hab.lt.succ_le.lt_of_not_le hba · exact hba.trans (le_succ _) alias _root_.WCovBy.le_succ := le_succ_of_wcovBy diff --git a/Mathlib/Order/SupClosed.lean b/Mathlib/Order/SupClosed.lean index 9008ccc271cbc..998929ca5eab3 100644 --- a/Mathlib/Order/SupClosed.lean +++ b/Mathlib/Order/SupClosed.lean @@ -315,8 +315,8 @@ protected lemma Set.Finite.supClosure (hs : s.Finite) : (supClosure s).Finite := @[simp] lemma supClosure_prod (s : Set α) (t : Set β) : supClosure (s ×ˢ t) = supClosure s ×ˢ supClosure t := - le_antisymm (supClosure_min (Set.prod_mono subset_supClosure subset_supClosure) $ - supClosed_supClosure.prod supClosed_supClosure) $ by + le_antisymm (supClosure_min (Set.prod_mono subset_supClosure subset_supClosure) <| + supClosed_supClosure.prod supClosed_supClosure) <| by rintro ⟨_, _⟩ ⟨⟨u, hu, hus, rfl⟩, v, hv, hvt, rfl⟩ refine ⟨u ×ˢ v, hu.product hv, ?_, ?_⟩ · simpa only [coe_product] using Set.prod_mono hus hvt @@ -387,8 +387,8 @@ protected lemma Set.Finite.infClosure (hs : s.Finite) : (infClosure s).Finite := @[simp] lemma infClosure_prod (s : Set α) (t : Set β) : infClosure (s ×ˢ t) = infClosure s ×ˢ infClosure t := - le_antisymm (infClosure_min (Set.prod_mono subset_infClosure subset_infClosure) $ - infClosed_infClosure.prod infClosed_infClosure) $ by + le_antisymm (infClosure_min (Set.prod_mono subset_infClosure subset_infClosure) <| + infClosed_infClosure.prod infClosed_infClosure) <| by rintro ⟨_, _⟩ ⟨⟨u, hu, hus, rfl⟩, v, hv, hvt, rfl⟩ refine ⟨u ×ˢ v, hu.product hv, ?_, ?_⟩ · simpa only [coe_product] using Set.prod_mono hus hvt diff --git a/Mathlib/Order/UpperLower/Basic.lean b/Mathlib/Order/UpperLower/Basic.lean index 3e55eda001a9c..efe287062c81e 100644 --- a/Mathlib/Order/UpperLower/Basic.lean +++ b/Mathlib/Order/UpperLower/Basic.lean @@ -536,7 +536,7 @@ instance completeLattice : CompleteLattice (UpperSet α) := (fun _ _ => rfl) (fun _ => rfl) (fun _ => rfl) rfl rfl instance completelyDistribLattice : CompletelyDistribLattice (UpperSet α) := - .ofMinimalAxioms $ + .ofMinimalAxioms <| (toDual.injective.comp SetLike.coe_injective).completelyDistribLatticeMinimalAxioms .of _ (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ => rfl) rfl rfl @@ -671,7 +671,7 @@ instance completeLattice : CompleteLattice (LowerSet α) := (fun _ => rfl) rfl rfl instance completelyDistribLattice : CompletelyDistribLattice (LowerSet α) := - .ofMinimalAxioms $ SetLike.coe_injective.completelyDistribLatticeMinimalAxioms .of _ + .ofMinimalAxioms <| SetLike.coe_injective.completelyDistribLatticeMinimalAxioms .of _ (fun _ _ => rfl) (fun _ _ => rfl) (fun _ => rfl) (fun _ => rfl) rfl rfl instance : Inhabited (LowerSet α) := diff --git a/Mathlib/RingTheory/Flat/Basic.lean b/Mathlib/RingTheory/Flat/Basic.lean index 80c9c48425e01..c495efa177414 100644 --- a/Mathlib/RingTheory/Flat/Basic.lean +++ b/Mathlib/RingTheory/Flat/Basic.lean @@ -281,7 +281,7 @@ lemma lTensor_exact [Small.{v} R] [flat : Flat R M] ⦃N N' N'' : Type v⦄ suffices exact1 : Function.Exact (f.lTensor M) (π.lTensor M) by rw [show g = ι.comp π by aesop, lTensor_comp] exact exact1.comp_injective - (inj := iff_lTensor_preserves_injective_linearMap R M |>.mp flat _ $ by + (inj := iff_lTensor_preserves_injective_linearMap R M |>.mp flat _ <| by simpa [ι] using Subtype.val_injective) (h0 := map_zero _) @@ -301,7 +301,7 @@ lemma rTensor_exact [Small.{v} R] [flat : Flat R M] ⦃N N' N'' : Type v⦄ suffices exact1 : Function.Exact (f.rTensor M) (π.rTensor M) by rw [show g = ι.comp π by aesop, rTensor_comp] exact exact1.comp_injective - (inj := iff_rTensor_preserves_injective_linearMap R M |>.mp flat _ $ by + (inj := iff_rTensor_preserves_injective_linearMap R M |>.mp flat _ <| by simpa [ι] using Subtype.val_injective) (h0 := map_zero _) diff --git a/Mathlib/RingTheory/Flat/CategoryTheory.lean b/Mathlib/RingTheory/Flat/CategoryTheory.lean index 3d88a691a2bac..ea7bc816d2f57 100644 --- a/Mathlib/RingTheory/Flat/CategoryTheory.lean +++ b/Mathlib/RingTheory/Flat/CategoryTheory.lean @@ -36,32 +36,32 @@ namespace Module.Flat variable {R : Type u} [CommRing R] (M : ModuleCat.{u} R) -lemma lTensor_shortComplex_exact [Flat R M] (C : ShortComplex $ ModuleCat R) (hC : C.Exact) : +lemma lTensor_shortComplex_exact [Flat R M] (C : ShortComplex <| ModuleCat R) (hC : C.Exact) : C.map (tensorLeft M) |>.Exact := by rw [moduleCat_exact_iff_function_exact] at hC ⊢ exact lTensor_exact M hC -lemma rTensor_shortComplex_exact [Flat R M] (C : ShortComplex $ ModuleCat R) (hC : C.Exact) : +lemma rTensor_shortComplex_exact [Flat R M] (C : ShortComplex <| ModuleCat R) (hC : C.Exact) : C.map (tensorRight M) |>.Exact := by rw [moduleCat_exact_iff_function_exact] at hC ⊢ exact rTensor_exact M hC lemma iff_lTensor_preserves_shortComplex_exact : Flat R M ↔ - ∀ (C : ShortComplex $ ModuleCat R) (_ : C.Exact), (C.map (tensorLeft M) |>.Exact) := - ⟨fun _ _ ↦ lTensor_shortComplex_exact _ _, fun H ↦ iff_lTensor_exact.2 $ + ∀ (C : ShortComplex <| ModuleCat R) (_ : C.Exact), (C.map (tensorLeft M) |>.Exact) := + ⟨fun _ _ ↦ lTensor_shortComplex_exact _ _, fun H ↦ iff_lTensor_exact.2 <| fun _ _ _ _ _ _ _ _ _ f g h ↦ - moduleCat_exact_iff_function_exact _ |>.1 $ + moduleCat_exact_iff_function_exact _ |>.1 <| H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) (DFunLike.ext _ _ h.apply_apply_eq_zero)) (moduleCat_exact_iff_function_exact _ |>.2 h)⟩ lemma iff_rTensor_preserves_shortComplex_exact : Flat R M ↔ - ∀ (C : ShortComplex $ ModuleCat R) (_ : C.Exact), (C.map (tensorRight M) |>.Exact) := - ⟨fun _ _ ↦ rTensor_shortComplex_exact _ _, fun H ↦ iff_rTensor_exact.2 $ + ∀ (C : ShortComplex <| ModuleCat R) (_ : C.Exact), (C.map (tensorRight M) |>.Exact) := + ⟨fun _ _ ↦ rTensor_shortComplex_exact _ _, fun H ↦ iff_rTensor_exact.2 <| fun _ _ _ _ _ _ _ _ _ f g h ↦ - moduleCat_exact_iff_function_exact _ |>.1 $ + moduleCat_exact_iff_function_exact _ |>.1 <| H (.mk (ModuleCat.ofHom f) (ModuleCat.ofHom g) (DFunLike.ext _ _ h.apply_apply_eq_zero)) (moduleCat_exact_iff_function_exact _ |>.2 h)⟩ diff --git a/Mathlib/RingTheory/HahnSeries/Summable.lean b/Mathlib/RingTheory/HahnSeries/Summable.lean index 33be95d2952fa..517632b481a87 100644 --- a/Mathlib/RingTheory/HahnSeries/Summable.lean +++ b/Mathlib/RingTheory/HahnSeries/Summable.lean @@ -505,7 +505,7 @@ instance instField [Field R] : Field (HahnSeries Γ R) where (single (-x.order)) (x.leadingCoeff)⁻¹ * (SummableFamily.powers _ (unit_aux x (inv_mul_cancel₀ (leadingCoeff_ne_iff.mpr x0)))).hsum inv_zero := dif_pos rfl - mul_inv_cancel x x0 := (congr rfl (dif_neg x0)).trans $ by + mul_inv_cancel x x0 := (congr rfl (dif_neg x0)).trans <| by have h := SummableFamily.one_sub_self_mul_hsum_powers (unit_aux x (inv_mul_cancel₀ (leadingCoeff_ne_iff.mpr x0))) diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 3a108417c37a7..a31ec62d4e7be 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -1093,7 +1093,7 @@ theorem subset_union_prime' {R : Type u} [CommRing R] {s : Finset ι} {f : ι · rw [Set.mem_iUnion₂] at ht rcases ht with ⟨j, hjt, hj⟩ simp only [Finset.inf_eq_iInf, SetLike.mem_coe, Submodule.mem_iInf] at hr - exact hs $ Or.inr $ Set.mem_biUnion hjt <| + exact hs <| Or.inr <| Set.mem_biUnion hjt <| add_sub_cancel_left r s ▸ (f j).sub_mem hj <| hr j hjt /-- Prime avoidance. Atiyah-Macdonald 1.11, Eisenbud 3.3, Stacks 00DS, Matsumura Ex.1.6. -/ diff --git a/Mathlib/RingTheory/Polynomial/Bernstein.lean b/Mathlib/RingTheory/Polynomial/Bernstein.lean index af0d3936dd937..d3efc2a2b5cb0 100644 --- a/Mathlib/RingTheory/Polynomial/Bernstein.lean +++ b/Mathlib/RingTheory/Polynomial/Bernstein.lean @@ -87,7 +87,7 @@ theorem eval_at_1 (n ν : ℕ) : (bernsteinPolynomial R n ν).eval 1 = if ν = n split_ifs with h · subst h; simp · obtain hνn | hnν := Ne.lt_or_lt h - · simp [zero_pow $ Nat.sub_ne_zero_of_lt hνn] + · simp [zero_pow <| Nat.sub_ne_zero_of_lt hνn] · simp [Nat.choose_eq_zero_of_lt hnν] theorem derivative_succ_aux (n ν : ℕ) : diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean index 7995fdbf7fc85..5f2544087a31a 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Expand.lean @@ -167,7 +167,7 @@ theorem isRoot_cyclotomic_prime_pow_mul_iff_of_charP {m k p : ℕ} {R : Type*} [ · rw [← isRoot_cyclotomic_iff, IsRoot.def] at h rw [cyclotomic_mul_prime_pow_eq R (NeZero.not_char_dvd R p m) hk, IsRoot.def, eval_pow, h, zero_pow] - exact Nat.sub_ne_zero_of_lt $ pow_right_strictMono hp.out.one_lt $ Nat.pred_lt hk.ne' + exact Nat.sub_ne_zero_of_lt <| pow_right_strictMono hp.out.one_lt <| Nat.pred_lt hk.ne' end CharP diff --git a/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean b/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean index 01229b99f3fe5..0358a0fe72767 100644 --- a/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean +++ b/Mathlib/RingTheory/TwoSidedIdeal/Basic.lean @@ -83,7 +83,7 @@ def orderIsoRingCon : TwoSidedIdeal R ≃o RingCon R where invFun := .mk left_inv _ := rfl right_inv _ := rfl - map_rel_iff' {I J} := Iff.symm $ le_iff.trans ⟨fun h x y r => by rw [rel_iff] at r ⊢; exact h r, + map_rel_iff' {I J} := Iff.symm <| le_iff.trans ⟨fun h x y r => by rw [rel_iff] at r ⊢; exact h r, fun h x hx => by rw [SetLike.mem_coe, mem_iff] at hx ⊢; exact h hx⟩ lemma ringCon_injective : Function.Injective (TwoSidedIdeal.ringCon (R := R)) := by diff --git a/Mathlib/RingTheory/TwoSidedIdeal/Lattice.lean b/Mathlib/RingTheory/TwoSidedIdeal/Lattice.lean index 9414b3c25d465..9d596c74f10e6 100644 --- a/Mathlib/RingTheory/TwoSidedIdeal/Lattice.lean +++ b/Mathlib/RingTheory/TwoSidedIdeal/Lattice.lean @@ -73,7 +73,7 @@ lemma mem_inf {I J : TwoSidedIdeal R} {x : R} : Iff.rfl instance : SupSet (TwoSidedIdeal R) where - sSup s := { ringCon := sSup $ TwoSidedIdeal.ringCon '' s } + sSup s := { ringCon := sSup <| TwoSidedIdeal.ringCon '' s } lemma sSup_ringCon (S : Set (TwoSidedIdeal R)) : (sSup S).ringCon = sSup (TwoSidedIdeal.ringCon '' S) := rfl @@ -83,11 +83,11 @@ lemma iSup_ringCon {ι : Type*} (I : ι → TwoSidedIdeal R) : simp only [iSup, sSup_ringCon]; congr; ext; simp instance : CompleteSemilatticeSup (TwoSidedIdeal R) where - sSup_le s I h := by simp_rw [ringCon_le_iff] at h ⊢; exact sSup_le $ by aesop - le_sSup s I hI := by rw [ringCon_le_iff]; exact le_sSup $ by aesop + sSup_le s I h := by simp_rw [ringCon_le_iff] at h ⊢; exact sSup_le <| by aesop + le_sSup s I hI := by rw [ringCon_le_iff]; exact le_sSup <| by aesop instance : InfSet (TwoSidedIdeal R) where - sInf s := { ringCon := sInf $ TwoSidedIdeal.ringCon '' s } + sInf s := { ringCon := sInf <| TwoSidedIdeal.ringCon '' s } lemma sInf_ringCon (S : Set (TwoSidedIdeal R)) : (sInf S).ringCon = sInf (TwoSidedIdeal.ringCon '' S) := rfl @@ -97,8 +97,8 @@ lemma iInf_ringCon {ι : Type*} (I : ι → TwoSidedIdeal R) : simp only [iInf, sInf_ringCon]; congr!; ext; simp instance : CompleteSemilatticeInf (TwoSidedIdeal R) where - le_sInf s I h := by simp_rw [ringCon_le_iff] at h ⊢; exact le_sInf $ by aesop - sInf_le s I hI := by rw [ringCon_le_iff]; exact sInf_le $ by aesop + le_sInf s I h := by simp_rw [ringCon_le_iff] at h ⊢; exact le_sInf <| by aesop + sInf_le s I hI := by rw [ringCon_le_iff]; exact sInf_le <| by aesop lemma mem_iInf {ι : Type*} {I : ι → TwoSidedIdeal R} {x : R} : x ∈ iInf I ↔ ∀ i, x ∈ I i := diff --git a/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean b/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean index 2dd43d8a231f6..7261a54d9e359 100644 --- a/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean +++ b/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean @@ -120,7 +120,7 @@ theorem exists_eq_pow_p_mul (a : 𝕎 k) (ha : a ≠ 0) : rw [← this] at hcm refine ⟨m, b, ?_, ?_⟩ · contrapose! hc - simp [hc, zero_pow $ pow_ne_zero _ hp.out.ne_zero] + simp [hc, zero_pow <| pow_ne_zero _ hp.out.ne_zero] · simp_rw [← mul_left_iterate (p : 𝕎 k) m] convert hcm using 2 ext1 x diff --git a/Mathlib/RingTheory/WittVector/InitTail.lean b/Mathlib/RingTheory/WittVector/InitTail.lean index c2c6e3cc9a3a5..ab53e4c4a1507 100644 --- a/Mathlib/RingTheory/WittVector/InitTail.lean +++ b/Mathlib/RingTheory/WittVector/InitTail.lean @@ -101,7 +101,7 @@ theorem select_add_select_not : ∀ x : 𝕎 R, select P x + select (fun i => ¬ refine fun m _ => mul_eq_mul_left_iff.mpr (Or.inl ?_) rw [ite_pow, zero_pow (pow_ne_zero _ hp.out.ne_zero)] by_cases Pm : P m - · rw [if_pos Pm, if_neg $ not_not_intro Pm, zero_pow Fin.size_pos'.ne', add_zero] + · rw [if_pos Pm, if_neg <| not_not_intro Pm, zero_pow Fin.size_pos'.ne', add_zero] · rwa [if_neg Pm, if_pos, zero_add] theorem coeff_add_of_disjoint (x y : 𝕎 R) (h : ∀ n, x.coeff n = 0 ∨ y.coeff n = 0) : diff --git a/Mathlib/SetTheory/Game/Basic.lean b/Mathlib/SetTheory/Game/Basic.lean index 2752abc297073..b8a4c1008298e 100644 --- a/Mathlib/SetTheory/Game/Basic.lean +++ b/Mathlib/SetTheory/Game/Basic.lean @@ -198,7 +198,7 @@ lemma bddAbove_range_of_small {ι : Type*} [Small.{u} ι] (f : ι → Game.{u}) BddAbove (Set.range f) := by obtain ⟨x, hx⟩ := PGame.bddAbove_range_of_small (Quotient.out ∘ f) refine ⟨⟦x⟧, Set.forall_mem_range.2 fun i ↦ ?_⟩ - simpa [PGame.le_iff_game_le] using hx $ Set.mem_range_self i + simpa [PGame.le_iff_game_le] using hx <| Set.mem_range_self i /-- A small set of games is bounded above. -/ lemma bddAbove_of_small (s : Set Game.{u}) [Small.{u} s] : BddAbove s := by @@ -209,7 +209,7 @@ lemma bddBelow_range_of_small {ι : Type*} [Small.{u} ι] (f : ι → Game.{u}) BddBelow (Set.range f) := by obtain ⟨x, hx⟩ := PGame.bddBelow_range_of_small (Quotient.out ∘ f) refine ⟨⟦x⟧, Set.forall_mem_range.2 fun i ↦ ?_⟩ - simpa [PGame.le_iff_game_le] using hx $ Set.mem_range_self i + simpa [PGame.le_iff_game_le] using hx <| Set.mem_range_self i /-- A small set of games is bounded below. -/ lemma bddBelow_of_small (s : Set Game.{u}) [Small.{u} s] : BddBelow s := by diff --git a/Mathlib/SetTheory/Game/PGame.lean b/Mathlib/SetTheory/Game/PGame.lean index c17af8b22e725..2f4a23900dd0c 100644 --- a/Mathlib/SetTheory/Game/PGame.lean +++ b/Mathlib/SetTheory/Game/PGame.lean @@ -638,7 +638,7 @@ theorem leftResponse_spec {x : PGame} (h : 0 ≤ x) (j : x.RightMoves) : /-- A small family of pre-games is bounded above. -/ lemma bddAbove_range_of_small {ι : Type*} [Small.{u} ι] (f : ι → PGame.{u}) : BddAbove (Set.range f) := by - let x : PGame.{u} := ⟨Σ i, (f $ (equivShrink.{u} ι).symm i).LeftMoves, PEmpty, + let x : PGame.{u} := ⟨Σ i, (f <| (equivShrink.{u} ι).symm i).LeftMoves, PEmpty, fun x ↦ moveLeft _ x.2, PEmpty.elim⟩ refine ⟨x, Set.forall_mem_range.2 fun i ↦ ?_⟩ rw [← (equivShrink ι).symm_apply_apply i, le_iff_forall_lf] @@ -651,7 +651,7 @@ lemma bddAbove_of_small (s : Set PGame.{u}) [Small.{u} s] : BddAbove s := by /-- A small family of pre-games is bounded below. -/ lemma bddBelow_range_of_small {ι : Type*} [Small.{u} ι] (f : ι → PGame.{u}) : BddBelow (Set.range f) := by - let x : PGame.{u} := ⟨PEmpty, Σ i, (f $ (equivShrink.{u} ι).symm i).RightMoves, PEmpty.elim, + let x : PGame.{u} := ⟨PEmpty, Σ i, (f <| (equivShrink.{u} ι).symm i).RightMoves, PEmpty.elim, fun x ↦ moveRight _ x.2⟩ refine ⟨x, Set.forall_mem_range.2 fun i ↦ ?_⟩ rw [← (equivShrink ι).symm_apply_apply i, le_iff_forall_lf] diff --git a/Mathlib/Tactic/Cases.lean b/Mathlib/Tactic/Cases.lean index b803e05583a44..9c917e8d070eb 100644 --- a/Mathlib/Tactic/Cases.lean +++ b/Mathlib/Tactic/Cases.lean @@ -57,7 +57,7 @@ def ElimApp.evalNames (elimInfo : ElimInfo) (alts : Array ElimApp.Alt) (withArg let (introduced, g) ← g.introNP generalized.size let subst := (generalized.zip introduced).foldl (init := subst) fun subst (a, b) => subst.insert a (.fvar b) - let g ← liftM $ toClear.foldlM (·.tryClear) g + let g ← liftM <| toClear.foldlM (·.tryClear) g g.withContext do for (stx, fvar) in toTag do Term.addLocalVarInfo stx (subst.get fvar) diff --git a/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean b/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean index c8bab48d68682..d5194c0971c73 100644 --- a/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean +++ b/Mathlib/Tactic/FunProp/RefinedDiscrTree.lean @@ -135,16 +135,16 @@ inductive Key where deriving Inhabited, BEq, Repr private nonrec def Key.hash : Key → UInt64 - | .star i => mixHash 7883 $ hash i + | .star i => mixHash 7883 <| hash i | .opaque => 342 - | .const n a => mixHash 5237 $ mixHash (hash n) (hash a) - | .fvar n a => mixHash 8765 $ mixHash (hash n) (hash a) - | .bvar i a => mixHash 4323 $ mixHash (hash i) (hash a) - | .lit v => mixHash 1879 $ hash v + | .const n a => mixHash 5237 <| mixHash (hash n) (hash a) + | .fvar n a => mixHash 8765 <| mixHash (hash n) (hash a) + | .bvar i a => mixHash 4323 <| mixHash (hash i) (hash a) + | .lit v => mixHash 1879 <| hash v | .sort => 2411 | .lam => 4742 | .«forall» => 9752 - | .proj s i a => mixHash (hash a) $ mixHash (hash s) (hash i) + | .proj s i a => mixHash (hash a) <| mixHash (hash s) (hash i) instance : Hashable Key := ⟨Key.hash⟩ @@ -242,7 +242,7 @@ def Trie.children! : Trie α → Array (Key × Trie α) | .values _ => panic! "did not expect .values constructor" private partial def Trie.format [ToFormat α] : Trie α → Format - | .node cs => Format.group $ Format.paren $ + | .node cs => Format.group <| Format.paren <| "node " ++ Format.join (cs.toList.map fun (k, c) => Format.line ++ Format.paren (format (prepend k c))) | .values vs => if vs.isEmpty then Format.nil else Std.format vs @@ -708,7 +708,7 @@ partial def mkDTExprAux (e : Expr) (root : Bool) : ReaderT Context MetaM DTExpr | _ => unreachable! -private abbrev M := StateListT (AssocList Expr DTExpr) $ ReaderT Context MetaM +private abbrev M := StateListT (AssocList Expr DTExpr) <| ReaderT Context MetaM /- Caching values is a bit dangerous, because when two expressions are be equal and they live under @@ -889,7 +889,7 @@ termination_by vs.size - i partial def insertInTrie [BEq α] (keys : Array Key) (v : α) (i : Nat) : Trie α → Trie α | .node cs => let k := keys[i]! - let c := Id.run $ cs.binInsertM + let c := Id.run <| cs.binInsertM (fun a b => a.1 < b.1) (fun (k', s) => (k', insertInTrie keys v (i+1) s)) (fun _ => (k, Trie.singleton keys v (i+1))) @@ -982,7 +982,7 @@ private structure State where mvarAssignments : Std.HashMap MVarId (Array Key) := {} -private abbrev M := ReaderT Context $ StateListM State +private abbrev M := ReaderT Context <| StateListM State /-- Return all values from `x` in an array, together with their scores. -/ private def M.run (unify : Bool) (config : WhnfCoreConfig) (x : M (Trie α)) : diff --git a/Mathlib/Tactic/Positivity/Basic.lean b/Mathlib/Tactic/Positivity/Basic.lean index f15f4b15daf14..14fd841e445d2 100644 --- a/Mathlib/Tactic/Positivity/Basic.lean +++ b/Mathlib/Tactic/Positivity/Basic.lean @@ -484,9 +484,9 @@ def evalRatNum : PositivityExt where eval {u α} _ _ e := do let pα : Q(PartialOrder ℚ) := q(inferInstance) assumeInstancesCommute match ← core zα pα a with - | .positive pa => pure $ .positive q(num_pos_of_pos $pa) - | .nonnegative pa => pure $ .nonnegative q(num_nonneg_of_nonneg $pa) - | .nonzero pa => pure $ .nonzero q(num_ne_zero_of_ne_zero $pa) + | .positive pa => pure <| .positive q(num_pos_of_pos $pa) + | .nonnegative pa => pure <| .nonnegative q(num_nonneg_of_nonneg $pa) + | .nonzero pa => pure <| .nonzero q(num_ne_zero_of_ne_zero $pa) | .none => pure .none | _, _ => throwError "not Rat.num" @@ -496,7 +496,7 @@ def evalRatDen : PositivityExt where eval {u α} _ _ e := do match u, α, e with | 0, ~q(ℕ), ~q(Rat.den $a) => assumeInstancesCommute - pure $ .positive q(den_pos $a) + pure <| .positive q(den_pos $a) | _, _ => throwError "not Rat.num" /-- Extension for `posPart`. `a⁺` is always nonegative, and positive if `a` is. -/ diff --git a/Mathlib/Tactic/Widget/InteractiveUnfold.lean b/Mathlib/Tactic/Widget/InteractiveUnfold.lean index 06ed44cc330f3..b88ea0416d05d 100644 --- a/Mathlib/Tactic/Widget/InteractiveUnfold.lean +++ b/Mathlib/Tactic/Widget/InteractiveUnfold.lean @@ -161,7 +161,7 @@ def renderUnfolds (e : Expr) (occ : Option Nat) (loc : Option Name) (range : Lsp #[ { Html.ofComponent MakeEditLink (.ofReplaceRange doc.meta range tactic) - #[.text $ Format.pretty $ (← Meta.ppExpr unfold)] } + #[.text <| Format.pretty <| (← Meta.ppExpr unfold)] } ] } return
@@ -221,7 +221,7 @@ Click on a suggestion to replace `unfold?` by a tactic that performs this rewrit elab stx:"unfold?" : tactic => do let some range := (← getFileMap).rangeOfStx? stx | return Widget.savePanelWidgetInfo (hash UnfoldComponent.javascript) - (pure $ json% { replaceRange : $range }) stx + (pure <| json% { replaceRange : $range }) stx /-- `#unfold? e` gives all unfolds of `e`. In tactic mode, use `unfold?` instead. -/ diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index d7d6a27b5b8c0..6c821ade241e9 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -1085,7 +1085,7 @@ theorem mem_closure_iff_nhdsWithin_neBot : x ∈ closure s ↔ NeBot (𝓝[s] x) lemma nhdsWithin_neBot : (𝓝[s] x).NeBot ↔ ∀ ⦃t⦄, t ∈ 𝓝 x → (t ∩ s).Nonempty := by rw [nhdsWithin, inf_neBot_iff] exact forall₂_congr fun U _ ↦ - ⟨fun h ↦ h (mem_principal_self _), fun h u hsu ↦ h.mono $ inter_subset_inter_right _ hsu⟩ + ⟨fun h ↦ h (mem_principal_self _), fun h u hsu ↦ h.mono <| inter_subset_inter_right _ hsu⟩ @[gcongr] theorem nhdsWithin_mono (x : X) {s t : Set X} (h : s ⊆ t) : 𝓝[s] x ≤ 𝓝[t] x := diff --git a/Mathlib/Topology/Inseparable.lean b/Mathlib/Topology/Inseparable.lean index d2089d77aed05..a58bfe2c76c2f 100644 --- a/Mathlib/Topology/Inseparable.lean +++ b/Mathlib/Topology/Inseparable.lean @@ -358,7 +358,7 @@ lemma specializingMap_iff_isClosed_image_closure_singleton (hf : Continuous f) : exact isClosed_closure lemma IsClosedMap.specializingMap (hf : IsClosedMap f) : SpecializingMap f := - specializingMap_iff_stableUnderSpecialization_image_singleton.mpr $ + specializingMap_iff_stableUnderSpecialization_image_singleton.mpr <| fun _ ↦ (hf _ isClosed_closure).stableUnderSpecialization lemma Inducing.specializingMap (hf : Inducing f) (h : StableUnderSpecialization (range f)) : diff --git a/Mathlib/Topology/Order/Bornology.lean b/Mathlib/Topology/Order/Bornology.lean index 0c7dfd80bd711..e2307a276340e 100644 --- a/Mathlib/Topology/Order/Bornology.lean +++ b/Mathlib/Topology/Order/Bornology.lean @@ -71,10 +71,10 @@ protected lemma BddAbove.isBounded (hs₀ : BddAbove s) (hs₁ : BddBelow s) : I isBounded_iff_bddBelow_bddAbove.2 ⟨hs₁, hs₀⟩ lemma BddBelow.isBounded_inter (hs : BddBelow s) (ht : BddAbove t) : IsBounded (s ∩ t) := - (hs.mono inter_subset_left).isBounded $ ht.mono inter_subset_right + (hs.mono inter_subset_left).isBounded <| ht.mono inter_subset_right lemma BddAbove.isBounded_inter (hs : BddAbove s) (ht : BddBelow t) : IsBounded (s ∩ t) := - (hs.mono inter_subset_left).isBounded $ ht.mono inter_subset_right + (hs.mono inter_subset_left).isBounded <| ht.mono inter_subset_right instance OrderDual.instIsOrderBornology : IsOrderBornology αᵒᵈ where isBounded_iff_bddBelow_bddAbove s := by diff --git a/Mathlib/Topology/Order/IntermediateValue.lean b/Mathlib/Topology/Order/IntermediateValue.lean index 00431baf1a323..06c7281cea152 100644 --- a/Mathlib/Topology/Order/IntermediateValue.lean +++ b/Mathlib/Topology/Order/IntermediateValue.lean @@ -641,7 +641,7 @@ theorem Continuous.strictMonoOn_of_inj_rigidity {f : α → δ} let t := max b y have hsa : s ≤ a := min_le_left a x have hbt : b ≤ t := le_max_left b y - have hst : s ≤ t := hsa.trans $ hbt.trans' hab.le + have hst : s ≤ t := hsa.trans <| hbt.trans' hab.le have hf_mono_st : StrictMonoOn f (Icc s t) ∨ StrictAntiOn f (Icc s t) := by letI := Icc.completeLinearOrder hst have := Continuous.strictMono_of_inj_boundedOrder' (f := Set.restrict (Icc s t) f) From fb7b6b7806b5be9b7a387c56c28f56008c59875d Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Sat, 17 Aug 2024 12:07:39 +0000 Subject: [PATCH 357/359] chore (Data.Multiset.Basic): avoid bundled ordered algebra for basic results on multisets (#14477) `Data.Multiset.Basic` is a fairly low-level file. Currently, it imports `Algebra.Order.Monoid.Defs` to establish that multisets on a type are an `StrictOrderedAddCancelCommMonoid` and uses that classes API in a dozen places. Unfortunately, this also adds, to this file and all that import it, the projections from the ordered monoids defined in `Algebra.Order.Monoid.Defs` as instances for Lean to try when synthesizing a `CommMonoid` impacting performance. This PR moves the ordered monoid instances to a new file `Data.Multiset.OrderedMonoid` and makes minor changes to `Basic` to avoid this import. Co-authored-by: Matthew Robert Ballard <100034030+mattrobball@users.noreply.github.com> --- Mathlib.lean | 4 +- .../Algebra/BigOperators/Group/Finset.lean | 10 +- Mathlib/Algebra/Order/Antidiag/Prod.lean | 1 + .../Order/BigOperators/Group/Finset.lean | 1 + Mathlib/Algebra/Order/Field/Pi.lean | 1 + Mathlib/Algebra/Order/Ring/Canonical.lean | 3 +- Mathlib/Algebra/Order/Sub/Basic.lean | 215 +++++++++++++++--- .../{Canonical.lean => Unbundled/Basic.lean} | 196 +--------------- Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean | 58 +++++ Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean | 2 +- Mathlib/Data/Finset/Basic.lean | 6 +- Mathlib/Data/Finset/Fold.lean | 5 +- Mathlib/Data/Finset/Lattice.lean | 12 +- Mathlib/Data/Finset/MulAntidiagonal.lean | 1 + Mathlib/Data/Finset/NatAntidiagonal.lean | 2 +- Mathlib/Data/Int/Interval.lean | 2 +- Mathlib/Data/Multiset/Antidiagonal.lean | 3 +- Mathlib/Data/Multiset/Basic.lean | 39 ++-- Mathlib/Data/Multiset/Nodup.lean | 2 +- Mathlib/Data/Multiset/OrderedMonoid.lean | 31 +++ Mathlib/Data/PNat/Factors.lean | 1 + Mathlib/Data/Set/MulAntidiagonal.lean | 9 +- .../GroupTheory/MonoidLocalization/Order.lean | 1 + Mathlib/RingTheory/HahnSeries/Basic.lean | 2 +- Mathlib/Tactic/Monotonicity/Lemmas.lean | 2 +- 25 files changed, 339 insertions(+), 270 deletions(-) rename Mathlib/Algebra/Order/Sub/{Canonical.lean => Unbundled/Basic.lean} (60%) create mode 100644 Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean create mode 100644 Mathlib/Data/Multiset/OrderedMonoid.lean diff --git a/Mathlib.lean b/Mathlib.lean index c288e312d7466..dfa5d02698df5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -645,9 +645,10 @@ import Mathlib.Algebra.Order.Ring.Unbundled.Basic import Mathlib.Algebra.Order.Ring.Unbundled.Nonneg import Mathlib.Algebra.Order.Ring.WithTop import Mathlib.Algebra.Order.Sub.Basic -import Mathlib.Algebra.Order.Sub.Canonical import Mathlib.Algebra.Order.Sub.Defs import Mathlib.Algebra.Order.Sub.Prod +import Mathlib.Algebra.Order.Sub.Unbundled.Basic +import Mathlib.Algebra.Order.Sub.Unbundled.Hom import Mathlib.Algebra.Order.Sub.WithTop import Mathlib.Algebra.Order.Sum import Mathlib.Algebra.Order.ToIntervalMod @@ -2290,6 +2291,7 @@ import Mathlib.Data.Multiset.Interval import Mathlib.Data.Multiset.Lattice import Mathlib.Data.Multiset.NatAntidiagonal import Mathlib.Data.Multiset.Nodup +import Mathlib.Data.Multiset.OrderedMonoid import Mathlib.Data.Multiset.Pi import Mathlib.Data.Multiset.Powerset import Mathlib.Data.Multiset.Range diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index 602f1900802b6..db06873e55a04 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -39,6 +39,7 @@ See the documentation of `to_additive.attr` for more information. -- assert_not_exists AddCommMonoidWithOne assert_not_exists MonoidWithZero assert_not_exists MulAction +assert_not_exists OrderedCommMonoid variable {ι κ α β γ : Type*} @@ -1455,11 +1456,12 @@ theorem eq_prod_range_div' {M : Type*} [CommGroup M] (f : ℕ → M) (n : ℕ) : reduces to the difference of the last and first terms when the function we are summing is monotone. -/ -theorem sum_range_tsub [CanonicallyOrderedAddCommMonoid α] [Sub α] [OrderedSub α] - [ContravariantClass α α (· + ·) (· ≤ ·)] {f : ℕ → α} (h : Monotone f) (n : ℕ) : +theorem sum_range_tsub [AddCommMonoid α] [PartialOrder α] [Sub α] [OrderedSub α] + [CovariantClass α α (· + ·) (· ≤ ·)] [ContravariantClass α α (· + ·) (· ≤ ·)] [ExistsAddOfLE α] + {f : ℕ → α} (h : Monotone f) (n : ℕ) : ∑ i ∈ range n, (f (i + 1) - f i) = f n - f 0 := by apply sum_range_induction - case base => apply tsub_self + case base => apply tsub_eq_of_eq_add; rw [zero_add] case step => intro n have h₁ : f n ≤ f (n + 1) := h (Nat.le_succ _) @@ -1684,7 +1686,7 @@ theorem prod_ite_one (s : Finset α) (p : α → Prop) [DecidablePred p] exact fun i hi => if_neg (h i hi) @[to_additive] -theorem prod_erase_lt_of_one_lt {γ : Type*} [DecidableEq α] [OrderedCommMonoid γ] +theorem prod_erase_lt_of_one_lt {γ : Type*} [DecidableEq α] [CommMonoid γ] [Preorder γ] [CovariantClass γ γ (· * ·) (· < ·)] {s : Finset α} {d : α} (hd : d ∈ s) {f : α → γ} (hdf : 1 < f d) : ∏ m ∈ s.erase d, f m < ∏ m ∈ s, f m := by conv in ∏ m ∈ s, f m => rw [← Finset.insert_erase hd] diff --git a/Mathlib/Algebra/Order/Antidiag/Prod.lean b/Mathlib/Algebra/Order/Antidiag/Prod.lean index dde92767a53bf..6f49ffaa10635 100644 --- a/Mathlib/Algebra/Order/Antidiag/Prod.lean +++ b/Mathlib/Algebra/Order/Antidiag/Prod.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Antoine Chambert-Loir and María Inés de Frutos-Fernández. Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir, María Inés de Frutos-Fernández, Bhavik Mehta, Eric Wieser -/ +import Mathlib.Algebra.Order.Monoid.Canonical.Defs import Mathlib.Algebra.Order.Sub.Defs import Mathlib.Data.Finset.Basic import Mathlib.Order.Interval.Finset.Defs diff --git a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean index ce3eb2fbc4a57..93535400087e6 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/Finset.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Order.BigOperators.Group.Multiset +import Mathlib.Data.Multiset.OrderedMonoid import Mathlib.Tactic.Bound.Attribute import Mathlib.Tactic.NormNum.Basic import Mathlib.Tactic.Positivity.Core diff --git a/Mathlib/Algebra/Order/Field/Pi.lean b/Mathlib/Algebra/Order/Field/Pi.lean index c0765788d3056..7100886df9c3c 100644 --- a/Mathlib/Algebra/Order/Field/Pi.lean +++ b/Mathlib/Algebra/Order/Field/Pi.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import Mathlib.Algebra.Order.Monoid.Defs import Mathlib.Data.Finset.Lattice import Mathlib.Data.Fintype.Card diff --git a/Mathlib/Algebra/Order/Ring/Canonical.lean b/Mathlib/Algebra/Order/Ring/Canonical.lean index 432f79a5c6202..1f64f689d1e79 100644 --- a/Mathlib/Algebra/Order/Ring/Canonical.lean +++ b/Mathlib/Algebra/Order/Ring/Canonical.lean @@ -3,8 +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 -/ +import Mathlib.Algebra.Order.Monoid.Canonical.Defs import Mathlib.Algebra.Order.Ring.Defs -import Mathlib.Algebra.Order.Sub.Canonical +import Mathlib.Algebra.Order.Sub.Basic import Mathlib.Algebra.Ring.Parity /-! diff --git a/Mathlib/Algebra/Order/Sub/Basic.lean b/Mathlib/Algebra/Order/Sub/Basic.lean index dbd995c294236..79046c54164c0 100644 --- a/Mathlib/Algebra/Order/Sub/Basic.lean +++ b/Mathlib/Algebra/Order/Sub/Basic.lean @@ -3,58 +3,199 @@ Copyright (c) 2021 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.Monoid.Unbundled.ExistsOfLE +import Mathlib.Algebra.Order.Monoid.Canonical.Defs +import Mathlib.Algebra.Order.Sub.Unbundled.Basic import Mathlib.Algebra.Group.Equiv.Basic -import Mathlib.Algebra.Ring.Basic -import Mathlib.Algebra.Order.Sub.Defs -import Mathlib.Order.Hom.Basic - +import Mathlib.Algebra.Group.Even /-! -# Additional results about ordered Subtraction - +# Lemmas about subtraction in unbundled canonically ordered monoids -/ +variable {α : Type*} + +section CanonicallyOrderedAddCommMonoid + +variable [CanonicallyOrderedAddCommMonoid α] [Sub α] [OrderedSub α] {a b c d : α} + +theorem add_tsub_cancel_iff_le : a + (b - a) = b ↔ a ≤ b := + ⟨fun h => le_iff_exists_add.mpr ⟨b - a, h.symm⟩, add_tsub_cancel_of_le⟩ + +theorem tsub_add_cancel_iff_le : b - a + a = b ↔ a ≤ b := by + rw [add_comm] + exact add_tsub_cancel_iff_le + +-- This was previously a `@[simp]` lemma, but it is not necessarily a good idea, e.g. in +-- `example (h : n - m = 0) : a + (n - m) = a := by simp_all` +theorem tsub_eq_zero_iff_le : a - b = 0 ↔ a ≤ b := by + rw [← nonpos_iff_eq_zero, tsub_le_iff_left, add_zero] + +alias ⟨_, tsub_eq_zero_of_le⟩ := tsub_eq_zero_iff_le + +attribute [simp] tsub_eq_zero_of_le + +theorem tsub_self (a : α) : a - a = 0 := + tsub_eq_zero_of_le le_rfl + +theorem tsub_le_self : a - b ≤ a := + tsub_le_iff_left.mpr <| le_add_left le_rfl + +theorem zero_tsub (a : α) : 0 - a = 0 := + tsub_eq_zero_of_le <| zero_le a + +theorem tsub_self_add (a b : α) : a - (a + b) = 0 := + tsub_eq_zero_of_le <| self_le_add_right _ _ + +theorem tsub_pos_iff_not_le : 0 < a - b ↔ ¬a ≤ b := by + rw [pos_iff_ne_zero, Ne, tsub_eq_zero_iff_le] + +theorem tsub_pos_of_lt (h : a < b) : 0 < b - a := + tsub_pos_iff_not_le.mpr h.not_le + +theorem tsub_lt_of_lt (h : a < b) : a - c < b := + lt_of_le_of_lt tsub_le_self h + +namespace AddLECancellable + +protected theorem tsub_le_tsub_iff_left (ha : AddLECancellable a) (hc : AddLECancellable c) + (h : c ≤ a) : a - b ≤ a - c ↔ c ≤ b := by + refine ⟨?_, fun h => tsub_le_tsub_left h a⟩ + rw [tsub_le_iff_left, ← hc.add_tsub_assoc_of_le h, hc.le_tsub_iff_right (h.trans le_add_self), + add_comm b] + apply ha + +protected theorem tsub_right_inj (ha : AddLECancellable a) (hb : AddLECancellable b) + (hc : AddLECancellable c) (hba : b ≤ a) (hca : c ≤ a) : a - b = a - c ↔ b = c := by + simp_rw [le_antisymm_iff, ha.tsub_le_tsub_iff_left hb hba, ha.tsub_le_tsub_iff_left hc hca, + and_comm] + +end AddLECancellable + +/-! #### Lemmas where addition is order-reflecting. -/ + + +section Contra + +variable [ContravariantClass α α (· + ·) (· ≤ ·)] + +theorem tsub_le_tsub_iff_left (h : c ≤ a) : a - b ≤ a - c ↔ c ≤ b := + Contravariant.AddLECancellable.tsub_le_tsub_iff_left Contravariant.AddLECancellable h + +theorem tsub_right_inj (hba : b ≤ a) (hca : c ≤ a) : a - b = a - c ↔ b = c := + Contravariant.AddLECancellable.tsub_right_inj Contravariant.AddLECancellable + Contravariant.AddLECancellable hba hca + +variable (α) + +/-- A `CanonicallyOrderedAddCommMonoid` with ordered subtraction and order-reflecting addition is +cancellative. This is not an instance as it would form a typeclass loop. + +See note [reducible non-instances]. -/ +abbrev CanonicallyOrderedAddCommMonoid.toAddCancelCommMonoid : AddCancelCommMonoid α := + { (by infer_instance : AddCommMonoid α) with + add_left_cancel := fun a b c h => by + simpa only [add_tsub_cancel_left] using congr_arg (fun x => x - a) h } + +end Contra + +end CanonicallyOrderedAddCommMonoid + +/-! ### Lemmas in a linearly canonically ordered monoid. -/ + + +section CanonicallyLinearOrderedAddCommMonoid + +variable [CanonicallyLinearOrderedAddCommMonoid α] [Sub α] [OrderedSub α] {a b c d : α} + +@[simp] +theorem tsub_pos_iff_lt : 0 < a - b ↔ b < a := by rw [tsub_pos_iff_not_le, not_le] + +theorem tsub_eq_tsub_min (a b : α) : a - b = a - min a b := by + rcases le_total a b with h | h + · rw [min_eq_left h, tsub_self, tsub_eq_zero_of_le h] + · rw [min_eq_right h] + +namespace AddLECancellable + +protected theorem lt_tsub_iff_right (hc : AddLECancellable c) : a < b - c ↔ a + c < b := + ⟨lt_imp_lt_of_le_imp_le tsub_le_iff_right.mpr, hc.lt_tsub_of_add_lt_right⟩ + +protected theorem lt_tsub_iff_left (hc : AddLECancellable c) : a < b - c ↔ c + a < b := + ⟨lt_imp_lt_of_le_imp_le tsub_le_iff_left.mpr, hc.lt_tsub_of_add_lt_left⟩ + +protected theorem tsub_lt_tsub_iff_right (hc : AddLECancellable c) (h : c ≤ a) : + a - c < b - c ↔ a < b := by rw [hc.lt_tsub_iff_left, add_tsub_cancel_of_le h] + +protected theorem tsub_lt_self (ha : AddLECancellable a) (h₁ : 0 < a) (h₂ : 0 < b) : a - b < a := by + refine tsub_le_self.lt_of_ne fun h => ?_ + rw [← h, tsub_pos_iff_lt] at h₁ + exact h₂.not_le (ha.add_le_iff_nonpos_left.1 <| add_le_of_le_tsub_left_of_le h₁.le h.ge) + +protected theorem tsub_lt_self_iff (ha : AddLECancellable a) : a - b < a ↔ 0 < a ∧ 0 < b := by + refine + ⟨fun h => ⟨(zero_le _).trans_lt h, (zero_le b).lt_of_ne ?_⟩, fun h => ha.tsub_lt_self h.1 h.2⟩ + rintro rfl + rw [tsub_zero] at h + exact h.false + +/-- See `lt_tsub_iff_left_of_le_of_le` for a weaker statement in a partial order. -/ +protected theorem tsub_lt_tsub_iff_left_of_le (ha : AddLECancellable a) (hb : AddLECancellable b) + (h : b ≤ a) : a - b < a - c ↔ c < b := + lt_iff_lt_of_le_iff_le <| ha.tsub_le_tsub_iff_left hb h + +end AddLECancellable + +section Contra + +variable [ContravariantClass α α (· + ·) (· ≤ ·)] -variable {α β : Type*} +/-- This lemma also holds for `ENNReal`, but we need a different proof for that. -/ +theorem tsub_lt_tsub_iff_right (h : c ≤ a) : a - c < b - c ↔ a < b := + Contravariant.AddLECancellable.tsub_lt_tsub_iff_right h -section Add +theorem tsub_lt_self : 0 < a → 0 < b → a - b < a := + Contravariant.AddLECancellable.tsub_lt_self -variable [Preorder α] [Add α] [Sub α] [OrderedSub α] {a b c d : α} +@[simp] theorem tsub_lt_self_iff : a - b < a ↔ 0 < a ∧ 0 < b := + Contravariant.AddLECancellable.tsub_lt_self_iff -theorem AddHom.le_map_tsub [Preorder β] [Add β] [Sub β] [OrderedSub β] (f : AddHom α β) - (hf : Monotone f) (a b : α) : f a - f b ≤ f (a - b) := by - rw [tsub_le_iff_right, ← f.map_add] - exact hf le_tsub_add +/-- See `lt_tsub_iff_left_of_le_of_le` for a weaker statement in a partial order. -/ +theorem tsub_lt_tsub_iff_left_of_le (h : b ≤ a) : a - b < a - c ↔ c < b := + Contravariant.AddLECancellable.tsub_lt_tsub_iff_left_of_le Contravariant.AddLECancellable h -theorem le_mul_tsub {R : Type*} [Distrib R] [Preorder R] [Sub R] [OrderedSub R] - [CovariantClass R R (· * ·) (· ≤ ·)] {a b c : R} : a * b - a * c ≤ a * (b - c) := - (AddHom.mulLeft a).le_map_tsub (monotone_id.const_mul' a) _ _ +lemma tsub_tsub_eq_min (a b : α) : a - (a - b) = min a b := by + rw [tsub_eq_tsub_min _ b, tsub_tsub_cancel_of_le (min_le_left a _)] -theorem le_tsub_mul {R : Type*} [CommSemiring R] [Preorder R] [Sub R] [OrderedSub R] - [CovariantClass R R (· * ·) (· ≤ ·)] {a b c : R} : a * c - b * c ≤ (a - b) * c := by - simpa only [mul_comm _ c] using le_mul_tsub +end Contra -end Add +/-! ### Lemmas about `max` and `min`. -/ -/-- An order isomorphism between types with ordered subtraction preserves subtraction provided that -it preserves addition. -/ -theorem OrderIso.map_tsub {M N : Type*} [Preorder M] [Add M] [Sub M] [OrderedSub M] - [PartialOrder N] [Add N] [Sub N] [OrderedSub N] (e : M ≃o N) - (h_add : ∀ a b, e (a + b) = e a + e b) (a b : M) : e (a - b) = e a - e b := by - let e_add : M ≃+ N := { e with map_add' := h_add } - refine le_antisymm ?_ (e_add.toAddHom.le_map_tsub e.monotone a b) - suffices e (e.symm (e a) - e.symm (e b)) ≤ e (e.symm (e a - e b)) by simpa - exact e.monotone (e_add.symm.toAddHom.le_map_tsub e.symm.monotone _ _) -/-! ### Preorder -/ +theorem tsub_add_eq_max : a - b + b = max a b := by + rcases le_total a b with h | h + · rw [max_eq_right h, tsub_eq_zero_of_le h, zero_add] + · rw [max_eq_left h, tsub_add_cancel_of_le h] +theorem add_tsub_eq_max : a + (b - a) = max a b := by rw [add_comm, max_comm, tsub_add_eq_max] -section Preorder +theorem tsub_min : a - min a b = a - b := by + rcases le_total a b with h | h + · rw [min_eq_left h, tsub_self, tsub_eq_zero_of_le h] + · rw [min_eq_right h] -variable [Preorder α] -variable [AddCommMonoid α] [Sub α] [OrderedSub α] {a b c d : α} +theorem tsub_add_min : a - b + min a b = a := by + rw [← tsub_min, @tsub_add_cancel_of_le] + apply min_le_left -theorem AddMonoidHom.le_map_tsub [Preorder β] [AddCommMonoid β] [Sub β] [OrderedSub β] (f : α →+ β) - (hf : Monotone f) (a b : α) : f a - f b ≤ f (a - b) := - f.toAddHom.le_map_tsub hf a b +-- `Odd.tsub` requires `CanonicallyLinearOrderedSemiring`, which we don't have +lemma Even.tsub + [ContravariantClass α α (· + ·) (· ≤ ·)] {m n : α} (hm : Even m) (hn : Even n) : + Even (m - n) := by + obtain ⟨a, rfl⟩ := hm + obtain ⟨b, rfl⟩ := hn + refine ⟨a - b, ?_⟩ + obtain h | h := le_total a b + · rw [tsub_eq_zero_of_le h, tsub_eq_zero_of_le (add_le_add h h), add_zero] + · exact (tsub_add_tsub_comm h h).symm -end Preorder +end CanonicallyLinearOrderedAddCommMonoid diff --git a/Mathlib/Algebra/Order/Sub/Canonical.lean b/Mathlib/Algebra/Order/Sub/Unbundled/Basic.lean similarity index 60% rename from Mathlib/Algebra/Order/Sub/Canonical.lean rename to Mathlib/Algebra/Order/Sub/Unbundled/Basic.lean index f8591c56f68d8..979e4903eb48f 100644 --- a/Mathlib/Algebra/Order/Sub/Canonical.lean +++ b/Mathlib/Algebra/Order/Sub/Unbundled/Basic.lean @@ -3,14 +3,15 @@ Copyright (c) 2021 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.Group.Even -import Mathlib.Algebra.Order.Monoid.Canonical.Defs import Mathlib.Algebra.Order.Sub.Defs +import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE /-! -# Lemmas about subtraction in canonically ordered monoids +# Lemmas about subtraction in an unbundled canonically ordered monoids -/ +-- These are about *unbundled* canonically ordered monoids +assert_not_exists OrderedCommMonoid variable {α : Type*} @@ -255,192 +256,3 @@ theorem tsub_tsub_eq_add_tsub_of_le end Contra end ExistsAddOfLE - -/-! ### Lemmas in a canonically ordered monoid. -/ - - -section CanonicallyOrderedAddCommMonoid - -variable [CanonicallyOrderedAddCommMonoid α] [Sub α] [OrderedSub α] {a b c d : α} - -theorem add_tsub_cancel_iff_le : a + (b - a) = b ↔ a ≤ b := - ⟨fun h => le_iff_exists_add.mpr ⟨b - a, h.symm⟩, add_tsub_cancel_of_le⟩ - -theorem tsub_add_cancel_iff_le : b - a + a = b ↔ a ≤ b := by - rw [add_comm] - exact add_tsub_cancel_iff_le - --- This was previously a `@[simp]` lemma, but it is not necessarily a good idea, e.g. in --- `example (h : n - m = 0) : a + (n - m) = a := by simp_all` -theorem tsub_eq_zero_iff_le : a - b = 0 ↔ a ≤ b := by - rw [← nonpos_iff_eq_zero, tsub_le_iff_left, add_zero] - -alias ⟨_, tsub_eq_zero_of_le⟩ := tsub_eq_zero_iff_le - -attribute [simp] tsub_eq_zero_of_le - -theorem tsub_self (a : α) : a - a = 0 := - tsub_eq_zero_of_le le_rfl - -theorem tsub_le_self : a - b ≤ a := - tsub_le_iff_left.mpr <| le_add_left le_rfl - -theorem zero_tsub (a : α) : 0 - a = 0 := - tsub_eq_zero_of_le <| zero_le a - -theorem tsub_self_add (a b : α) : a - (a + b) = 0 := - tsub_eq_zero_of_le <| self_le_add_right _ _ - -theorem tsub_pos_iff_not_le : 0 < a - b ↔ ¬a ≤ b := by - rw [pos_iff_ne_zero, Ne, tsub_eq_zero_iff_le] - -theorem tsub_pos_of_lt (h : a < b) : 0 < b - a := - tsub_pos_iff_not_le.mpr h.not_le - -theorem tsub_lt_of_lt (h : a < b) : a - c < b := - lt_of_le_of_lt tsub_le_self h - -namespace AddLECancellable - -protected theorem tsub_le_tsub_iff_left (ha : AddLECancellable a) (hc : AddLECancellable c) - (h : c ≤ a) : a - b ≤ a - c ↔ c ≤ b := by - refine ⟨?_, fun h => tsub_le_tsub_left h a⟩ - rw [tsub_le_iff_left, ← hc.add_tsub_assoc_of_le h, hc.le_tsub_iff_right (h.trans le_add_self), - add_comm b] - apply ha - -protected theorem tsub_right_inj (ha : AddLECancellable a) (hb : AddLECancellable b) - (hc : AddLECancellable c) (hba : b ≤ a) (hca : c ≤ a) : a - b = a - c ↔ b = c := by - simp_rw [le_antisymm_iff, ha.tsub_le_tsub_iff_left hb hba, ha.tsub_le_tsub_iff_left hc hca, - and_comm] - -end AddLECancellable - -/-! #### Lemmas where addition is order-reflecting. -/ - - -section Contra - -variable [ContravariantClass α α (· + ·) (· ≤ ·)] - -theorem tsub_le_tsub_iff_left (h : c ≤ a) : a - b ≤ a - c ↔ c ≤ b := - Contravariant.AddLECancellable.tsub_le_tsub_iff_left Contravariant.AddLECancellable h - -theorem tsub_right_inj (hba : b ≤ a) (hca : c ≤ a) : a - b = a - c ↔ b = c := - Contravariant.AddLECancellable.tsub_right_inj Contravariant.AddLECancellable - Contravariant.AddLECancellable hba hca - -variable (α) - -/-- A `CanonicallyOrderedAddCommMonoid` with ordered subtraction and order-reflecting addition is -cancellative. This is not an instance as it would form a typeclass loop. - -See note [reducible non-instances]. -/ -abbrev CanonicallyOrderedAddCommMonoid.toAddCancelCommMonoid : AddCancelCommMonoid α := - { (by infer_instance : AddCommMonoid α) with - add_left_cancel := fun a b c h => by - simpa only [add_tsub_cancel_left] using congr_arg (fun x => x - a) h } - -end Contra - -end CanonicallyOrderedAddCommMonoid - -/-! ### Lemmas in a linearly canonically ordered monoid. -/ - - -section CanonicallyLinearOrderedAddCommMonoid - -variable [CanonicallyLinearOrderedAddCommMonoid α] [Sub α] [OrderedSub α] {a b c d : α} - -@[simp] -theorem tsub_pos_iff_lt : 0 < a - b ↔ b < a := by rw [tsub_pos_iff_not_le, not_le] - -theorem tsub_eq_tsub_min (a b : α) : a - b = a - min a b := by - rcases le_total a b with h | h - · rw [min_eq_left h, tsub_self, tsub_eq_zero_of_le h] - · rw [min_eq_right h] - -namespace AddLECancellable - -protected theorem lt_tsub_iff_right (hc : AddLECancellable c) : a < b - c ↔ a + c < b := - ⟨lt_imp_lt_of_le_imp_le tsub_le_iff_right.mpr, hc.lt_tsub_of_add_lt_right⟩ - -protected theorem lt_tsub_iff_left (hc : AddLECancellable c) : a < b - c ↔ c + a < b := - ⟨lt_imp_lt_of_le_imp_le tsub_le_iff_left.mpr, hc.lt_tsub_of_add_lt_left⟩ - -protected theorem tsub_lt_tsub_iff_right (hc : AddLECancellable c) (h : c ≤ a) : - a - c < b - c ↔ a < b := by rw [hc.lt_tsub_iff_left, add_tsub_cancel_of_le h] - -protected theorem tsub_lt_self (ha : AddLECancellable a) (h₁ : 0 < a) (h₂ : 0 < b) : a - b < a := by - refine tsub_le_self.lt_of_ne fun h => ?_ - rw [← h, tsub_pos_iff_lt] at h₁ - exact h₂.not_le (ha.add_le_iff_nonpos_left.1 <| add_le_of_le_tsub_left_of_le h₁.le h.ge) - -protected theorem tsub_lt_self_iff (ha : AddLECancellable a) : a - b < a ↔ 0 < a ∧ 0 < b := by - refine - ⟨fun h => ⟨(zero_le _).trans_lt h, (zero_le b).lt_of_ne ?_⟩, fun h => ha.tsub_lt_self h.1 h.2⟩ - rintro rfl - rw [tsub_zero] at h - exact h.false - -/-- See `lt_tsub_iff_left_of_le_of_le` for a weaker statement in a partial order. -/ -protected theorem tsub_lt_tsub_iff_left_of_le (ha : AddLECancellable a) (hb : AddLECancellable b) - (h : b ≤ a) : a - b < a - c ↔ c < b := - lt_iff_lt_of_le_iff_le <| ha.tsub_le_tsub_iff_left hb h - -end AddLECancellable - -section Contra - -variable [ContravariantClass α α (· + ·) (· ≤ ·)] - -/-- This lemma also holds for `ENNReal`, but we need a different proof for that. -/ -theorem tsub_lt_tsub_iff_right (h : c ≤ a) : a - c < b - c ↔ a < b := - Contravariant.AddLECancellable.tsub_lt_tsub_iff_right h - -theorem tsub_lt_self : 0 < a → 0 < b → a - b < a := - Contravariant.AddLECancellable.tsub_lt_self - -@[simp] theorem tsub_lt_self_iff : a - b < a ↔ 0 < a ∧ 0 < b := - Contravariant.AddLECancellable.tsub_lt_self_iff - -/-- See `lt_tsub_iff_left_of_le_of_le` for a weaker statement in a partial order. -/ -theorem tsub_lt_tsub_iff_left_of_le (h : b ≤ a) : a - b < a - c ↔ c < b := - Contravariant.AddLECancellable.tsub_lt_tsub_iff_left_of_le Contravariant.AddLECancellable h - -lemma tsub_tsub_eq_min (a b : α) : a - (a - b) = min a b := by - rw [tsub_eq_tsub_min _ b, tsub_tsub_cancel_of_le (min_le_left a _)] - -end Contra - -/-! ### Lemmas about `max` and `min`. -/ - - -theorem tsub_add_eq_max : a - b + b = max a b := by - rcases le_total a b with h | h - · rw [max_eq_right h, tsub_eq_zero_of_le h, zero_add] - · rw [max_eq_left h, tsub_add_cancel_of_le h] - -theorem add_tsub_eq_max : a + (b - a) = max a b := by rw [add_comm, max_comm, tsub_add_eq_max] - -theorem tsub_min : a - min a b = a - b := by - rcases le_total a b with h | h - · rw [min_eq_left h, tsub_self, tsub_eq_zero_of_le h] - · rw [min_eq_right h] - -theorem tsub_add_min : a - b + min a b = a := by - rw [← tsub_min, @tsub_add_cancel_of_le] - apply min_le_left - --- `Odd.tsub` requires `CanonicallyLinearOrderedSemiring`, which we don't have -lemma Even.tsub - [ContravariantClass α α (· + ·) (· ≤ ·)] {m n : α} (hm : Even m) (hn : Even n) : - Even (m - n) := by - obtain ⟨a, rfl⟩ := hm - obtain ⟨b, rfl⟩ := hn - refine ⟨a - b, ?_⟩ - obtain h | h := le_total a b - · rw [tsub_eq_zero_of_le h, tsub_eq_zero_of_le (add_le_add h h), add_zero] - · exact (tsub_add_tsub_comm h h).symm - -end CanonicallyLinearOrderedAddCommMonoid diff --git a/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean b/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean new file mode 100644 index 0000000000000..2bf9585fa56a6 --- /dev/null +++ b/Mathlib/Algebra/Order/Sub/Unbundled/Hom.lean @@ -0,0 +1,58 @@ +/- +Copyright (c) 2021 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.Sub.Defs +import Mathlib.Algebra.Group.Equiv.Basic +import Mathlib.Algebra.Ring.Basic +import Mathlib.Order.Hom.Basic +/-! +# Lemmas about subtraction in unbundled canonically ordered monoids +-/ + + +variable {α β : Type*} + +section Add + +variable [Preorder α] [Add α] [Sub α] [OrderedSub α] {a b c d : α} + +theorem AddHom.le_map_tsub [Preorder β] [Add β] [Sub β] [OrderedSub β] (f : AddHom α β) + (hf : Monotone f) (a b : α) : f a - f b ≤ f (a - b) := by + rw [tsub_le_iff_right, ← f.map_add] + exact hf le_tsub_add + +theorem le_mul_tsub {R : Type*} [Distrib R] [Preorder R] [Sub R] [OrderedSub R] + [CovariantClass R R (· * ·) (· ≤ ·)] {a b c : R} : a * b - a * c ≤ a * (b - c) := + (AddHom.mulLeft a).le_map_tsub (monotone_id.const_mul' a) _ _ + +theorem le_tsub_mul {R : Type*} [CommSemiring R] [Preorder R] [Sub R] [OrderedSub R] + [CovariantClass R R (· * ·) (· ≤ ·)] {a b c : R} : a * c - b * c ≤ (a - b) * c := by + simpa only [mul_comm _ c] using le_mul_tsub + +end Add + +/-- An order isomorphism between types with ordered subtraction preserves subtraction provided that +it preserves addition. -/ +theorem OrderIso.map_tsub {M N : Type*} [Preorder M] [Add M] [Sub M] [OrderedSub M] + [PartialOrder N] [Add N] [Sub N] [OrderedSub N] (e : M ≃o N) + (h_add : ∀ a b, e (a + b) = e a + e b) (a b : M) : e (a - b) = e a - e b := by + let e_add : M ≃+ N := { e with map_add' := h_add } + refine le_antisymm ?_ (e_add.toAddHom.le_map_tsub e.monotone a b) + suffices e (e.symm (e a) - e.symm (e b)) ≤ e (e.symm (e a - e b)) by simpa + exact e.monotone (e_add.symm.toAddHom.le_map_tsub e.symm.monotone _ _) + +/-! ### Preorder -/ + + +section Preorder + +variable [Preorder α] +variable [AddCommMonoid α] [Sub α] [OrderedSub α] {a b c d : α} + +theorem AddMonoidHom.le_map_tsub [Preorder β] [AddCommMonoid β] [Sub β] [OrderedSub β] (f : α →+ β) + (hf : Monotone f) (a b : α) : f a - f b ≤ f (a - b) := + f.toAddHom.le_map_tsub hf a b + +end Preorder diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index a2b9e02f3bc43..4cef85d6ad0d3 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -120,7 +120,7 @@ theorem antidiagonalTuple_zero_right : ∀ k, antidiagonalTuple k 0 = [0] @[simp] theorem antidiagonalTuple_one (n : ℕ) : antidiagonalTuple 1 n = [![n]] := by simp_rw [antidiagonalTuple, antidiagonal, List.range_succ, List.map_append, List.map_singleton, - tsub_self, List.bind_append, List.bind_singleton, List.bind_map] + Nat.sub_self, List.bind_append, List.bind_singleton, List.bind_map] conv_rhs => rw [← List.nil_append [![n]]] congr 1 simp_rw [List.bind_eq_nil, List.mem_range, List.map_eq_nil] diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index bd963644daca7..e421e88c289c7 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -122,6 +122,8 @@ assert_not_exists Multiset.powerset assert_not_exists CompleteLattice +assert_not_exists OrderedCommMonoid + open Multiset Subtype Nat Function universe u @@ -1745,7 +1747,7 @@ variable [DecidableEq α] {s t u v : Finset α} {a b : α} /-- `s \ t` is the set consisting of the elements of `s` that are not in `t`. -/ instance : SDiff (Finset α) := - ⟨fun s₁ s₂ => ⟨s₁.1 - s₂.1, nodup_of_le tsub_le_self s₁.2⟩⟩ + ⟨fun s₁ s₂ => ⟨s₁.1 - s₂.1, nodup_of_le (Multiset.sub_le_self ..) s₁.2⟩⟩ @[simp] theorem sdiff_val (s₁ s₂ : Finset α) : (s₁ \ s₂).val = s₁.val - s₂.val := @@ -2580,7 +2582,7 @@ theorem range_filter_eq {n m : ℕ} : (range n).filter (· = m) = if m < n then lemma range_nontrivial {n : ℕ} (hn : 1 < n) : (Finset.range n).Nontrivial := by rw [Finset.Nontrivial, Finset.coe_range] - exact ⟨0, Nat.zero_lt_one.trans hn, 1, hn, zero_ne_one⟩ + exact ⟨0, Nat.zero_lt_one.trans hn, 1, hn, Nat.zero_ne_one⟩ end Range diff --git a/Mathlib/Data/Finset/Fold.lean b/Mathlib/Data/Finset/Fold.lean index 06a0145f30d7f..39befbf6c8239 100644 --- a/Mathlib/Data/Finset/Fold.lean +++ b/Mathlib/Data/Finset/Fold.lean @@ -4,16 +4,15 @@ 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 -import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop /-! # The fold operation for a commutative associative operation over a finset. -/ --- TODO: --- assert_not_exists OrderedCommMonoid +assert_not_exists OrderedCommMonoid assert_not_exists MonoidWithZero namespace Finset diff --git a/Mathlib/Data/Finset/Lattice.lean b/Mathlib/Data/Finset/Lattice.lean index 16acc98aceb59..9cb9fd5e3090c 100644 --- a/Mathlib/Data/Finset/Lattice.lean +++ b/Mathlib/Data/Finset/Lattice.lean @@ -19,7 +19,7 @@ import Mathlib.Order.Nat -/ -- TODO: --- assert_not_exists OrderedCommMonoid +assert_not_exists OrderedCommMonoid assert_not_exists MonoidWithZero open Function Multiset OrderDual @@ -812,8 +812,9 @@ 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' {α'} [LinearOrderedAddCommMonoid β] {s : Finset α'} - (hs : s.Nonempty) (f : α' → β) (n : ℕ) : +lemma nsmul_sup' {α β : Type*} [AddMonoid β] [LinearOrder β] + [CovariantClass β β (· + ·) (· ≤ ·)] [CovariantClass β β (swap (· + ·)) (· ≤ ·)] + {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 @@ -963,8 +964,9 @@ 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' {α'} [LinearOrderedAddCommMonoid β] {s : Finset α'} - (hs : s.Nonempty) (f : α' → β) (n : ℕ) : +lemma nsmul_inf' {α β : Type*} [AddMonoid β] [LinearOrder β] + [CovariantClass β β (· + ·) (· ≤ ·)] [CovariantClass β β (swap (· + ·)) (· ≤ ·)] + {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 diff --git a/Mathlib/Data/Finset/MulAntidiagonal.lean b/Mathlib/Data/Finset/MulAntidiagonal.lean index 6d9a95ba9eb9e..683d02323b1ff 100644 --- a/Mathlib/Data/Finset/MulAntidiagonal.lean +++ b/Mathlib/Data/Finset/MulAntidiagonal.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Yaël Dillies -/ +import Mathlib.Algebra.Order.Monoid.Defs import Mathlib.Data.Set.Pointwise.Basic import Mathlib.Data.Set.MulAntidiagonal diff --git a/Mathlib/Data/Finset/NatAntidiagonal.lean b/Mathlib/Data/Finset/NatAntidiagonal.lean index 012eccf69ede9..c5cb7f832ea29 100644 --- a/Mathlib/Data/Finset/NatAntidiagonal.lean +++ b/Mathlib/Data/Finset/NatAntidiagonal.lean @@ -151,7 +151,7 @@ theorem antidiagonal.snd_lt {n : ℕ} {kl : ℕ × ℕ} (hlk : kl ∈ antidiagon def antidiagonalEquivFin (n : ℕ) : antidiagonal n ≃ Fin (n + 1) where toFun := fun ⟨⟨i, j⟩, h⟩ ↦ ⟨i, antidiagonal.fst_lt h⟩ invFun := fun ⟨i, h⟩ ↦ ⟨⟨i, n - i⟩, by - rw [mem_antidiagonal, add_comm, tsub_add_cancel_iff_le] + rw [mem_antidiagonal, add_comm, Nat.sub_add_cancel] exact Nat.le_of_lt_succ h⟩ left_inv := by rintro ⟨⟨i, j⟩, h⟩; ext; rfl right_inv x := rfl diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index 39cd170eb6188..a0d03d9652740 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -182,7 +182,7 @@ theorem image_Ico_emod (n a : ℤ) (h : 0 ≤ a) : (Ico n (n + a)).image (· % a · exact hn.symm.le.trans (add_le_add_right hi _) · rw [add_comm n a] refine add_lt_add_of_lt_of_le hia.right (le_trans ?_ hn.le) - simp only [zero_le, le_add_iff_nonneg_left] + simp only [Nat.zero_le, le_add_iff_nonneg_left] exact Int.emod_nonneg n (ne_of_gt ha) · rw [Int.add_mul_emod_self_left, Int.emod_eq_of_lt hia.left hia.right] diff --git a/Mathlib/Data/Multiset/Antidiagonal.lean b/Mathlib/Data/Multiset/Antidiagonal.lean index f74172f629ecb..df02ceb0f2f41 100644 --- a/Mathlib/Data/Multiset/Antidiagonal.lean +++ b/Mathlib/Data/Multiset/Antidiagonal.lean @@ -12,6 +12,7 @@ The antidiagonal of a multiset `s` consists of all pairs `(t₁, t₂)` such that `t₁ + t₂ = s`. These pairs are counted with multiplicities. -/ +assert_not_exists OrderedCommMonoid assert_not_exists Ring universe u @@ -80,7 +81,7 @@ theorem antidiagonal_cons (a : α) (s) : theorem antidiagonal_eq_map_powerset [DecidableEq α] (s : Multiset α) : s.antidiagonal = s.powerset.map fun t ↦ (s - t, t) := by induction' s using Multiset.induction_on with a s hs - · simp only [antidiagonal_zero, powerset_zero, zero_tsub, map_singleton] + · simp only [antidiagonal_zero, powerset_zero, Multiset.zero_sub, map_singleton] · simp_rw [antidiagonal_cons, powerset_cons, map_add, hs, map_map, Function.comp, Prod.map_mk, id, sub_cons, erase_cons_head] rw [add_comm] diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 6889f2116d1eb..2c9d3259722e0 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Algebra.Group.Nat -import Mathlib.Algebra.Order.Sub.Canonical +import Mathlib.Algebra.Order.Sub.Unbundled.Basic import Mathlib.Data.List.Perm import Mathlib.Data.Set.List import Mathlib.Init.Quot @@ -17,6 +17,9 @@ These are implemented as the quotient of a list by permutations. We define the global infix notation `::ₘ` for `Multiset.cons`. -/ +-- No bundled ordered algebra should be required +assert_not_exists OrderedCommMonoid + universe v open List Subtype Nat Function @@ -585,16 +588,14 @@ instance : CovariantClass (Multiset α) (Multiset α) (· + ·) (· ≤ ·) := instance : ContravariantClass (Multiset α) (Multiset α) (· + ·) (· ≤ ·) := ⟨fun _s _t _u => add_le_add_iff_left'.1⟩ -instance : OrderedCancelAddCommMonoid (Multiset α) where - zero := 0 - add := (· + ·) +instance instAddCommMonoid : AddCancelCommMonoid (Multiset α) where add_comm := fun s t => Quotient.inductionOn₂ s t fun l₁ l₂ => Quot.sound perm_append_comm add_assoc := fun s₁ s₂ s₃ => Quotient.inductionOn₃ s₁ s₂ s₃ fun l₁ l₂ l₃ => congr_arg _ <| append_assoc l₁ l₂ l₃ zero_add := fun s => Quot.inductionOn s fun l => rfl add_zero := fun s => Quotient.inductionOn s fun l => congr_arg _ <| append_nil l - add_le_add_left := fun s₁ s₂ => add_le_add_left - le_of_add_le_add_left := fun s₁ s₂ s₃ => le_of_add_le_add_left + add_left_cancel := fun _ _ _ h => + le_antisymm (Multiset.add_le_add_iff_left'.mp h.le) (Multiset.add_le_add_iff_left'.mp h.ge) nsmul := nsmulRec theorem le_add_right (s t : Multiset α) : s ≤ s + t := by simpa using add_le_add_left (zero_le t) s @@ -608,13 +609,6 @@ theorem le_iff_exists_add {s t : Multiset α} : s ≤ t ↔ ∃ u, t = s + u := ⟨l, Quot.sound p⟩, fun ⟨_u, e⟩ => e.symm ▸ le_add_right _ _⟩ -instance : CanonicallyOrderedAddCommMonoid (Multiset α) where - __ := inferInstanceAs (OrderBot (Multiset α)) - le_self_add := le_add_right - exists_add_of_le h := leInductionOn h fun s => - let ⟨l, p⟩ := s.exists_perm_append - ⟨l, Quot.sound p⟩ - @[simp] theorem cons_add (a : α) (s t : Multiset α) : a ::ₘ s + t = a ::ₘ (s + t) := by rw [← singleton_add, ← singleton_add, add_assoc] @@ -1446,6 +1440,9 @@ protected theorem sub_zero (s : Multiset α) : s - 0 = s := theorem sub_cons (a : α) (s t : Multiset α) : s - a ::ₘ t = s.erase a - t := Quotient.inductionOn₂ s t fun _l₁ _l₂ => congr_arg _ <| diff_cons _ _ _ +protected theorem zero_sub (t : Multiset α) : 0 - t = 0 := + Multiset.induction_on t rfl fun a s ih => by simp [ih] + /-- This is a special case of `tsub_le_iff_right`, which should be used instead of this. This is needed to prove `OrderedSub (Multiset α)`. -/ protected theorem sub_le_iff_le_add : s - t ≤ u ↔ s ≤ u + t := by @@ -1453,9 +1450,18 @@ protected theorem sub_le_iff_le_add : s - t ≤ u ↔ s ≤ u + t := by exact @(Multiset.induction_on t (by simp [Multiset.sub_zero]) fun a t IH s => by simp [IH, erase_le_iff_le_cons]) +protected theorem sub_le_self (s t : Multiset α) : s - t ≤ s := by + rw [Multiset.sub_le_iff_le_add] + exact le_add_right _ _ + instance : OrderedSub (Multiset α) := ⟨fun _n _m _k => Multiset.sub_le_iff_le_add⟩ +instance : ExistsAddOfLE (Multiset α) where + exists_add_of_le h := leInductionOn h fun s => + let ⟨l, p⟩ := s.exists_perm_append + ⟨l, Quot.sound p⟩ + theorem cons_sub_of_le (a : α) {s t : Multiset α} (h : t ≤ s) : a ::ₘ s - t = a ::ₘ (s - t) := by rw [← singleton_add, ← singleton_add, add_tsub_assoc_of_le h] @@ -1500,9 +1506,10 @@ theorem union_le_union_right (h : s ≤ t) (u) : s ∪ u ≤ t ∪ u := theorem union_le (h₁ : s ≤ u) (h₂ : t ≤ u) : s ∪ t ≤ u := by rw [← eq_union_left h₂]; exact union_le_union_right h₁ t + @[simp] theorem mem_union : a ∈ s ∪ t ↔ a ∈ s ∨ a ∈ t := - ⟨fun h => (mem_add.1 h).imp_left (mem_of_le tsub_le_self), + ⟨fun h => (mem_add.1 h).imp_left (mem_of_le <| Multiset.sub_le_self _ _), (Or.elim · (mem_of_le <| le_union_left _ _) (mem_of_le <| le_union_right _ _))⟩ @[simp] @@ -1512,7 +1519,7 @@ theorem map_union [DecidableEq β] {f : α → β} (finj : Function.Injective f) congr_arg ofList (by rw [List.map_append f, List.map_diff finj]) @[simp] theorem zero_union : 0 ∪ s = s := by - simp [union_def] + simp [union_def, Multiset.zero_sub] @[simp] theorem union_zero : s ∪ 0 = s := by simp [union_def] @@ -1554,7 +1561,7 @@ theorem inter_le_right (s : Multiset α) : ∀ t, s ∩ t ≤ t := theorem le_inter (h₁ : s ≤ t) (h₂ : s ≤ u) : s ≤ t ∩ u := by revert s u; refine @(Multiset.induction_on t ?_ fun a t IH => ?_) <;> intros s u h₁ h₂ - · simpa only [zero_inter, nonpos_iff_eq_zero] using h₁ + · simpa only [zero_inter] using h₁ by_cases h : a ∈ u · rw [cons_inter_of_pos _ h, ← erase_le_iff_le_cons] exact IH (erase_le_iff_le_cons.2 h₁) (erase_le_erase _ h₂) diff --git a/Mathlib/Data/Multiset/Nodup.lean b/Mathlib/Data/Multiset/Nodup.lean index 7e24031fb1f8c..2ea7505ae7caa 100644 --- a/Mathlib/Data/Multiset/Nodup.lean +++ b/Mathlib/Data/Multiset/Nodup.lean @@ -192,7 +192,7 @@ theorem range_le {m n : ℕ} : range m ≤ range n ↔ m ≤ n := theorem mem_sub_of_nodup [DecidableEq α] {a : α} {s t : Multiset α} (d : Nodup s) : a ∈ s - t ↔ a ∈ s ∧ a ∉ t := ⟨fun h => - ⟨mem_of_le tsub_le_self h, fun h' => by + ⟨mem_of_le (Multiset.sub_le_self ..) h, fun h' => by refine count_eq_zero.1 ?_ h rw [count_sub a s t, Nat.sub_eq_zero_iff_le] exact le_trans (nodup_iff_count_le_one.1 d _) (count_pos.2 h')⟩, diff --git a/Mathlib/Data/Multiset/OrderedMonoid.lean b/Mathlib/Data/Multiset/OrderedMonoid.lean new file mode 100644 index 0000000000000..bf40bc0b213eb --- /dev/null +++ b/Mathlib/Data/Multiset/OrderedMonoid.lean @@ -0,0 +1,31 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import Mathlib.Data.Multiset.Basic +import Mathlib.Algebra.Order.Monoid.Canonical.Defs + +/-! +# Multisets as ordered monoids + +The `OrderedCancelAddCommMonoid` and `CanonicallyOrderedAddCommMonoid` instances on `Multiset α` + +-/ + +variable {α : Type*} + +namespace Multiset + +open List + +instance : OrderedCancelAddCommMonoid (Multiset α) where + add_le_add_left := fun _ _ => add_le_add_left + le_of_add_le_add_left := fun _ _ _ => le_of_add_le_add_left + +instance : CanonicallyOrderedAddCommMonoid (Multiset α) where + __ := inferInstanceAs (OrderBot (Multiset α)) + le_self_add := le_add_right + exists_add_of_le h := exists_add_of_le h + +end Multiset diff --git a/Mathlib/Data/PNat/Factors.lean b/Mathlib/Data/PNat/Factors.lean index 32cae0329ae6c..103a0f6a9d279 100644 --- a/Mathlib/Data/PNat/Factors.lean +++ b/Mathlib/Data/PNat/Factors.lean @@ -6,6 +6,7 @@ Authors: Neil Strickland import Mathlib.Algebra.BigOperators.Group.Multiset import Mathlib.Data.PNat.Prime import Mathlib.Data.Nat.Factors +import Mathlib.Data.Multiset.OrderedMonoid import Mathlib.Data.Multiset.Sort /-! diff --git a/Mathlib/Data/Set/MulAntidiagonal.lean b/Mathlib/Data/Set/MulAntidiagonal.lean index b6b66e2d2c290..487d0d5850eb3 100644 --- a/Mathlib/Data/Set/MulAntidiagonal.lean +++ b/Mathlib/Data/Set/MulAntidiagonal.lean @@ -83,7 +83,9 @@ end CancelCommMonoid section OrderedCancelCommMonoid -variable [OrderedCancelCommMonoid α] (s t : Set α) (a : α) {x y : mulAntidiagonal s t a} +variable [CancelCommMonoid α] [PartialOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] + [CovariantClass α α (Function.swap (· * ·)) (· < ·)] + (s t : Set α) (a : α) {x y : mulAntidiagonal s t a} @[to_additive Set.AddAntidiagonal.eq_of_fst_le_fst_of_snd_le_snd] theorem eq_of_fst_le_fst_of_snd_le_snd (h₁ : (x : α × α).1 ≤ (y : α × α).1) @@ -110,8 +112,11 @@ theorem finite_of_isPWO (hs : s.IsPWO) (ht : t.IsPWO) (a) : (mulAntidiagonal s t end OrderedCancelCommMonoid +variable [CancelCommMonoid α] [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)] + [CovariantClass α α (Function.swap (· * ·)) (· < ·)] + @[to_additive Set.AddAntidiagonal.finite_of_isWF] -theorem finite_of_isWF [LinearOrderedCancelCommMonoid α] {s t : Set α} (hs : s.IsWF) (ht : t.IsWF) +theorem finite_of_isWF {s t : Set α} (hs : s.IsWF) (ht : t.IsWF) (a) : (mulAntidiagonal s t a).Finite := finite_of_isPWO hs.isPWO ht.isPWO a diff --git a/Mathlib/GroupTheory/MonoidLocalization/Order.lean b/Mathlib/GroupTheory/MonoidLocalization/Order.lean index c25bc39fbfb7a..ccdd1707f6e4d 100644 --- a/Mathlib/GroupTheory/MonoidLocalization/Order.lean +++ b/Mathlib/GroupTheory/MonoidLocalization/Order.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Amelia Livingston. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston -/ +import Mathlib.Algebra.Order.Monoid.Defs import Mathlib.GroupTheory.MonoidLocalization.Basic /-! diff --git a/Mathlib/RingTheory/HahnSeries/Basic.lean b/Mathlib/RingTheory/HahnSeries/Basic.lean index b189d32690b7d..b063204680d7f 100644 --- a/Mathlib/RingTheory/HahnSeries/Basic.lean +++ b/Mathlib/RingTheory/HahnSeries/Basic.lean @@ -321,7 +321,7 @@ theorem coeff_order_ne_zero {x : HahnSeries Γ R} (hx : x ≠ 0) : x.coeff x.ord rw [order_of_ne hx] exact x.isWF_support.min_mem (support_nonempty_iff.2 hx) -theorem order_le_of_coeff_ne_zero {Γ} [LinearOrderedCancelAddCommMonoid Γ] {x : HahnSeries Γ R} +theorem order_le_of_coeff_ne_zero {Γ} [AddMonoid Γ] [LinearOrder Γ] {x : HahnSeries Γ R} {g : Γ} (h : x.coeff g ≠ 0) : x.order ≤ g := le_trans (le_of_eq (order_of_ne (ne_zero_of_coeff_ne_zero h))) (Set.IsWF.min_le _ _ ((mem_support _ _).2 h)) diff --git a/Mathlib/Tactic/Monotonicity/Lemmas.lean b/Mathlib/Tactic/Monotonicity/Lemmas.lean index 3ad48a50d8db3..a0ad5266eac62 100644 --- a/Mathlib/Tactic/Monotonicity/Lemmas.lean +++ b/Mathlib/Tactic/Monotonicity/Lemmas.lean @@ -5,7 +5,7 @@ Authors: Simon Hudon -/ import Mathlib.Algebra.Order.Group.Abs import Mathlib.Algebra.Order.Ring.Defs -import Mathlib.Algebra.Order.Sub.Canonical +import Mathlib.Algebra.Order.Sub.Unbundled.Basic import Mathlib.Data.Set.Lattice import Mathlib.Tactic.Monotonicity.Attr From 67d7502bb773ddf0384a3ac2309b8acd3530bede Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Sat, 17 Aug 2024 13:18:00 +0000 Subject: [PATCH 358/359] chore(Algebra.Order.Ring.Rat): split into unbundled and bundled ordered algebra results (#15071) We split off all but the `LinearOrderedCommRing` and the inferred `OrderedAddCommMonoid` instances on `Rat` into `Algebra.Order.Ring.Unbundled.Rat` as these results do not require bundled ordered algebra classes. --- Mathlib.lean | 1 + Mathlib/Algebra/Order/Ring/Rat.lean | 194 +-------------- Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean | 226 ++++++++++++++++++ Mathlib/Data/FP/Basic.lean | 2 +- Mathlib/Data/Int/Star.lean | 1 + Mathlib/Data/NNRat/Defs.lean | 9 +- Mathlib/Data/Rat/Cast/Lemmas.lean | 4 +- 7 files changed, 239 insertions(+), 198 deletions(-) create mode 100644 Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean diff --git a/Mathlib.lean b/Mathlib.lean index dfa5d02698df5..7227416de3c6a 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -643,6 +643,7 @@ import Mathlib.Algebra.Order.Ring.Star import Mathlib.Algebra.Order.Ring.Synonym import Mathlib.Algebra.Order.Ring.Unbundled.Basic import Mathlib.Algebra.Order.Ring.Unbundled.Nonneg +import Mathlib.Algebra.Order.Ring.Unbundled.Rat import Mathlib.Algebra.Order.Ring.WithTop import Mathlib.Algebra.Order.Sub.Basic import Mathlib.Algebra.Order.Sub.Defs diff --git a/Mathlib/Algebra/Order/Ring/Rat.lean b/Mathlib/Algebra/Order/Ring/Rat.lean index 00c5c90fcde79..88cca6d79229f 100644 --- a/Mathlib/Algebra/Order/Ring/Rat.lean +++ b/Mathlib/Algebra/Order/Ring/Rat.lean @@ -3,8 +3,8 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import Mathlib.Algebra.Order.Ring.Int -import Mathlib.Algebra.Ring.Rat +import Mathlib.Algebra.Order.Ring.Defs +import Mathlib.Algebra.Order.Ring.Unbundled.Rat /-! # The rational numbers form a linear ordered field @@ -27,168 +27,6 @@ assert_not_exists GaloisConnection namespace Rat -variable {a b c p q : ℚ} - -@[simp] lemma divInt_nonneg_iff_of_pos_right {a b : ℤ} (hb : 0 < b) : 0 ≤ a /. b ↔ 0 ≤ a := by - cases' hab : a /. b with n d hd hnd - rw [mk'_eq_divInt, divInt_eq_iff hb.ne' (mod_cast hd)] at hab - rw [← num_nonneg, ← mul_nonneg_iff_of_pos_right hb, ← hab, - mul_nonneg_iff_of_pos_right (mod_cast Nat.pos_of_ne_zero hd)] - -@[simp] lemma divInt_nonneg {a b : ℤ} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a /. b := by - obtain rfl | hb := hb.eq_or_lt - · simp - rfl - rwa [divInt_nonneg_iff_of_pos_right hb] - -@[simp] lemma mkRat_nonneg {a : ℤ} (ha : 0 ≤ a) (b : ℕ) : 0 ≤ mkRat a b := by - simpa using divInt_nonneg ha (Int.natCast_nonneg _) - -theorem ofScientific_nonneg (m : ℕ) (s : Bool) (e : ℕ) : - 0 ≤ Rat.ofScientific m s e := by - rw [Rat.ofScientific] - cases s - · rw [if_neg (by decide)] - refine num_nonneg.mp ?_ - rw [num_natCast] - exact Int.natCast_nonneg _ - · rw [if_pos rfl, normalize_eq_mkRat] - exact Rat.mkRat_nonneg (Int.natCast_nonneg _) _ - -instance _root_.NNRatCast.toOfScientific {K} [NNRatCast K] : OfScientific K where - ofScientific (m : ℕ) (b : Bool) (d : ℕ) := - NNRat.cast ⟨Rat.ofScientific m b d, ofScientific_nonneg m b d⟩ - -/-- Casting a scientific literal via `ℚ≥0` is the same as casting directly. -/ -@[simp, norm_cast] -theorem _root_.NNRat.cast_ofScientific {K} [NNRatCast K] (m : ℕ) (s : Bool) (e : ℕ) : - (OfScientific.ofScientific m s e : ℚ≥0) = (OfScientific.ofScientific m s e : K) := - rfl - -protected lemma add_nonneg : 0 ≤ a → 0 ≤ b → 0 ≤ a + b := - numDenCasesOn' a fun n₁ d₁ h₁ ↦ numDenCasesOn' b fun n₂ d₂ h₂ ↦ by - have d₁0 : 0 < (d₁ : ℤ) := mod_cast Nat.pos_of_ne_zero h₁ - have d₂0 : 0 < (d₂ : ℤ) := mod_cast Nat.pos_of_ne_zero h₂ - simp only [d₁0, d₂0, h₁, h₂, mul_pos, divInt_nonneg_iff_of_pos_right, divInt_add_divInt, Ne, - Nat.cast_eq_zero, not_false_iff] - intro n₁0 n₂0 - apply add_nonneg <;> apply mul_nonneg <;> · first |assumption|apply Int.ofNat_zero_le - -protected lemma mul_nonneg : 0 ≤ a → 0 ≤ b → 0 ≤ a * b := - numDenCasesOn' a fun n₁ d₁ h₁ => - numDenCasesOn' b fun n₂ d₂ h₂ => by - have d₁0 : 0 < (d₁ : ℤ) := mod_cast Nat.pos_of_ne_zero h₁ - have d₂0 : 0 < (d₂ : ℤ) := mod_cast Nat.pos_of_ne_zero h₂ - simp only [d₁0, d₂0, mul_pos, divInt_nonneg_iff_of_pos_right, - divInt_mul_divInt _ _ d₁0.ne' d₂0.ne'] - apply mul_nonneg - --- Porting note (#11215): TODO can this be shortened? -protected theorem le_iff_sub_nonneg (a b : ℚ) : a ≤ b ↔ 0 ≤ b - a := - numDenCasesOn'' a fun na da ha hared => - numDenCasesOn'' b fun nb db hb hbred => by - change Rat.blt _ _ = false ↔ _ - unfold Rat.blt - simp only [Bool.and_eq_true, decide_eq_true_eq, Bool.ite_eq_false_distrib, - decide_eq_false_iff_not, not_lt, ite_eq_left_iff, not_and, not_le, ← num_nonneg] - split_ifs with h h' - · rw [Rat.sub_def] - simp only [false_iff, not_le] - simp only [normalize_eq] - apply Int.ediv_neg' - · rw [sub_neg] - apply lt_of_lt_of_le - · apply mul_neg_of_neg_of_pos h.1 - rwa [Int.natCast_pos, Nat.pos_iff_ne_zero] - · apply mul_nonneg h.2 (Int.natCast_nonneg _) - · simp only [Int.natCast_pos, Nat.pos_iff_ne_zero] - exact Nat.gcd_ne_zero_right (Nat.mul_ne_zero hb ha) - · simp only [divInt_ofNat, ← zero_iff_num_zero, mkRat_eq_zero hb] at h' - simp [h'] - · simp only [Rat.sub_def, normalize_eq] - refine ⟨fun H => ?_, fun H _ => ?_⟩ - · refine Int.ediv_nonneg ?_ (Int.natCast_nonneg _) - rw [sub_nonneg] - push_neg at h - obtain hb|hb := Ne.lt_or_lt h' - · apply H - intro H' - exact (hb.trans H').false.elim - · obtain ha|ha := le_or_lt na 0 - · apply le_trans <| mul_nonpos_of_nonpos_of_nonneg ha (Int.natCast_nonneg _) - exact mul_nonneg hb.le (Int.natCast_nonneg _) - · exact H (fun _ => ha) - · rw [← sub_nonneg] - contrapose! H - apply Int.ediv_neg' H - simp only [Int.natCast_pos, Nat.pos_iff_ne_zero] - exact Nat.gcd_ne_zero_right (Nat.mul_ne_zero hb ha) - -protected lemma divInt_le_divInt {a b c d : ℤ} (b0 : 0 < b) (d0 : 0 < d) : - a /. b ≤ c /. d ↔ a * d ≤ c * b := by - rw [Rat.le_iff_sub_nonneg, ← sub_nonneg (α := ℤ)] - simp [sub_eq_add_neg, ne_of_gt b0, ne_of_gt d0, mul_pos d0 b0] - -protected lemma le_total : a ≤ b ∨ b ≤ a := by - simpa only [← Rat.le_iff_sub_nonneg, neg_sub] using Rat.nonneg_total (b - a) - -protected theorem not_le {a b : ℚ} : ¬a ≤ b ↔ b < a := (Bool.not_eq_false _).to_iff - -instance linearOrder : LinearOrder ℚ where - le_refl a := by rw [Rat.le_iff_sub_nonneg, ← num_nonneg]; simp - le_trans a b c hab hbc := by - rw [Rat.le_iff_sub_nonneg] at hab hbc - have := Rat.add_nonneg hab hbc - simp_rw [sub_eq_add_neg, add_left_comm (b + -a) c (-b), add_comm (b + -a) (-b), - add_left_comm (-b) b (-a), add_comm (-b) (-a), add_neg_cancel_comm_assoc, - ← sub_eq_add_neg] at this - rwa [Rat.le_iff_sub_nonneg] - le_antisymm a b hab hba := by - rw [Rat.le_iff_sub_nonneg] at hab hba - rw [sub_eq_add_neg] at hba - rw [← neg_sub, sub_eq_add_neg] at hab - have := eq_neg_of_add_eq_zero_left (Rat.nonneg_antisymm hba hab) - rwa [neg_neg] at this - le_total _ _ := Rat.le_total - decidableEq := inferInstance - decidableLE := inferInstance - decidableLT := inferInstance - lt_iff_le_not_le _ _ := by rw [← Rat.not_le, and_iff_right_of_imp Rat.le_total.resolve_left] - -/-! -### Extra instances to short-circuit type class resolution - -These also prevent non-computable instances being used to construct these instances non-computably. --/ - -instance instDistribLattice : DistribLattice ℚ := inferInstance -instance instLattice : Lattice ℚ := inferInstance -instance instSemilatticeInf : SemilatticeInf ℚ := inferInstance -instance instSemilatticeSup : SemilatticeSup ℚ := inferInstance -instance instInf : Inf ℚ := inferInstance -instance instSup : Sup ℚ := inferInstance -instance instPartialOrder : PartialOrder ℚ := inferInstance -instance instPreorder : Preorder ℚ := inferInstance - -/-! ### Miscellaneous lemmas -/ - -protected lemma le_def : p ≤ q ↔ p.num * q.den ≤ q.num * p.den := by - rw [← num_divInt_den q, ← num_divInt_den p] - conv_rhs => simp only [num_divInt_den] - exact Rat.divInt_le_divInt (mod_cast p.pos) (mod_cast q.pos) - -protected lemma lt_def : p < q ↔ p.num * q.den < q.num * p.den := by - rw [lt_iff_le_and_ne, Rat.le_def] - suffices p ≠ q ↔ p.num * q.den ≠ q.num * p.den by - constructor <;> intro h - · exact lt_iff_le_and_ne.mpr ⟨h.left, this.mp h.right⟩ - · have tmp := lt_iff_le_and_ne.mp h - exact ⟨tmp.left, this.mpr tmp.right⟩ - exact not_iff_not.mpr eq_iff_mul_eq_mul - -protected theorem add_le_add_left {a b c : ℚ} : c + a ≤ c + b ↔ a ≤ b := by - rw [Rat.le_iff_sub_nonneg, add_sub_add_left_eq_sub, ← Rat.le_iff_sub_nonneg] - instance instLinearOrderedCommRing : LinearOrderedCommRing ℚ where __ := Rat.linearOrder __ := Rat.commRing @@ -213,32 +51,4 @@ instance : OrderedCancelAddCommMonoid ℚ := by infer_instance instance : OrderedAddCommMonoid ℚ := by infer_instance -@[simp] lemma num_nonpos {a : ℚ} : a.num ≤ 0 ↔ a ≤ 0 := by simpa using @num_nonneg (-a) -@[simp] lemma num_pos {a : ℚ} : 0 < a.num ↔ 0 < a := lt_iff_lt_of_le_iff_le num_nonpos -@[simp] lemma num_neg {a : ℚ} : a.num < 0 ↔ a < 0 := lt_iff_lt_of_le_iff_le num_nonneg - -@[deprecated (since := "2024-02-16")] alias num_nonneg_iff_zero_le := num_nonneg -@[deprecated (since := "2024-02-16")] alias num_pos_iff_pos := num_pos - -theorem div_lt_div_iff_mul_lt_mul {a b c d : ℤ} (b_pos : 0 < b) (d_pos : 0 < d) : - (a : ℚ) / b < c / d ↔ a * d < c * b := by - simp only [lt_iff_le_not_le] - apply and_congr - · simp [div_def', Rat.divInt_le_divInt b_pos d_pos] - · apply not_congr - simp [div_def', Rat.divInt_le_divInt d_pos b_pos] - -theorem lt_one_iff_num_lt_denom {q : ℚ} : q < 1 ↔ q.num < q.den := by simp [Rat.lt_def] - -theorem abs_def (q : ℚ) : |q| = q.num.natAbs /. q.den := by - rcases le_total q 0 with hq | hq - · rw [abs_of_nonpos hq] - rw [← num_divInt_den q, ← zero_divInt, Rat.divInt_le_divInt (mod_cast q.pos) zero_lt_one, - mul_one, zero_mul] at hq - rw [Int.ofNat_natAbs_of_nonpos hq, ← neg_def] - · rw [abs_of_nonneg hq] - rw [← num_divInt_den q, ← zero_divInt, Rat.divInt_le_divInt zero_lt_one (mod_cast q.pos), - mul_one, zero_mul] at hq - rw [Int.natAbs_of_nonneg hq, num_divInt_den] - end Rat diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean new file mode 100644 index 0000000000000..8b5c91a55c468 --- /dev/null +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Rat.lean @@ -0,0 +1,226 @@ +/- +Copyright (c) 2019 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro +-/ +import Mathlib.Algebra.Order.Group.Unbundled.Abs +import Mathlib.Algebra.Order.Group.Unbundled.Basic +import Mathlib.Algebra.Ring.Rat +import Mathlib.Init.Data.Int.Order + +/-! +# The rational numbers possess a linear order + +This file constructs the order on `ℚ` and proves various facts relating the order to +ring structure on `ℚ`. This only used unbundled type classes relating the order structure +and algebra structure on `ℚ`. For the bundled `LinearOrderedCommRing` instance on `Q`, +see `Algebra.Order.Ring.Rat`. + +## Tags + +rat, rationals, field, ℚ, numerator, denominator, num, denom, order, ordering +-/ + +assert_not_exists OrderedCommMonoid +assert_not_exists Field +assert_not_exists Finset +assert_not_exists Set.Icc +assert_not_exists GaloisConnection + +namespace Rat + +variable {a b c p q : ℚ} + +@[simp] lemma divInt_nonneg_iff_of_pos_right {a b : ℤ} (hb : 0 < b) : 0 ≤ a /. b ↔ 0 ≤ a := by + cases' hab : a /. b with n d hd hnd + rw [mk'_eq_divInt, divInt_eq_iff hb.ne' (mod_cast hd)] at hab + rw [← num_nonneg, ← Int.mul_nonneg_iff_of_pos_right hb, ← hab, + Int.mul_nonneg_iff_of_pos_right (mod_cast Nat.pos_of_ne_zero hd)] + +@[simp] lemma divInt_nonneg {a b : ℤ} (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a /. b := by + obtain rfl | hb := hb.eq_or_lt + · simp + rfl + rwa [divInt_nonneg_iff_of_pos_right hb] + +@[simp] lemma mkRat_nonneg {a : ℤ} (ha : 0 ≤ a) (b : ℕ) : 0 ≤ mkRat a b := by + simpa using divInt_nonneg ha (Int.natCast_nonneg _) + +theorem ofScientific_nonneg (m : ℕ) (s : Bool) (e : ℕ) : + 0 ≤ Rat.ofScientific m s e := by + rw [Rat.ofScientific] + cases s + · rw [if_neg (by decide)] + refine num_nonneg.mp ?_ + rw [num_natCast] + exact Int.natCast_nonneg _ + · rw [if_pos rfl, normalize_eq_mkRat] + exact Rat.mkRat_nonneg (Int.natCast_nonneg _) _ + +instance _root_.NNRatCast.toOfScientific {K} [NNRatCast K] : OfScientific K where + ofScientific (m : ℕ) (b : Bool) (d : ℕ) := + NNRat.cast ⟨Rat.ofScientific m b d, ofScientific_nonneg m b d⟩ + +/-- Casting a scientific literal via `ℚ≥0` is the same as casting directly. -/ +@[simp, norm_cast] +theorem _root_.NNRat.cast_ofScientific {K} [NNRatCast K] (m : ℕ) (s : Bool) (e : ℕ) : + (OfScientific.ofScientific m s e : ℚ≥0) = (OfScientific.ofScientific m s e : K) := + rfl + +protected lemma add_nonneg : 0 ≤ a → 0 ≤ b → 0 ≤ a + b := + numDenCasesOn' a fun n₁ d₁ h₁ ↦ numDenCasesOn' b fun n₂ d₂ h₂ ↦ by + have d₁0 : 0 < (d₁ : ℤ) := mod_cast Nat.pos_of_ne_zero h₁ + have d₂0 : 0 < (d₂ : ℤ) := mod_cast Nat.pos_of_ne_zero h₂ + simp only [d₁0, d₂0, h₁, h₂, Int.mul_pos, divInt_nonneg_iff_of_pos_right, divInt_add_divInt, Ne, + Nat.cast_eq_zero, not_false_iff] + intro n₁0 n₂0 + apply Int.add_nonneg <;> apply Int.mul_nonneg <;> · first | assumption | apply Int.ofNat_zero_le + +protected lemma mul_nonneg : 0 ≤ a → 0 ≤ b → 0 ≤ a * b := + numDenCasesOn' a fun n₁ d₁ h₁ => + numDenCasesOn' b fun n₂ d₂ h₂ => by + have d₁0 : 0 < (d₁ : ℤ) := mod_cast Nat.pos_of_ne_zero h₁ + have d₂0 : 0 < (d₂ : ℤ) := mod_cast Nat.pos_of_ne_zero h₂ + simp only [d₁0, d₂0, Int.mul_pos, divInt_nonneg_iff_of_pos_right, + divInt_mul_divInt _ _ d₁0.ne' d₂0.ne'] + apply Int.mul_nonneg + +-- Porting note (#11215): TODO can this be shortened? +protected theorem le_iff_sub_nonneg (a b : ℚ) : a ≤ b ↔ 0 ≤ b - a := + numDenCasesOn'' a fun na da ha hared => + numDenCasesOn'' b fun nb db hb hbred => by + change Rat.blt _ _ = false ↔ _ + unfold Rat.blt + simp only [Bool.and_eq_true, decide_eq_true_eq, Bool.ite_eq_false_distrib, + decide_eq_false_iff_not, not_lt, ite_eq_left_iff, not_and, not_le, ← num_nonneg] + split_ifs with h h' + · rw [Rat.sub_def] + simp only [false_iff, not_le] + simp only [normalize_eq] + apply Int.ediv_neg' + · rw [sub_neg] + apply lt_of_lt_of_le + · apply Int.mul_neg_of_neg_of_pos h.1 + rwa [Int.natCast_pos, Nat.pos_iff_ne_zero] + · apply Int.mul_nonneg h.2 (Int.natCast_nonneg _) + · simp only [Int.natCast_pos, Nat.pos_iff_ne_zero] + exact Nat.gcd_ne_zero_right (Nat.mul_ne_zero hb ha) + · simp only [divInt_ofNat, ← zero_iff_num_zero, mkRat_eq_zero hb] at h' + simp [h'] + · simp only [Rat.sub_def, normalize_eq] + refine ⟨fun H => ?_, fun H _ => ?_⟩ + · refine Int.ediv_nonneg ?_ (Int.natCast_nonneg _) + rw [Int.sub_nonneg] + push_neg at h + obtain hb|hb := Ne.lt_or_lt h' + · apply H + intro H' + exact (hb.trans H').false.elim + · obtain ha|ha := le_or_lt na 0 + · apply le_trans <| Int.mul_nonpos_of_nonpos_of_nonneg ha (Int.natCast_nonneg _) + exact Int.mul_nonneg hb.le (Int.natCast_nonneg _) + · exact H (fun _ => ha) + · rw [← Int.sub_nonneg] + contrapose! H + apply Int.ediv_neg' H + simp only [Int.natCast_pos, Nat.pos_iff_ne_zero] + exact Nat.gcd_ne_zero_right (Nat.mul_ne_zero hb ha) + +protected lemma divInt_le_divInt {a b c d : ℤ} (b0 : 0 < b) (d0 : 0 < d) : + a /. b ≤ c /. d ↔ a * d ≤ c * b := by + rw [Rat.le_iff_sub_nonneg, ← Int.sub_nonneg] + simp [sub_eq_add_neg, ne_of_gt b0, ne_of_gt d0, Int.mul_pos d0 b0] + +protected lemma le_total : a ≤ b ∨ b ≤ a := by + simpa only [← Rat.le_iff_sub_nonneg, neg_sub] using Rat.nonneg_total (b - a) + +protected theorem not_le {a b : ℚ} : ¬a ≤ b ↔ b < a := (Bool.not_eq_false _).to_iff + +instance linearOrder : LinearOrder ℚ where + le_refl a := by rw [Rat.le_iff_sub_nonneg, ← num_nonneg]; simp + le_trans a b c hab hbc := by + rw [Rat.le_iff_sub_nonneg] at hab hbc + have := Rat.add_nonneg hab hbc + simp_rw [sub_eq_add_neg, add_left_comm (b + -a) c (-b), add_comm (b + -a) (-b), + add_left_comm (-b) b (-a), add_comm (-b) (-a), add_neg_cancel_comm_assoc, + ← sub_eq_add_neg] at this + rwa [Rat.le_iff_sub_nonneg] + le_antisymm a b hab hba := by + rw [Rat.le_iff_sub_nonneg] at hab hba + rw [sub_eq_add_neg] at hba + rw [← neg_sub, sub_eq_add_neg] at hab + have := eq_neg_of_add_eq_zero_left (Rat.nonneg_antisymm hba hab) + rwa [neg_neg] at this + le_total _ _ := Rat.le_total + decidableEq := inferInstance + decidableLE := inferInstance + decidableLT := inferInstance + lt_iff_le_not_le _ _ := by rw [← Rat.not_le, and_iff_right_of_imp Rat.le_total.resolve_left] + +/-! +### Extra instances to short-circuit type class resolution + +These also prevent non-computable instances being used to construct these instances non-computably. +-/ + +instance instDistribLattice : DistribLattice ℚ := inferInstance +instance instLattice : Lattice ℚ := inferInstance +instance instSemilatticeInf : SemilatticeInf ℚ := inferInstance +instance instSemilatticeSup : SemilatticeSup ℚ := inferInstance +instance instInf : Inf ℚ := inferInstance +instance instSup : Sup ℚ := inferInstance +instance instPartialOrder : PartialOrder ℚ := inferInstance +instance instPreorder : Preorder ℚ := inferInstance + +/-! ### Miscellaneous lemmas -/ + +protected lemma le_def : p ≤ q ↔ p.num * q.den ≤ q.num * p.den := by + rw [← num_divInt_den q, ← num_divInt_den p] + conv_rhs => simp only [num_divInt_den] + exact Rat.divInt_le_divInt (mod_cast p.pos) (mod_cast q.pos) + +protected lemma lt_def : p < q ↔ p.num * q.den < q.num * p.den := by + rw [lt_iff_le_and_ne, Rat.le_def] + suffices p ≠ q ↔ p.num * q.den ≠ q.num * p.den by + constructor <;> intro h + · exact lt_iff_le_and_ne.mpr ⟨h.left, this.mp h.right⟩ + · have tmp := lt_iff_le_and_ne.mp h + exact ⟨tmp.left, this.mpr tmp.right⟩ + exact not_iff_not.mpr eq_iff_mul_eq_mul + +protected theorem add_le_add_left {a b c : ℚ} : c + a ≤ c + b ↔ a ≤ b := by + rw [Rat.le_iff_sub_nonneg, add_sub_add_left_eq_sub, ← Rat.le_iff_sub_nonneg] + +instance : CovariantClass ℚ ℚ (· + ·) (· ≤ ·) where + elim := fun _ _ _ h => Rat.add_le_add_left.2 h + +@[simp] lemma num_nonpos {a : ℚ} : a.num ≤ 0 ↔ a ≤ 0 := by + simp [Int.le_iff_lt_or_eq, instLE, Rat.blt, Int.not_lt] +@[simp] lemma num_pos {a : ℚ} : 0 < a.num ↔ 0 < a := lt_iff_lt_of_le_iff_le num_nonpos +@[simp] lemma num_neg {a : ℚ} : a.num < 0 ↔ a < 0 := lt_iff_lt_of_le_iff_le num_nonneg + +@[deprecated (since := "2024-02-16")] alias num_nonneg_iff_zero_le := num_nonneg +@[deprecated (since := "2024-02-16")] alias num_pos_iff_pos := num_pos + +theorem div_lt_div_iff_mul_lt_mul {a b c d : ℤ} (b_pos : 0 < b) (d_pos : 0 < d) : + (a : ℚ) / b < c / d ↔ a * d < c * b := by + simp only [lt_iff_le_not_le] + apply and_congr + · simp [div_def', Rat.divInt_le_divInt b_pos d_pos] + · apply not_congr + simp [div_def', Rat.divInt_le_divInt d_pos b_pos] + +theorem lt_one_iff_num_lt_denom {q : ℚ} : q < 1 ↔ q.num < q.den := by simp [Rat.lt_def] + +theorem abs_def (q : ℚ) : |q| = q.num.natAbs /. q.den := by + rcases le_total q 0 with hq | hq + · rw [abs_of_nonpos hq] + rw [← num_divInt_den q, ← zero_divInt, Rat.divInt_le_divInt (mod_cast q.pos) Int.zero_lt_one, + mul_one, zero_mul] at hq + rw [Int.ofNat_natAbs_of_nonpos hq, ← neg_def] + · rw [abs_of_nonneg hq] + rw [← num_divInt_den q, ← zero_divInt, Rat.divInt_le_divInt Int.zero_lt_one (mod_cast q.pos), + mul_one, zero_mul] at hq + rw [Int.natAbs_of_nonneg hq, num_divInt_den] + +end Rat diff --git a/Mathlib/Data/FP/Basic.lean b/Mathlib/Data/FP/Basic.lean index 8ba466a35a5cb..c854edb12daef 100644 --- a/Mathlib/Data/FP/Basic.lean +++ b/Mathlib/Data/FP/Basic.lean @@ -86,7 +86,7 @@ theorem Float.Zero.valid : ValidFinite emin 0 := rw [mul_comm] assumption le_trans C.precMax (Nat.le_mul_of_pos_left _ two_pos), - by (rw [max_eq_right]; simp [sub_eq_add_neg])⟩ + by (rw [max_eq_right]; simp [sub_eq_add_neg, Int.ofNat_zero_le])⟩ @[nolint docBlame] def Float.zero (s : Bool) : Float := diff --git a/Mathlib/Data/Int/Star.lean b/Mathlib/Data/Int/Star.lean index 9492664978ba1..1304983eb851c 100644 --- a/Mathlib/Data/Int/Star.lean +++ b/Mathlib/Data/Int/Star.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Star.Order import Mathlib.Algebra.Order.Group.Abs import Mathlib.Algebra.Order.Monoid.Submonoid import Mathlib.Algebra.Order.Ring.Basic +import Mathlib.Algebra.Order.Ring.Int /-! # Star ordered ring structure on `ℤ` diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index dc43e243e47d4..48552f3ade343 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -4,7 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Bhavik Mehta -/ import Mathlib.Algebra.Order.Nonneg.Ring +import Mathlib.Algebra.Order.Group.Unbundled.Int import Mathlib.Algebra.Order.Ring.Rat +import Mathlib.Algebra.Order.Ring.Unbundled.Rat import Mathlib.Data.Nat.Cast.Order.Ring /-! @@ -296,7 +298,7 @@ namespace NNRat variable {p q : ℚ≥0} @[norm_cast] lemma num_coe (q : ℚ≥0) : (q : ℚ).num = q.num := by - simp [num, abs_of_nonneg, Rat.num_nonneg, q.2] + simp only [num, Int.natCast_natAbs, Rat.num_nonneg, coe_nonneg, abs_of_nonneg] theorem natAbs_num_coe : (q : ℚ).num.natAbs = q.num := rfl @@ -328,14 +330,15 @@ theorem ext_num_den_iff : p = q ↔ p.num = q.num ∧ p.den = q.den := /-- Form the quotient `n / d` where `n d : ℕ`. See also `Rat.divInt` and `mkRat`. -/ -def divNat (n d : ℕ) : ℚ≥0 := ⟨.divInt n d, Rat.divInt_nonneg n.cast_nonneg d.cast_nonneg⟩ +def divNat (n d : ℕ) : ℚ≥0 := + ⟨.divInt n d, Rat.divInt_nonneg (Int.ofNat_zero_le n) (Int.ofNat_zero_le d)⟩ variable {n₁ n₂ d₁ d₂ d : ℕ} @[simp, norm_cast] lemma coe_divNat (n d : ℕ) : (divNat n d : ℚ) = .divInt n d := rfl lemma mk_divInt (n d : ℕ) : - ⟨.divInt n d, Rat.divInt_nonneg n.cast_nonneg d.cast_nonneg⟩ = divNat n d := rfl + ⟨.divInt n d, Rat.divInt_nonneg (Int.ofNat_zero_le n) (Int.ofNat_zero_le d)⟩ = divNat n d := rfl lemma divNat_inj (h₁ : d₁ ≠ 0) (h₂ : d₂ ≠ 0) : divNat n₁ d₁ = divNat n₂ d₂ ↔ n₁ * d₂ = n₂ * d₁ := by rw [← coe_inj]; simp [Rat.mkRat_eq_iff, h₁, h₂]; norm_cast diff --git a/Mathlib/Data/Rat/Cast/Lemmas.lean b/Mathlib/Data/Rat/Cast/Lemmas.lean index 9dea343405326..4089d6b56d2e4 100644 --- a/Mathlib/Data/Rat/Cast/Lemmas.lean +++ b/Mathlib/Data/Rat/Cast/Lemmas.lean @@ -32,7 +32,7 @@ theorem cast_inv_nat (n : ℕ) : ((n⁻¹ : ℚ) : α) = (n : α)⁻¹ := by cases' n with n · simp rw [cast_def, inv_natCast_num, inv_natCast_den, if_neg n.succ_ne_zero, - Int.sign_eq_one_of_pos (Nat.cast_pos.mpr n.succ_pos), Int.cast_one, one_div] + Int.sign_eq_one_of_pos (Int.ofNat_succ_pos n), Int.cast_one, one_div] -- Porting note: proof got a lot easier - is this still the intended statement? @[simp] @@ -47,7 +47,7 @@ theorem cast_nnratCast {K} [DivisionRing K] (q : ℚ≥0) : rw [Rat.cast_def, NNRat.cast_def, NNRat.cast_def] have hn := @num_div_eq_of_coprime q.num q.den ?hdp q.coprime_num_den on_goal 1 => have hd := @den_div_eq_of_coprime q.num q.den ?hdp q.coprime_num_den - case hdp => simpa only [Nat.cast_pos] using q.den_pos + case hdp => simpa only [Int.ofNat_pos] using q.den_pos simp only [Int.cast_natCast, Nat.cast_inj] at hn hd rw [hn, hd, Int.cast_natCast] From aae0e915a7ea38179de2a1dfb4c1eeb830e1be4b Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Sat, 17 Aug 2024 15:02:13 +0000 Subject: [PATCH 359/359] chore(Algebra/Lie): rename weightSpace to genWeightspace (#15911) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit There are two variants of the concept of weight space, similar to the two concepts of eigenspace and generalized eigenspace. This PR renames `weightSpace` to `genWeightSpace` so that the name `weightSpace` can be reused for the version of weight spaces using non-generalized eigenspaces. Moves: - coe_lie_shiftedWeightSpace_apply -> coe_lie_shiftedgenWeightSpace_apply - coe_weightSpaceOf_zero -> coe_genWeightSpaceOf_zero - coe_weightSpace_of_top -> coe_genWeightSpace_of_top - comap_weightSpace_eq_of_injective -> comap_genWeightSpace_eq_of_injective - disjoint_weightSpace -> disjoint_genWeightSpace - disjoint_weightSpaceOf -> disjoint_genWeightSpaceOf - eq_zero_of_mem_weightSpace_mem_posFitting -> eq_zero_of_mem_genWeightSpace_mem_posFitting - eventually_weightSpace_smul_add_eq_bot -> eventually_genWeightSpace_smul_add_eq_bot - exists_weightSpace_le_ker_of_isNoetherian -> exists_genWeightSpace_le_ker_of_isNoetherian - exists_weightSpace_smul_add_eq_bot -> exists_genWeightSpace_smul_add_eq_bot - exists_weightSpace_zero_le_ker_of_isNoetherian -> exists_genWeightSpace_zero_le_ker_of_isNoetherian - exists₂_weightSpace_smul_add_eq_bot -> exists₂_genWeightSpace_smul_add_eq_bot - finite_weightSpaceOf_ne_bot -> finite_genWeightSpaceOf_ne_bot - finite_weightSpace_ne_bot -> finite_genWeightSpace_ne_bot - iSup_ucs_eq_weightSpace_zero -> iSup_ucs_eq_genWeightSpace_zero - iSup_ucs_le_weightSpace_zero -> iSup_ucs_le_genWeightSpace_zero - iSup_weightSpaceOf_eq_top -> iSup_genWeightSpaceOf_eq_top - iSup_weightSpace_eq_top -> iSup_genWeightSpace_eq_top - iSup_weightSpace_eq_top' -> iSup_genWeightSpace_eq_top' - independent_weightSpace -> independent_genWeightSpace - independent_weightSpace' -> independent_genWeightSpace' - independent_weightSpaceOf -> independent_genWeightSpaceOf - injOn_weightSpace -> injOn_genWeightSpace - isCompl_weightSpaceOf_zero_posFittingCompOf -> isCompl_genWeightSpaceOf_zero_posFittingCompOf - isCompl_weightSpace_zero_posFittingComp -> isCompl_genWeightSpace_zero_posFittingComp - isNilpotent_toEnd_weightSpace_zero -> isNilpotent_toEnd_genWeightSpace_zero - lie_mem_weightSpaceChain_of_weightSpace_eq_bot_left -> lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_left - lie_mem_weightSpaceChain_of_weightSpace_eq_bot_right -> lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_right - lie_mem_weightSpace_of_mem_weightSpace -> lie_mem_genWeightSpace_of_mem_genWeightSpace - map_weightSpace_eq -> map_genWeightSpace_eq - map_weightSpace_eq_of_injective -> map_genWeightSpace_eq_of_injective - map_weightSpace_le -> map_genWeightSpace_le - mapsTo_toEnd_weightSpace_add_of_mem_rootSpace -> mapsTo_toEnd_genWeightSpace_add_of_mem_rootSpace - mem_weightSpace -> mem_genWeightSpace - mem_weightSpaceOf -> mem_genWeightSpaceOf - rootSpace_comap_eq_weightSpace -> rootSpace_comap_eq_genWeightSpace - shiftedWeightSpace -> shiftedgenWeightSpace - traceForm_eq_sum_weightSpaceOf -> traceForm_eq_sum_genWeightSpaceOf - traceForm_weightSpace_eq -> traceForm_genWeightSpace_eq - trace_comp_toEnd_weightSpace_eq -> trace_comp_toEnd_genWeightSpace_eq - trace_toEnd_weightSpace -> trace_toEnd_genWeightSpace - trace_toEnd_weightSpaceChain_eq_zero -> trace_toEnd_genWeightSpaceChain_eq_zero - weightSpace -> genWeightSpace - weightSpaceChain -> genWeightSpaceChain - weightSpaceChain_def -> genWeightSpaceChain_def - weightSpaceChain_def' -> genWeightSpaceChain_def' - weightSpaceChain_neg -> genWeightSpaceChain_neg - weightSpaceOf -> genWeightSpaceOf - weightSpaceOf_ne_bot -> genWeightSpaceOf_ne_bot - weightSpace_add_chainTop -> genWeightSpace_add_chainTop - weightSpace_chainBotCoeff_sub_one_zsmul_sub -> genWeightSpace_chainBotCoeff_sub_one_zsmul_sub - weightSpace_chainTopCoeff_add_one_nsmul_add -> genWeightSpace_chainTopCoeff_add_one_nsmul_add - weightSpace_chainTopCoeff_add_one_zsmul_add -> genWeightSpace_chainTopCoeff_add_one_zsmul_add - weightSpace_le_weightSpaceChain -> genWeightSpace_le_genWeightSpaceChain - weightSpace_le_weightSpaceOf -> genWeightSpace_le_genWeightSpaceOf - weightSpace_ne_bot -> genWeightSpace_ne_bot - weightSpace_neg_add_chainBot -> genWeightSpace_neg_add_chainBot - weightSpace_neg_zsmul_add_ne_bot -> genWeightSpace_neg_zsmul_add_ne_bot - weightSpace_nsmul_add_ne_bot_of_le -> genWeightSpace_nsmul_add_ne_bot_of_le - weightSpace_weightSpaceOf_map_incl -> genWeightSpace_genWeightSpaceOf_map_incl - weightSpace_zero_normalizer_eq_self -> genWeightSpace_zero_normalizer_eq_self - weightSpace_zsmul_add_ne_bot -> genWeightSpace_zsmul_add_ne_bot - zero_lt_finrank_weightSpace -> zero_lt_finrank_genWeightSpace - zero_weightSpace_eq_top_of_nilpotent -> zero_genWeightSpace_eq_top_of_nilpotent - zero_weightSpace_eq_top_of_nilpotent' -> zero_genWeightSpace_eq_top_of_nilpotent' --- Mathlib/Algebra/Lie/TraceForm.lean | 49 +-- Mathlib/Algebra/Lie/Weights/Basic.lean | 374 ++++++++++---------- Mathlib/Algebra/Lie/Weights/Cartan.lean | 54 +-- Mathlib/Algebra/Lie/Weights/Chain.lean | 199 +++++------ Mathlib/Algebra/Lie/Weights/Killing.lean | 32 +- Mathlib/Algebra/Lie/Weights/Linear.lean | 109 +++--- Mathlib/Algebra/Lie/Weights/RootSystem.lean | 30 +- 7 files changed, 429 insertions(+), 418 deletions(-) diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index 735f6856fa8fe..0837cc5f32d0f 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -103,22 +103,22 @@ lemma traceForm_lieInvariant : (traceForm R L M).lieInvariant L := by exact isNilpotent_toEnd_of_isNilpotent₂ R L M x y @[simp] -lemma traceForm_weightSpace_eq [Module.Free R M] +lemma traceForm_genWeightSpace_eq [Module.Free R M] [IsDomain R] [IsPrincipalIdealRing R] [LieAlgebra.IsNilpotent R L] [IsNoetherian R M] [LinearWeights R L M] (χ : L → R) (x y : L) : - traceForm R L (weightSpace M χ) x y = finrank R (weightSpace M χ) • (χ x * χ y) := by - set d := finrank R (weightSpace M χ) + traceForm R L (genWeightSpace M χ) x y = finrank R (genWeightSpace M χ) • (χ x * χ y) := by + set d := finrank R (genWeightSpace M χ) have h₁ : χ y • d • χ x - χ y • χ x • (d : R) = 0 := by simp [mul_comm (χ x)] have h₂ : χ x • d • χ y = d • (χ x * χ y) := by simpa [nsmul_eq_mul, smul_eq_mul] using mul_left_comm (χ x) d (χ y) - have := traceForm_eq_zero_of_isNilpotent R L (shiftedWeightSpace R L M χ) + have := traceForm_eq_zero_of_isNilpotent R L (shiftedGenWeightSpace R L M χ) replace this := LinearMap.congr_fun (LinearMap.congr_fun this x) y rwa [LinearMap.zero_apply, LinearMap.zero_apply, traceForm_apply_apply, - shiftedWeightSpace.toEnd_eq, shiftedWeightSpace.toEnd_eq, + shiftedGenWeightSpace.toEnd_eq, shiftedGenWeightSpace.toEnd_eq, ← LinearEquiv.conj_comp, LinearMap.trace_conj', LinearMap.comp_sub, LinearMap.sub_comp, LinearMap.sub_comp, map_sub, map_sub, map_sub, LinearMap.comp_smul, LinearMap.smul_comp, LinearMap.comp_id, LinearMap.id_comp, LinearMap.map_smul, LinearMap.map_smul, - trace_toEnd_weightSpace, trace_toEnd_weightSpace, + trace_toEnd_genWeightSpace, trace_toEnd_genWeightSpace, LinearMap.comp_smul, LinearMap.smul_comp, LinearMap.id_comp, map_smul, map_smul, LinearMap.trace_id, ← traceForm_apply_apply, h₁, h₂, sub_zero, sub_eq_zero] at this @@ -159,9 +159,9 @@ lemma traceForm_apply_eq_zero_of_mem_lcs_of_mem_center {x y : L} /-- Given a bilinear form `B` on a representation `M` of a nilpotent Lie algebra `L`, if `B` is invariant (in the sense that the action of `L` is skew-adjoint wrt `B`) then components of the Fitting decomposition of `M` are orthogonal wrt `B`. -/ -lemma eq_zero_of_mem_weightSpace_mem_posFitting [LieAlgebra.IsNilpotent R L] +lemma eq_zero_of_mem_genWeightSpace_mem_posFitting [LieAlgebra.IsNilpotent R L] {B : LinearMap.BilinForm R M} (hB : ∀ (x : L) (m n : M), B ⁅x, m⁆ n = - B m ⁅x, n⁆) - {m₀ m₁ : M} (hm₀ : m₀ ∈ weightSpace M (0 : L → R)) (hm₁ : m₁ ∈ posFittingComp R L M) : + {m₀ m₁ : M} (hm₀ : m₀ ∈ genWeightSpace M (0 : L → R)) (hm₁ : m₁ ∈ posFittingComp R L M) : B m₀ m₁ = 0 := by replace hB : ∀ x (k : ℕ) m n, B m ((φ x ^ k) n) = (- 1 : R) ^ k • B ((φ x ^ k) m) n := by intro x k @@ -177,7 +177,7 @@ lemma eq_zero_of_mem_weightSpace_mem_posFitting [LieAlgebra.IsNilpotent R L] apply LieSubmodule.iSup_induction _ hm₁ this (map_zero _) aesop clear hm₁ m₁; intro x m₁ hm₁ - simp only [mem_weightSpace, Pi.zero_apply, zero_smul, sub_zero] at hm₀ + simp only [mem_genWeightSpace, Pi.zero_apply, zero_smul, sub_zero] at hm₀ obtain ⟨k, hk⟩ := hm₀ x obtain ⟨m, rfl⟩ := (mem_posFittingCompOf R x m₁).mp hm₁ k simp [hB, hk] @@ -211,20 +211,21 @@ open TensorProduct variable [LieAlgebra.IsNilpotent R L] [IsDomain R] [IsPrincipalIdealRing R] -lemma traceForm_eq_sum_weightSpaceOf +lemma traceForm_eq_sum_genWeightSpaceOf [NoZeroSMulDivisors R M] [IsNoetherian R M] [IsTriangularizable R L M] (z : L) : traceForm R L M = - ∑ χ ∈ (finite_weightSpaceOf_ne_bot R L M z).toFinset, traceForm R L (weightSpaceOf M χ z) := by + ∑ χ ∈ (finite_genWeightSpaceOf_ne_bot R L M z).toFinset, + traceForm R L (genWeightSpaceOf M χ z) := by ext x y have hxy : ∀ χ : R, MapsTo ((toEnd R L M x).comp (toEnd R L M y)) - (weightSpaceOf M χ z) (weightSpaceOf M χ z) := + (genWeightSpaceOf M χ z) (genWeightSpaceOf M χ z) := fun χ m hm ↦ LieSubmodule.lie_mem _ <| LieSubmodule.lie_mem _ hm - have hfin : {χ : R | (weightSpaceOf M χ z : Submodule R M) ≠ ⊥}.Finite := by - convert finite_weightSpaceOf_ne_bot R L M z - exact LieSubmodule.coeSubmodule_eq_bot_iff (weightSpaceOf M _ _) + have hfin : {χ : R | (genWeightSpaceOf M χ z : Submodule R M) ≠ ⊥}.Finite := by + convert finite_genWeightSpaceOf_ne_bot R L M z + exact LieSubmodule.coeSubmodule_eq_bot_iff (genWeightSpaceOf M _ _) classical have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_weightSpaceOf R L M z) + (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_genWeightSpaceOf R L M z) (IsTriangularizable.iSup_eq_top z) simp only [LinearMap.coeFn_sum, Finset.sum_apply, traceForm_apply_apply, LinearMap.trace_eq_sum_trace_restrict' hds hfin hxy] @@ -275,7 +276,7 @@ lemma lowerCentralSeries_one_inf_center_le_ker_traceForm [Module.Free R M] [Modu apply LinearMap.trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero · exact IsTriangularizable.iSup_eq_top (1 ⊗ₜ[R] x) · exact fun μ ↦ trace_toEnd_eq_zero_of_mem_lcs A (A ⊗[R] L) - (weightSpaceOf (A ⊗[R] M) μ (1 ⊗ₜ x)) (le_refl 1) hz + (genWeightSpaceOf (A ⊗[R] M) μ (1 ⊗ₜ x)) (le_refl 1) hz · exact commute_toEnd_of_mem_center_right (A ⊗[R] M) hzc (1 ⊗ₜ x) /-- A nilpotent Lie algebra with a representation whose trace form is non-singular is Abelian. -/ @@ -342,7 +343,7 @@ lemma killingForm_eq_zero_of_mem_zeroRoot_mem_posFitting (hx₀ : x₀ ∈ LieAlgebra.zeroRootSubalgebra R L H) (hx₁ : x₁ ∈ LieModule.posFittingComp R H L) : killingForm R L x₀ x₁ = 0 := - LieModule.eq_zero_of_mem_weightSpace_mem_posFitting R H L + LieModule.eq_zero_of_mem_genWeightSpace_mem_posFitting R H L (fun x y z ↦ LieModule.traceForm_apply_lie_apply' R L L x y z) hx₀ hx₁ namespace LieIdeal @@ -395,20 +396,20 @@ variable [Field K] [LieAlgebra K L] [Module K M] [LieModule K L M] [FiniteDimens variable [LieAlgebra.IsNilpotent K L] [LinearWeights K L M] [IsTriangularizable K L M] lemma traceForm_eq_sum_finrank_nsmul_mul (x y : L) : - traceForm K L M x y = ∑ χ : Weight K L M, finrank K (weightSpace M χ) • (χ x * χ y) := by + traceForm K L M x y = ∑ χ : Weight K L M, finrank K (genWeightSpace M χ) • (χ x * χ y) := by have hxy : ∀ χ : Weight K L M, MapsTo (toEnd K L M x ∘ₗ toEnd K L M y) - (weightSpace M χ) (weightSpace M χ) := + (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_weightSpace' K L M) - (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_weightSpace_eq_top' K L M) + (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_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_weightSpace_eq K L M _ x y] + ← traceForm_genWeightSpace_eq K L M _ x y] rfl lemma traceForm_eq_sum_finrank_nsmul : - traceForm K L M = ∑ χ : Weight K L M, finrank K (weightSpace M χ) • + traceForm K L M = ∑ χ : Weight K L M, finrank K (genWeightSpace M χ) • (χ : L →ₗ[K] K).smulRight (χ : L →ₗ[K] K) := by ext rw [traceForm_eq_sum_finrank_nsmul_mul, ← Finset.sum_attach] diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index 9b309c2715513..24ac8b5a7f5c9 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -25,16 +25,16 @@ Basic definitions and properties of the above ideas are provided in this file. ## Main definitions - * `LieModule.weightSpaceOf` - * `LieModule.weightSpace` + * `LieModule.genWeightSpaceOf` + * `LieModule.genWeightSpace` * `LieModule.Weight` * `LieModule.posFittingCompOf` * `LieModule.posFittingComp` - * `LieModule.iSup_ucs_eq_weightSpace_zero` + * `LieModule.iSup_ucs_eq_genWeightSpace_zero` * `LieModule.iInf_lowerCentralSeries_eq_posFittingComp` - * `LieModule.isCompl_weightSpace_zero_posFittingComp` - * `LieModule.independent_weightSpace` - * `LieModule.iSup_weightSpace_eq_top` + * `LieModule.isCompl_genWeightSpace_zero_posFittingComp` + * `LieModule.independent_genWeightSpace` + * `LieModule.iSup_genWeightSpace_eq_top` ## References @@ -53,9 +53,9 @@ namespace LieModule open Set Function LieAlgebra TensorProduct TensorProduct.LieModule open scoped TensorProduct -section notation_weightSpaceOf +section notation_genWeightSpaceOf -/-- Until we define `LieModule.weightSpaceOf`, it is useful to have some notation as follows: -/ +/-- Until we define `LieModule.genWeightSpaceOf`, it is useful to have some notation as follows: -/ local notation3 "𝕎("M", " χ", " x")" => (toEnd R L M x).maxGenEigenspace χ /-- See also `bourbaki1975b` Chapter VII §1.1, Proposition 2 (ii). -/ @@ -143,10 +143,10 @@ lemma lie_mem_maxGenEigenspace_toEnd variable (M) /-- If `M` is a representation of a nilpotent Lie algebra `L`, `χ` is a scalar, and `x : L`, then -`weightSpaceOf M χ x` is the maximal generalized `χ`-eigenspace of the action of `x` on `M`. +`genWeightSpaceOf M χ x` is the maximal generalized `χ`-eigenspace of the action of `x` on `M`. It is a Lie submodule because `L` is nilpotent. -/ -def weightSpaceOf [LieAlgebra.IsNilpotent R L] (χ : R) (x : L) : LieSubmodule R L M := +def genWeightSpaceOf [LieAlgebra.IsNilpotent R L] (χ : R) (x : L) : LieSubmodule R L M := { 𝕎(M, χ, x) with lie_mem := by intro y m hm @@ -155,33 +155,34 @@ def weightSpaceOf [LieAlgebra.IsNilpotent R L] (χ : R) (x : L) : LieSubmodule R rw [← zero_add χ] exact lie_mem_maxGenEigenspace_toEnd (by simp) hm } -end notation_weightSpaceOf +end notation_genWeightSpaceOf variable (M) variable [LieAlgebra.IsNilpotent R L] -theorem mem_weightSpaceOf (χ : R) (x : L) (m : M) : - m ∈ weightSpaceOf M χ x ↔ ∃ k : ℕ, ((toEnd R L M x - χ • ↑1) ^ k) m = 0 := by - simp [weightSpaceOf] +theorem mem_genWeightSpaceOf (χ : R) (x : L) (m : M) : + m ∈ genWeightSpaceOf M χ x ↔ ∃ k : ℕ, ((toEnd R L M x - χ • ↑1) ^ k) m = 0 := by + simp [genWeightSpaceOf] -theorem coe_weightSpaceOf_zero (x : L) : - ↑(weightSpaceOf M (0 : R) x) = ⨆ k, LinearMap.ker (toEnd R L M x ^ k) := by - simp [weightSpaceOf, Module.End.maxGenEigenspace] +theorem coe_genWeightSpaceOf_zero (x : L) : + ↑(genWeightSpaceOf M (0 : R) x) = ⨆ k, LinearMap.ker (toEnd R L M x ^ k) := by + simp [genWeightSpaceOf, Module.End.maxGenEigenspace] -/-- If `M` is a representation of a nilpotent Lie algebra `L` and `χ : L → R` is a family of -scalars, then `weightSpace M χ` is the intersection of the maximal generalized `χ x`-eigenspaces of -the action of `x` on `M` as `x` ranges over `L`. +/-- If `M` is a representation of a nilpotent Lie algebra `L` +and `χ : L → R` is a family of scalars, +then `genWeightSpace M χ` is the intersection of the maximal generalized `χ x`-eigenspaces +of the action of `x` on `M` as `x` ranges over `L`. It is a Lie submodule because `L` is nilpotent. -/ -def weightSpace (χ : L → R) : LieSubmodule R L M := - ⨅ x, weightSpaceOf M (χ x) x +def genWeightSpace (χ : L → R) : LieSubmodule R L M := + ⨅ x, genWeightSpaceOf M (χ x) x -theorem mem_weightSpace (χ : L → R) (m : M) : - m ∈ weightSpace M χ ↔ ∀ x, ∃ k : ℕ, ((toEnd R L M x - χ x • ↑1) ^ k) m = 0 := by - simp [weightSpace, mem_weightSpaceOf] +theorem mem_genWeightSpace (χ : L → R) (m : M) : + m ∈ genWeightSpace M χ ↔ ∀ x, ∃ k : ℕ, ((toEnd R L M x - χ x • ↑1) ^ k) m = 0 := by + simp [genWeightSpace, mem_genWeightSpaceOf] -lemma weightSpace_le_weightSpaceOf (x : L) (χ : L → R) : - weightSpace M χ ≤ weightSpaceOf M (χ x) x := +lemma genWeightSpace_le_genWeightSpaceOf (x : L) (χ : L → R) : + genWeightSpace M χ ≤ genWeightSpaceOf M (χ x) x := iInf_le _ x variable (R L) in @@ -190,7 +191,7 @@ non-trivial. -/ structure Weight where /-- The family of eigenvalues corresponding to a weight. -/ toFun : L → R - weightSpace_ne_bot' : weightSpace M toFun ≠ ⊥ + genWeightSpace_ne_bot' : genWeightSpace M toFun ≠ ⊥ namespace Weight @@ -202,7 +203,7 @@ instance instFunLike : FunLike (Weight R L M) L R where (↑(⟨χ, h⟩ : Weight R L M) : L → R) = χ := rfl -lemma weightSpace_ne_bot (χ : Weight R L M) : weightSpace M χ ≠ ⊥ := χ.weightSpace_ne_bot' +lemma genWeightSpace_ne_bot (χ : Weight R L M) : genWeightSpace M χ ≠ ⊥ := χ.genWeightSpace_ne_bot' variable {M} @@ -212,19 +213,19 @@ variable {M} lemma ext_iff' {χ₁ χ₂ : Weight R L M} : (χ₁ : L → R) = χ₂ ↔ χ₁ = χ₂ := by aesop lemma exists_ne_zero (χ : Weight R L M) : - ∃ x ∈ weightSpace M χ, x ≠ 0 := by - simpa [LieSubmodule.eq_bot_iff] using χ.weightSpace_ne_bot + ∃ x ∈ genWeightSpace M χ, x ≠ 0 := by + simpa [LieSubmodule.eq_bot_iff] using χ.genWeightSpace_ne_bot instance [Subsingleton M] : IsEmpty (Weight R L M) := ⟨fun h ↦ h.2 (Subsingleton.elim _ _)⟩ -instance [Nontrivial (weightSpace M (0 : L → R))] : Zero (Weight R L M) := +instance [Nontrivial (genWeightSpace M (0 : L → R))] : Zero (Weight R L M) := ⟨0, fun e ↦ not_nontrivial (⊥ : LieSubmodule R L M) (e ▸ ‹_›)⟩ @[simp] -lemma coe_zero [Nontrivial (weightSpace M (0 : L → R))] : ((0 : Weight R L M) : L → R) = 0 := rfl +lemma coe_zero [Nontrivial (genWeightSpace M (0 : L → R))] : ((0 : Weight R L M) : L → R) = 0 := rfl -lemma zero_apply [Nontrivial (weightSpace M (0 : L → R))] (x) : (0 : Weight R L M) x = 0 := rfl +lemma zero_apply [Nontrivial (genWeightSpace M (0 : L → R))] (x) : (0 : Weight R L M) x = 0 := rfl /-- The proposition that a weight of a Lie module is zero. @@ -236,28 +237,28 @@ def IsZero (χ : Weight R L M) := (χ : L → R) = 0 @[simp] lemma coe_eq_zero_iff (χ : Weight R L M) : (χ : L → R) = 0 ↔ χ.IsZero := Iff.rfl -lemma isZero_iff_eq_zero [Nontrivial (weightSpace M (0 : L → R))] {χ : Weight R L M} : +lemma isZero_iff_eq_zero [Nontrivial (genWeightSpace M (0 : L → R))] {χ : Weight R L M} : χ.IsZero ↔ χ = 0 := Weight.ext_iff' (χ₂ := 0) -lemma isZero_zero [Nontrivial (weightSpace M (0 : L → R))] : IsZero (0 : Weight R L M) := rfl +lemma isZero_zero [Nontrivial (genWeightSpace M (0 : L → R))] : IsZero (0 : Weight R L M) := rfl /-- The proposition that a weight of a Lie module is non-zero. -/ abbrev IsNonZero (χ : Weight R L M) := ¬ IsZero (χ : Weight R L M) -lemma isNonZero_iff_ne_zero [Nontrivial (weightSpace M (0 : L → R))] {χ : Weight R L M} : +lemma isNonZero_iff_ne_zero [Nontrivial (genWeightSpace M (0 : L → R))] {χ : Weight R L M} : χ.IsNonZero ↔ χ ≠ 0 := isZero_iff_eq_zero.not variable (R L M) in /-- The set of weights is equivalent to a subtype. -/ -def equivSetOf : Weight R L M ≃ {χ : L → R | weightSpace M χ ≠ ⊥} where +def equivSetOf : Weight R L M ≃ {χ : L → R | genWeightSpace M χ ≠ ⊥} where toFun w := ⟨w.1, w.2⟩ invFun w := ⟨w.1, w.2⟩ left_inv w := by simp right_inv w := by simp -lemma weightSpaceOf_ne_bot (χ : Weight R L M) (x : L) : - weightSpaceOf M (χ x) x ≠ ⊥ := by - have : ⨅ x, weightSpaceOf M (χ x) x ≠ ⊥ := χ.weightSpace_ne_bot +lemma genWeightSpaceOf_ne_bot (χ : Weight R L M) (x : L) : + genWeightSpaceOf M (χ x) x ≠ ⊥ := by + have : ⨅ x, genWeightSpaceOf M (χ x) x ≠ ⊥ := χ.genWeightSpace_ne_bot contrapose! this rw [eq_bot_iff] exact le_of_le_of_eq (iInf_le _ _) this @@ -265,7 +266,7 @@ lemma weightSpaceOf_ne_bot (χ : Weight R L M) (x : L) : lemma hasEigenvalueAt (χ : Weight R L M) (x : L) : (toEnd R L M x).HasEigenvalue (χ x) := by obtain ⟨k : ℕ, hk : (toEnd R L M x).genEigenspace (χ x) k ≠ ⊥⟩ := by - simpa [Module.End.maxGenEigenspace, weightSpaceOf] using χ.weightSpaceOf_ne_bot x + simpa [Module.End.maxGenEigenspace, genWeightSpaceOf] using χ.genWeightSpaceOf_ne_bot x exact Module.End.hasEigenvalue_of_hasGenEigenvalue hk lemma apply_eq_zero_of_isNilpotent [NoZeroSMulDivisors R M] [IsReduced R] @@ -275,91 +276,92 @@ lemma apply_eq_zero_of_isNilpotent [NoZeroSMulDivisors R M] [IsReduced R] end Weight -/-- See also the more useful form `LieModule.zero_weightSpace_eq_top_of_nilpotent`. -/ +/-- See also the more useful form `LieModule.zero_genWeightSpace_eq_top_of_nilpotent`. -/ @[simp] -theorem zero_weightSpace_eq_top_of_nilpotent' [IsNilpotent R L M] : - weightSpace M (0 : L → R) = ⊤ := by +theorem zero_genWeightSpace_eq_top_of_nilpotent' [IsNilpotent R L M] : + genWeightSpace M (0 : L → R) = ⊤ := by ext - simp [weightSpace, weightSpaceOf] + simp [genWeightSpace, genWeightSpaceOf] -theorem coe_weightSpace_of_top (χ : L → R) : - (weightSpace M (χ ∘ (⊤ : LieSubalgebra R L).incl) : Submodule R M) = weightSpace M χ := by +theorem coe_genWeightSpace_of_top (χ : L → R) : + (genWeightSpace M (χ ∘ (⊤ : LieSubalgebra R L).incl) : Submodule R M) = genWeightSpace M χ := by ext m - simp only [mem_weightSpace, LieSubmodule.mem_coeSubmodule, Subtype.forall] + simp only [mem_genWeightSpace, LieSubmodule.mem_coeSubmodule, Subtype.forall] apply forall_congr' simp @[simp] -theorem zero_weightSpace_eq_top_of_nilpotent [IsNilpotent R L M] : - weightSpace M (0 : (⊤ : LieSubalgebra R L) → R) = ⊤ := by +theorem zero_genWeightSpace_eq_top_of_nilpotent [IsNilpotent R L M] : + genWeightSpace M (0 : (⊤ : LieSubalgebra R L) → R) = ⊤ := by ext m - simp only [mem_weightSpace, Pi.zero_apply, zero_smul, sub_zero, Subtype.forall, forall_true_left, - LieSubalgebra.toEnd_mk, LieSubalgebra.mem_top, LieSubmodule.mem_top, iff_true] + simp only [mem_genWeightSpace, Pi.zero_apply, zero_smul, sub_zero, Subtype.forall, + forall_true_left, LieSubalgebra.toEnd_mk, LieSubalgebra.mem_top, LieSubmodule.mem_top, iff_true] intro x obtain ⟨k, hk⟩ := exists_forall_pow_toEnd_eq_zero R L M exact ⟨k, by simp [hk x]⟩ -theorem exists_weightSpace_le_ker_of_isNoetherian [IsNoetherian R M] (χ : L → R) (x : L) : +theorem exists_genWeightSpace_le_ker_of_isNoetherian [IsNoetherian R M] (χ : L → R) (x : L) : ∃ k : ℕ, - weightSpace M χ ≤ LinearMap.ker ((toEnd R L M x - algebraMap R _ (χ x)) ^ k) := by + genWeightSpace M χ ≤ LinearMap.ker ((toEnd R L M x - algebraMap R _ (χ x)) ^ k) := by use (toEnd R L M x).maxGenEigenspaceIndex (χ x) intro m hm replace hm : m ∈ (toEnd R L M x).maxGenEigenspace (χ x) := - weightSpace_le_weightSpaceOf M x χ hm + genWeightSpace_le_genWeightSpaceOf M x χ hm rwa [Module.End.maxGenEigenspace_eq] at hm variable (R) in -theorem exists_weightSpace_zero_le_ker_of_isNoetherian +theorem exists_genWeightSpace_zero_le_ker_of_isNoetherian [IsNoetherian R M] (x : L) : - ∃ k : ℕ, weightSpace M (0 : L → R) ≤ LinearMap.ker (toEnd R L M x ^ k) := by - simpa using exists_weightSpace_le_ker_of_isNoetherian M (0 : L → R) x + ∃ k : ℕ, genWeightSpace M (0 : L → R) ≤ LinearMap.ker (toEnd R L M x ^ k) := by + simpa using exists_genWeightSpace_le_ker_of_isNoetherian M (0 : L → R) x lemma isNilpotent_toEnd_sub_algebraMap [IsNoetherian R M] (χ : L → R) (x : L) : - _root_.IsNilpotent <| toEnd R L (weightSpace M χ) x - algebraMap R _ (χ x) := by - have : toEnd R L (weightSpace M χ) x - algebraMap R _ (χ x) = + _root_.IsNilpotent <| toEnd R L (genWeightSpace M χ) x - algebraMap R _ (χ x) := by + have : toEnd R L (genWeightSpace M χ) x - algebraMap R _ (χ x) = (toEnd R L M x - algebraMap R _ (χ x)).restrict (fun m hm ↦ sub_mem (LieSubmodule.lie_mem _ hm) (Submodule.smul_mem _ _ hm)) := by rfl - obtain ⟨k, hk⟩ := exists_weightSpace_le_ker_of_isNoetherian M χ x + obtain ⟨k, hk⟩ := exists_genWeightSpace_le_ker_of_isNoetherian M χ x use k ext ⟨m, hm⟩ simpa [this, LinearMap.pow_restrict _, LinearMap.restrict_apply] using hk hm /-- A (nilpotent) Lie algebra acts nilpotently on the zero weight space of a Noetherian Lie module. -/ -theorem isNilpotent_toEnd_weightSpace_zero [IsNoetherian R M] (x : L) : - _root_.IsNilpotent <| toEnd R L (weightSpace M (0 : L → R)) x := by +theorem isNilpotent_toEnd_genWeightSpace_zero [IsNoetherian R M] (x : L) : + _root_.IsNilpotent <| toEnd R L (genWeightSpace M (0 : L → R)) x := by simpa using isNilpotent_toEnd_sub_algebraMap M (0 : L → R) x /-- By Engel's theorem, the zero weight space of a Noetherian Lie module is nilpotent. -/ instance [IsNoetherian R M] : - IsNilpotent R L (weightSpace M (0 : L → R)) := - isNilpotent_iff_forall'.mpr <| isNilpotent_toEnd_weightSpace_zero M + IsNilpotent R L (genWeightSpace M (0 : L → R)) := + isNilpotent_iff_forall'.mpr <| isNilpotent_toEnd_genWeightSpace_zero M variable (R L) @[simp] -lemma weightSpace_zero_normalizer_eq_self : - (weightSpace M (0 : L → R)).normalizer = weightSpace M 0 := by +lemma genWeightSpace_zero_normalizer_eq_self : + (genWeightSpace M (0 : L → R)).normalizer = genWeightSpace M 0 := by refine le_antisymm ?_ (LieSubmodule.le_normalizer _) intro m hm rw [LieSubmodule.mem_normalizer] at hm - simp only [mem_weightSpace, Pi.zero_apply, zero_smul, sub_zero] at hm ⊢ + simp only [mem_genWeightSpace, Pi.zero_apply, zero_smul, sub_zero] at hm ⊢ intro y obtain ⟨k, hk⟩ := hm y y use k + 1 simpa [pow_succ, LinearMap.mul_eq_comp] -lemma iSup_ucs_le_weightSpace_zero : - ⨆ k, (⊥ : LieSubmodule R L M).ucs k ≤ weightSpace M (0 : L → R) := by - simpa using LieSubmodule.ucs_le_of_normalizer_eq_self (weightSpace_zero_normalizer_eq_self R L M) +lemma iSup_ucs_le_genWeightSpace_zero : + ⨆ k, (⊥ : LieSubmodule R L M).ucs k ≤ genWeightSpace M (0 : L → R) := by + simpa using + LieSubmodule.ucs_le_of_normalizer_eq_self (genWeightSpace_zero_normalizer_eq_self R L M) /-- See also `LieModule.iInf_lowerCentralSeries_eq_posFittingComp`. -/ -lemma iSup_ucs_eq_weightSpace_zero [IsNoetherian R M] : - ⨆ k, (⊥ : LieSubmodule R L M).ucs k = weightSpace M (0 : L → R) := by +lemma iSup_ucs_eq_genWeightSpace_zero [IsNoetherian R M] : + ⨆ k, (⊥ : LieSubmodule R L M).ucs k = genWeightSpace M (0 : L → R) := by obtain ⟨k, hk⟩ := (LieSubmodule.isNilpotent_iff_exists_self_le_ucs - <| weightSpace M (0 : L → R)).mp inferInstance - refine le_antisymm (iSup_ucs_le_weightSpace_zero R L M) (le_trans hk ?_) + <| genWeightSpace M (0 : L → R)).mp inferInstance + refine le_antisymm (iSup_ucs_le_genWeightSpace_zero R L M) (le_trans hk ?_) exact le_iSup (fun k ↦ (⊥ : LieSubmodule R L M).ucs k) k variable {L} @@ -369,7 +371,7 @@ variable {L} `range φₓ ⊇ range φₓ² ⊇ range φₓ³ ⊇ ⋯` where `φₓ : End R M := toEnd R L M x`. We call this the "positive Fitting component" because with appropriate assumptions (e.g., `R` is a field and `M` is finite-dimensional) `φₓ` induces the so-called Fitting decomposition: `M = M₀ ⊕ M₁` where -`M₀ = weightSpaceOf M 0 x` and `M₁ = posFittingCompOf R M x`. +`M₀ = genWeightSpaceOf M 0 x` and `M₁ = posFittingCompOf R M x`. It is a Lie submodule because `L` is nilpotent. -/ def posFittingCompOf (x : L) : LieSubmodule R L M := @@ -439,7 +441,7 @@ lemma posFittingComp_le_iInf_lowerCentralSeries : posFittingComp R L M ≤ ⨅ k, lowerCentralSeries R L M k := by simp [posFittingComp] -/-- See also `LieModule.iSup_ucs_eq_weightSpace_zero`. -/ +/-- See also `LieModule.iSup_ucs_eq_genWeightSpace_zero`. -/ @[simp] lemma iInf_lowerCentralSeries_eq_posFittingComp [IsNoetherian R M] [IsArtinian R M] : ⨅ k, lowerCentralSeries R L M k = posFittingComp R L M := by @@ -485,24 +487,24 @@ lemma map_posFittingComp_le : use f n rw [LieModule.toEnd_pow_apply_map, hn] -lemma map_weightSpace_le : - (weightSpace M χ).map f ≤ weightSpace M₂ χ := by +lemma map_genWeightSpace_le : + (genWeightSpace M χ).map f ≤ genWeightSpace M₂ χ := by rw [LieSubmodule.map_le_iff_le_comap] intro m hm - simp only [LieSubmodule.mem_comap, mem_weightSpace] + simp only [LieSubmodule.mem_comap, mem_genWeightSpace] intro x have : (toEnd R L M₂ x - χ x • ↑1) ∘ₗ f = f ∘ₗ (toEnd R L M x - χ x • ↑1) := by ext; simp - obtain ⟨k, h⟩ := (mem_weightSpace _ _ _).mp hm x + obtain ⟨k, h⟩ := (mem_genWeightSpace _ _ _).mp hm x exact ⟨k, by simpa [h] using LinearMap.congr_fun (LinearMap.commute_pow_left_of_commute this k) m⟩ variable {f} -lemma comap_weightSpace_eq_of_injective (hf : Injective f) : - (weightSpace M₂ χ).comap f = weightSpace M χ := by +lemma comap_genWeightSpace_eq_of_injective (hf : Injective f) : + (genWeightSpace M₂ χ).comap f = genWeightSpace M χ := by refine le_antisymm (fun m hm ↦ ?_) ?_ - · simp only [LieSubmodule.mem_comap, mem_weightSpace] at hm - simp only [mem_weightSpace] + · simp only [LieSubmodule.mem_comap, mem_genWeightSpace] at hm + simp only [mem_genWeightSpace] intro x have h : (toEnd R L M₂ x - χ x • ↑1) ∘ₗ f = f ∘ₗ (toEnd R L M x - χ x • ↑1) := by ext; simp @@ -512,18 +514,19 @@ lemma comap_weightSpace_eq_of_injective (hf : Injective f) : rw [← f.map_zero] at this; exact hf this simpa [hk] using (LinearMap.congr_fun (LinearMap.commute_pow_left_of_commute h k) m).symm · rw [← LieSubmodule.map_le_iff_le_comap] - exact map_weightSpace_le f + exact map_genWeightSpace_le f -lemma map_weightSpace_eq_of_injective (hf : Injective f) : - (weightSpace M χ).map f = weightSpace M₂ χ ⊓ f.range := by - refine le_antisymm (le_inf_iff.mpr ⟨map_weightSpace_le f, LieSubmodule.map_le_range f⟩) ?_ +lemma map_genWeightSpace_eq_of_injective (hf : Injective f) : + (genWeightSpace M χ).map f = genWeightSpace M₂ χ ⊓ f.range := by + refine le_antisymm (le_inf_iff.mpr ⟨map_genWeightSpace_le f, LieSubmodule.map_le_range f⟩) ?_ rintro - ⟨hm, ⟨m, rfl⟩⟩ - simp only [← comap_weightSpace_eq_of_injective hf, LieSubmodule.mem_map, LieSubmodule.mem_comap] + simp only [← comap_genWeightSpace_eq_of_injective hf, LieSubmodule.mem_map, + LieSubmodule.mem_comap] exact ⟨m, hm, rfl⟩ -lemma map_weightSpace_eq (e : M ≃ₗ⁅R,L⁆ M₂) : - (weightSpace M χ).map e = weightSpace M₂ χ := by - simp [map_weightSpace_eq_of_injective e.injective] +lemma map_genWeightSpace_eq (e : M ≃ₗ⁅R,L⁆ M₂) : + (genWeightSpace M χ).map e = genWeightSpace M₂ χ := by + simp [map_genWeightSpace_eq_of_injective e.injective] lemma map_posFittingComp_eq (e : M ≃ₗ⁅R,L⁆ M₂) : (posFittingComp R L M).map e = posFittingComp R L M₂ := by @@ -549,11 +552,11 @@ lemma posFittingComp_map_incl_sup_of_codisjoint [IsNoetherian R M] [IsArtinian R LieSubmodule.lowerCentralSeries_map_eq_lcs, ← LieSubmodule.lcs_sup, lowerCentralSeries, h.eq_top] -lemma weightSpace_weightSpaceOf_map_incl (x : L) (χ : L → R) : - (weightSpace (weightSpaceOf M (χ x) x) χ).map (weightSpaceOf M (χ x) x).incl = - weightSpace M χ := by - simpa [map_weightSpace_eq_of_injective (weightSpaceOf M (χ x) x).injective_incl] - using weightSpace_le_weightSpaceOf M x χ +lemma genWeightSpace_genWeightSpaceOf_map_incl (x : L) (χ : L → R) : + (genWeightSpace (genWeightSpaceOf M (χ x) x) χ).map (genWeightSpaceOf M (χ x) x).incl = + genWeightSpace M χ := by + simpa [map_genWeightSpace_eq_of_injective (genWeightSpaceOf M (χ x) x).injective_incl] + using genWeightSpace_le_genWeightSpaceOf M x χ end map_comap @@ -561,34 +564,35 @@ section fitting_decomposition variable [IsNoetherian R M] [IsArtinian R M] -lemma isCompl_weightSpaceOf_zero_posFittingCompOf (x : L) : - IsCompl (weightSpaceOf M 0 x) (posFittingCompOf R M x) := by +lemma isCompl_genWeightSpaceOf_zero_posFittingCompOf (x : L) : + IsCompl (genWeightSpaceOf M 0 x) (posFittingCompOf R M x) := by simpa only [isCompl_iff, codisjoint_iff, disjoint_iff, ← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.sup_coe_toSubmodule, LieSubmodule.inf_coe_toSubmodule, - LieSubmodule.top_coeSubmodule, LieSubmodule.bot_coeSubmodule, coe_weightSpaceOf_zero] using + LieSubmodule.top_coeSubmodule, LieSubmodule.bot_coeSubmodule, coe_genWeightSpaceOf_zero] using (toEnd R L M x).isCompl_iSup_ker_pow_iInf_range_pow /-- This lemma exists only to simplify the proof of -`LieModule.isCompl_weightSpace_zero_posFittingComp`. -/ -private lemma isCompl_weightSpace_zero_posFittingComp_aux - (h : ∀ N < (⊤ : LieSubmodule R L M), IsCompl (weightSpace N 0) (posFittingComp R L N)) : - IsCompl (weightSpace M 0) (posFittingComp R L M) := by - set M₀ := weightSpace M (0 : L → R) +`LieModule.isCompl_genWeightSpace_zero_posFittingComp`. -/ +private lemma isCompl_genWeightSpace_zero_posFittingComp_aux + (h : ∀ N < (⊤ : LieSubmodule R L M), IsCompl (genWeightSpace N 0) (posFittingComp R L N)) : + IsCompl (genWeightSpace M 0) (posFittingComp R L M) := by + set M₀ := genWeightSpace M (0 : L → R) set M₁ := posFittingComp R L M - rcases forall_or_exists_not (fun (x : L) ↦ weightSpaceOf M (0 : R) x = ⊤) - with h | ⟨x, hx : weightSpaceOf M (0 : R) x ≠ ⊤⟩ + rcases forall_or_exists_not (fun (x : L) ↦ genWeightSpaceOf M (0 : R) x = ⊤) + with h | ⟨x, hx : genWeightSpaceOf M (0 : R) x ≠ ⊤⟩ · suffices IsNilpotent R L M by simp [M₀, M₁, isCompl_top_bot] - replace h : M₀ = ⊤ := by simpa [M₀, weightSpace] + replace h : M₀ = ⊤ := by simpa [M₀, genWeightSpace] rw [← LieModule.isNilpotent_of_top_iff', ← h] infer_instance - · set M₀ₓ := weightSpaceOf M (0 : R) x + · set M₀ₓ := genWeightSpaceOf M (0 : R) x set M₁ₓ := posFittingCompOf R M x - set M₀ₓ₀ := weightSpace M₀ₓ (0 : L → R) + set M₀ₓ₀ := genWeightSpace M₀ₓ (0 : L → R) set M₀ₓ₁ := posFittingComp R L M₀ₓ - have h₁ : IsCompl M₀ₓ M₁ₓ := isCompl_weightSpaceOf_zero_posFittingCompOf R L M x + have h₁ : IsCompl M₀ₓ M₁ₓ := isCompl_genWeightSpaceOf_zero_posFittingCompOf R L M x have h₂ : IsCompl M₀ₓ₀ M₀ₓ₁ := h M₀ₓ hx.lt_top have h₃ : M₀ₓ₀.map M₀ₓ.incl = M₀ := by - rw [map_weightSpace_eq_of_injective M₀ₓ.injective_incl, inf_eq_left, LieSubmodule.range_incl] + rw [map_genWeightSpace_eq_of_injective M₀ₓ.injective_incl, inf_eq_left, + LieSubmodule.range_incl] exact iInf_le _ x have h₄ : M₀ₓ₁.map M₀ₓ.incl ⊔ M₁ₓ = M₁ := by apply le_antisymm <| sup_le_iff.mpr @@ -602,49 +606,49 @@ private lemma isCompl_weightSpace_zero_posFittingComp_aux · rwa [← LieSubmodule.map_sup, h₂.sup_eq_top, LieModuleHom.map_top, LieSubmodule.range_incl] /-- This is the Fitting decomposition of the Lie module `M`. -/ -lemma isCompl_weightSpace_zero_posFittingComp : - IsCompl (weightSpace M 0) (posFittingComp R L M) := by - let P : LieSubmodule R L M → Prop := fun N ↦ IsCompl (weightSpace N 0) (posFittingComp R L N) +lemma isCompl_genWeightSpace_zero_posFittingComp : + IsCompl (genWeightSpace M 0) (posFittingComp R L M) := by + let P : LieSubmodule R L M → Prop := fun N ↦ IsCompl (genWeightSpace N 0) (posFittingComp R L N) suffices P ⊤ by let e := LieModuleEquiv.ofTop R L M - rw [← map_weightSpace_eq e, ← map_posFittingComp_eq e] + rw [← map_genWeightSpace_eq e, ← map_posFittingComp_eq e] exact (LieSubmodule.orderIsoMapComap e).isCompl_iff.mp this refine (LieSubmodule.wellFounded_of_isArtinian R L M).induction (C := P) _ fun N hN ↦ ?_ - refine isCompl_weightSpace_zero_posFittingComp_aux R L N fun N' hN' ↦ ?_ - suffices IsCompl (weightSpace (N'.map N.incl) 0) (posFittingComp R L (N'.map N.incl)) by + refine isCompl_genWeightSpace_zero_posFittingComp_aux R L N fun N' hN' ↦ ?_ + suffices IsCompl (genWeightSpace (N'.map N.incl) 0) (posFittingComp R L (N'.map N.incl)) by let e := LieSubmodule.equivMapOfInjective N' N.injective_incl - rw [← map_weightSpace_eq e, ← map_posFittingComp_eq e] at this + rw [← map_genWeightSpace_eq e, ← map_posFittingComp_eq e] at this exact (LieSubmodule.orderIsoMapComap e).isCompl_iff.mpr this exact hN _ (LieSubmodule.map_incl_lt_iff_lt_top.mpr hN') end fitting_decomposition -lemma disjoint_weightSpaceOf [NoZeroSMulDivisors R M] {x : L} {φ₁ φ₂ : R} (h : φ₁ ≠ φ₂) : - Disjoint (weightSpaceOf M φ₁ x) (weightSpaceOf M φ₂ x) := by +lemma disjoint_genWeightSpaceOf [NoZeroSMulDivisors R M] {x : L} {φ₁ φ₂ : R} (h : φ₁ ≠ φ₂) : + Disjoint (genWeightSpaceOf M φ₁ x) (genWeightSpaceOf M φ₂ x) := by rw [LieSubmodule.disjoint_iff_coe_toSubmodule] exact Module.End.disjoint_iSup_genEigenspace _ h -lemma disjoint_weightSpace [NoZeroSMulDivisors R M] {χ₁ χ₂ : L → R} (h : χ₁ ≠ χ₂) : - Disjoint (weightSpace M χ₁) (weightSpace M χ₂) := by +lemma disjoint_genWeightSpace [NoZeroSMulDivisors R M] {χ₁ χ₂ : L → R} (h : χ₁ ≠ χ₂) : + Disjoint (genWeightSpace M χ₁) (genWeightSpace M χ₂) := by obtain ⟨x, hx⟩ : ∃ x, χ₁ x ≠ χ₂ x := Function.ne_iff.mp h - exact (disjoint_weightSpaceOf R L M hx).mono - (weightSpace_le_weightSpaceOf M x χ₁) (weightSpace_le_weightSpaceOf M x χ₂) + exact (disjoint_genWeightSpaceOf R L M hx).mono + (genWeightSpace_le_genWeightSpaceOf M x χ₁) (genWeightSpace_le_genWeightSpaceOf M x χ₂) -lemma injOn_weightSpace [NoZeroSMulDivisors R M] : - InjOn (fun (χ : L → R) ↦ weightSpace M χ) {χ | weightSpace M χ ≠ ⊥} := by - rintro χ₁ _ χ₂ hχ₂ (hχ₁₂ : weightSpace M χ₁ = weightSpace M χ₂) +lemma injOn_genWeightSpace [NoZeroSMulDivisors R M] : + InjOn (fun (χ : L → R) ↦ genWeightSpace M χ) {χ | genWeightSpace M χ ≠ ⊥} := by + rintro χ₁ _ χ₂ hχ₂ (hχ₁₂ : genWeightSpace M χ₁ = genWeightSpace M χ₂) contrapose! hχ₂ - simpa [hχ₁₂] using disjoint_weightSpace R L M hχ₂ + simpa [hχ₁₂] using disjoint_genWeightSpace R L M hχ₂ /-- Lie module weight spaces are independent. -See also `LieModule.independent_weightSpace'`. -/ -lemma independent_weightSpace [NoZeroSMulDivisors R M] : - CompleteLattice.Independent fun (χ : L → R) ↦ weightSpace M χ := by +See also `LieModule.independent_genWeightSpace'`. -/ +lemma independent_genWeightSpace [NoZeroSMulDivisors R M] : + CompleteLattice.Independent fun (χ : L → R) ↦ genWeightSpace M χ := by classical suffices ∀ χ (s : Finset (L → R)) (_ : χ ∉ s), - Disjoint (weightSpace M χ) (s.sup fun (χ : L → R) ↦ weightSpace M χ) by - simpa only [CompleteLattice.independent_iff_supIndep_of_injOn (injOn_weightSpace R L M), + Disjoint (genWeightSpace M χ) (s.sup fun (χ : L → R) ↦ genWeightSpace M χ) by + simpa only [CompleteLattice.independent_iff_supIndep_of_injOn (injOn_genWeightSpace R L M), Finset.supIndep_iff_disjoint_erase] using fun s χ _ ↦ this _ _ (s.not_mem_erase χ) intro χ₁ s induction' s using Finset.induction_on with χ₂ s _ ih @@ -655,22 +659,22 @@ lemma independent_weightSpace [NoZeroSMulDivisors R M] : rw [Finset.sup_insert, disjoint_iff, LieSubmodule.eq_bot_iff] rintro x ⟨hx, hx'⟩ simp only [SetLike.mem_coe, LieSubmodule.mem_coeSubmodule] at hx hx' - suffices x ∈ weightSpace M χ₂ by - rw [← LieSubmodule.mem_bot (R := R) (L := L), ← (disjoint_weightSpace R L M hχ₁₂).eq_bot] + suffices x ∈ genWeightSpace M χ₂ by + rw [← LieSubmodule.mem_bot (R := R) (L := L), ← (disjoint_genWeightSpace R L M hχ₁₂).eq_bot] exact ⟨hx, this⟩ obtain ⟨y, hy, z, hz, rfl⟩ := (LieSubmodule.mem_sup _ _ _).mp hx'; clear hx' suffices ∀ l, ∃ (k : ℕ), ((toEnd R L M l - algebraMap R (Module.End R M) (χ₂ l)) ^ k) (y + z) ∈ - weightSpace M χ₁ ⊓ Finset.sup s fun χ ↦ weightSpace M χ by - simpa only [ih.eq_bot, LieSubmodule.mem_bot, mem_weightSpace] using this + genWeightSpace M χ₁ ⊓ Finset.sup s fun χ ↦ genWeightSpace M χ by + simpa only [ih.eq_bot, LieSubmodule.mem_bot, mem_genWeightSpace] using this intro l let g : Module.End R M := toEnd R L M l - algebraMap R (Module.End R M) (χ₂ l) - obtain ⟨k, hk : (g ^ k) y = 0⟩ := (mem_weightSpace _ _ _).mp hy l + obtain ⟨k, hk : (g ^ k) y = 0⟩ := (mem_genWeightSpace _ _ _).mp hy l refine ⟨k, (LieSubmodule.mem_inf _ _ _).mp ⟨?_, ?_⟩⟩ · exact LieSubmodule.mapsTo_pow_toEnd_sub_algebraMap _ hx · rw [map_add, hk, zero_add] - suffices (s.sup fun χ ↦ weightSpace M χ : Submodule R M).map (g ^ k) ≤ - s.sup fun χ ↦ weightSpace M χ by + suffices (s.sup fun χ ↦ genWeightSpace M χ : Submodule R M).map (g ^ k) ≤ + s.sup fun χ ↦ genWeightSpace M χ by refine this (Submodule.mem_map_of_mem ?_) simp_rw [← LieSubmodule.mem_coeSubmodule, Finset.sup_eq_iSup, LieSubmodule.iSup_coe_toSubmodule, ← Finset.sup_eq_iSup] at hz @@ -681,29 +685,29 @@ lemma independent_weightSpace [NoZeroSMulDivisors R M] : rintro - ⟨u, hu, rfl⟩ exact LieSubmodule.mapsTo_pow_toEnd_sub_algebraMap _ hu -lemma independent_weightSpace' [NoZeroSMulDivisors R M] : - CompleteLattice.Independent fun χ : Weight R L M ↦ weightSpace M χ := - (independent_weightSpace R L M).comp <| +lemma independent_genWeightSpace' [NoZeroSMulDivisors R M] : + CompleteLattice.Independent fun χ : Weight R L M ↦ genWeightSpace M χ := + (independent_genWeightSpace R L M).comp <| Subtype.val_injective.comp (Weight.equivSetOf R L M).injective -lemma independent_weightSpaceOf [NoZeroSMulDivisors R M] (x : L) : - CompleteLattice.Independent fun (χ : R) ↦ weightSpaceOf M χ x := by +lemma independent_genWeightSpaceOf [NoZeroSMulDivisors R M] (x : L) : + CompleteLattice.Independent fun (χ : R) ↦ genWeightSpaceOf M χ x := by rw [LieSubmodule.independent_iff_coe_toSubmodule] exact (toEnd R L M x).independent_genEigenspace -lemma finite_weightSpaceOf_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] (x : L) : - {χ : R | weightSpaceOf M χ x ≠ ⊥}.Finite := +lemma finite_genWeightSpaceOf_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] (x : L) : + {χ : R | genWeightSpaceOf M χ x ≠ ⊥}.Finite := CompleteLattice.WellFounded.finite_ne_bot_of_independent - (LieSubmodule.wellFounded_of_noetherian R L M) (independent_weightSpaceOf R L M x) + (LieSubmodule.wellFounded_of_noetherian R L M) (independent_genWeightSpaceOf R L M x) -lemma finite_weightSpace_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] : - {χ : L → R | weightSpace M χ ≠ ⊥}.Finite := +lemma finite_genWeightSpace_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] : + {χ : L → R | genWeightSpace M χ ≠ ⊥}.Finite := CompleteLattice.WellFounded.finite_ne_bot_of_independent - (LieSubmodule.wellFounded_of_noetherian R L M) (independent_weightSpace R L M) + (LieSubmodule.wellFounded_of_noetherian R L M) (independent_genWeightSpace R L M) instance Weight.instFinite [NoZeroSMulDivisors R M] [IsNoetherian R M] : Finite (Weight R L M) := by - have : Finite {χ : L → R | weightSpace M χ ≠ ⊥} := finite_weightSpace_ne_bot R L M + have : Finite {χ : L → R | genWeightSpace M χ ≠ ⊥} := finite_genWeightSpace_ne_bot R L M exact Finite.of_injective (equivSetOf R L M) (equivSetOf R L M).injective noncomputable instance Weight.instFintype [NoZeroSMulDivisors R M] [IsNoetherian R M] : @@ -725,18 +729,18 @@ instance [IsTriangularizable R L M] : IsTriangularizable R (LieModule.toEnd R L iSup_eq_top := by rintro ⟨-, x, rfl⟩; exact IsTriangularizable.iSup_eq_top x @[simp] -lemma iSup_weightSpaceOf_eq_top [IsTriangularizable R L M] (x : L) : - ⨆ (φ : R), weightSpaceOf M φ x = ⊤ := by +lemma iSup_genWeightSpaceOf_eq_top [IsTriangularizable R L M] (x : L) : + ⨆ (φ : R), genWeightSpaceOf M φ x = ⊤ := by rw [← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.iSup_coe_toSubmodule, LieSubmodule.top_coeSubmodule] exact IsTriangularizable.iSup_eq_top x open LinearMap FiniteDimensional in @[simp] -lemma trace_toEnd_weightSpace [IsDomain R] [IsPrincipalIdealRing R] +lemma trace_toEnd_genWeightSpace [IsDomain R] [IsPrincipalIdealRing R] [Module.Free R M] [Module.Finite R M] (χ : L → R) (x : L) : - trace R _ (toEnd R L (weightSpace M χ) x) = finrank R (weightSpace M χ) • χ x := by - suffices _root_.IsNilpotent ((toEnd R L (weightSpace M χ) x) - χ x • LinearMap.id) by + trace R _ (toEnd R L (genWeightSpace M χ) x) = finrank R (genWeightSpace M χ) • χ x := by + suffices _root_.IsNilpotent ((toEnd R L (genWeightSpace M χ) x) - χ x • LinearMap.id) by replace this := (isNilpotent_trace_of_isNilpotent this).eq_zero rwa [map_sub, map_smul, trace_id, sub_eq_zero, smul_eq_mul, mul_comm, ← nsmul_eq_mul] at this @@ -761,41 +765,43 @@ instance (N : LieSubmodule K L M) [IsTriangularizable K L M] : IsTriangularizabl /-- For a triangularizable Lie module in finite dimensions, the weight spaces span the entire space. -See also `LieModule.iSup_weightSpace_eq_top'`. -/ -lemma iSup_weightSpace_eq_top [IsTriangularizable K L M] : - ⨆ χ : L → K, weightSpace M χ = ⊤ := by +See also `LieModule.iSup_genWeightSpace_eq_top'`. -/ +lemma iSup_genWeightSpace_eq_top [IsTriangularizable K L M] : + ⨆ χ : L → K, genWeightSpace M χ = ⊤ := by induction' h_dim : finrank K M using Nat.strong_induction_on with n ih generalizing M - obtain h' | ⟨y : L, hy : ¬ ∃ φ, weightSpaceOf M φ y = ⊤⟩ := - forall_or_exists_not (fun (x : L) ↦ ∃ (φ : K), weightSpaceOf M φ x = ⊤) + obtain h' | ⟨y : L, hy : ¬ ∃ φ, genWeightSpaceOf M φ y = ⊤⟩ := + forall_or_exists_not (fun (x : L) ↦ ∃ (φ : K), genWeightSpaceOf M φ x = ⊤) · choose χ hχ using h' - replace hχ : weightSpace M χ = ⊤ := by simpa only [weightSpace, hχ] using iInf_top - exact eq_top_iff.mpr <| hχ ▸ le_iSup (weightSpace M) χ - · replace hy : ∀ φ, finrank K (weightSpaceOf M φ y) < n := fun φ ↦ by + replace hχ : genWeightSpace M χ = ⊤ := by simpa only [genWeightSpace, hχ] using iInf_top + exact eq_top_iff.mpr <| hχ ▸ le_iSup (genWeightSpace M) χ + · replace hy : ∀ φ, finrank K (genWeightSpaceOf M φ y) < n := fun φ ↦ by simp_rw [not_exists, ← lt_top_iff_ne_top] at hy; exact h_dim ▸ Submodule.finrank_lt (hy φ) - replace ih : ∀ φ, ⨆ χ : L → K, weightSpace (weightSpaceOf M φ y) χ = ⊤ := - fun φ ↦ ih _ (hy φ) (weightSpaceOf M φ y) rfl - replace ih : ∀ φ, ⨆ (χ : L → K) (_ : χ y = φ), weightSpace (weightSpaceOf M φ y) χ = ⊤ := by + replace ih : ∀ φ, ⨆ χ : L → K, genWeightSpace (genWeightSpaceOf M φ y) χ = ⊤ := + fun φ ↦ ih _ (hy φ) (genWeightSpaceOf M φ y) rfl + replace ih : ∀ φ, ⨆ (χ : L → K) (_ : χ y = φ), + genWeightSpace (genWeightSpaceOf M φ y) χ = ⊤ := by intro φ - suffices ∀ χ : L → K, χ y ≠ φ → weightSpace (weightSpaceOf M φ y) χ = ⊥ by + suffices ∀ χ : L → K, χ y ≠ φ → genWeightSpace (genWeightSpaceOf M φ y) χ = ⊥ by specialize ih φ; rw [iSup_split, biSup_congr this] at ih; simpa using ih intro χ hχ - rw [eq_bot_iff, ← (weightSpaceOf M φ y).ker_incl, LieModuleHom.ker, - ← LieSubmodule.map_le_iff_le_comap, map_weightSpace_eq_of_injective - (weightSpaceOf M φ y).injective_incl, LieSubmodule.range_incl, ← disjoint_iff_inf_le] - exact (disjoint_weightSpaceOf K L M hχ).mono_left (weightSpace_le_weightSpaceOf M y χ) - replace ih : ∀ φ, ⨆ (χ : L → K) (_ : χ y = φ), weightSpace M χ = weightSpaceOf M φ y := by + rw [eq_bot_iff, ← (genWeightSpaceOf M φ y).ker_incl, LieModuleHom.ker, + ← LieSubmodule.map_le_iff_le_comap, map_genWeightSpace_eq_of_injective + (genWeightSpaceOf M φ y).injective_incl, LieSubmodule.range_incl, ← disjoint_iff_inf_le] + exact (disjoint_genWeightSpaceOf K L M hχ).mono_left + (genWeightSpace_le_genWeightSpaceOf M y χ) + replace ih : ∀ φ, ⨆ (χ : L → K) (_ : χ y = φ), genWeightSpace M χ = genWeightSpaceOf M φ y := by intro φ - have : ∀ (χ : L → K) (_ : χ y = φ), weightSpace M χ = - (weightSpace (weightSpaceOf M φ y) χ).map (weightSpaceOf M φ y).incl := fun χ hχ ↦ by - rw [← hχ, weightSpace_weightSpaceOf_map_incl] + have (χ : L → K) (hχ : χ y = φ) : genWeightSpace M χ = + (genWeightSpace (genWeightSpaceOf M φ y) χ).map (genWeightSpaceOf M φ y).incl := by + rw [← hχ, genWeightSpace_genWeightSpaceOf_map_incl] simp_rw [biSup_congr this, ← LieSubmodule.map_iSup, ih, LieModuleHom.map_top, LieSubmodule.range_incl] simpa only [← ih, iSup_comm (ι := K), iSup_iSup_eq_right] using - iSup_weightSpaceOf_eq_top K L M y + iSup_genWeightSpaceOf_eq_top K L M y -lemma iSup_weightSpace_eq_top' [IsTriangularizable K L M] : - ⨆ χ : Weight K L M, weightSpace M χ = ⊤ := by - have := iSup_weightSpace_eq_top K L M +lemma iSup_genWeightSpace_eq_top' [IsTriangularizable K L M] : + ⨆ χ : Weight K L M, genWeightSpace M χ = ⊤ := by + have := iSup_genWeightSpace_eq_top K L M erw [← iSup_ne_bot_subtype, ← (Weight.equivSetOf K L M).iSup_comp] at this exact this diff --git a/Mathlib/Algebra/Lie/Weights/Cartan.lean b/Mathlib/Algebra/Lie/Weights/Cartan.lean index a78eaf08e3214..6dbc65cb2e75f 100644 --- a/Mathlib/Algebra/Lie/Weights/Cartan.lean +++ b/Mathlib/Algebra/Lie/Weights/Cartan.lean @@ -42,38 +42,39 @@ open TensorProduct.LieModule LieModule /-- Given a nilpotent Lie subalgebra `H ⊆ L`, the root space of a map `χ : H → R` is the weight space of `L` regarded as a module of `H` via the adjoint action. -/ abbrev rootSpace (χ : H → R) : LieSubmodule R H L := - weightSpace L χ + genWeightSpace L χ theorem zero_rootSpace_eq_top_of_nilpotent [IsNilpotent R L] : rootSpace (⊤ : LieSubalgebra R L) 0 = ⊤ := - zero_weightSpace_eq_top_of_nilpotent L + zero_genWeightSpace_eq_top_of_nilpotent L @[simp] -theorem rootSpace_comap_eq_weightSpace (χ : H → R) : - (rootSpace H χ).comap H.incl' = weightSpace H χ := - comap_weightSpace_eq_of_injective Subtype.coe_injective +theorem rootSpace_comap_eq_genWeightSpace (χ : H → R) : + (rootSpace H χ).comap H.incl' = genWeightSpace H χ := + comap_genWeightSpace_eq_of_injective Subtype.coe_injective variable {H} -theorem lie_mem_weightSpace_of_mem_weightSpace {χ₁ χ₂ : H → R} {x : L} {m : M} - (hx : x ∈ rootSpace H χ₁) (hm : m ∈ weightSpace M χ₂) : ⁅x, m⁆ ∈ weightSpace M (χ₁ + χ₂) := by - rw [weightSpace, LieSubmodule.mem_iInf] +theorem lie_mem_genWeightSpace_of_mem_genWeightSpace {χ₁ χ₂ : H → R} {x : L} {m : M} + (hx : x ∈ rootSpace H χ₁) (hm : m ∈ genWeightSpace M χ₂) : + ⁅x, m⁆ ∈ genWeightSpace M (χ₁ + χ₂) := by + rw [genWeightSpace, LieSubmodule.mem_iInf] intro y - replace hx : x ∈ weightSpaceOf L (χ₁ y) y := by - rw [rootSpace, weightSpace, LieSubmodule.mem_iInf] at hx; exact hx y - replace hm : m ∈ weightSpaceOf M (χ₂ y) y := by - rw [weightSpace, LieSubmodule.mem_iInf] at hm; exact hm y + replace hx : x ∈ genWeightSpaceOf L (χ₁ y) y := by + rw [rootSpace, genWeightSpace, LieSubmodule.mem_iInf] at hx; exact hx y + replace hm : m ∈ genWeightSpaceOf M (χ₂ y) y := by + rw [genWeightSpace, LieSubmodule.mem_iInf] at hm; exact hm y exact lie_mem_maxGenEigenspace_toEnd hx hm lemma toEnd_pow_apply_mem {χ₁ χ₂ : H → R} {x : L} {m : M} - (hx : x ∈ rootSpace H χ₁) (hm : m ∈ weightSpace M χ₂) (n) : - (toEnd R L M x ^ n : Module.End R M) m ∈ weightSpace M (n • χ₁ + χ₂) := by + (hx : x ∈ rootSpace H χ₁) (hm : m ∈ genWeightSpace M χ₂) (n) : + (toEnd R L M x ^ n : Module.End R M) m ∈ genWeightSpace M (n • χ₁ + χ₂) := by induction n with | zero => simpa using hm | succ n IH => simp only [pow_succ', LinearMap.mul_apply, toEnd_apply_apply, Nat.cast_add, Nat.cast_one, rootSpace] - convert lie_mem_weightSpace_of_mem_weightSpace hx IH using 2 + convert lie_mem_genWeightSpace_of_mem_genWeightSpace hx IH using 2 rw [succ_nsmul, ← add_assoc, add_comm (n • _)] variable (R L H M) @@ -82,10 +83,11 @@ variable (R L H M) which is close to the deterministic timeout limit. -/ def rootSpaceWeightSpaceProductAux {χ₁ χ₂ χ₃ : H → R} (hχ : χ₁ + χ₂ = χ₃) : - rootSpace H χ₁ →ₗ[R] weightSpace M χ₂ →ₗ[R] weightSpace M χ₃ where + rootSpace H χ₁ →ₗ[R] genWeightSpace M χ₂ →ₗ[R] genWeightSpace M χ₃ where toFun x := { toFun := fun m => - ⟨⁅(x : L), (m : M)⁆, hχ ▸ lie_mem_weightSpace_of_mem_weightSpace x.property m.property⟩ + ⟨⁅(x : L), (m : M)⁆, + hχ ▸ lie_mem_genWeightSpace_of_mem_genWeightSpace x.property m.property⟩ map_add' := fun m n => by simp only [LieSubmodule.coe_add, lie_add]; rfl map_smul' := fun t m => by dsimp only @@ -108,8 +110,8 @@ def rootSpaceWeightSpaceProductAux {χ₁ χ₂ χ₃ : H → R} (hχ : χ₁ + /-- Given a nilpotent Lie subalgebra `H ⊆ L` together with `χ₁ χ₂ : H → R`, there is a natural `R`-bilinear product of root vectors and weight vectors, compatible with the actions of `H`. -/ def rootSpaceWeightSpaceProduct (χ₁ χ₂ χ₃ : H → R) (hχ : χ₁ + χ₂ = χ₃) : - rootSpace H χ₁ ⊗[R] weightSpace M χ₂ →ₗ⁅R,H⁆ weightSpace M χ₃ := - liftLie R H (rootSpace H χ₁) (weightSpace M χ₂) (weightSpace M χ₃) + rootSpace H χ₁ ⊗[R] genWeightSpace M χ₂ →ₗ⁅R,H⁆ genWeightSpace M χ₃ := + liftLie R H (rootSpace H χ₁) (genWeightSpace M χ₂) (genWeightSpace M χ₃) { toLinearMap := rootSpaceWeightSpaceProductAux R L H M hχ map_lie' := fun {x y} => by ext m @@ -119,18 +121,18 @@ def rootSpaceWeightSpaceProduct (χ₁ χ₂ χ₃ : H → R) (hχ : χ₁ + χ @[simp] theorem coe_rootSpaceWeightSpaceProduct_tmul (χ₁ χ₂ χ₃ : H → R) (hχ : χ₁ + χ₂ = χ₃) - (x : rootSpace H χ₁) (m : weightSpace M χ₂) : + (x : rootSpace H χ₁) (m : genWeightSpace M χ₂) : (rootSpaceWeightSpaceProduct R L H M χ₁ χ₂ χ₃ hχ (x ⊗ₜ m) : M) = ⁅(x : L), (m : M)⁆ := by simp only [rootSpaceWeightSpaceProduct, rootSpaceWeightSpaceProductAux, coe_liftLie_eq_lift_coe, AddHom.toFun_eq_coe, LinearMap.coe_toAddHom, lift_apply, LinearMap.coe_mk, AddHom.coe_mk, Submodule.coe_mk] -theorem mapsTo_toEnd_weightSpace_add_of_mem_rootSpace (α χ : H → R) +theorem mapsTo_toEnd_genWeightSpace_add_of_mem_rootSpace (α χ : H → R) {x : L} (hx : x ∈ rootSpace H α) : - MapsTo (toEnd R L M x) (weightSpace M χ) (weightSpace M (α + χ)) := by + MapsTo (toEnd R L M x) (genWeightSpace M χ) (genWeightSpace M (α + χ)) := by intro m hm let x' : rootSpace H α := ⟨x, hx⟩ - let m' : weightSpace M χ := ⟨m, hm⟩ + let m' : genWeightSpace M χ := ⟨m, hm⟩ exact (rootSpaceWeightSpaceProduct R L H M α χ (α + χ) rfl (x' ⊗ₜ m')).property /-- Given a nilpotent Lie subalgebra `H ⊆ L` together with `χ₁ χ₂ : H → R`, there is a natural @@ -163,12 +165,12 @@ theorem coe_zeroRootSubalgebra : (zeroRootSubalgebra R L H : Submodule R L) = ro theorem mem_zeroRootSubalgebra (x : L) : x ∈ zeroRootSubalgebra R L H ↔ ∀ y : H, ∃ k : ℕ, (toEnd R H L y ^ k) x = 0 := by change x ∈ rootSpace H 0 ↔ _ - simp only [mem_weightSpace, Pi.zero_apply, zero_smul, sub_zero] + simp only [mem_genWeightSpace, Pi.zero_apply, zero_smul, sub_zero] theorem toLieSubmodule_le_rootSpace_zero : H.toLieSubmodule ≤ rootSpace H 0 := by intro x hx simp only [LieSubalgebra.mem_toLieSubmodule] at hx - simp only [mem_weightSpace, Pi.zero_apply, sub_zero, zero_smul] + simp only [mem_genWeightSpace, Pi.zero_apply, sub_zero, zero_smul] intro y obtain ⟨k, hk⟩ := (inferInstance : IsNilpotent R H) use k @@ -189,7 +191,7 @@ theorem toLieSubmodule_le_rootSpace_zero : H.toLieSubmodule ≤ rootSpace H 0 := exact h /-- This enables the instance `Zero (Weight R H L)`. -/ -instance [Nontrivial H] : Nontrivial (weightSpace L (0 : H → R)) := by +instance [Nontrivial H] : Nontrivial (genWeightSpace L (0 : H → R)) := by obtain ⟨⟨x, hx⟩, ⟨y, hy⟩, e⟩ := exists_pair_ne H exact ⟨⟨x, toLieSubmodule_le_rootSpace_zero R L H hx⟩, ⟨y, toLieSubmodule_le_rootSpace_zero R L H hy⟩, by simpa using e⟩ diff --git a/Mathlib/Algebra/Lie/Weights/Chain.lean b/Mathlib/Algebra/Lie/Weights/Chain.lean index beb8dcaeaa05e..e8a85aaf3e5ea 100644 --- a/Mathlib/Algebra/Lie/Weights/Chain.lean +++ b/Mathlib/Algebra/Lie/Weights/Chain.lean @@ -25,12 +25,12 @@ We provide basic definitions and results to support `α`-chain techniques in thi ## Main definitions / results - * `LieModule.exists₂_weightSpace_smul_add_eq_bot`: given weights `χ₁`, `χ₂` if `χ₁ ≠ 0`, we can + * `LieModule.exists₂_genWeightSpace_smul_add_eq_bot`: given weights `χ₁`, `χ₂` if `χ₁ ≠ 0`, we can find `p < 0` and `q > 0` such that the weight spaces `p • χ₁ + χ₂` and `q • χ₁ + χ₂` are both trivial. - * `LieModule.weightSpaceChain`: given weights `χ₁`, `χ₂` together with integers `p` and `q`, this - is the sum of the weight spaces `k • χ₁ + χ₂` for `p < k < q`. - * `LieModule.trace_toEnd_weightSpaceChain_eq_zero`: given a root `α` relative to a Cartan + * `LieModule.genWeightSpaceChain`: given weights `χ₁`, `χ₂` together with integers `p` and `q`, + this is the sum of the weight spaces `k • χ₁ + χ₂` for `p < k < q`. + * `LieModule.trace_toEnd_genWeightSpaceChain_eq_zero`: given a root `α` relative to a Cartan subalgebra `H`, there is a natural ideal `corootSpace α` in `H`. This lemma states that this ideal acts by trace-zero endomorphisms on the sum of root spaces of any `α`-chain, provided the weight spaces at the endpoints are both trivial. @@ -57,28 +57,28 @@ section variable [NoZeroSMulDivisors ℤ R] [NoZeroSMulDivisors R M] [IsNoetherian R M] (hχ₁ : χ₁ ≠ 0) include hχ₁ -lemma eventually_weightSpace_smul_add_eq_bot : - ∀ᶠ (k : ℕ) in Filter.atTop, weightSpace M (k • χ₁ + χ₂) = ⊥ := by +lemma eventually_genWeightSpace_smul_add_eq_bot : + ∀ᶠ (k : ℕ) in Filter.atTop, genWeightSpace M (k • χ₁ + χ₂) = ⊥ := by let f : ℕ → L → R := fun k ↦ k • χ₁ + χ₂ suffices Injective f by rw [← Nat.cofinite_eq_atTop, Filter.eventually_cofinite, ← finite_image_iff this.injOn] - apply (finite_weightSpace_ne_bot R L M).subset + apply (finite_genWeightSpace_ne_bot R L M).subset simp [f] intro k l hkl replace hkl : (k : ℤ) • χ₁ = (l : ℤ) • χ₁ := by simpa only [f, add_left_inj, natCast_zsmul] using hkl exact Nat.cast_inj.mp <| smul_left_injective ℤ hχ₁ hkl -lemma exists_weightSpace_smul_add_eq_bot : - ∃ k > 0, weightSpace M (k • χ₁ + χ₂) = ⊥ := - (Nat.eventually_pos.and <| eventually_weightSpace_smul_add_eq_bot M χ₁ χ₂ hχ₁).exists +lemma exists_genWeightSpace_smul_add_eq_bot : + ∃ k > 0, genWeightSpace M (k • χ₁ + χ₂) = ⊥ := + (Nat.eventually_pos.and <| eventually_genWeightSpace_smul_add_eq_bot M χ₁ χ₂ hχ₁).exists -lemma exists₂_weightSpace_smul_add_eq_bot : +lemma exists₂_genWeightSpace_smul_add_eq_bot : ∃ᵉ (p < (0 : ℤ)) (q > (0 : ℤ)), - weightSpace M (p • χ₁ + χ₂) = ⊥ ∧ - weightSpace M (q • χ₁ + χ₂) = ⊥ := by - obtain ⟨q, hq₀, hq⟩ := exists_weightSpace_smul_add_eq_bot M χ₁ χ₂ hχ₁ - obtain ⟨p, hp₀, hp⟩ := exists_weightSpace_smul_add_eq_bot M (-χ₁) χ₂ (neg_ne_zero.mpr hχ₁) + genWeightSpace M (p • χ₁ + χ₂) = ⊥ ∧ + genWeightSpace M (q • χ₁ + χ₂) = ⊥ := by + obtain ⟨q, hq₀, hq⟩ := exists_genWeightSpace_smul_add_eq_bot M χ₁ χ₂ hχ₁ + obtain ⟨p, hp₀, hp⟩ := exists_genWeightSpace_smul_add_eq_bot M (-χ₁) χ₂ (neg_ne_zero.mpr hχ₁) refine ⟨-(p : ℤ), by simpa, q, by simpa, ?_, ?_⟩ · rw [neg_smul, ← smul_neg, natCast_zsmul] exact hp @@ -90,28 +90,28 @@ end /-- Given two (potential) weights `χ₁` and `χ₂` together with integers `p` and `q`, it is often useful to study the sum of weight spaces associated to the family of weights `k • χ₁ + χ₂` for `p < k < q`. -/ -def weightSpaceChain : LieSubmodule R L M := - ⨆ k ∈ Ioo p q, weightSpace M (k • χ₁ + χ₂) +def genWeightSpaceChain : LieSubmodule R L M := + ⨆ k ∈ Ioo p q, genWeightSpace M (k • χ₁ + χ₂) -lemma weightSpaceChain_def : - weightSpaceChain M χ₁ χ₂ p q = ⨆ k ∈ Ioo p q, weightSpace M (k • χ₁ + χ₂) := +lemma genWeightSpaceChain_def : + genWeightSpaceChain M χ₁ χ₂ p q = ⨆ k ∈ Ioo p q, genWeightSpace M (k • χ₁ + χ₂) := rfl -lemma weightSpaceChain_def' : - weightSpaceChain M χ₁ χ₂ p q = ⨆ k ∈ Finset.Ioo p q, weightSpace M (k • χ₁ + χ₂) := by +lemma genWeightSpaceChain_def' : + genWeightSpaceChain M χ₁ χ₂ p q = ⨆ k ∈ Finset.Ioo p q, genWeightSpace M (k • χ₁ + χ₂) := by have : ∀ (k : ℤ), k ∈ Ioo p q ↔ k ∈ Finset.Ioo p q := by simp - simp_rw [weightSpaceChain_def, this] + simp_rw [genWeightSpaceChain_def, this] @[simp] -lemma weightSpaceChain_neg : - weightSpaceChain M (-χ₁) χ₂ (-q) (-p) = weightSpaceChain M χ₁ χ₂ p q := by +lemma genWeightSpaceChain_neg : + genWeightSpaceChain M (-χ₁) χ₂ (-q) (-p) = genWeightSpaceChain M χ₁ χ₂ p q := by let e : ℤ ≃ ℤ := neg_involutive.toPerm - simp_rw [weightSpaceChain, ← e.biSup_comp (Ioo p q)] + simp_rw [genWeightSpaceChain, ← e.biSup_comp (Ioo p q)] simp [e, -mem_Ioo, neg_mem_Ioo_iff] -lemma weightSpace_le_weightSpaceChain {k : ℤ} (hk : k ∈ Ioo p q) : - weightSpace M (k • χ₁ + χ₂) ≤ weightSpaceChain M χ₁ χ₂ p q := - le_biSup (fun i ↦ weightSpace M (i • χ₁ + χ₂)) hk +lemma genWeightSpace_le_genWeightSpaceChain {k : ℤ} (hk : k ∈ Ioo p q) : + genWeightSpace M (k • χ₁ + χ₂) ≤ genWeightSpaceChain M χ₁ χ₂ p q := + le_biSup (fun i ↦ genWeightSpace M (i • χ₁ + χ₂)) hk end IsNilpotent @@ -121,56 +121,56 @@ open LieAlgebra variable {H : LieSubalgebra R L} (α χ : H → R) (p q : ℤ) -lemma lie_mem_weightSpaceChain_of_weightSpace_eq_bot_right [LieAlgebra.IsNilpotent R H] - (hq : weightSpace M (q • α + χ) = ⊥) +lemma lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_right [LieAlgebra.IsNilpotent R H] + (hq : genWeightSpace M (q • α + χ) = ⊥) {x : L} (hx : x ∈ rootSpace H α) - {y : M} (hy : y ∈ weightSpaceChain M α χ p q) : - ⁅x, y⁆ ∈ weightSpaceChain M α χ p q := by - rw [weightSpaceChain, iSup_subtype'] at hy + {y : M} (hy : y ∈ genWeightSpaceChain M α χ p q) : + ⁅x, y⁆ ∈ genWeightSpaceChain M α χ p q := by + rw [genWeightSpaceChain, iSup_subtype'] at hy induction' hy using LieSubmodule.iSup_induction' with k z hz z₁ z₂ _ _ hz₁ hz₂ · obtain ⟨k, hk⟩ := k - suffices weightSpace M ((k + 1) • α + χ) ≤ weightSpaceChain M α χ p q by + suffices genWeightSpace M ((k + 1) • α + χ) ≤ genWeightSpaceChain M α χ p q by apply this simpa using (rootSpaceWeightSpaceProduct R L H M α (k • α + χ) ((k + 1) • α + χ) (by rw [add_smul]; abel) (⟨x, hx⟩ ⊗ₜ ⟨z, hz⟩)).property - rw [weightSpaceChain] + rw [genWeightSpaceChain] rcases eq_or_ne (k + 1) q with rfl | hk'; · simp only [hq, bot_le] replace hk' : k + 1 ∈ Ioo p q := ⟨by linarith [hk.1], lt_of_le_of_ne hk.2 hk'⟩ - exact le_biSup (fun k ↦ weightSpace M (k • α + χ)) hk' + exact le_biSup (fun k ↦ genWeightSpace M (k • α + χ)) hk' · simp · rw [lie_add] exact add_mem hz₁ hz₂ -lemma lie_mem_weightSpaceChain_of_weightSpace_eq_bot_left [LieAlgebra.IsNilpotent R H] - (hp : weightSpace M (p • α + χ) = ⊥) +lemma lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_left [LieAlgebra.IsNilpotent R H] + (hp : genWeightSpace M (p • α + χ) = ⊥) {x : L} (hx : x ∈ rootSpace H (-α)) - {y : M} (hy : y ∈ weightSpaceChain M α χ p q) : - ⁅x, y⁆ ∈ weightSpaceChain M α χ p q := by - replace hp : weightSpace M ((-p) • (-α) + χ) = ⊥ := by rwa [smul_neg, neg_smul, neg_neg] - rw [← weightSpaceChain_neg] at hy ⊢ - exact lie_mem_weightSpaceChain_of_weightSpace_eq_bot_right M (-α) χ (-q) (-p) hp hx hy + {y : M} (hy : y ∈ genWeightSpaceChain M α χ p q) : + ⁅x, y⁆ ∈ genWeightSpaceChain M α χ p q := by + replace hp : genWeightSpace M ((-p) • (-α) + χ) = ⊥ := by rwa [smul_neg, neg_smul, neg_neg] + rw [← genWeightSpaceChain_neg] at hy ⊢ + exact lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_right M (-α) χ (-q) (-p) hp hx hy section IsCartanSubalgebra variable [H.IsCartanSubalgebra] [IsNoetherian R L] -lemma trace_toEnd_weightSpaceChain_eq_zero - (hp : weightSpace M (p • α + χ) = ⊥) - (hq : weightSpace M (q • α + χ) = ⊥) +lemma trace_toEnd_genWeightSpaceChain_eq_zero + (hp : genWeightSpace M (p • α + χ) = ⊥) + (hq : genWeightSpace M (q • α + χ) = ⊥) {x : H} (hx : x ∈ corootSpace α) : - LinearMap.trace R _ (toEnd R H (weightSpaceChain M α χ p q) x) = 0 := by + LinearMap.trace R _ (toEnd R H (genWeightSpaceChain M α χ p q) x) = 0 := by rw [LieAlgebra.mem_corootSpace'] at hx induction hx using Submodule.span_induction' · next u hu => obtain ⟨y, hy, z, hz, hyz⟩ := hu - let f : Module.End R (weightSpaceChain M α χ p q) := + let f : Module.End R (genWeightSpaceChain M α χ p q) := { toFun := fun ⟨m, hm⟩ ↦ ⟨⁅(y : L), m⁆, - lie_mem_weightSpaceChain_of_weightSpace_eq_bot_right M α χ p q hq hy hm⟩ + lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_right M α χ p q hq hy hm⟩ map_add' := fun _ _ ↦ by simp map_smul' := fun t m ↦ by simp } - let g : Module.End R (weightSpaceChain M α χ p q) := + let g : Module.End R (genWeightSpaceChain M α χ p q) := { toFun := fun ⟨m, hm⟩ ↦ ⟨⁅(z : L), m⁆, - lie_mem_weightSpaceChain_of_weightSpace_eq_bot_left M α χ p q hp hz hm⟩ + lie_mem_genWeightSpaceChain_of_genWeightSpace_eq_bot_left M α χ p q hp hz hm⟩ map_add' := fun _ _ ↦ by simp map_smul' := fun t m ↦ by simp } have hfg : toEnd R H _ u = ⁅f, g⁆ := by ext; simp [f, g, ← hyz] @@ -188,30 +188,30 @@ semisimple Lie algebra form a root system. It shows that the restriction of `α` the restriction of every root to `I` vanishes (which cannot happen in a semisimple Lie algebra). -/ lemma exists_forall_mem_corootSpace_smul_add_eq_zero [IsDomain R] [IsPrincipalIdealRing R] [CharZero R] [NoZeroSMulDivisors R M] [IsNoetherian R M] - (hα : α ≠ 0) (hχ : weightSpace M χ ≠ ⊥) : + (hα : α ≠ 0) (hχ : genWeightSpace M χ ≠ ⊥) : ∃ a b : ℤ, 0 < b ∧ ∀ x ∈ corootSpace α, (a • α + b • χ) x = 0 := by - obtain ⟨p, hp₀, q, hq₀, hp, hq⟩ := exists₂_weightSpace_smul_add_eq_bot M α χ hα - let a := ∑ i ∈ Finset.Ioo p q, finrank R (weightSpace M (i • α + χ)) • i - let b := ∑ i ∈ Finset.Ioo p q, finrank R (weightSpace M (i • α + χ)) + obtain ⟨p, hp₀, q, hq₀, hp, hq⟩ := exists₂_genWeightSpace_smul_add_eq_bot M α χ hα + let a := ∑ i ∈ Finset.Ioo p q, finrank R (genWeightSpace M (i • α + χ)) • i + let b := ∑ i ∈ Finset.Ioo p q, finrank R (genWeightSpace M (i • α + χ)) have hb : 0 < b := by - replace hχ : Nontrivial (weightSpace M χ) := by rwa [LieSubmodule.nontrivial_iff_ne_bot] + replace hχ : Nontrivial (genWeightSpace M χ) := by rwa [LieSubmodule.nontrivial_iff_ne_bot] refine Finset.sum_pos' (fun _ _ ↦ zero_le _) ⟨0, Finset.mem_Ioo.mpr ⟨hp₀, hq₀⟩, ?_⟩ rw [zero_smul, zero_add] exact finrank_pos refine ⟨a, b, Int.ofNat_pos.mpr hb, fun x hx ↦ ?_⟩ - let N : ℤ → Submodule R M := fun k ↦ weightSpace M (k • α + χ) + 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_weightSpace R H M).comp fun i j hij ↦ ?_ + refine (independent_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₃ : weightSpaceChain M α χ p q = ⨆ i ∈ Finset.Ioo p q, N i := by - simp_rw [weightSpaceChain_def', LieSubmodule.iSup_coe_toSubmodule] - rw [← trace_toEnd_weightSpaceChain_eq_zero M α χ p q hp hq hx, + have h₃ : genWeightSpaceChain M α χ p q = ⨆ i ∈ Finset.Ioo p q, N i := by + simp_rw [genWeightSpaceChain_def', LieSubmodule.iSup_coe_toSubmodule] + rw [← trace_toEnd_genWeightSpaceChain_eq_zero M α χ p q hp hq hx, ← LieSubmodule.toEnd_restrict_eq_toEnd, - LinearMap.trace_eq_sum_trace_restrict_of_eq_biSup _ h₁ h₂ (weightSpaceChain M α χ p q) h₃] + LinearMap.trace_eq_sum_trace_restrict_of_eq_biSup _ h₁ h₂ (genWeightSpaceChain M α χ p q) h₃] simp_rw [LieSubmodule.toEnd_restrict_eq_toEnd, - trace_toEnd_weightSpace, Pi.add_apply, Pi.smul_apply, smul_add, ← smul_assoc, + trace_toEnd_genWeightSpace, Pi.add_apply, Pi.smul_apply, smul_add, ← smul_assoc, Finset.sum_add_distrib, ← Finset.sum_smul, natCast_zsmul] end IsCartanSubalgebra @@ -230,8 +230,8 @@ noncomputable def chainTopCoeff : ℕ := letI := Classical.propDecidable if hα : α = 0 then 0 else - Nat.pred <| Nat.find (show ∃ n, weightSpace M (n • α + β : L → R) = ⊥ from - (eventually_weightSpace_smul_add_eq_bot M α β hα).exists) + Nat.pred <| Nat.find (show ∃ n, genWeightSpace M (n • α + β : L → R) = ⊥ from + (eventually_genWeightSpace_smul_add_eq_bot M α β hα).exists) /-- This is the largest `n : ℕ` such that `-i • α + β` is a weight for all `0 ≤ i ≤ n`. -/ noncomputable @@ -251,66 +251,66 @@ include hα lemma chainTopCoeff_add_one : letI := Classical.propDecidable chainTopCoeff α β + 1 = - Nat.find (eventually_weightSpace_smul_add_eq_bot M α β hα).exists := by + Nat.find (eventually_genWeightSpace_smul_add_eq_bot M α β hα).exists := by classical rw [chainTopCoeff, dif_neg hα] apply Nat.succ_pred_eq_of_pos rw [zero_lt_iff] intro e - have : weightSpace M (0 • α + β : L → R) = ⊥ := by + have : genWeightSpace M (0 • α + β : L → R) = ⊥ := by rw [← e] - exact Nat.find_spec (eventually_weightSpace_smul_add_eq_bot M α β hα).exists - exact β.weightSpace_ne_bot _ (by simpa only [zero_smul, zero_add] using this) + exact Nat.find_spec (eventually_genWeightSpace_smul_add_eq_bot M α β hα).exists + exact β.genWeightSpace_ne_bot _ (by simpa only [zero_smul, zero_add] using this) -lemma weightSpace_chainTopCoeff_add_one_nsmul_add : - weightSpace M ((chainTopCoeff α β + 1) • α + β : L → R) = ⊥ := by +lemma genWeightSpace_chainTopCoeff_add_one_nsmul_add : + genWeightSpace M ((chainTopCoeff α β + 1) • α + β : L → R) = ⊥ := by classical rw [chainTopCoeff_add_one _ _ hα] - exact Nat.find_spec (eventually_weightSpace_smul_add_eq_bot M α β hα).exists + exact Nat.find_spec (eventually_genWeightSpace_smul_add_eq_bot M α β hα).exists -lemma weightSpace_chainTopCoeff_add_one_zsmul_add : - weightSpace M ((chainTopCoeff α β + 1 : ℤ) • α + β : L → R) = ⊥ := by - rw [← weightSpace_chainTopCoeff_add_one_nsmul_add α β hα, ← Nat.cast_smul_eq_nsmul ℤ, +lemma genWeightSpace_chainTopCoeff_add_one_zsmul_add : + genWeightSpace M ((chainTopCoeff α β + 1 : ℤ) • α + β : L → R) = ⊥ := by + rw [← genWeightSpace_chainTopCoeff_add_one_nsmul_add α β hα, ← Nat.cast_smul_eq_nsmul ℤ, Nat.cast_add, Nat.cast_one] -lemma weightSpace_chainBotCoeff_sub_one_zsmul_sub : - weightSpace M ((-chainBotCoeff α β - 1 : ℤ) • α + β : L → R) = ⊥ := by +lemma genWeightSpace_chainBotCoeff_sub_one_zsmul_sub : + genWeightSpace M ((-chainBotCoeff α β - 1 : ℤ) • α + β : L → R) = ⊥ := by rw [sub_eq_add_neg, ← neg_add, neg_smul, ← smul_neg, chainBotCoeff, - weightSpace_chainTopCoeff_add_one_zsmul_add _ _ (by simpa using hα)] + genWeightSpace_chainTopCoeff_add_one_zsmul_add _ _ (by simpa using hα)] end -lemma weightSpace_nsmul_add_ne_bot_of_le {n} (hn : n ≤ chainTopCoeff α β) : - weightSpace M (n • α + β : L → R) ≠ ⊥ := by +lemma genWeightSpace_nsmul_add_ne_bot_of_le {n} (hn : n ≤ chainTopCoeff α β) : + genWeightSpace M (n • α + β : L → R) ≠ ⊥ := by by_cases hα : α = 0 - · rw [hα, smul_zero, zero_add]; exact β.weightSpace_ne_bot + · rw [hα, smul_zero, zero_add]; exact β.genWeightSpace_ne_bot classical rw [← Nat.lt_succ, Nat.succ_eq_add_one, chainTopCoeff_add_one _ _ hα] at hn - exact Nat.find_min (eventually_weightSpace_smul_add_eq_bot M α β hα).exists hn + exact Nat.find_min (eventually_genWeightSpace_smul_add_eq_bot M α β hα).exists hn -lemma weightSpace_zsmul_add_ne_bot {n : ℤ} +lemma genWeightSpace_zsmul_add_ne_bot {n : ℤ} (hn : -chainBotCoeff α β ≤ n) (hn' : n ≤ chainTopCoeff α β) : - weightSpace M (n • α + β : L → R) ≠ ⊥ := by + genWeightSpace M (n • α + β : L → R) ≠ ⊥ := by rcases n with (n | n) · simp only [Int.ofNat_eq_coe, Nat.cast_le, Nat.cast_smul_eq_nsmul] at hn' ⊢ - exact weightSpace_nsmul_add_ne_bot_of_le α β hn' + exact genWeightSpace_nsmul_add_ne_bot_of_le α β hn' · simp only [Int.negSucc_eq, ← Nat.cast_succ, neg_le_neg_iff, Nat.cast_le] at hn ⊢ rw [neg_smul, ← smul_neg, Nat.cast_smul_eq_nsmul] - exact weightSpace_nsmul_add_ne_bot_of_le (-α) β hn + exact genWeightSpace_nsmul_add_ne_bot_of_le (-α) β hn -lemma weightSpace_neg_zsmul_add_ne_bot {n : ℕ} (hn : n ≤ chainBotCoeff α β) : - weightSpace M ((-n : ℤ) • α + β : L → R) ≠ ⊥ := by - apply weightSpace_zsmul_add_ne_bot α β <;> omega +lemma genWeightSpace_neg_zsmul_add_ne_bot {n : ℕ} (hn : n ≤ chainBotCoeff α β) : + genWeightSpace M ((-n : ℤ) • α + β : L → R) ≠ ⊥ := by + apply genWeightSpace_zsmul_add_ne_bot α β <;> omega /-- The last weight in an `α`-chain through `β`. -/ noncomputable def chainTop (α : L → R) (β : Weight R L M) : Weight R L M := - ⟨chainTopCoeff α β • α + β, weightSpace_nsmul_add_ne_bot_of_le α β le_rfl⟩ + ⟨chainTopCoeff α β • α + β, genWeightSpace_nsmul_add_ne_bot_of_le α β le_rfl⟩ /-- The first weight in an `α`-chain through `β`. -/ noncomputable def chainBot (α : L → R) (β : Weight R L M) : Weight R L M := - ⟨(- chainBotCoeff α β : ℤ) • α + β, weightSpace_neg_zsmul_add_ne_bot α β le_rfl⟩ + ⟨(- chainBotCoeff α β : ℤ) • α + β, genWeightSpace_neg_zsmul_add_ne_bot α β le_rfl⟩ lemma coe_chainTop' : (chainTop α β : L → R) = chainTopCoeff α β • α + β := rfl @@ -328,19 +328,20 @@ section variable (hα : α ≠ 0) include hα -lemma weightSpace_add_chainTop : - weightSpace M (α + chainTop α β : L → R) = ⊥ := by - rw [coe_chainTop', ← add_assoc, ← succ_nsmul', weightSpace_chainTopCoeff_add_one_nsmul_add _ _ hα] +lemma genWeightSpace_add_chainTop : + genWeightSpace M (α + chainTop α β : L → R) = ⊥ := by + rw [coe_chainTop', ← add_assoc, ← succ_nsmul', + genWeightSpace_chainTopCoeff_add_one_nsmul_add _ _ hα] -lemma weightSpace_neg_add_chainBot : - weightSpace M (-α + chainBot α β : L → R) = ⊥ := by - rw [← chainTop_neg, weightSpace_add_chainTop _ _ (by simpa using hα)] +lemma genWeightSpace_neg_add_chainBot : + genWeightSpace M (-α + chainBot α β : L → R) = ⊥ := by + rw [← chainTop_neg, genWeightSpace_add_chainTop _ _ (by simpa using hα)] -lemma chainTop_isNonZero' (hα' : weightSpace M α ≠ ⊥) : +lemma chainTop_isNonZero' (hα' : genWeightSpace M α ≠ ⊥) : (chainTop α β).IsNonZero := by by_contra e apply hα' - rw [← add_zero (α : L → R), ← e, weightSpace_add_chainTop _ _ hα] + rw [← add_zero (α : L → R), ← e, genWeightSpace_add_chainTop _ _ hα] end diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean index cee61eb9c3bc1..03e0421bfd2dc 100644 --- a/Mathlib/Algebra/Lie/Weights/Killing.lean +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -52,7 +52,7 @@ lemma ker_restrict_eq_bot_of_isCartanSubalgebra [IsNoetherian R L] [IsArtinian R L] (H : LieSubalgebra R L) [H.IsCartanSubalgebra] : LinearMap.ker ((killingForm R L).restrict H) = ⊥ := by have h : Codisjoint (rootSpace H 0) (LieModule.posFittingComp R H L) := - (LieModule.isCompl_weightSpace_zero_posFittingComp R H L).codisjoint + (LieModule.isCompl_genWeightSpace_zero_posFittingComp R H L).codisjoint replace h : Codisjoint (H : Submodule R L) (LieModule.posFittingComp R H L : Submodule R L) := by rwa [codisjoint_iff, ← LieSubmodule.coe_toSubmodule_eq_iff, LieSubmodule.sup_coe_toSubmodule, LieSubmodule.top_coeSubmodule, rootSpace_zero_eq R L H, LieSubalgebra.coe_toLieSubmodule, @@ -109,12 +109,12 @@ lemma killingForm_apply_eq_zero_of_mem_rootSpace_of_add_ne_zero {α β : H → K have hσ : ∀ γ, σ γ ≠ γ := fun γ ↦ by simpa only [σ, ← add_assoc] using add_left_ne_self.mpr hαβ let f : Module.End K L := (ad K L x) ∘ₗ (ad K L y) have hf : ∀ γ, MapsTo f (rootSpace H γ) (rootSpace H (σ γ)) := fun γ ↦ - (mapsTo_toEnd_weightSpace_add_of_mem_rootSpace K L H L α (β + γ) hx).comp <| - mapsTo_toEnd_weightSpace_add_of_mem_rootSpace K L H L β γ hy + (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_weightSpace K H L) - (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_weightSpace_eq_top K H L) + (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_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 /-- Elements of the `α` root space which are Killing-orthogonal to the `-α` root space are @@ -125,7 +125,7 @@ lemma mem_ker_killingForm_of_mem_rootSpace_of_forall_rootSpace_neg x ∈ LinearMap.ker (killingForm K L) := by rw [LinearMap.mem_ker] ext y - have hy : y ∈ ⨆ β, rootSpace H β := by simp [iSup_weightSpace_eq_top K H L] + have hy : y ∈ ⨆ β, rootSpace H β := by simp [iSup_genWeightSpace_eq_top K H L] induction hy using LieSubmodule.iSup_induction' with | hN β y hy => by_cases hαβ : α + β = 0 @@ -197,7 +197,7 @@ lemma lie_eq_killingForm_smul_of_mem_rootSpace_of_mem_rootSpace_neg_aux apply mem_ker_killingForm_of_mem_rootSpace_of_forall_rootSpace_neg (α := (0 : H → K)) · simp only [rootSpace_zero_eq, LieSubalgebra.mem_toLieSubmodule] refine sub_mem ?_ (H.smul_mem _ α'.property) - simpa using mapsTo_toEnd_weightSpace_add_of_mem_rootSpace K L H L α (-α) heα hfα + simpa using mapsTo_toEnd_genWeightSpace_add_of_mem_rootSpace K L H L α (-α) heα hfα · intro z hz replace hz : z ∈ H := by simpa using hz have he : ⁅z, e⁆ = α ⟨z, hz⟩ • e := aux ⟨z, hz⟩ @@ -211,7 +211,7 @@ assuming `K` has characteristic zero). -/ lemma cartanEquivDual_symm_apply_mem_corootSpace (α : Weight K H L) : (cartanEquivDual H).symm α ∈ corootSpace α := by obtain ⟨e : L, he₀ : e ≠ 0, he : ∀ x, ⁅x, e⁆ = α x • e⟩ := exists_forall_lie_eq_smul K H L α - have heα : e ∈ rootSpace H α := (mem_weightSpace L α e).mpr fun x ↦ ⟨1, by simp [← he x]⟩ + have heα : e ∈ rootSpace H α := (mem_genWeightSpace L α e).mpr fun x ↦ ⟨1, by simp [← he x]⟩ obtain ⟨f, hfα, hf⟩ : ∃ f ∈ rootSpace H (-α), killingForm K L e f ≠ 0 := by contrapose! he₀ simpa using mem_ker_killingForm_of_mem_rootSpace_of_forall_rootSpace_neg K L H heα he₀ @@ -264,7 +264,7 @@ lemma isSemisimple_ad_of_mem_isCartanSubalgebra {x : L} (hx : x ∈ H) : /- Note that the semisimple part `S` is just a scalar action on each root space. -/ have aux {α : H → K} {y : L} (hy : y ∈ rootSpace H α) : S y = α x' • y := by replace hy : y ∈ (ad K L x).maxGenEigenspace (α x') := - (weightSpace_le_weightSpaceOf L x' α) hy + (genWeightSpace_le_genWeightSpaceOf L x' α) hy rw [maxGenEigenspace_eq] at hy set k := maxGenEigenspaceIndex (ad K L x) (α x') rw [apply_eq_of_mem_genEigenspace_of_comm_of_isSemisimple_of_isNilpotent_sub hy hS₀ hS hN] @@ -272,12 +272,12 @@ lemma isSemisimple_ad_of_mem_isCartanSubalgebra {x : L} (hx : x ∈ H) : have h_der (y z : L) (α β : H → K) (hy : y ∈ rootSpace H α) (hz : z ∈ rootSpace H β) : S ⁅y, z⁆ = ⁅S y, z⁆ + ⁅y, S z⁆ := by have hyz : ⁅y, z⁆ ∈ rootSpace H (α + β) := - mapsTo_toEnd_weightSpace_add_of_mem_rootSpace K L H L α β hy hz + mapsTo_toEnd_genWeightSpace_add_of_mem_rootSpace K L H L α β hy hz rw [aux hy, aux hz, aux hyz, smul_lie, lie_smul, ← add_smul, ← Pi.add_apply] /- Thus `S` is a derivation since root spaces span. -/ replace h_der (y z : L) : S ⁅y, z⁆ = ⁅S y, z⁆ + ⁅y, S z⁆ := by - have hy : y ∈ ⨆ α : H → K, rootSpace H α := by simp [iSup_weightSpace_eq_top] - have hz : z ∈ ⨆ α : H → K, rootSpace H α := by simp [iSup_weightSpace_eq_top] + have hy : y ∈ ⨆ α : H → K, rootSpace H α := by simp [iSup_genWeightSpace_eq_top] + have hz : z ∈ ⨆ α : H → K, rootSpace H α := by simp [iSup_genWeightSpace_eq_top] induction hy using LieSubmodule.iSup_induction' with | hN α y hy => induction hz using LieSubmodule.iSup_induction' with @@ -313,7 +313,7 @@ lemma isSemisimple_ad_of_mem_isCartanSubalgebra {x : L} (hx : x ∈ H) : lemma lie_eq_smul_of_mem_rootSpace {α : H → K} {x : L} (hx : x ∈ rootSpace H α) (h : H) : ⁅h, x⁆ = α h • x := by replace hx : x ∈ (ad K L h).maxGenEigenspace (α h) := - weightSpace_le_weightSpaceOf L h α hx + genWeightSpace_le_genWeightSpaceOf L h α hx rw [(isSemisimple_ad_of_mem_isCartanSubalgebra h.property).maxGenEigenspace_eq_eigenspace, Module.End.mem_eigenspace_iff] at hx simpa using hx @@ -360,7 +360,7 @@ lemma eq_zero_of_apply_eq_zero_of_mem_corootSpace replace hx : x ∈ ⨅ β : Weight K H L, β.ker := by refine (Submodule.mem_iInf _).mpr fun β ↦ ?_ obtain ⟨a, b, hb, hab⟩ := - exists_forall_mem_corootSpace_smul_add_eq_zero L α β hα β.weightSpace_ne_bot + exists_forall_mem_corootSpace_smul_add_eq_zero L α β hα β.genWeightSpace_ne_bot simpa [hαx, hb.ne'] using hab _ hx simpa using hx @@ -528,7 +528,7 @@ lemma _root_.IsSl2Triple.h_eq_coroot {α : Weight K H L} (hα : α.IsNonZero) lemma finrank_rootSpace_eq_one (α : Weight K H L) (hα : α.IsNonZero) : finrank K (rootSpace H α) = 1 := by suffices ¬ 1 < finrank K (rootSpace H α) by - have h₀ : finrank K (rootSpace H α) ≠ 0 := by simpa using α.weightSpace_ne_bot + have h₀ : finrank K (rootSpace H α) ≠ 0 := by simpa using α.genWeightSpace_ne_bot omega intro contra obtain ⟨h, e, f, ht, heα, hfα⟩ := exists_isSl2Triple_of_weight_isNonZero hα @@ -571,7 +571,7 @@ variable {α : Weight K H L} instance : InvolutiveNeg (Weight K H L) where neg α := ⟨-α, by by_cases hα : α.IsZero - · convert α.weightSpace_ne_bot; rw [hα, neg_zero] + · convert α.genWeightSpace_ne_bot; rw [hα, neg_zero] · intro e obtain ⟨x, hx, x_ne0⟩ := α.exists_ne_zero have := mem_ker_killingForm_of_mem_rootSpace_of_forall_rootSpace_neg K L H hx diff --git a/Mathlib/Algebra/Lie/Weights/Linear.lean b/Mathlib/Algebra/Lie/Weights/Linear.lean index 39d847bba73c0..dc963d3db3a9c 100644 --- a/Mathlib/Algebra/Lie/Weights/Linear.lean +++ b/Mathlib/Algebra/Lie/Weights/Linear.lean @@ -12,7 +12,7 @@ import Mathlib.LinearAlgebra.FreeModule.PID Given a Lie module `M` over a nilpotent Lie algebra `L` with coefficients in `R`, one frequently studies `M` via its weights. These are functions `χ : L → R` whose corresponding weight space -`LieModule.weightSpace M χ`, is non-trivial. If `L` is Abelian or if `R` has characteristic zero +`LieModule.genWeightSpace M χ`, is non-trivial. If `L` is Abelian or if `R` has characteristic zero (and `M` is finite-dimensional) then such `χ` are necessarily `R`-linear. However in general non-linear weights do exist. For example if we take: * `R`: the field with two elements (or indeed any perfect field of characteristic two), @@ -49,9 +49,9 @@ namespace LieModule /-- A typeclass encoding the fact that a given Lie module has linear weights, vanishing on the derived ideal. -/ class LinearWeights [LieAlgebra.IsNilpotent R L] : Prop := - map_add : ∀ χ : L → R, weightSpace M χ ≠ ⊥ → ∀ x y, χ (x + y) = χ x + χ y - map_smul : ∀ χ : L → R, weightSpace M χ ≠ ⊥ → ∀ (t : R) x, χ (t • x) = t • χ x - map_lie : ∀ χ : L → R, weightSpace M χ ≠ ⊥ → ∀ x y : L, χ ⁅x, y⁆ = 0 + map_add : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ x y, χ (x + y) = χ x + χ y + map_smul : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ (t : R) x, χ (t • x) = t • χ x + map_lie : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ x y : L, χ ⁅x, y⁆ = 0 namespace Weight @@ -61,22 +61,22 @@ variable [LieAlgebra.IsNilpotent R L] [LinearWeights R L M] (χ : Weight R L M) @[simps] def toLinear : L →ₗ[R] R where toFun := χ - map_add' := LinearWeights.map_add χ χ.weightSpace_ne_bot - map_smul' := LinearWeights.map_smul χ χ.weightSpace_ne_bot + map_add' := LinearWeights.map_add χ χ.genWeightSpace_ne_bot + map_smul' := LinearWeights.map_smul χ χ.genWeightSpace_ne_bot instance instCoeLinearMap : CoeOut (Weight R L M) (L →ₗ[R] R) where coe := Weight.toLinear R L M instance instLinearMapClass : LinearMapClass (Weight R L M) R L R where - map_add χ := LinearWeights.map_add χ χ.weightSpace_ne_bot - map_smulₛₗ χ := LinearWeights.map_smul χ χ.weightSpace_ne_bot + map_add χ := LinearWeights.map_add χ χ.genWeightSpace_ne_bot + map_smulₛₗ χ := LinearWeights.map_smul χ χ.genWeightSpace_ne_bot variable {R L M χ} @[simp] lemma apply_lie (x y : L) : χ ⁅x, y⁆ = 0 := - LinearWeights.map_lie χ χ.weightSpace_ne_bot x y + LinearWeights.map_lie χ χ.genWeightSpace_ne_bot x y @[simp] lemma coe_coe : (↑(χ : L →ₗ[R] R) : L → R) = (χ : L → R) := rfl @@ -93,17 +93,17 @@ end Weight /-- For an Abelian Lie algebra, the weights of any Lie module are linear. -/ instance instLinearWeightsOfIsLieAbelian [IsLieAbelian L] [NoZeroSMulDivisors R M] : LinearWeights R L M := - have aux : ∀ (χ : L → R), weightSpace M χ ≠ ⊥ → ∀ (x y : L), χ (x + y) = χ x + χ y := by + have aux : ∀ (χ : L → R), genWeightSpace M χ ≠ ⊥ → ∀ (x y : L), χ (x + y) = χ x + χ y := by have h : ∀ x y, Commute (toEnd R L M x) (toEnd R L M y) := fun x y ↦ by rw [commute_iff_lie_eq, ← LieHom.map_lie, trivial_lie_zero, LieHom.map_zero] intro χ hχ x y - simp_rw [Ne, ← LieSubmodule.coe_toSubmodule_eq_iff, weightSpace, weightSpaceOf, + simp_rw [Ne, ← LieSubmodule.coe_toSubmodule_eq_iff, genWeightSpace, genWeightSpaceOf, LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.bot_coeSubmodule] at hχ exact Module.End.map_add_of_iInf_genEigenspace_ne_bot_of_commute (toEnd R L M).toLinearMap χ hχ h x y { map_add := aux map_smul := fun χ hχ t x ↦ by - simp_rw [Ne, ← LieSubmodule.coe_toSubmodule_eq_iff, weightSpace, weightSpaceOf, + simp_rw [Ne, ← LieSubmodule.coe_toSubmodule_eq_iff, genWeightSpace, genWeightSpaceOf, LieSubmodule.iInf_coe_toSubmodule, LieSubmodule.bot_coeSubmodule] at hχ exact Module.End.map_smul_of_iInf_genEigenspace_ne_bot (toEnd R L M).toLinearMap χ hχ t x @@ -117,23 +117,23 @@ open FiniteDimensional variable [IsDomain R] [IsPrincipalIdealRing R] [Module.Free R M] [Module.Finite R M] [LieAlgebra.IsNilpotent R L] -lemma trace_comp_toEnd_weightSpace_eq (χ : L → R) : - LinearMap.trace R _ ∘ₗ (toEnd R L (weightSpace M χ)).toLinearMap = - finrank R (weightSpace M χ) • χ := by +lemma trace_comp_toEnd_genWeightSpace_eq (χ : L → R) : + LinearMap.trace R _ ∘ₗ (toEnd R L (genWeightSpace M χ)).toLinearMap = + finrank R (genWeightSpace M χ) • χ := by ext x - let n := toEnd R L (weightSpace M χ) x - χ x • LinearMap.id - have h₁ : toEnd R L (weightSpace M χ) x = n + χ x • LinearMap.id := eq_add_of_sub_eq rfl + let n := toEnd R L (genWeightSpace M χ) x - χ x • LinearMap.id + have h₁ : toEnd R L (genWeightSpace M χ) x = n + χ x • LinearMap.id := eq_add_of_sub_eq rfl have h₂ : LinearMap.trace R _ n = 0 := IsReduced.eq_zero _ <| LinearMap.isNilpotent_trace_of_isNilpotent <| isNilpotent_toEnd_sub_algebraMap M χ x rw [LinearMap.comp_apply, LieHom.coe_toLinearMap, h₁, map_add, h₂] simp [mul_comm (χ x)] @[deprecated (since := "2024-04-06")] -alias trace_comp_toEnd_weight_space_eq := trace_comp_toEnd_weightSpace_eq +alias trace_comp_toEnd_weight_space_eq := trace_comp_toEnd_genWeightSpace_eq variable {R L M} in -lemma zero_lt_finrank_weightSpace {χ : L → R} (hχ : weightSpace M χ ≠ ⊥) : - 0 < finrank R (weightSpace M χ) := by +lemma zero_lt_finrank_genWeightSpace {χ : L → R} (hχ : genWeightSpace M χ ≠ ⊥) : + 0 < finrank R (genWeightSpace M χ) := by rwa [← LieSubmodule.nontrivial_iff_ne_bot, ← rank_pos_iff_nontrivial (R := R), ← finrank_eq_rank, Nat.cast_pos] at hχ @@ -142,90 +142,91 @@ on the derived ideal. -/ instance instLinearWeightsOfCharZero [CharZero R] : LinearWeights R L M where map_add χ hχ x y := by - rw [← smul_right_inj (zero_lt_finrank_weightSpace hχ).ne', smul_add, ← Pi.smul_apply, - ← Pi.smul_apply, ← Pi.smul_apply, ← trace_comp_toEnd_weightSpace_eq, map_add] + rw [← smul_right_inj (zero_lt_finrank_genWeightSpace hχ).ne', smul_add, ← Pi.smul_apply, + ← Pi.smul_apply, ← Pi.smul_apply, ← trace_comp_toEnd_genWeightSpace_eq, map_add] map_smul χ hχ t x := by - rw [← smul_right_inj (zero_lt_finrank_weightSpace hχ).ne', smul_comm, ← Pi.smul_apply, - ← Pi.smul_apply (finrank R _), ← trace_comp_toEnd_weightSpace_eq, map_smul] + rw [← smul_right_inj (zero_lt_finrank_genWeightSpace hχ).ne', smul_comm, ← Pi.smul_apply, + ← Pi.smul_apply (finrank R _), ← trace_comp_toEnd_genWeightSpace_eq, map_smul] map_lie χ hχ x y := by - rw [← smul_right_inj (zero_lt_finrank_weightSpace hχ).ne', nsmul_zero, ← Pi.smul_apply, - ← trace_comp_toEnd_weightSpace_eq, LinearMap.comp_apply, LieHom.coe_toLinearMap, + rw [← smul_right_inj (zero_lt_finrank_genWeightSpace hχ).ne', nsmul_zero, ← Pi.smul_apply, + ← trace_comp_toEnd_genWeightSpace_eq, LinearMap.comp_apply, LieHom.coe_toLinearMap, LieHom.map_lie, Ring.lie_def, map_sub, LinearMap.trace_mul_comm, sub_self] end FiniteDimensional variable [LieAlgebra.IsNilpotent R L] (χ : L → R) -/-- A type synonym for the `χ`-weight space but with the action of `x : L` on `m : weightSpace M χ`, -shifted to act as `⁅x, m⁆ - χ x • m`. -/ -def shiftedWeightSpace := weightSpace M χ +/-- A type synonym for the `χ`-weight space but with the action of `x : L` +on `m : genWeightSpace M χ`, shifted to act as `⁅x, m⁆ - χ x • m`. -/ +def shiftedGenWeightSpace := genWeightSpace M χ -namespace shiftedWeightSpace +namespace shiftedGenWeightSpace -private lemma aux [h : Nontrivial (shiftedWeightSpace R L M χ)] : weightSpace M χ ≠ ⊥ := +private lemma aux [h : Nontrivial (shiftedGenWeightSpace R L M χ)] : genWeightSpace M χ ≠ ⊥ := (LieSubmodule.nontrivial_iff_ne_bot _ _ _).mp h variable [LinearWeights R L M] -instance : LieRingModule L (shiftedWeightSpace R L M χ) where +instance : LieRingModule L (shiftedGenWeightSpace R L M χ) where bracket x m := ⁅x, m⁆ - χ x • m add_lie x y m := by - nontriviality shiftedWeightSpace R L M χ + nontriviality shiftedGenWeightSpace R L M χ simp only [add_lie, LinearWeights.map_add χ (aux R L M χ), add_smul] abel lie_add x m n := by - nontriviality shiftedWeightSpace R L M χ + nontriviality shiftedGenWeightSpace R L M χ simp only [lie_add, LinearWeights.map_add χ (aux R L M χ), smul_add] abel leibniz_lie x y m := by - nontriviality shiftedWeightSpace R L M χ + nontriviality shiftedGenWeightSpace R L M χ simp only [lie_sub, lie_smul, lie_lie, LinearWeights.map_lie χ (aux R L M χ), zero_smul, sub_zero, smul_sub, smul_comm (χ x)] abel -@[simp] lemma coe_lie_shiftedWeightSpace_apply (x : L) (m : shiftedWeightSpace R L M χ) : +@[simp] lemma coe_lie_shiftedGenWeightSpace_apply (x : L) (m : shiftedGenWeightSpace R L M χ) : ⁅x, m⁆ = ⁅x, (m : M)⁆ - χ x • m := rfl -instance : LieModule R L (shiftedWeightSpace R L M χ) where +instance : LieModule R L (shiftedGenWeightSpace R L M χ) where smul_lie t x m := by - nontriviality shiftedWeightSpace R L M χ + nontriviality shiftedGenWeightSpace R L M χ apply Subtype.ext - simp only [coe_lie_shiftedWeightSpace_apply, smul_lie, LinearWeights.map_smul χ (aux R L M χ), - SetLike.val_smul, smul_sub, sub_right_inj, smul_assoc t] + simp only [coe_lie_shiftedGenWeightSpace_apply, smul_lie, + LinearWeights.map_smul χ (aux R L M χ), smul_assoc t, SetLike.val_smul, smul_sub] lie_smul t x m := by - nontriviality shiftedWeightSpace R L M χ + nontriviality shiftedGenWeightSpace R L M χ apply Subtype.ext - simp only [coe_lie_shiftedWeightSpace_apply, lie_smul, LinearWeights.map_smul χ (aux R L M χ), - SetLike.val_smul, smul_sub, sub_right_inj, smul_comm t] + simp only [coe_lie_shiftedGenWeightSpace_apply, SetLike.val_smul, lie_smul, smul_sub, + smul_comm t] -/-- Forgetting the action of `L`, the spaces `weightSpace M χ` and `shiftedWeightSpace R L M χ` are -equivalent. -/ -@[simps!] def shift : weightSpace M χ ≃ₗ[R] shiftedWeightSpace R L M χ := LinearEquiv.refl R _ +/-- Forgetting the action of `L`, +the spaces `genWeightSpace M χ` and `shiftedGenWeightSpace R L M χ` are equivalent. -/ +@[simps!] def shift : genWeightSpace M χ ≃ₗ[R] shiftedGenWeightSpace R L M χ := LinearEquiv.refl R _ lemma toEnd_eq (x : L) : - toEnd R L (shiftedWeightSpace R L M χ) x = - (shift R L M χ).conj (toEnd R L (weightSpace M χ) x - χ x • LinearMap.id) := by + toEnd R L (shiftedGenWeightSpace R L M χ) x = + (shift R L M χ).conj (toEnd R L (genWeightSpace M χ) x - χ x • LinearMap.id) := by ext; simp [LinearEquiv.conj_apply] /-- By Engel's theorem, if `M` is Noetherian, the shifted action `⁅x, m⁆ - χ x • m` makes the `χ`-weight space into a nilpotent Lie module. -/ -instance [IsNoetherian R M] : IsNilpotent R L (shiftedWeightSpace R L M χ) := +instance [IsNoetherian R M] : IsNilpotent R L (shiftedGenWeightSpace R L M χ) := LieModule.isNilpotent_iff_forall'.mpr fun x ↦ isNilpotent_toEnd_sub_algebraMap M χ x -end shiftedWeightSpace +end shiftedGenWeightSpace /-- Given a Lie module `M` of a Lie algebra `L` with coefficients in `R`, if a function `χ : L → R` has a simultaneous generalized eigenvector for the action of `L` then it has a simultaneous true eigenvector, provided `M` is Noetherian and has linear weights. -/ lemma exists_forall_lie_eq_smul [LinearWeights R L M] [IsNoetherian R M] (χ : Weight R L M) : ∃ m : M, m ≠ 0 ∧ ∀ x : L, ⁅x, m⁆ = χ x • m := by - replace hχ : Nontrivial (shiftedWeightSpace R L M χ) := - (LieSubmodule.nontrivial_iff_ne_bot R L M).mpr χ.weightSpace_ne_bot + replace hχ : Nontrivial (shiftedGenWeightSpace R L M χ) := + (LieSubmodule.nontrivial_iff_ne_bot R L M).mpr χ.genWeightSpace_ne_bot obtain ⟨⟨⟨m, _⟩, hm₁⟩, hm₂⟩ := - @exists_ne _ (nontrivial_max_triv_of_isNilpotent R L (shiftedWeightSpace R L M χ)) 0 + @exists_ne _ (nontrivial_max_triv_of_isNilpotent R L (shiftedGenWeightSpace R L M χ)) 0 simp_rw [LieSubmodule.mem_coeSubmodule, mem_maxTrivSubmodule, Subtype.ext_iff, - shiftedWeightSpace.coe_lie_shiftedWeightSpace_apply, ZeroMemClass.coe_zero, sub_eq_zero] at hm₁ + shiftedGenWeightSpace.coe_lie_shiftedGenWeightSpace_apply, + ZeroMemClass.coe_zero, sub_eq_zero] at hm₁ exact ⟨m, by simpa using hm₂, hm₁⟩ end LieModule diff --git a/Mathlib/Algebra/Lie/Weights/RootSystem.lean b/Mathlib/Algebra/Lie/Weights/RootSystem.lean index e4cecabbee665..7f73bf98624fa 100644 --- a/Mathlib/Algebra/Lie/Weights/RootSystem.lean +++ b/Mathlib/Algebra/Lie/Weights/RootSystem.lean @@ -50,9 +50,9 @@ private lemma chainLength_aux (hα : α.IsNonZero) {x} (hx : x ∈ rootSpace H ( obtain ⟨h, e, f, isSl2, he, hf⟩ := exists_isSl2Triple_of_weight_isNonZero hα obtain rfl := isSl2.h_eq_coroot hα he hf have : isSl2.HasPrimitiveVectorWith x (chainTop α β (coroot α)) := - have := lie_mem_weightSpace_of_mem_weightSpace he hx + have := lie_mem_genWeightSpace_of_mem_genWeightSpace he hx ⟨hx', by rw [← lie_eq_smul_of_mem_rootSpace hx]; rfl, - by rwa [weightSpace_add_chainTop α β hα] at this⟩ + by rwa [genWeightSpace_add_chainTop α β hα] at this⟩ obtain ⟨μ, hμ⟩ := this.exists_nat exact ⟨μ, by rw [← Nat.cast_smul_eq_nsmul K, ← hμ, lie_eq_smul_of_mem_rootSpace hx]⟩ @@ -101,8 +101,8 @@ lemma rootSpace_neg_nsmul_add_chainTop_of_le {n : ℕ} (hn : n ≤ chainLength obtain ⟨h, e, f, isSl2, he, hf⟩ := exists_isSl2Triple_of_weight_isNonZero hα obtain rfl := isSl2.h_eq_coroot hα he hf have prim : isSl2.HasPrimitiveVectorWith x (chainLength α β : K) := - have := lie_mem_weightSpace_of_mem_weightSpace he hx - ⟨x_ne0, (chainLength_smul _ _ hx).symm, by rwa [weightSpace_add_chainTop _ _ hα] at this⟩ + have := lie_mem_genWeightSpace_of_mem_genWeightSpace he hx + ⟨x_ne0, (chainLength_smul _ _ hx).symm, by rwa [genWeightSpace_add_chainTop _ _ hα] at this⟩ simp only [← smul_neg, ne_eq, LieSubmodule.eq_bot_iff, not_forall] exact ⟨_, toEnd_pow_apply_mem hf hx n, prim.pow_toEnd_f_ne_zero_of_eq_nat rfl hn⟩ @@ -127,14 +127,14 @@ lemma rootSpace_neg_nsmul_add_chainTop_of_lt (hα : α.IsNonZero) {n : ℕ} (hn ring have := rootSpace_neg_nsmul_add_chainTop_of_le (-α) W H₁ rw [Weight.coe_neg, ← smul_neg, neg_neg, ← Weight.coe_neg, H₂] at this - exact this (weightSpace_chainTopCoeff_add_one_nsmul_add α β hα) + exact this (genWeightSpace_chainTopCoeff_add_one_nsmul_add α β hα) lemma chainTopCoeff_le_chainLength : chainTopCoeff α β ≤ chainLength α β := by by_cases hα : α.IsZero · simp only [hα.eq, chainTopCoeff_zero, zero_le] rw [← not_lt, ← Nat.succ_le] intro e - apply weightSpace_nsmul_add_ne_bot_of_le α β + apply genWeightSpace_nsmul_add_ne_bot_of_le α β (Nat.sub_le (chainTopCoeff α β) (chainLength α β).succ) rw [← Nat.cast_smul_eq_nsmul ℤ, Nat.cast_sub e, sub_smul, sub_eq_neg_add, add_assoc, ← coe_chainTop, Nat.cast_smul_eq_nsmul] @@ -148,7 +148,7 @@ lemma chainBotCoeff_add_chainTopCoeff : · rw [← Nat.le_sub_iff_add_le (chainTopCoeff_le_chainLength α β), ← not_lt, ← Nat.succ_le, chainBotCoeff, ← Weight.coe_neg] intro e - apply weightSpace_nsmul_add_ne_bot_of_le _ _ e + apply genWeightSpace_nsmul_add_ne_bot_of_le _ _ e rw [← Nat.cast_smul_eq_nsmul ℤ, Nat.cast_succ, Nat.cast_sub (chainTopCoeff_le_chainLength α β), LieModule.Weight.coe_neg, smul_neg, ← neg_smul, neg_add_rev, neg_sub, sub_eq_neg_add, ← add_assoc, ← neg_add_rev, add_smul, add_assoc, ← coe_chainTop, neg_smul, @@ -160,7 +160,7 @@ lemma chainBotCoeff_add_chainTopCoeff : rw [← Nat.succ_add, ← Nat.cast_smul_eq_nsmul ℤ, ← neg_smul, coe_chainTop, ← add_assoc, ← add_smul, Nat.cast_add, neg_add, add_assoc, neg_add_cancel, add_zero, neg_smul, ← smul_neg, Nat.cast_smul_eq_nsmul] - exact weightSpace_chainTopCoeff_add_one_nsmul_add (-α) β (Weight.IsNonZero.neg hα) + exact genWeightSpace_chainTopCoeff_add_one_nsmul_add (-α) β (Weight.IsNonZero.neg hα) lemma chainTopCoeff_add_chainBotCoeff : chainTopCoeff α β + chainBotCoeff α β = chainLength α β := by @@ -254,13 +254,13 @@ lemma chainTopCoeff_zero_right [Nontrivial L] (hα : α.IsNonZero) : · rw [Nat.one_le_iff_ne_zero] intro e exact α.2 (by simpa [e, Weight.coe_zero] using - weightSpace_chainTopCoeff_add_one_nsmul_add α (0 : Weight K H L) hα) + genWeightSpace_chainTopCoeff_add_one_nsmul_add α (0 : Weight K H L) hα) obtain ⟨x, hx, x_ne0⟩ := (chainTop α (0 : Weight K H L)).exists_ne_zero obtain ⟨h, e, f, isSl2, he, hf⟩ := exists_isSl2Triple_of_weight_isNonZero hα obtain rfl := isSl2.h_eq_coroot hα he hf have prim : isSl2.HasPrimitiveVectorWith x (chainLength α (0 : Weight K H L) : K) := - have := lie_mem_weightSpace_of_mem_weightSpace he hx - ⟨x_ne0, (chainLength_smul _ _ hx).symm, by rwa [weightSpace_add_chainTop _ _ hα] at this⟩ + have := lie_mem_genWeightSpace_of_mem_genWeightSpace he hx + ⟨x_ne0, (chainLength_smul _ _ hx).symm, by rwa [genWeightSpace_add_chainTop _ _ hα] at this⟩ obtain ⟨k, hk⟩ : ∃ k : K, k • f = (toEnd K L L f ^ (chainTopCoeff α (0 : Weight K H L) + 1)) x := by have : (toEnd K L L f ^ (chainTopCoeff α (0 : Weight K H L) + 1)) x ∈ rootSpace H (-α) := by @@ -290,7 +290,7 @@ lemma rootSpace_two_smul (hα : α.IsNonZero) : rootSpace H (2 • α) = ⊥ := cases subsingleton_or_nontrivial L · exact IsEmpty.elim inferInstance α simpa [chainTopCoeff_zero_right α hα] using - weightSpace_chainTopCoeff_add_one_nsmul_add α (0 : Weight K H L) hα + genWeightSpace_chainTopCoeff_add_one_nsmul_add α (0 : Weight K H L) hα lemma rootSpace_one_div_two_smul (hα : α.IsNonZero) : rootSpace H ((2⁻¹ : K) • α) = ⊥ := by by_contra h @@ -298,7 +298,7 @@ lemma rootSpace_one_div_two_smul (hα : α.IsNonZero) : rootSpace H ((2⁻¹ : K have hW : 2 • (W : H → K) = α := by show 2 • (2⁻¹ : K) • (α : H → K) = α rw [← Nat.cast_smul_eq_nsmul K, smul_smul]; simp - apply α.weightSpace_ne_bot + apply α.genWeightSpace_ne_bot have := rootSpace_two_smul W (fun (e : (W : H → K) = 0) ↦ hα <| by apply_fun (2 • ·) at e; simpa [hW] using e) rwa [hW] at this @@ -353,9 +353,9 @@ lemma eq_neg_or_eq_of_eq_smul (hβ : β.IsNonZero) (k : K) (h : (β : H → K) = /-- The reflection of a root along another. -/ def reflectRoot (α β : Weight K H L) : Weight K H L where toFun := β - β (coroot α) • α - weightSpace_ne_bot' := by + genWeightSpace_ne_bot' := by by_cases hα : α.IsZero - · simpa [hα.eq] using β.weightSpace_ne_bot + · simpa [hα.eq] using β.genWeightSpace_ne_bot rw [sub_eq_neg_add, apply_coroot_eq_cast α β, ← neg_smul, ← Int.cast_neg, Int.cast_smul_eq_zsmul, rootSpace_zsmul_add_ne_bot_iff α β hα] omega